aboutsummaryrefslogtreecommitdiffstats
path: root/kvx
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2020-10-16 15:52:19 +0200
committerCyril SIX <cyril.six@kalray.eu>2020-10-16 15:52:19 +0200
commitae3e1ed0515236e25924b12d475bc991a33b7632 (patch)
tree154428527771fbdec77f23846652d252400f81c8 /kvx
parentfae36491fa22adaaf447e189988848483eb01dcd (diff)
parent53544f625ac6ed6ddb000f5ae8f28faac0da7a7b (diff)
downloadcompcert-kvx-ae3e1ed0515236e25924b12d475bc991a33b7632.tar.gz
compcert-kvx-ae3e1ed0515236e25924b12d475bc991a33b7632.zip
Merge remote-tracking branch 'origin/kvx-test-prepass' into mppa-RTLpathSE-verif
Diffstat (limited to 'kvx')
-rw-r--r--kvx/Asm.v40
-rw-r--r--kvx/Asmblock.v52
-rw-r--r--kvx/Asmblockprops.v6
-rw-r--r--kvx/Asmgenproof.v4
-rw-r--r--kvx/Asmvliw.v16
-rw-r--r--kvx/CSE2depsproof.v6
-rw-r--r--kvx/CombineOpproof.v56
-rw-r--r--kvx/ConstpropOpproof.v196
-rw-r--r--kvx/Conventions1.v34
-rw-r--r--kvx/ExtValues.v72
-rw-r--r--kvx/InstructionScheduler.ml1254
-rw-r--r--kvx/InstructionScheduler.mli113
-rw-r--r--kvx/NeedOp.v54
-rw-r--r--kvx/Op.v484
-rw-r--r--kvx/OpWeights.ml4
-rw-r--r--kvx/Peephole.v2
-rw-r--r--kvx/PostpassSchedulingOracle.ml27
l---------kvx/PrepassSchedulingOracle.ml1
-rw-r--r--kvx/Stacklayout.v6
-rw-r--r--kvx/ValueAOp.v313
-rw-r--r--kvx/abstractbb/AbstractBasicBlocksDef.v94
-rw-r--r--kvx/abstractbb/ImpSimuTest.v82
-rw-r--r--kvx/abstractbb/Impure/ImpConfig.v85
-rw-r--r--kvx/abstractbb/Impure/ImpCore.v196
-rw-r--r--kvx/abstractbb/Impure/ImpExtern.v7
-rw-r--r--kvx/abstractbb/Impure/ImpHCons.v199
-rw-r--r--kvx/abstractbb/Impure/ImpIO.v159
-rw-r--r--kvx/abstractbb/Impure/ImpLoops.v123
-rw-r--r--kvx/abstractbb/Impure/ImpMonads.v148
-rw-r--r--kvx/abstractbb/Impure/ImpPrelude.v206
-rw-r--r--kvx/abstractbb/Impure/LICENSE165
-rw-r--r--kvx/abstractbb/Impure/README.md31
-rw-r--r--kvx/abstractbb/Impure/ocaml/ImpHConsOracles.ml72
-rw-r--r--kvx/abstractbb/Impure/ocaml/ImpHConsOracles.mli5
-rw-r--r--kvx/abstractbb/Impure/ocaml/ImpIOOracles.ml142
-rw-r--r--kvx/abstractbb/Impure/ocaml/ImpIOOracles.mli33
-rw-r--r--kvx/abstractbb/Impure/ocaml/ImpLoopOracles.ml78
-rw-r--r--kvx/abstractbb/Impure/ocaml/ImpLoopOracles.mli8
-rw-r--r--kvx/abstractbb/Parallelizability.v149
-rw-r--r--kvx/abstractbb/SeqSimuTheory.v77
-rw-r--r--kvx/lib/ForwardSimulationBlock.v30
-rw-r--r--kvx/lib/Machblock.v14
-rw-r--r--kvx/lib/Machblockgen.v14
-rw-r--r--kvx/lib/Machblockgenproof.v138
-rw-r--r--kvx/lib/PrepassSchedulingOracle.ml432
-rw-r--r--kvx/lib/RTLpath.v1067
-rw-r--r--kvx/lib/RTLpathLivegen.v290
-rw-r--r--kvx/lib/RTLpathLivegenaux.ml309
-rw-r--r--kvx/lib/RTLpathLivegenproof.v736
-rw-r--r--kvx/lib/RTLpathSE_impl.v1529
-rw-r--r--kvx/lib/RTLpathSE_theory.v1854
-rw-r--r--kvx/lib/RTLpathScheduler.v330
-rw-r--r--kvx/lib/RTLpathScheduleraux.ml368
-rw-r--r--kvx/lib/RTLpathSchedulerproof.v363
-rw-r--r--kvx/lib/RTLpathproof.v50
55 files changed, 859 insertions, 11464 deletions
diff --git a/kvx/Asm.v b/kvx/Asm.v
index 30aafc55..515e13e0 100644
--- a/kvx/Asm.v
+++ b/kvx/Asm.v
@@ -611,15 +611,15 @@ Program Definition genv_trans (ge: genv) : Asmvliw.genv :=
Genv.genv_defs := PTree.map1 globdef_proj (Genv.genv_defs ge);
Genv.genv_next := Genv.genv_next ge |}.
Next Obligation.
- destruct ge. simpl in *. eauto.
+ destruct ge. cbn in *. eauto.
Qed. Next Obligation.
- destruct ge; simpl in *.
+ destruct ge; cbn in *.
rewrite PTree.gmap1 in H.
destruct (genv_defs ! b) eqn:GEN.
- eauto.
- discriminate.
Qed. Next Obligation.
- destruct ge; simpl in *.
+ destruct ge; cbn in *.
eauto.
Qed.
@@ -655,14 +655,14 @@ Program Definition transf_function (f: Asmvliw.function) : function :=
Lemma transf_function_proj: forall f, function_proj (transf_function f) = f.
Proof.
- intros f. destruct f as [sig blks]. unfold function_proj. simpl. auto.
+ intros f. destruct f as [sig blks]. unfold function_proj. cbn. auto.
Qed.
Definition transf_fundef : Asmvliw.fundef -> fundef := AST.transf_fundef transf_function.
Lemma transf_fundef_proj: forall f, fundef_proj (transf_fundef f) = f.
Proof.
- intros f. destruct f as [f|e]; simpl; auto.
+ intros f. destruct f as [f|e]; cbn; auto.
rewrite transf_function_proj. auto.
Qed.
@@ -674,18 +674,18 @@ Lemma program_equals {A B: Type} : forall (p1 p2: AST.program A B),
prog_main p1 = prog_main p2 ->
p1 = p2.
Proof.
- intros. destruct p1. destruct p2. simpl in *. subst. auto.
+ intros. destruct p1. destruct p2. cbn in *. subst. auto.
Qed.
Lemma transf_program_proj: forall p, program_proj (transf_program p) = p.
Proof.
- intros p. destruct p as [defs pub main]. unfold program_proj. simpl.
- apply program_equals; simpl; auto.
+ intros p. destruct p as [defs pub main]. unfold program_proj. cbn.
+ apply program_equals; cbn; auto.
induction defs.
- - simpl; auto.
- - simpl. rewrite IHdefs.
- destruct a as [id gd]; simpl.
- destruct gd as [f|v]; simpl; auto.
+ - cbn; auto.
+ - cbn. rewrite IHdefs.
+ destruct a as [id gd]; cbn.
+ destruct gd as [f|v]; cbn; auto.
rewrite transf_fundef_proj. auto.
Qed.
@@ -707,16 +707,16 @@ Lemma match_program_transf:
forall p tp, match_prog p tp -> transf_program p = tp.
Proof.
intros p tp H. inversion_clear H. inv H1.
- destruct p as [defs pub main]. destruct tp as [tdefs tpub tmain]. simpl in *.
- subst. unfold transf_program. unfold transform_program. simpl.
- apply program_equals; simpl; auto.
- induction H0; simpl; auto.
+ destruct p as [defs pub main]. destruct tp as [tdefs tpub tmain]. cbn in *.
+ subst. unfold transf_program. unfold transform_program. cbn.
+ apply program_equals; cbn; auto.
+ induction H0; cbn; auto.
rewrite IHlist_forall2. apply cons_extract.
destruct a1 as [ida gda]. destruct b1 as [idb gdb].
- simpl in *.
+ cbn in *.
inv H. inv H2.
- - simpl in *. subst. auto.
- - simpl in *. subst. inv H. auto.
+ - cbn in *. subst. auto.
+ - cbn in *. subst. inv H. auto.
Qed.
Section PRESERVATION.
@@ -744,7 +744,7 @@ Proof.
pose proof (match_program_transf prog tprog TRANSF) as TR.
subst. unfold semantics. rewrite transf_program_proj.
- eapply forward_simulation_step with (match_states := match_states); simpl; auto.
+ eapply forward_simulation_step with (match_states := match_states); cbn; auto.
- intros. exists s1. split; auto. congruence.
- intros. inv H. auto.
- intros. exists s1'. inv H0. split; auto. congruence.
diff --git a/kvx/Asmblock.v b/kvx/Asmblock.v
index 9c8e4cc3..64b2c535 100644
--- a/kvx/Asmblock.v
+++ b/kvx/Asmblock.v
@@ -78,7 +78,7 @@ Fixpoint code_to_basics (c: code) :=
Lemma code_to_basics_id: forall c, code_to_basics (basics_to_code c) = Some c.
Proof.
- intros. induction c as [|i c]; simpl; auto.
+ intros. induction c as [|i c]; cbn; auto.
rewrite IHc. auto.
Qed.
@@ -88,8 +88,8 @@ Lemma code_to_basics_dist:
code_to_basics c' = Some l' ->
code_to_basics (c ++ c') = Some (l ++ l').
Proof.
- induction c as [|i c]; simpl; auto.
- - intros. inv H. simpl. auto.
+ induction c as [|i c]; cbn; auto.
+ - intros. inv H. cbn. auto.
- intros. destruct i; try discriminate. destruct (code_to_basics c) eqn:CTB; try discriminate.
inv H. erewrite IHc; eauto. auto.
Qed.
@@ -138,9 +138,9 @@ Lemma non_empty_bblock_refl:
Proof.
intros. split.
- destruct body; destruct exit.
- all: simpl; auto. intros. inversion H; contradiction.
+ all: cbn; auto. intros. inversion H; contradiction.
- destruct body; destruct exit.
- all: simpl; auto.
+ all: cbn; auto.
all: intros; try (right; discriminate); try (left; discriminate).
contradiction.
Qed.
@@ -155,14 +155,14 @@ Lemma builtin_alone_refl:
Proof.
intros. split.
- destruct body; destruct exit.
- all: simpl; auto.
- all: exploreInst; simpl; auto.
+ all: cbn; auto.
+ all: exploreInst; cbn; auto.
unfold builtin_alone. intros. assert (Some (Pbuiltin e l b0) = Some (Pbuiltin e l b0)); auto.
assert (b :: body = nil). eapply H; eauto. discriminate.
- destruct body; destruct exit.
- all: simpl; auto; try constructor.
+ all: cbn; auto; try constructor.
+ exploreInst; try discriminate.
- simpl. contradiction.
+ cbn. contradiction.
+ intros. discriminate.
Qed.
@@ -185,14 +185,14 @@ Ltac bblock_auto_correct := (apply non_empty_bblock_refl; try discriminate; try
Lemma Istrue_proof_irrelevant (b: bool): forall (p1 p2:Is_true b), p1=p2.
Proof.
- destruct b; simpl; auto.
+ destruct b; cbn; auto.
- destruct p1, p2; auto.
- destruct p1.
Qed.
Lemma bblock_equality bb1 bb2: header bb1=header bb2 -> body bb1 = body bb2 -> exit bb1 = exit bb2 -> bb1 = bb2.
Proof.
- destruct bb1 as [h1 b1 e1 c1], bb2 as [h2 b2 e2 c2]; simpl.
+ destruct bb1 as [h1 b1 e1 c1], bb2 as [h2 b2 e2 c2]; cbn.
intros; subst.
rewrite (Istrue_proof_irrelevant _ c1 c2).
auto.
@@ -212,51 +212,51 @@ Qed.
Lemma length_nonil {A: Type} : forall l:(list A), l <> nil -> (length l > 0)%nat.
Proof.
intros. destruct l; try (contradict H; auto; fail).
- simpl. omega.
+ cbn. omega.
Qed.
Lemma to_nat_pos : forall z:Z, (Z.to_nat z > 0)%nat -> z > 0.
Proof.
intros. destruct z; auto.
- - contradict H. simpl. apply gt_irrefl.
+ - contradict H. cbn. apply gt_irrefl.
- apply Zgt_pos_0.
- - contradict H. simpl. apply gt_irrefl.
+ - contradict H. cbn. apply gt_irrefl.
Qed.
Lemma size_positive (b:bblock): size b > 0.
Proof.
- unfold size. destruct b as [hd bdy ex cor]. simpl.
- destruct ex; destruct bdy; try (apply to_nat_pos; rewrite Nat2Z.id; simpl; omega).
- inversion cor; contradict H; simpl; auto.
+ unfold size. destruct b as [hd bdy ex cor]. cbn.
+ destruct ex; destruct bdy; try (apply to_nat_pos; rewrite Nat2Z.id; cbn; omega).
+ inversion cor; contradict H; cbn; auto.
Qed.
Program Definition no_header (bb : bblock) := {| header := nil; body := body bb; exit := exit bb |}.
Next Obligation.
- destruct bb; simpl. assumption.
+ destruct bb; cbn. assumption.
Defined.
Lemma no_header_size:
forall bb, size (no_header bb) = size bb.
Proof.
- intros. destruct bb as [hd bdy ex COR]. unfold no_header. simpl. reflexivity.
+ intros. destruct bb as [hd bdy ex COR]. unfold no_header. cbn. reflexivity.
Qed.
Program Definition stick_header (h : list label) (bb : bblock) := {| header := h; body := body bb; exit := exit bb |}.
Next Obligation.
- destruct bb; simpl. assumption.
+ destruct bb; cbn. assumption.
Defined.
Lemma stick_header_size:
forall h bb, size (stick_header h bb) = size bb.
Proof.
- intros. destruct bb. unfold stick_header. simpl. reflexivity.
+ intros. destruct bb. unfold stick_header. cbn. reflexivity.
Qed.
Lemma stick_header_no_header:
forall bb, stick_header (header bb) (no_header bb) = bb.
Proof.
- intros. destruct bb as [hd bdy ex COR]. simpl. unfold no_header; unfold stick_header; simpl. reflexivity.
+ intros. destruct bb as [hd bdy ex COR]. cbn. unfold no_header; unfold stick_header; cbn. reflexivity.
Qed.
(** * Sequential Semantics of basic blocks *)
@@ -308,7 +308,7 @@ Fixpoint exec_body (body: list basic) (rs: regset) (m: mem): outcome :=
Theorem builtin_body_nil:
forall bb ef args res, exit bb = Some (PExpand (Pbuiltin ef args res)) -> body bb = nil.
Proof.
- intros. destruct bb as [hd bdy ex WF]. simpl in *.
+ intros. destruct bb as [hd bdy ex WF]. cbn in *.
apply wf_bblock_refl in WF. inv WF. unfold builtin_alone in H1.
eapply H1; eauto.
Qed.
@@ -321,11 +321,11 @@ Theorem exec_body_app:
/\ exec_body l' rs' m' = Next rs'' m''.
Proof.
induction l.
- - intros. simpl in H. repeat eexists. auto.
- - intros. rewrite <- app_comm_cons in H. simpl in H.
+ - intros. cbn in H. repeat eexists. auto.
+ - intros. rewrite <- app_comm_cons in H. cbn in H.
destruct (exec_basic_instr a rs m) eqn:EXEBI.
+ apply IHl in H. destruct H as (rs1 & m1 & EXEB1 & EXEB2).
- repeat eexists. simpl. rewrite EXEBI. eauto. auto.
+ repeat eexists. cbn. rewrite EXEBI. eauto. auto.
+ discriminate.
Qed.
diff --git a/kvx/Asmblockprops.v b/kvx/Asmblockprops.v
index bc14b231..c3929be5 100644
--- a/kvx/Asmblockprops.v
+++ b/kvx/Asmblockprops.v
@@ -53,7 +53,7 @@ Qed.
Lemma preg_of_not_SP:
forall r, preg_of r <> SP.
Proof.
- intros. unfold preg_of; destruct r; simpl; congruence.
+ intros. unfold preg_of; destruct r; cbn; congruence.
Qed.
Hint Resolve preg_of_not_SP preg_of_not_PC: asmgen.
@@ -233,7 +233,7 @@ Proof.
destruct (ireg_eq rd2 ra); try discriminate.
*)
rewrite Pregmap.gso; try discriminate.
- simpl in *.
+ cbn in *.
destruct (Mem.loadv _ _ _); try discriminate.
destruct (Mem.loadv _ _ _); try discriminate.
destruct (Mem.loadv _ _ _); try discriminate.
@@ -264,7 +264,7 @@ Lemma exec_store_q_offset_pc_var:
exec_store_q_offset rs # PC <- v m rd ra ofs = Next rs' # PC <- v m'.
Proof.
intros. unfold exec_store_q_offset in *. unfold parexec_store_q_offset in *. rewrite Pregmap.gso; try discriminate.
- simpl in *.
+ cbn in *.
destruct (gpreg_q_expand _) as [s0 s1].
destruct (Mem.storev _ _ _); try discriminate.
destruct (Mem.storev _ _ _); try discriminate.
diff --git a/kvx/Asmgenproof.v b/kvx/Asmgenproof.v
index 9e35e268..636c105f 100644
--- a/kvx/Asmgenproof.v
+++ b/kvx/Asmgenproof.v
@@ -39,7 +39,7 @@ Proof.
unfold Asmgen.transf_program in H. apply bind_inversion in H. destruct H.
inversion_clear H. apply bind_inversion in H1. destruct H1.
inversion_clear H. inversion H2. unfold time, Compopts.time in *. remember (Machblockgen.transf_program p) as mbp.
- unfold match_prog; simpl.
+ unfold match_prog; cbn.
exists mbp; split. apply Machblockgenproof.transf_program_match; auto.
exists x; split. apply Asmblockgenproof.transf_program_match; auto.
exists x0; split. apply PostpassSchedulingproof.transf_program_match; auto.
@@ -72,7 +72,7 @@ Let tge := Genv.globalenv tprog.
Theorem transf_program_correct:
forward_simulation (Mach.semantics return_address_offset prog) (Asm.semantics tprog).
Proof.
- unfold match_prog in TRANSF. simpl in TRANSF.
+ unfold match_prog in TRANSF. cbn in TRANSF.
inv TRANSF. inv H. inv H1. inv H. inv H2. inv H. inv H3. inv H.
eapply compose_forward_simulations.
exploit Machblockgenproof.transf_program_correct; eauto.
diff --git a/kvx/Asmvliw.v b/kvx/Asmvliw.v
index 296963a7..66b468d7 100644
--- a/kvx/Asmvliw.v
+++ b/kvx/Asmvliw.v
@@ -849,7 +849,7 @@ Lemma Val_cmpu_correct (m:mem) (cmp: comparison) (v1 v2: val):
Proof.
unfold Val.cmpu, Val_cmpu.
remember (Val.cmpu_bool (Mem.valid_pointer m) cmp v1 v2) as ob.
- destruct ob; simpl.
+ destruct ob; cbn.
- erewrite Val_cmpu_bool_correct; eauto.
econstructor.
- econstructor.
@@ -873,9 +873,9 @@ Lemma Val_cmplu_correct (m:mem) (cmp: comparison) (v1 v2: val):
Proof.
unfold Val.cmplu, Val_cmplu.
remember (Val.cmplu_bool (Mem.valid_pointer m) cmp v1 v2) as ob.
- destruct ob as [b|]; simpl.
+ destruct ob as [b|]; cbn.
- erewrite Val_cmplu_bool_correct; eauto.
- simpl. econstructor.
+ cbn. econstructor.
- econstructor.
Qed.
@@ -1426,13 +1426,13 @@ Definition is_label (lbl: label) (bb: bblock) : bool :=
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)); simpl; intuition.
+ 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)); simpl; intuition.
+ unfold is_label; destruct (in_dec lbl (header bb)); cbn; intuition.
Qed.
@@ -1667,7 +1667,7 @@ Proof.
constructor 1.
- rewrite app_nil_r; auto.
- unfold parexec_wio_bblock.
- destruct (parexec_wio f _ _ _); simpl; auto.
+ destruct (parexec_wio f _ _ _); cbn; auto.
Qed.
@@ -1739,7 +1739,7 @@ Ltac Det_WIO X :=
exploit det_parexec_write_in_order; [ eapply H | idtac]; clear H; intro X
| _ => idtac
end.
- intros; constructor; simpl.
+ intros; constructor; cbn.
- (* determ *) intros s t1 s1 t2 s2 H H0. inv H; Det_WIO X1;
inv H0; Det_WIO X2; Equalities.
+ split. constructor. auto.
@@ -1754,7 +1754,7 @@ Ltac Det_WIO X :=
exploit external_call_determ. eexact H3. eexact H8. intros [A B].
split. auto. intros. destruct B; auto. subst. auto.
- (* trace length *)
- red; intros. inv H; simpl.
+ red; intros. inv H; cbn.
omega.
eapply external_call_trace_length; eauto.
eapply external_call_trace_length; eauto.
diff --git a/kvx/CSE2depsproof.v b/kvx/CSE2depsproof.v
index f283c8ac..6c584450 100644
--- a/kvx/CSE2depsproof.v
+++ b/kvx/CSE2depsproof.v
@@ -71,7 +71,7 @@ Section MEMORY_WRITE.
unfold largest_size_chunk in *.
rewrite ptrofs_modulus in *.
- simpl in *.
+ cbn in *.
inv ADDRR.
inv ADDRW.
destruct base; try discriminate.
@@ -126,12 +126,12 @@ Proof.
{ (* Aindexed / Aindexed *)
destruct args as [ | base [ | ]]. 1,3: discriminate.
destruct args' as [ | base' [ | ]]. 1,3: discriminate.
- simpl in OVERLAP.
+ cbn in OVERLAP.
destruct (peq base base'). 2: discriminate.
subst base'.
destruct (can_swap_accesses_ofs (Ptrofs.unsigned i0) chunk' (Ptrofs.unsigned i) chunk) eqn:SWAP.
2: discriminate.
- simpl in *.
+ cbn in *.
eapply load_store_away with (F:=F) (V:=V) (genv:=genv) (sp:=sp); eassumption.
}
Qed.
diff --git a/kvx/CombineOpproof.v b/kvx/CombineOpproof.v
index dafc90df..5dffc565 100644
--- a/kvx/CombineOpproof.v
+++ b/kvx/CombineOpproof.v
@@ -46,7 +46,7 @@ Qed.
Ltac UseGetSound :=
match goal with
| [ H: get _ = Some _ |- _ ] =>
- let x := fresh "EQ" in (generalize (get_op_sound _ _ _ H); intros x; simpl in x; FuncInv)
+ let x := fresh "EQ" in (generalize (get_op_sound _ _ _ H); intros x; cbn in x; FuncInv)
end.
Lemma combine_compimm_ne_0_sound:
@@ -58,7 +58,7 @@ Proof.
intros until args. functional induction (combine_compimm_ne_0 get x); intros EQ; inv EQ.
(* of cmp *)
UseGetSound. rewrite <- H.
- destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto.
+ destruct (eval_condition cond (map valu args) m); cbn; auto. destruct b; auto.
Qed.
Lemma combine_compimm_eq_0_sound:
@@ -71,7 +71,7 @@ Proof.
(* of cmp *)
UseGetSound. rewrite <- H.
rewrite eval_negate_condition.
- destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto.
+ destruct (eval_condition c (map valu args) m); cbn; auto. destruct b; auto.
Qed.
Lemma combine_compimm_eq_1_sound:
@@ -83,7 +83,7 @@ Proof.
intros until args. functional induction (combine_compimm_eq_1 get x); intros EQ; inv EQ.
(* of cmp *)
UseGetSound. rewrite <- H.
- destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto.
+ destruct (eval_condition cond (map valu args) m); cbn; auto. destruct b; auto.
Qed.
Lemma combine_compimm_ne_1_sound:
@@ -96,7 +96,7 @@ Proof.
(* of cmp *)
UseGetSound. rewrite <- H.
rewrite eval_negate_condition.
- destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto.
+ destruct (eval_condition c (map valu args) m); cbn; auto. destruct b; auto.
Qed.
Theorem combine_cond_sound:
@@ -106,21 +106,21 @@ Theorem combine_cond_sound:
Proof.
intros. functional inversion H; subst.
(* compimm ne zero *)
- - simpl; eapply combine_compimm_ne_0_sound; eauto.
+ - cbn; eapply combine_compimm_ne_0_sound; eauto.
(* compimm ne one *)
- - simpl; eapply combine_compimm_ne_1_sound; eauto.
+ - cbn; eapply combine_compimm_ne_1_sound; eauto.
(* compimm eq zero *)
- - simpl; eapply combine_compimm_eq_0_sound; eauto.
+ - cbn; eapply combine_compimm_eq_0_sound; eauto.
(* compimm eq one *)
- - simpl; eapply combine_compimm_eq_1_sound; eauto.
+ - cbn; eapply combine_compimm_eq_1_sound; eauto.
(* compuimm ne zero *)
- - simpl; eapply combine_compimm_ne_0_sound; eauto.
+ - cbn; eapply combine_compimm_ne_0_sound; eauto.
(* compuimm ne one *)
- - simpl; eapply combine_compimm_ne_1_sound; eauto.
+ - cbn; eapply combine_compimm_ne_1_sound; eauto.
(* compuimm eq zero *)
- - simpl; eapply combine_compimm_eq_0_sound; eauto.
+ - cbn; eapply combine_compimm_eq_0_sound; eauto.
(* compuimm eq one *)
- - simpl; eapply combine_compimm_eq_1_sound; eauto.
+ - cbn; eapply combine_compimm_eq_1_sound; eauto.
Qed.
Theorem combine_addr_sound:
@@ -130,10 +130,10 @@ Theorem combine_addr_sound:
Proof.
intros. functional inversion H; subst.
- (* indexed - addimm *)
- UseGetSound. simpl. rewrite <- H0. destruct v; auto. simpl; rewrite H7; simpl.
+ UseGetSound. cbn. rewrite <- H0. destruct v; auto. cbn; rewrite H7; cbn.
rewrite Ptrofs.add_assoc. auto.
- (* indexed - addimml *)
- UseGetSound. simpl. rewrite <- H0. destruct v; auto. simpl; rewrite H7; simpl.
+ UseGetSound. cbn. rewrite <- H0. destruct v; auto. cbn; rewrite H7; cbn.
rewrite Ptrofs.add_assoc. auto.
Qed.
@@ -144,33 +144,33 @@ Theorem combine_op_sound:
Proof.
intros. functional inversion H; subst.
(* addimm - addimm *)
- - UseGetSound. FuncInv. simpl.
+ - UseGetSound. FuncInv. cbn.
rewrite <- H0. rewrite Val.add_assoc. auto.
(* andimm - andimm *)
- - UseGetSound; simpl.
+ - UseGetSound; cbn.
generalize (Int.eq_spec p m0); rewrite H7; intros.
- rewrite <- H0. rewrite Val.and_assoc. simpl. fold p. rewrite H1. auto.
- - UseGetSound; simpl.
+ rewrite <- H0. rewrite Val.and_assoc. cbn. fold p. rewrite H1. auto.
+ - UseGetSound; cbn.
rewrite <- H0. rewrite Val.and_assoc. auto.
(* orimm - orimm *)
- - UseGetSound. simpl. rewrite <- H0. rewrite Val.or_assoc. auto.
+ - UseGetSound. cbn. rewrite <- H0. rewrite Val.or_assoc. auto.
(* xorimm - xorimm *)
- - UseGetSound. simpl. rewrite <- H0. rewrite Val.xor_assoc. auto.
+ - UseGetSound. cbn. rewrite <- H0. rewrite Val.xor_assoc. auto.
(* addlimm - addlimm *)
- - UseGetSound. FuncInv. simpl.
+ - UseGetSound. FuncInv. cbn.
rewrite <- H0. rewrite Val.addl_assoc. auto.
(* andlimm - andlimm *)
- - UseGetSound; simpl.
+ - UseGetSound; cbn.
generalize (Int64.eq_spec p m0); rewrite H7; intros.
- rewrite <- H0. rewrite Val.andl_assoc. simpl. fold p. rewrite H1. auto.
- - UseGetSound; simpl.
+ rewrite <- H0. rewrite Val.andl_assoc. cbn. fold p. rewrite H1. auto.
+ - UseGetSound; cbn.
rewrite <- H0. rewrite Val.andl_assoc. auto.
(* orlimm - orlimm *)
- - UseGetSound. simpl. rewrite <- H0. rewrite Val.orl_assoc. auto.
+ - UseGetSound. cbn. rewrite <- H0. rewrite Val.orl_assoc. auto.
(* xorlimm - xorlimm *)
- - UseGetSound. simpl. rewrite <- H0. rewrite Val.xorl_assoc. auto.
+ - UseGetSound. cbn. rewrite <- H0. rewrite Val.xorl_assoc. auto.
(* cmp *)
- - simpl. decEq; decEq. eapply combine_cond_sound; eauto.
+ - cbn. decEq; decEq. eapply combine_cond_sound; eauto.
Qed.
End COMBINE.
diff --git a/kvx/ConstpropOpproof.v b/kvx/ConstpropOpproof.v
index 05bbdde1..ffd35bcc 100644
--- a/kvx/ConstpropOpproof.v
+++ b/kvx/ConstpropOpproof.v
@@ -105,7 +105,7 @@ Proof.
+ (* global *)
inv H2. exists (Genv.symbol_address ge id ofs); auto.
+ (* stack *)
- inv H2. exists (Vptr sp ofs); split; auto. simpl. rewrite Ptrofs.add_zero_l; auto.
+ inv H2. exists (Vptr sp ofs); split; auto. cbn. rewrite Ptrofs.add_zero_l; auto.
Qed.
Lemma cond_strength_reduction_correct:
@@ -115,7 +115,7 @@ Lemma cond_strength_reduction_correct:
eval_condition cond' e##args' m = eval_condition cond e##args m.
Proof.
intros until vl. unfold cond_strength_reduction.
- case (cond_strength_reduction_match cond args vl); simpl; intros; InvApproxRegs; SimplVM.
+ case (cond_strength_reduction_match cond args vl); cbn; intros; InvApproxRegs; SimplVM.
- apply Val.swap_cmp_bool.
- auto.
- apply Val.swap_cmpu_bool.
@@ -137,7 +137,7 @@ Proof.
intros. unfold make_cmp_base.
generalize (cond_strength_reduction_correct c args vl H).
destruct (cond_strength_reduction c args vl) as [c' args']. intros EQ.
- econstructor; split. simpl; eauto. rewrite EQ. auto.
+ econstructor; split. cbn; eauto. rewrite EQ. auto.
Qed.
Lemma make_cmp_correct:
@@ -154,43 +154,43 @@ Proof.
unfold make_cmp. case (make_cmp_match c args vl); intros.
- unfold make_cmp_imm_eq.
destruct (Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1)) eqn:E1.
-+ simpl in H; inv H. InvBooleans. subst n.
- exists (e#r1); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
++ cbn in H; inv H. InvBooleans. subst n.
+ exists (e#r1); split; auto. cbn.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; cbn; auto.
+ destruct (Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1)) eqn:E0.
-* simpl in H; inv H. InvBooleans. subst n.
- exists (Val.xor e#r1 (Vint Int.one)); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
+* cbn in H; inv H. InvBooleans. subst n.
+ exists (Val.xor e#r1 (Vint Int.one)); split; auto. cbn.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; cbn; auto.
* apply make_cmp_base_correct; auto.
- unfold make_cmp_imm_ne.
destruct (Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1)) eqn:E0.
-+ simpl in H; inv H. InvBooleans. subst n.
- exists (e#r1); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
++ cbn in H; inv H. InvBooleans. subst n.
+ exists (e#r1); split; auto. cbn.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; cbn; auto.
+ destruct (Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1)) eqn:E1.
-* simpl in H; inv H. InvBooleans. subst n.
- exists (Val.xor e#r1 (Vint Int.one)); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
+* cbn in H; inv H. InvBooleans. subst n.
+ exists (Val.xor e#r1 (Vint Int.one)); split; auto. cbn.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; cbn; auto.
* apply make_cmp_base_correct; auto.
- unfold make_cmp_imm_eq.
destruct (Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1)) eqn:E1.
-+ simpl in H; inv H. InvBooleans. subst n.
- exists (e#r1); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
++ cbn in H; inv H. InvBooleans. subst n.
+ exists (e#r1); split; auto. cbn.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; cbn; auto.
+ destruct (Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1)) eqn:E0.
-* simpl in H; inv H. InvBooleans. subst n.
- exists (Val.xor e#r1 (Vint Int.one)); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
+* cbn in H; inv H. InvBooleans. subst n.
+ exists (Val.xor e#r1 (Vint Int.one)); split; auto. cbn.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; cbn; auto.
* apply make_cmp_base_correct; auto.
- unfold make_cmp_imm_ne.
destruct (Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1)) eqn:E0.
-+ simpl in H; inv H. InvBooleans. subst n.
- exists (e#r1); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
++ cbn in H; inv H. InvBooleans. subst n.
+ exists (e#r1); split; auto. cbn.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; cbn; auto.
+ destruct (Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1)) eqn:E1.
-* simpl in H; inv H. InvBooleans. subst n.
- exists (Val.xor e#r1 (Vint Int.one)); split; auto. simpl.
- exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
+* cbn in H; inv H. InvBooleans. subst n.
+ exists (Val.xor e#r1 (Vint Int.one)); split; auto. cbn.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; cbn; auto.
* apply make_cmp_base_correct; auto.
- apply make_cmp_base_correct; auto.
Qed.
@@ -203,7 +203,7 @@ Proof.
intros. unfold make_addimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros.
subst. exists (e#r); split; auto.
- destruct (e#r); simpl; auto; rewrite ?Int.add_zero, ?Ptrofs.add_zero; auto.
+ destruct (e#r); cbn; auto; rewrite ?Int.add_zero, ?Ptrofs.add_zero; auto.
exists (Val.add e#r (Vint n)); split; auto.
Qed.
@@ -215,10 +215,10 @@ Lemma make_shlimm_correct:
Proof.
intros; unfold make_shlimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shl_zero. auto.
+ exists (e#r1); split; auto. destruct (e#r1); cbn; auto. rewrite Int.shl_zero. auto.
destruct (Int.ltu n Int.iwordsize).
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
+ econstructor; split. cbn. eauto. auto.
+ econstructor; split. cbn. eauto. rewrite H; auto.
Qed.
Lemma make_shrimm_correct:
@@ -229,10 +229,10 @@ Lemma make_shrimm_correct:
Proof.
intros; unfold make_shrimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shr_zero. auto.
+ exists (e#r1); split; auto. destruct (e#r1); cbn; auto. rewrite Int.shr_zero. auto.
destruct (Int.ltu n Int.iwordsize).
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
+ econstructor; split. cbn. eauto. auto.
+ econstructor; split. cbn. eauto. rewrite H; auto.
Qed.
Lemma make_shruimm_correct:
@@ -243,10 +243,10 @@ Lemma make_shruimm_correct:
Proof.
intros; unfold make_shruimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shru_zero. auto.
+ exists (e#r1); split; auto. destruct (e#r1); cbn; auto. rewrite Int.shru_zero. auto.
destruct (Int.ltu n Int.iwordsize).
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
+ econstructor; split. cbn. eauto. auto.
+ econstructor; split. cbn. eauto. rewrite H; auto.
Qed.
Lemma make_mulimm_correct:
@@ -257,12 +257,12 @@ Lemma make_mulimm_correct:
Proof.
intros; unfold make_mulimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (Vint Int.zero); split; auto. destruct (e#r1); simpl; auto. rewrite Int.mul_zero; auto.
+ exists (Vint Int.zero); split; auto. destruct (e#r1); cbn; auto. rewrite Int.mul_zero; auto.
predSpec Int.eq Int.eq_spec n Int.one; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.mul_one; auto.
+ exists (e#r1); split; auto. destruct (e#r1); cbn; auto. rewrite Int.mul_one; auto.
destruct (Int.is_power2 n) eqn:?; intros.
- rewrite (Val.mul_pow2 e#r1 _ _ Heqo). econstructor; split. simpl; eauto. auto.
- econstructor; split; eauto. simpl. rewrite H; auto.
+ rewrite (Val.mul_pow2 e#r1 _ _ Heqo). econstructor; split. cbn; eauto. auto.
+ econstructor; split; eauto. cbn. rewrite H; auto.
Qed.
Lemma make_divimm_correct:
@@ -275,11 +275,11 @@ Proof.
intros; unfold make_divimm.
predSpec Int.eq Int.eq_spec n Int.one; intros. subst. rewrite H0 in H.
destruct (e#r1) eqn:?;
- try (rewrite Val.divs_one in H; exists (Vint i); split; simpl; try rewrite Heqv0; auto);
+ try (rewrite Val.divs_one in H; exists (Vint i); split; cbn; try rewrite Heqv0; auto);
inv H; auto.
destruct (Int.is_power2 n) eqn:?.
destruct (Int.ltu i (Int.repr 31)) eqn:?.
- exists v; split; auto. simpl.
+ exists v; split; auto. cbn.
erewrite Val.divs_pow2; eauto. reflexivity. congruence.
exists v; auto.
exists v; auto.
@@ -295,10 +295,10 @@ Proof.
intros; unfold make_divuimm.
predSpec Int.eq Int.eq_spec n Int.one; intros. subst. rewrite H0 in H.
destruct (e#r1) eqn:?;
- try (rewrite Val.divu_one in H; exists (Vint i); split; simpl; try rewrite Heqv0; auto);
+ try (rewrite Val.divu_one in H; exists (Vint i); split; cbn; try rewrite Heqv0; auto);
inv H; auto.
destruct (Int.is_power2 n) eqn:?.
- econstructor; split. simpl; eauto.
+ econstructor; split. cbn; eauto.
rewrite H0 in H. erewrite Val.divu_pow2 by eauto. auto.
exists v; auto.
Qed.
@@ -312,7 +312,7 @@ Lemma make_moduimm_correct:
Proof.
intros; unfold make_moduimm.
destruct (Int.is_power2 n) eqn:?.
- exists v; split; auto. simpl. decEq. eapply Val.modu_pow2; eauto. congruence.
+ exists v; split; auto. cbn. decEq. eapply Val.modu_pow2; eauto. congruence.
exists v; auto.
Qed.
@@ -324,18 +324,18 @@ Lemma make_andimm_correct:
Proof.
intros; unfold make_andimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros.
- subst n. exists (Vint Int.zero); split; auto. destruct (e#r); simpl; auto. rewrite Int.and_zero; auto.
+ subst n. exists (Vint Int.zero); split; auto. destruct (e#r); cbn; auto. rewrite Int.and_zero; auto.
predSpec Int.eq Int.eq_spec n Int.mone; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.and_mone; auto.
+ subst n. exists (e#r); split; auto. destruct (e#r); cbn; auto. rewrite Int.and_mone; auto.
destruct (match x with Uns _ k => Int.eq (Int.zero_ext k (Int.not n)) Int.zero
| _ => false end) eqn:UNS.
destruct x; try congruence.
exists (e#r); split; auto.
- inv H; auto. simpl. replace (Int.and i n) with i; auto.
+ inv H; auto. cbn. replace (Int.and i n) with i; auto.
generalize (Int.eq_spec (Int.zero_ext n0 (Int.not n)) Int.zero); rewrite UNS; intro EQ.
Int.bit_solve. destruct (zlt i0 n0).
replace (Int.testbit n i0) with (negb (Int.testbit Int.zero i0)).
- rewrite Int.bits_zero. simpl. rewrite andb_true_r. auto.
+ rewrite Int.bits_zero. cbn. rewrite andb_true_r. auto.
rewrite <- EQ. rewrite Int.bits_zero_ext by omega. rewrite zlt_true by auto.
rewrite Int.bits_not by auto. apply negb_involutive.
rewrite H6 by auto. auto.
@@ -349,9 +349,9 @@ Lemma make_orimm_correct:
Proof.
intros; unfold make_orimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.or_zero; auto.
+ subst n. exists (e#r); split; auto. destruct (e#r); cbn; auto. rewrite Int.or_zero; auto.
predSpec Int.eq Int.eq_spec n Int.mone; intros.
- subst n. exists (Vint Int.mone); split; auto. destruct (e#r); simpl; auto. rewrite Int.or_mone; auto.
+ subst n. exists (Vint Int.mone); split; auto. destruct (e#r); cbn; auto. rewrite Int.or_mone; auto.
econstructor; split; eauto. auto.
Qed.
@@ -362,7 +362,7 @@ Lemma make_xorimm_correct:
Proof.
intros; unfold make_xorimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.xor_zero; auto.
+ subst n. exists (e#r); split; auto. destruct (e#r); cbn; auto. rewrite Int.xor_zero; auto.
predSpec Int.eq Int.eq_spec n Int.mone; intros.
subst n. exists (Val.notint e#r); split; auto.
econstructor; split; eauto. auto.
@@ -376,7 +376,7 @@ Proof.
intros. unfold make_addlimm.
predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
subst. exists (e#r); split; auto.
- destruct (e#r); simpl; auto; rewrite ? Int64.add_zero, ? Ptrofs.add_zero; auto.
+ destruct (e#r); cbn; auto; rewrite ? Int64.add_zero, ? Ptrofs.add_zero; auto.
exists (Val.addl e#r (Vlong n)); split; auto.
Qed.
@@ -388,11 +388,11 @@ Lemma make_shllimm_correct:
Proof.
intros; unfold make_shllimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
+ exists (e#r1); split; auto. destruct (e#r1); cbn; auto.
unfold Int64.shl'. rewrite Z.shiftl_0_r, Int64.repr_unsigned. auto.
destruct (Int.ltu n Int64.iwordsize').
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
+ econstructor; split. cbn. eauto. auto.
+ econstructor; split. cbn. eauto. rewrite H; auto.
Qed.
Lemma make_shrlimm_correct:
@@ -403,11 +403,11 @@ Lemma make_shrlimm_correct:
Proof.
intros; unfold make_shrlimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
+ exists (e#r1); split; auto. destruct (e#r1); cbn; auto.
unfold Int64.shr'. rewrite Z.shiftr_0_r, Int64.repr_signed. auto.
destruct (Int.ltu n Int64.iwordsize').
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
+ econstructor; split. cbn. eauto. auto.
+ econstructor; split. cbn. eauto. rewrite H; auto.
Qed.
Lemma make_shrluimm_correct:
@@ -418,11 +418,11 @@ Lemma make_shrluimm_correct:
Proof.
intros; unfold make_shrluimm.
predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
+ exists (e#r1); split; auto. destruct (e#r1); cbn; auto.
unfold Int64.shru'. rewrite Z.shiftr_0_r, Int64.repr_unsigned. auto.
destruct (Int.ltu n Int64.iwordsize').
- econstructor; split. simpl. eauto. auto.
- econstructor; split. simpl. eauto. rewrite H; auto.
+ econstructor; split. cbn. eauto. auto.
+ econstructor; split. cbn. eauto. rewrite H; auto.
Qed.
Lemma make_mullimm_correct:
@@ -433,15 +433,15 @@ Lemma make_mullimm_correct:
Proof.
intros; unfold make_mullimm.
predSpec Int64.eq Int64.eq_spec n Int64.zero; intros. subst.
- exists (Vlong Int64.zero); split; auto. destruct (e#r1); simpl; auto. rewrite Int64.mul_zero; auto.
+ exists (Vlong Int64.zero); split; auto. destruct (e#r1); cbn; auto. rewrite Int64.mul_zero; auto.
predSpec Int64.eq Int64.eq_spec n Int64.one; intros. subst.
- exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int64.mul_one; auto.
+ exists (e#r1); split; auto. destruct (e#r1); cbn; auto. rewrite Int64.mul_one; auto.
destruct (Int64.is_power2' n) eqn:?; intros.
exists (Val.shll e#r1 (Vint i)); split; auto.
- destruct (e#r1); simpl; auto.
+ destruct (e#r1); cbn; auto.
erewrite Int64.is_power2'_range by eauto.
erewrite Int64.mul_pow2' by eauto. auto.
- econstructor; split; eauto. simpl; rewrite H; auto.
+ econstructor; split; eauto. cbn; rewrite H; auto.
Qed.
Lemma make_divlimm_correct:
@@ -453,7 +453,7 @@ Lemma make_divlimm_correct:
Proof.
intros; unfold make_divlimm.
destruct (Int64.is_power2' n) eqn:?. destruct (Int.ltu i (Int.repr 63)) eqn:?.
- rewrite H0 in H. econstructor; split. simpl; eauto.
+ rewrite H0 in H. econstructor; split. cbn; eauto.
erewrite Val.divls_pow2; eauto. auto.
exists v; auto.
exists v; auto.
@@ -468,9 +468,9 @@ Lemma make_divluimm_correct:
Proof.
intros; unfold make_divluimm.
destruct (Int64.is_power2' n) eqn:?.
- econstructor; split. simpl; eauto.
+ econstructor; split. cbn; eauto.
rewrite H0 in H. destruct (e#r1); inv H. destruct (Int64.eq n Int64.zero); inv H2.
- simpl.
+ cbn.
erewrite Int64.is_power2'_range by eauto.
erewrite Int64.divu_pow2' by eauto. auto.
exists v; auto.
@@ -485,9 +485,9 @@ Lemma make_modluimm_correct:
Proof.
intros; unfold make_modluimm.
destruct (Int64.is_power2 n) eqn:?.
- exists v; split; auto. simpl. decEq.
+ exists v; split; auto. cbn. decEq.
rewrite H0 in H. destruct (e#r1); inv H. destruct (Int64.eq n Int64.zero); inv H2.
- simpl. erewrite Int64.modu_and by eauto. auto.
+ cbn. erewrite Int64.modu_and by eauto. auto.
exists v; auto.
Qed.
@@ -498,9 +498,9 @@ Lemma make_andlimm_correct:
Proof.
intros; unfold make_andlimm.
predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
- subst n. exists (Vlong Int64.zero); split; auto. destruct (e#r); simpl; auto. rewrite Int64.and_zero; auto.
+ subst n. exists (Vlong Int64.zero); split; auto. destruct (e#r); cbn; auto. rewrite Int64.and_zero; auto.
predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.and_mone; auto.
+ subst n. exists (e#r); split; auto. destruct (e#r); cbn; auto. rewrite Int64.and_mone; auto.
econstructor; split; eauto. auto.
Qed.
@@ -511,9 +511,9 @@ Lemma make_orlimm_correct:
Proof.
intros; unfold make_orlimm.
predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.or_zero; auto.
+ subst n. exists (e#r); split; auto. destruct (e#r); cbn; auto. rewrite Int64.or_zero; auto.
predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
- subst n. exists (Vlong Int64.mone); split; auto. destruct (e#r); simpl; auto. rewrite Int64.or_mone; auto.
+ subst n. exists (Vlong Int64.mone); split; auto. destruct (e#r); cbn; auto. rewrite Int64.or_mone; auto.
econstructor; split; eauto. auto.
Qed.
@@ -524,7 +524,7 @@ Lemma make_xorlimm_correct:
Proof.
intros; unfold make_xorlimm.
predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
- subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.xor_zero; auto.
+ subst n. exists (e#r); split; auto. destruct (e#r); cbn; auto. rewrite Int64.xor_zero; auto.
predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
subst n. exists (Val.notl e#r); split; auto.
econstructor; split; eauto. auto.
@@ -538,9 +538,9 @@ Lemma make_mulfimm_correct:
Proof.
intros; unfold make_mulfimm.
destruct (Float.eq_dec n (Float.of_int (Int.repr 2))); intros.
- simpl. econstructor; split. eauto. rewrite H; subst n.
- destruct (e#r1); simpl; auto. rewrite Float.mul2_add; auto.
- simpl. econstructor; split; eauto.
+ cbn. econstructor; split. eauto. rewrite H; subst n.
+ destruct (e#r1); cbn; auto. rewrite Float.mul2_add; auto.
+ cbn. econstructor; split; eauto.
Qed.
Lemma make_mulfimm_correct_2:
@@ -551,10 +551,10 @@ Lemma make_mulfimm_correct_2:
Proof.
intros; unfold make_mulfimm.
destruct (Float.eq_dec n (Float.of_int (Int.repr 2))); intros.
- simpl. econstructor; split. eauto. rewrite H; subst n.
- destruct (e#r2); simpl; auto. rewrite Float.mul2_add; auto.
+ cbn. econstructor; split. eauto. rewrite H; subst n.
+ destruct (e#r2); cbn; auto. rewrite Float.mul2_add; auto.
rewrite Float.mul_commut; auto.
- simpl. econstructor; split; eauto.
+ cbn. econstructor; split; eauto.
Qed.
Lemma make_mulfsimm_correct:
@@ -565,9 +565,9 @@ Lemma make_mulfsimm_correct:
Proof.
intros; unfold make_mulfsimm.
destruct (Float32.eq_dec n (Float32.of_int (Int.repr 2))); intros.
- simpl. econstructor; split. eauto. rewrite H; subst n.
- destruct (e#r1); simpl; auto. rewrite Float32.mul2_add; auto.
- simpl. econstructor; split; eauto.
+ cbn. econstructor; split. eauto. rewrite H; subst n.
+ destruct (e#r1); cbn; auto. rewrite Float32.mul2_add; auto.
+ cbn. econstructor; split; eauto.
Qed.
Lemma make_mulfsimm_correct_2:
@@ -578,10 +578,10 @@ Lemma make_mulfsimm_correct_2:
Proof.
intros; unfold make_mulfsimm.
destruct (Float32.eq_dec n (Float32.of_int (Int.repr 2))); intros.
- simpl. econstructor; split. eauto. rewrite H; subst n.
- destruct (e#r2); simpl; auto. rewrite Float32.mul2_add; auto.
+ cbn. econstructor; split. eauto. rewrite H; subst n.
+ destruct (e#r2); cbn; auto. rewrite Float32.mul2_add; auto.
rewrite Float32.mul_commut; auto.
- simpl. econstructor; split; eauto.
+ cbn. econstructor; split; eauto.
Qed.
Lemma make_cast8signed_correct:
@@ -594,8 +594,8 @@ Proof.
exists e#r; split; auto.
assert (V: vmatch bc e#r (Sgn Ptop 8)).
{ eapply vmatch_ge; eauto. apply vincl_ge; auto. }
- inv V; simpl; auto. rewrite is_sgn_sign_ext in H4 by auto. rewrite H4; auto.
- econstructor; split; simpl; eauto.
+ inv V; cbn; auto. rewrite is_sgn_sign_ext in H4 by auto. rewrite H4; auto.
+ econstructor; split; cbn; eauto.
Qed.
Lemma make_cast16signed_correct:
@@ -608,8 +608,8 @@ Proof.
exists e#r; split; auto.
assert (V: vmatch bc e#r (Sgn Ptop 16)).
{ eapply vmatch_ge; eauto. apply vincl_ge; auto. }
- inv V; simpl; auto. rewrite is_sgn_sign_ext in H4 by auto. rewrite H4; auto.
- econstructor; split; simpl; eauto.
+ inv V; cbn; auto. rewrite is_sgn_sign_ext in H4 by auto. rewrite H4; auto.
+ econstructor; split; cbn; eauto.
Qed.
Lemma op_strength_reduction_correct:
@@ -620,7 +620,7 @@ Lemma op_strength_reduction_correct:
exists w, eval_operation ge (Vptr sp Ptrofs.zero) op' e##args' m = Some w /\ Val.lessdef v w.
Proof.
intros until v; unfold op_strength_reduction;
- case (op_strength_reduction_match op args vl); simpl; intros.
+ case (op_strength_reduction_match op args vl); cbn; intros.
- (* cast8signed *)
InvApproxRegs; SimplVM; inv H0. apply make_cast8signed_correct; auto.
- (* cast16signed *)
@@ -733,15 +733,15 @@ Lemma addr_strength_reduction_correct:
exists res', eval_addressing ge (Vptr sp Ptrofs.zero) addr' e##args' = Some res' /\ Val.lessdef res res'.
Proof.
intros until res. unfold addr_strength_reduction.
- destruct (addr_strength_reduction_match addr args vl); simpl;
+ destruct (addr_strength_reduction_match addr args vl); cbn;
intros VL EA; InvApproxRegs; SimplVM; try (inv EA).
- destruct (orb _ _).
+ exists (Val.offset_ptr e#r1 n); auto.
-+ simpl. rewrite Genv.shift_symbol_address. econstructor; split; eauto.
- inv H0; simpl; auto.
++ cbn. rewrite Genv.shift_symbol_address. econstructor; split; eauto.
+ inv H0; cbn; auto.
- rewrite Ptrofs.add_zero_l. econstructor; split; eauto.
change (Vptr sp (Ptrofs.add n1 n)) with (Val.offset_ptr (Vptr sp n1) n).
- inv H0; simpl; auto.
+ inv H0; cbn; auto.
- exists res; auto.
Qed.
diff --git a/kvx/Conventions1.v b/kvx/Conventions1.v
index ab30ded9..0b2cf406 100644
--- a/kvx/Conventions1.v
+++ b/kvx/Conventions1.v
@@ -108,7 +108,7 @@ Lemma loc_result_type:
subtype (proj_sig_res sig) (typ_rpair mreg_type (loc_result sig)) = true.
Proof.
intros. unfold proj_sig_res, loc_result, mreg_type.
- destruct (sig_res sig); try destruct Archi.ptr64; simpl; trivial; destruct t; trivial.
+ destruct (sig_res sig); try destruct Archi.ptr64; cbn; trivial; destruct t; trivial.
Qed.
(** The result locations are caller-save registers *)
@@ -118,7 +118,7 @@ Lemma loc_result_caller_save:
forall_rpair (fun r => is_callee_save r = false) (loc_result s).
Proof.
intros. unfold loc_result, is_callee_save;
- destruct (sig_res s); simpl; auto; try destruct Archi.ptr64; simpl; auto; try destruct t; simpl; auto.
+ destruct (sig_res s); cbn; auto; try destruct Archi.ptr64; cbn; auto; try destruct t; cbn; auto.
Qed.
(** If the result is in a pair of registers, those registers are distinct and have type [Tint] at least. *)
@@ -296,9 +296,9 @@ Proof.
OKREGS regs -> OKF f -> ofs >= 0 -> OK (one_arg regs rn ofs ty f)).
{ intros until f; intros OR OF OO; red; unfold one_arg; intros.
destruct (list_nth_z regs rn) as [r|] eqn:NTH; destruct H.
- - subst p; simpl. apply OR. eapply list_nth_z_in; eauto.
+ - subst p; cbn. apply OR. eapply list_nth_z_in; eauto.
- eapply OF; eauto.
- - subst p; simpl. auto using align_divides, typealign_pos.
+ - subst p; cbn. auto using align_divides, typealign_pos.
- eapply OF; [idtac|eauto].
generalize (AL ofs ty OO) (SKK ty); omega.
}
@@ -310,16 +310,16 @@ Proof.
assert (OO': ofs' >= 0) by (apply (AL ofs Tlong); auto).
assert (DFL: OK (Twolong (S Outgoing (ofs' + 1) Tint) (S Outgoing ofs' Tint)
:: f rn' (ofs' + 2))).
- { red; simpl; intros. destruct H.
- - subst p; simpl.
+ { red; cbn; intros. destruct H.
+ - subst p; cbn.
repeat split; auto using Z.divide_1_l. omega.
- eapply OF; [idtac|eauto]. omega.
}
destruct (list_nth_z regs rn') as [r1|] eqn:NTH1;
destruct (list_nth_z regs (rn' + 1)) as [r2|] eqn:NTH2;
try apply DFL.
- red; simpl; intros; destruct H.
- - subst p; simpl. split; apply OR; eauto using list_nth_z_in.
+ red; cbn; intros; destruct H.
+ - subst p; cbn. split; apply OR; eauto using list_nth_z_in.
- eapply OF; [idtac|eauto]. auto.
}
assert (C: forall regs rn ofs ty f,
@@ -327,10 +327,10 @@ Proof.
{ intros until f; intros OR OF OO OTY; unfold hybrid_arg; red; intros.
set (rn' := align rn 2) in *.
destruct (list_nth_z regs rn') as [r|] eqn:NTH; destruct H.
- - subst p; simpl. apply OR. eapply list_nth_z_in; eauto.
+ - subst p; cbn. apply OR. eapply list_nth_z_in; eauto.
- eapply OF; eauto.
- - subst p; simpl. rewrite OTY. split. apply (AL ofs Tlong OO). apply Z.divide_1_l.
- - eapply OF; [idtac|eauto]. generalize (AL ofs Tlong OO); simpl; omega.
+ - subst p; cbn. rewrite OTY. split. apply (AL ofs Tlong OO). apply Z.divide_1_l.
+ - eapply OF; [idtac|eauto]. generalize (AL ofs Tlong OO); cbn; omega.
}
assert (D: OKREGS param_regs).
{ red. decide_goal. }
@@ -339,8 +339,8 @@ Proof.
cut (forall va tyl rn ofs, ofs >= 0 -> OK (loc_arguments_rec va tyl rn ofs)).
unfold OK. eauto.
- induction tyl as [ | ty1 tyl]; intros until ofs; intros OO; simpl.
- - red; simpl; tauto.
+ induction tyl as [ | ty1 tyl]; intros until ofs; intros OO; cbn.
+ - red; cbn; tauto.
- destruct ty1.
+ (* int *) apply A; auto.
+ (* float *)
@@ -369,10 +369,10 @@ Remark fold_max_outgoing_above:
Proof.
assert (A: forall n l, max_outgoing_1 n l >= n).
{ intros; unfold max_outgoing_1. destruct l as [_ | []]; xomega. }
- induction l; simpl; intros.
+ induction l; cbn; intros.
- omega.
- eapply Zge_trans. eauto.
- destruct a; simpl. apply A. eapply Zge_trans; eauto.
+ destruct a; cbn. apply A. eapply Zge_trans; eauto.
Qed.
Lemma size_arguments_above:
@@ -392,14 +392,14 @@ Proof.
assert (B: forall p n,
In (S Outgoing ofs ty) (regs_of_rpair p) ->
ofs + typesize ty <= max_outgoing_2 n p).
- { intros. destruct p; simpl in H; intuition; subst; simpl.
+ { intros. destruct p; cbn in H; intuition; subst; cbn.
- xomega.
- eapply Z.le_trans. 2: apply A. xomega.
- xomega. }
assert (C: forall l n,
In (S Outgoing ofs ty) (regs_of_rpairs l) ->
ofs + typesize ty <= fold_left max_outgoing_2 l n).
- { induction l; simpl; intros.
+ { induction l; cbn; intros.
- contradiction.
- rewrite in_app_iff in H. destruct H.
+ eapply Z.le_trans. eapply B; eauto. apply Z.ge_le. apply fold_max_outgoing_above.
diff --git a/kvx/ExtValues.v b/kvx/ExtValues.v
index 3664c00a..a0c10ddd 100644
--- a/kvx/ExtValues.v
+++ b/kvx/ExtValues.v
@@ -62,10 +62,10 @@ Lemma shift1_4_of_z_correct :
end.
Proof.
intro. unfold shift1_4_of_z.
- destruct (Z.eq_dec _ _); simpl; try congruence.
- destruct (Z.eq_dec _ _); simpl; try congruence.
- destruct (Z.eq_dec _ _); simpl; try congruence.
- destruct (Z.eq_dec _ _); simpl; try congruence.
+ destruct (Z.eq_dec _ _); cbn; try congruence.
+ destruct (Z.eq_dec _ _); cbn; try congruence.
+ destruct (Z.eq_dec _ _); cbn; try congruence.
+ destruct (Z.eq_dec _ _); cbn; try congruence.
trivial.
Qed.
@@ -215,19 +215,19 @@ Theorem divu_is_divlu: forall v1 v2 : val,
end.
Proof.
intros.
- destruct v1; simpl; trivial.
- destruct v2; simpl; trivial.
+ destruct v1; cbn; trivial.
+ destruct v2; cbn; trivial.
destruct i as [i_val i_range].
destruct i0 as [i0_val i0_range].
- simpl.
+ cbn.
unfold Int.eq, Int64.eq, Int.zero, Int64.zero.
- simpl.
+ cbn.
rewrite Int.unsigned_repr by (compute; split; discriminate).
rewrite (Int64.unsigned_repr 0) by (compute; split; discriminate).
rewrite (unsigned64_repr i0_val) by assumption.
- destruct (zeq i0_val 0) as [ | Hnot0]; simpl; trivial.
+ destruct (zeq i0_val 0) as [ | Hnot0]; cbn; trivial.
f_equal. f_equal.
- unfold Int.divu, Int64.divu. simpl.
+ unfold Int.divu, Int64.divu. cbn.
rewrite (unsigned64_repr i_val) by assumption.
rewrite (unsigned64_repr i0_val) by assumption.
unfold Int64.loword.
@@ -260,19 +260,19 @@ Theorem modu_is_modlu: forall v1 v2 : val,
end.
Proof.
intros.
- destruct v1; simpl; trivial.
- destruct v2; simpl; trivial.
+ destruct v1; cbn; trivial.
+ destruct v2; cbn; trivial.
destruct i as [i_val i_range].
destruct i0 as [i0_val i0_range].
- simpl.
+ cbn.
unfold Int.eq, Int64.eq, Int.zero, Int64.zero.
- simpl.
+ cbn.
rewrite Int.unsigned_repr by (compute; split; discriminate).
rewrite (Int64.unsigned_repr 0) by (compute; split; discriminate).
rewrite (unsigned64_repr i0_val) by assumption.
- destruct (zeq i0_val 0) as [ | Hnot0]; simpl; trivial.
+ destruct (zeq i0_val 0) as [ | Hnot0]; cbn; trivial.
f_equal. f_equal.
- unfold Int.modu, Int64.modu. simpl.
+ unfold Int.modu, Int64.modu. cbn.
rewrite (unsigned64_repr i_val) by assumption.
rewrite (unsigned64_repr i0_val) by assumption.
unfold Int64.loword.
@@ -347,19 +347,19 @@ Theorem divs_is_divls: forall v1 v2 : val,
end.
Proof.
intros.
- destruct v1; simpl; trivial.
- destruct v2; simpl; trivial.
+ destruct v1; cbn; trivial.
+ destruct v2; cbn; trivial.
destruct i as [i_val i_range].
destruct i0 as [i0_val i0_range].
- simpl.
+ cbn.
unfold Int.eq, Int64.eq, Int.zero, Int64.zero.
- simpl.
+ cbn.
replace (Int.unsigned (Int.repr 0)) with 0 in * by reflexivity.
- destruct (zeq _ _) as [H0' | Hnot0]; simpl; trivial.
- destruct (zeq i_val (Int.unsigned (Int.repr Int.min_signed))) as [Hmin | Hnotmin]; simpl.
+ destruct (zeq _ _) as [H0' | Hnot0]; cbn; trivial.
+ destruct (zeq i_val (Int.unsigned (Int.repr Int.min_signed))) as [Hmin | Hnotmin]; cbn.
{ subst.
destruct (zeq i0_val (Int.unsigned Int.mone)) as [Hmone | Hnotmone]; trivial.
- unfold Int.signed. simpl.
+ unfold Int.signed. cbn.
replace (Int64.unsigned (Int64.repr 0)) with 0 in * by reflexivity.
rewrite if_zlt_min_signed_half_modulus.
replace (if
@@ -370,7 +370,7 @@ Proof.
(Int64.unsigned (Int64.repr Int64.min_signed))
then true
else false) with false by reflexivity.
- simpl.
+ cbn.
rewrite orb_false_r.
destruct (zlt i0_val Int.half_modulus) as [Hlt_half | Hge_half].
{
@@ -380,7 +380,7 @@ Proof.
unfold Val.loword.
f_equal.
unfold Int64.divs, Int.divs, Int64.loword.
- unfold Int.signed, Int64.signed. simpl.
+ unfold Int.signed, Int64.signed. cbn.
rewrite if_zlt_min_signed_half_modulus.
change Int.half_modulus with 2147483648 in *.
destruct (zlt _ _) as [discard|]; try omega. clear discard.
@@ -390,7 +390,7 @@ Proof.
with 18446744071562067968.
change Int64.half_modulus with 9223372036854775808.
change Int64.modulus with 18446744073709551616.
- simpl.
+ cbn.
rewrite (Int64.unsigned_repr i0_val) by (change Int64.max_unsigned with 18446744073709551615; omega).
destruct (zlt i0_val 9223372036854775808) as [discard |]; try omega.
clear discard.
@@ -449,7 +449,7 @@ Lemma big_unsigned_signed:
Proof.
destruct x as [xval xrange].
intro BIG.
- unfold Int.signed, Int.unsigned in *. simpl in *.
+ unfold Int.signed, Int.unsigned in *. cbn in *.
destruct (zlt _ _).
omega.
trivial.
@@ -499,10 +499,10 @@ Lemma divs_is_quot: forall v1 v2 : val,
end.
Proof.
- destruct v1; destruct v2; simpl; trivial.
+ destruct v1; destruct v2; cbn; trivial.
unfold Int.divs.
rewrite signed_0_eqb.
- destruct (Int.eq i0 Int.zero) eqn:Eeq0; simpl; trivial.
+ destruct (Int.eq i0 Int.zero) eqn:Eeq0; cbn; trivial.
destruct (Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone) eqn:EXCEPTION.
{ replace (Int.signed i0) with (-1).
replace (Int.signed i) with Int.min_signed.
@@ -523,7 +523,7 @@ Proof.
unfold Int.eq in EXCEPTION.
destruct (zeq _ _) in EXCEPTION; try discriminate.
destruct (zeq _ _) as [Hmone | ] in EXCEPTION; try discriminate.
- destruct i0 as [i0val i0range]; unfold Int.signed in *; simpl in *.
+ destruct i0 as [i0val i0range]; unfold Int.signed in *; cbn in *.
rewrite Hmone.
reflexivity.
}
@@ -651,7 +651,7 @@ Qed.
Lemma sub_add_neg :
forall x y, Val.sub x y = Val.add x (Val.neg y).
Proof.
- destruct x; destruct y; simpl; trivial.
+ destruct x; destruct y; cbn; trivial.
f_equal.
apply Int.sub_add_opp.
Qed.
@@ -659,7 +659,7 @@ Qed.
Lemma neg_mul_distr_r :
forall x y, Val.neg (Val.mul x y) = Val.mul x (Val.neg y).
Proof.
- destruct x; destruct y; simpl; trivial.
+ destruct x; destruct y; cbn; trivial.
f_equal.
apply Int.neg_mul_distr_r.
Qed.
@@ -668,7 +668,7 @@ Qed.
Lemma sub_addl_negl :
forall x y, Val.subl x y = Val.addl x (Val.negl y).
Proof.
- destruct x; destruct y; simpl; trivial.
+ destruct x; destruct y; cbn; trivial.
+ f_equal. apply Int64.sub_add_opp.
+ destruct (Archi.ptr64) eqn:ARCHI64; trivial.
f_equal. rewrite Ptrofs.sub_add_opp.
@@ -681,15 +681,15 @@ Proof.
rewrite Hagree2.
reflexivity.
exact (Ptrofs.agree64_of_int ARCHI64 i0).
- + destruct (Archi.ptr64) eqn:ARCHI64; simpl; trivial.
- destruct (eq_block _ _); simpl; trivial.
+ + destruct (Archi.ptr64) eqn:ARCHI64; cbn; trivial.
+ destruct (eq_block _ _); cbn; trivial.
Qed.
*)
Lemma negl_mull_distr_r :
forall x y, Val.negl (Val.mull x y) = Val.mull x (Val.negl y).
Proof.
- destruct x; destruct y; simpl; trivial.
+ destruct x; destruct y; cbn; trivial.
f_equal.
apply Int64.neg_mul_distr_r.
Qed.
diff --git a/kvx/InstructionScheduler.ml b/kvx/InstructionScheduler.ml
deleted file mode 100644
index 32f394b1..00000000
--- a/kvx/InstructionScheduler.ml
+++ /dev/null
@@ -1,1254 +0,0 @@
-(* *************************************************************)
-(* *)
-(* 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. *)
-(* *)
-(* *************************************************************)
-
-(** 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;
- 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_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;;
-
-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;
- 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"
- (* needed only if not using stdout and filename_out = "problem.sol" *) in
- let opb_problem = open_out filename_in in
- let mapper = pseudo_boolean_print_problem opb_problem problem pb_type in
- close_out opb_problem;
-
- let opb_solution = Unix.open_process_in (!pseudo_boolean_solver ^ " " ^ filename_in) in
- let ret = adjust_check_solution mapper (pseudo_boolean_read_solution mapper opb_solution) in
- close_in opb_solution;
- Some ret
- 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 opb_problem = open_out filename_in in
- let mapper = ilp_print_problem opb_problem problem pb_type in
- close_out opb_problem;
-
- begin
- match Unix.system (!ilp_solver ^ " " ^ filename_in ^ " " ^ filename_out) with
- | Unix.WEXITED 0 ->
- let opb_solution = open_in filename_out in
- let ret = adjust_check_solution mapper (ilp_read_solution mapper opb_solution) in
- close_in opb_solution;
- Some ret
- | 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
- | "greedy" -> greedy_scheduler
- | s -> failwith ("unknown scheduler: " ^ s);;
diff --git a/kvx/InstructionScheduler.mli b/kvx/InstructionScheduler.mli
deleted file mode 100644
index 6d27b30c..00000000
--- a/kvx/InstructionScheduler.mli
+++ /dev/null
@@ -1,113 +0,0 @@
-(** 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. *)
-
- 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;;
-
-(** 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
-
-(** 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. *)
-val schedule_reversed : scheduler -> problem -> int array option
-
-(** Schedule a problem from the end using a list scheduler. *)
-val reverse_list_scheduler : problem -> int array option
-
-(** Check that a problem is well-formed.
-@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/kvx/NeedOp.v b/kvx/NeedOp.v
index 4c354d5a..f636336d 100644
--- a/kvx/NeedOp.v
+++ b/kvx/NeedOp.v
@@ -229,7 +229,7 @@ Lemma needs_of_condition0_sound:
Proof.
intros until arg2.
intros Hcond Hagree.
- apply eval_condition0_inj with (f := inject_id) (m1 := m1) (v1 := arg1); simpl; auto.
+ apply eval_condition0_inj with (f := inject_id) (m1 := m1) (v1 := arg1); cbn; auto.
apply val_inject_lessdef. apply lessdef_vagree. assumption.
Qed.
@@ -239,7 +239,7 @@ Lemma addl_sound:
vagree (Val.addl v1 v2) (Val.addl w1 w2) x.
Proof.
unfold default; intros.
- destruct x; simpl in *; trivial.
+ destruct x; cbn in *; trivial.
- unfold Val.addl.
destruct v1; destruct v2; trivial; destruct Archi.ptr64; trivial.
- apply Val.addl_lessdef; trivial.
@@ -249,7 +249,7 @@ Lemma subl_lessdef:
forall v1 v1' v2 v2',
Val.lessdef v1 v1' -> Val.lessdef v2 v2' -> Val.lessdef (Val.subl v1 v2) (Val.subl v1' v2').
Proof.
- intros. inv H. inv H0. auto. destruct v1'; simpl; auto. simpl; auto.
+ intros. inv H. inv H0. auto. destruct v1'; cbn; auto. cbn; auto.
Qed.
Lemma subl_sound:
@@ -258,10 +258,10 @@ Lemma subl_sound:
vagree (Val.subl v1 v2) (Val.subl w1 w2) x.
Proof.
unfold default; intros.
- destruct x; simpl in *; trivial.
+ destruct x; cbn in *; trivial.
- unfold Val.subl.
- destruct v1; destruct v2; trivial; destruct Archi.ptr64; simpl; trivial.
- destruct (eq_block _ _) ; simpl; trivial.
+ destruct v1; destruct v2; trivial; destruct Archi.ptr64; cbn; trivial.
+ destruct (eq_block _ _) ; cbn; trivial.
- apply subl_lessdef; trivial.
Qed.
@@ -272,7 +272,7 @@ Lemma mull_sound:
vagree (Val.mull v1 v2) (Val.mull w1 w2) x.
Proof.
unfold default; intros.
- destruct x; simpl in *; trivial.
+ destruct x; cbn in *; trivial.
- unfold Val.mull.
destruct v1; destruct v2; trivial.
- unfold Val.mull.
@@ -284,7 +284,7 @@ Qed.
Remark default_idem: forall nv, default (default nv) = default nv.
Proof.
- destruct nv; simpl; trivial.
+ destruct nv; cbn; trivial.
Qed.
Lemma vagree_triple_op_float :
@@ -298,14 +298,14 @@ Proof.
induction nv;
intros Hax Hby Hcz.
- trivial.
- - simpl in *. destruct a; simpl; trivial.
- destruct b; simpl; trivial.
- destruct c; simpl; trivial.
- - simpl in *. destruct a; simpl; trivial.
- destruct b; simpl; trivial.
- destruct c; simpl; trivial.
+ - cbn in *. destruct a; cbn; trivial.
+ destruct b; cbn; trivial.
+ destruct c; cbn; trivial.
+ - cbn in *. destruct a; cbn; trivial.
+ destruct b; cbn; trivial.
+ destruct c; cbn; trivial.
inv Hax. inv Hby. inv Hcz.
- simpl.
+ cbn.
constructor.
Qed.
@@ -320,14 +320,14 @@ Proof.
induction nv;
intros Hax Hby Hcz.
- trivial.
- - simpl in *. destruct a; simpl; trivial.
- destruct b; simpl; trivial.
- destruct c; simpl; trivial.
- - simpl in *. destruct a; simpl; trivial.
- destruct b; simpl; trivial.
- destruct c; simpl; trivial.
+ - cbn in *. destruct a; cbn; trivial.
+ destruct b; cbn; trivial.
+ destruct c; cbn; trivial.
+ - cbn in *. destruct a; cbn; trivial.
+ destruct b; cbn; trivial.
+ destruct c; cbn; trivial.
inv Hax. inv Hby. inv Hcz.
- simpl.
+ cbn.
constructor.
Qed.
@@ -343,7 +343,7 @@ Lemma needs_of_operation_sound:
/\ vagree v v' nv.
Proof.
unfold needs_of_operation; intros; destruct op; try (eapply default_needs_of_operation_sound; eauto; fail);
- simpl in *; FuncInv; InvAgree; TrivialExists.
+ cbn in *; FuncInv; InvAgree; TrivialExists.
- apply sign_ext_sound; auto. compute; auto.
- apply sign_ext_sound; auto. compute; auto.
- apply add_sound; auto.
@@ -384,17 +384,17 @@ Proof.
- destruct (eval_condition0 _ _ _) as [b|] eqn:EC.
erewrite needs_of_condition0_sound by eauto.
apply select_sound; auto.
- simpl; auto with na.
+ cbn; auto with na.
(* select imm *)
- destruct (eval_condition0 _ _ _) as [b|] eqn:EC.
{ erewrite needs_of_condition0_sound by eauto.
apply select_sound; auto with na. }
- simpl; auto with na.
+ cbn; auto with na.
(* select long imm *)
- destruct (eval_condition0 _ _ _) as [b|] eqn:EC.
{ erewrite needs_of_condition0_sound by eauto.
apply select_sound; auto with na. }
- simpl; auto with na.
+ cbn; auto with na.
Qed.
Lemma operation_is_redundant_sound:
@@ -404,7 +404,7 @@ Lemma operation_is_redundant_sound:
vagree_list (arg1 :: args) (arg1' :: args') (needs_of_operation op nv) ->
vagree v arg1' nv.
Proof.
- intros. destruct op; simpl in *; try discriminate; inv H1; FuncInv; subst.
+ intros. destruct op; cbn in *; try discriminate; inv H1; FuncInv; subst.
- apply sign_ext_redundant_sound; auto. omega.
- apply sign_ext_redundant_sound; auto. omega.
- apply andimm_redundant_sound; auto.
diff --git a/kvx/Op.v b/kvx/Op.v
index b78b7b97..cda5ef78 100644
--- a/kvx/Op.v
+++ b/kvx/Op.v
@@ -508,9 +508,9 @@ Qed.
Ltac FuncInv :=
match goal with
| H: (match ?x with nil => _ | _ :: _ => _ end = Some _) |- _ =>
- destruct x; simpl in H; FuncInv
+ destruct x; cbn in H; FuncInv
| H: (match ?v with Vundef => _ | Vint _ => _ | Vfloat _ => _ | Vptr _ _ => _ end = Some _) |- _ =>
- destruct v; simpl in H; FuncInv
+ destruct v; cbn in H; FuncInv
| H: (if Archi.ptr64 then _ else _) = Some _ |- _ =>
destruct Archi.ptr64 eqn:?; FuncInv
| H: (Some _ = Some _) |- _ =>
@@ -727,27 +727,27 @@ Qed.
Remark type_sub:
forall v1 v2, Val.has_type (Val.sub v1 v2) Tint.
Proof.
- intros. unfold Val.has_type, Val.sub. destruct Archi.ptr64, v1, v2; simpl; auto.
+ intros. unfold Val.has_type, Val.sub. destruct Archi.ptr64, v1, v2; cbn; auto.
destruct (eq_block _ _); auto.
Qed.
Remark type_subl:
forall v1 v2, Val.has_type (Val.subl v1 v2) Tlong.
Proof.
- intros. unfold Val.has_type, Val.subl. destruct Archi.ptr64, v1, v2; simpl; auto.
+ intros. unfold Val.has_type, Val.subl. destruct Archi.ptr64, v1, v2; cbn; auto.
destruct (eq_block _ _); auto.
Qed.
Remark type_shl:
forall v1 v2, Val.has_type (Val.shl v1 v2) Tint.
Proof.
- destruct v1, v2; simpl; trivial; destruct (Int.ltu _ _); simpl; trivial.
+ destruct v1, v2; cbn; trivial; destruct (Int.ltu _ _); cbn; trivial.
Qed.
Remark type_shll:
forall v1 v2, Val.has_type (Val.shll v1 v2) Tlong.
Proof.
- destruct v1, v2; simpl; trivial; destruct (Int.ltu _ _); simpl; trivial.
+ destruct v1, v2; cbn; trivial; destruct (Int.ltu _ _); cbn; trivial.
Qed.
Lemma type_of_operation_sound:
@@ -757,7 +757,7 @@ Lemma type_of_operation_sound:
Val.has_type v (snd (type_of_operation op)).
Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
intros.
- destruct op; simpl; simpl in H0; FuncInv; subst; simpl.
+ destruct op; cbn; cbn in H0; FuncInv; subst; cbn.
(* move *)
- congruence.
(* intconst, longconst, floatconst, singleconst *)
@@ -777,30 +777,30 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- apply type_add.
(* addx, addximm *)
- apply type_add.
- - destruct v0; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
+ - destruct v0; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
(* neg, sub *)
- destruct v0...
- apply type_sub.
(* revsubimm, revsubx, revsubximm *)
- destruct v0...
- apply type_sub.
- - destruct v0; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
+ - destruct v0; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
(* mul, mulimm, mulhs, mulhu *)
- destruct v0; destruct v1...
- destruct v0...
- destruct v0; destruct v1...
- destruct v0; destruct v1...
(* div, divu *)
- - destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int.eq i0 Int.zero || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2...
- - destruct v0; destruct v1; simpl in *; inv H0.
+ - destruct v0; destruct v1; cbn in *; inv H0.
+ destruct (_ || _); inv H2...
+ - destruct v0; destruct v1; cbn in *; inv H0.
destruct (Int.eq i0 Int.zero); inv H2...
(* mod, modu *)
- - destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int.eq i0 Int.zero || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2...
- - destruct v0; destruct v1; simpl in *; inv H0.
+ - destruct v0; destruct v1; cbn in *; inv H0.
+ destruct (_ || _); inv H2...
+ - destruct v0; destruct v1; cbn in *; inv H0.
destruct (Int.eq i0 Int.zero); inv H2...
(* and, andimm *)
- destruct v0; destruct v1...
@@ -829,18 +829,18 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0; destruct v1...
- destruct v0...
(* shl, shlimm *)
- - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- - destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
+ - destruct v0; destruct v1; cbn... destruct (Int.ltu i0 Int.iwordsize)...
+ - destruct v0; cbn... destruct (Int.ltu n Int.iwordsize)...
(* shr, shrimm *)
- - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- - destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
+ - destruct v0; destruct v1; cbn... destruct (Int.ltu i0 Int.iwordsize)...
+ - destruct v0; cbn... destruct (Int.ltu n Int.iwordsize)...
(* shru, shruimm *)
- - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
- - destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
+ - destruct v0; destruct v1; cbn... destruct (Int.ltu i0 Int.iwordsize)...
+ - destruct v0; cbn... destruct (Int.ltu n Int.iwordsize)...
(* shrx *)
- - destruct v0; simpl... destruct (Int.ltu n (Int.repr 31)); simpl; trivial.
+ - destruct v0; cbn... destruct (Int.ltu n (Int.repr 31)); cbn; trivial.
(* shrimm *)
- - destruct v0; simpl...
+ - destruct v0; cbn...
(* madd *)
- apply type_add.
- apply type_add.
@@ -858,13 +858,13 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- apply type_addl.
(* addxl addxlimm *)
- apply type_addl.
- - destruct v0; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
+ - destruct v0; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
(* negl, subl *)
- destruct v0...
- apply type_subl.
- - destruct v0; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
+ - destruct v0; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
- destruct v0...
- apply type_subl.
(* mull, mullhs, mullhu *)
@@ -873,14 +873,14 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0; destruct v1...
- destruct v0; destruct v1...
(* divl, divlu *)
- - destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2...
- - destruct v0; destruct v1; simpl in *; inv H0.
+ - destruct v0; destruct v1; cbn in *; inv H0.
+ destruct (_ || _); inv H2...
+ - destruct v0; destruct v1; cbn in *; inv H0.
destruct (Int64.eq i0 Int64.zero); inv H2...
(* modl, modlu *)
- - destruct v0; destruct v1; simpl in *; inv H0.
- destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2...
- - destruct v0; destruct v1; simpl in *; inv H0.
+ - destruct v0; destruct v1; cbn in *; inv H0.
+ destruct (_ || _); inv H2...
+ - destruct v0; destruct v1; cbn in *; inv H0.
destruct (Int64.eq i0 Int64.zero); inv H2...
(* andl, andlimm *)
- destruct v0; destruct v1...
@@ -909,16 +909,16 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0; destruct v1...
- destruct v0...
(* shll, shllimm *)
- - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
- - destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
+ - destruct v0; destruct v1; cbn... destruct (Int.ltu i0 Int64.iwordsize')...
+ - destruct v0; cbn... destruct (Int.ltu n Int64.iwordsize')...
(* shr, shrimm *)
- - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
- - destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
+ - destruct v0; destruct v1; cbn... destruct (Int.ltu i0 Int64.iwordsize')...
+ - destruct v0; cbn... destruct (Int.ltu n Int64.iwordsize')...
(* shru, shruimm *)
- - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
- - destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
+ - destruct v0; destruct v1; cbn... destruct (Int.ltu i0 Int64.iwordsize')...
+ - destruct v0; cbn... destruct (Int.ltu n Int64.iwordsize')...
(* shrxl *)
- - destruct v0; simpl... destruct (Int.ltu n (Int.repr 63)); simpl; trivial.
+ - destruct v0; cbn... destruct (Int.ltu n (Int.repr 63)); cbn; trivial.
(* maddl, maddlim *)
- apply type_addl.
- apply type_addl.
@@ -960,59 +960,59 @@ Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
- destruct v0...
- destruct v0...
(* intoffloat, intuoffloat *)
- - destruct v0; simpl... destruct (Float.to_int f); simpl; trivial.
- - destruct v0; simpl... destruct (Float.to_intu f); simpl; trivial.
+ - destruct v0; cbn... destruct (Float.to_int f); cbn; trivial.
+ - destruct v0; cbn... destruct (Float.to_intu f); cbn; trivial.
(* intofsingle, intuofsingle *)
- - destruct v0; simpl... destruct (Float32.to_int f); simpl; trivial.
- - destruct v0; simpl... destruct (Float32.to_intu f); simpl; trivial.
+ - destruct v0; cbn... destruct (Float32.to_int f); cbn; trivial.
+ - destruct v0; cbn... destruct (Float32.to_intu f); cbn; trivial.
(* singleofint, singleofintu *)
- - destruct v0; simpl...
- - destruct v0; simpl...
+ - destruct v0; cbn...
+ - destruct v0; cbn...
(* longoffloat, longuoffloat *)
- - destruct v0; simpl... destruct (Float.to_long f); simpl; trivial.
- - destruct v0; simpl... destruct (Float.to_longu f); simpl; trivial.
+ - destruct v0; cbn... destruct (Float.to_long f); cbn; trivial.
+ - destruct v0; cbn... destruct (Float.to_longu f); cbn; trivial.
(* floatoflong, floatoflongu *)
- - destruct v0; simpl...
- - destruct v0; simpl...
+ - destruct v0; cbn...
+ - destruct v0; cbn...
(* longofsingle, longuofsingle *)
- - destruct v0; simpl... destruct (Float32.to_long f); simpl; trivial.
- - destruct v0; simpl... destruct (Float32.to_longu f); simpl; trivial.
+ - destruct v0; cbn... destruct (Float32.to_long f); cbn; trivial.
+ - destruct v0; cbn... destruct (Float32.to_longu f); cbn; trivial.
(* singleoflong, singleoflongu *)
- - destruct v0; simpl...
- - destruct v0; simpl...
+ - destruct v0; cbn...
+ - destruct v0; cbn...
(* cmp *)
- destruct (eval_condition cond vl m)... destruct b...
(* extfz *)
- unfold extfz.
destruct (is_bitfield _ _).
- + destruct v0; simpl; trivial.
+ + destruct v0; cbn; trivial.
+ constructor.
(* extfs *)
- unfold extfs.
destruct (is_bitfield _ _).
- + destruct v0; simpl; trivial.
+ + destruct v0; cbn; trivial.
+ constructor.
(* extfzl *)
- unfold extfzl.
destruct (is_bitfieldl _ _).
- + destruct v0; simpl; trivial.
+ + destruct v0; cbn; trivial.
+ constructor.
(* extfsl *)
- unfold extfsl.
destruct (is_bitfieldl _ _).
- + destruct v0; simpl; trivial.
+ + destruct v0; cbn; trivial.
+ constructor.
(* insf *)
- unfold insf, bitfield_mask.
destruct (is_bitfield _ _).
- + destruct v0; destruct v1; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
+ + destruct v0; destruct v1; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
+ constructor.
(* insf *)
- unfold insfl, bitfield_mask.
destruct (is_bitfieldl _ _).
- + destruct v0; destruct v1; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
+ + destruct v0; destruct v1; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
+ constructor.
(* Osel *)
- unfold Val.select. destruct (eval_condition0 _ _ m).
@@ -1047,7 +1047,7 @@ Lemma is_trapping_op_sound:
eval_operation genv sp op vl m <> None.
Proof.
unfold args_of_operation.
- destruct op; destruct eq_operation; intros; simpl in *; try congruence.
+ destruct op; destruct eq_operation; intros; cbn in *; try congruence.
all: try (destruct vl as [ | vh1 vl1]; try discriminate).
all: try (destruct vl1 as [ | vh2 vl2]; try discriminate).
all: try (destruct vl2 as [ | vh3 vl3]; try discriminate).
@@ -1101,7 +1101,7 @@ Lemma eval_negate_condition:
forall cond vl m,
eval_condition (negate_condition cond) vl m = option_map negb (eval_condition cond vl m).
Proof.
- intros. destruct cond; simpl.
+ intros. destruct cond; cbn.
repeat (destruct vl; auto). apply Val.negate_cmp_bool.
repeat (destruct vl; auto). apply Val.negate_cmpu_bool.
repeat (destruct vl; auto). apply Val.negate_cmp_bool.
@@ -1147,7 +1147,7 @@ Lemma eval_shift_stack_addressing:
eval_addressing ge (Vptr sp Ptrofs.zero) (shift_stack_addressing delta addr) vl =
eval_addressing ge (Vptr sp (Ptrofs.repr delta)) addr vl.
Proof.
- intros. destruct addr; simpl; auto. destruct vl; auto.
+ intros. destruct addr; cbn; auto. destruct vl; auto.
rewrite Ptrofs.add_zero_l, Ptrofs.add_commut; auto.
Qed.
@@ -1156,7 +1156,7 @@ Lemma eval_shift_stack_operation:
eval_operation ge (Vptr sp Ptrofs.zero) (shift_stack_operation delta op) vl m =
eval_operation ge (Vptr sp (Ptrofs.repr delta)) op vl m.
Proof.
- intros. destruct op; simpl; auto. destruct vl; auto.
+ intros. destruct op; cbn; auto. destruct vl; auto.
rewrite Ptrofs.add_zero_l, Ptrofs.add_commut; auto.
Qed.
@@ -1183,12 +1183,12 @@ Proof.
assert (A: forall x n,
Val.offset_ptr x (Ptrofs.add n (Ptrofs.repr delta)) =
Val.add (Val.offset_ptr x n) (Vint (Int.repr delta))).
- { intros; destruct x; simpl; auto. rewrite H1.
+ { intros; destruct x; cbn; auto. rewrite H1.
rewrite Ptrofs.add_assoc. f_equal; f_equal; f_equal. symmetry; auto with ptrofs. }
- destruct addr; simpl in H; inv H; simpl in *; FuncInv; subst.
+ destruct addr; cbn in H; inv H; cbn in *; FuncInv; subst.
- rewrite A; auto.
- unfold Genv.symbol_address. destruct (Genv.find_symbol ge i); auto.
- simpl. rewrite H1. f_equal; f_equal; f_equal. symmetry; auto with ptrofs.
+ cbn. rewrite H1. f_equal; f_equal; f_equal. symmetry; auto with ptrofs.
- rewrite A; auto.
Qed.
@@ -1223,17 +1223,17 @@ Lemma op_depends_on_memory_correct:
op_depends_on_memory op = false ->
eval_operation ge sp op args m1 = eval_operation ge sp op args m2.
Proof.
- intros until m2. destruct op; simpl; try congruence.
- - destruct cond; simpl; try congruence;
+ intros until m2. destruct op; cbn; try congruence.
+ - destruct cond; cbn; try congruence;
intros SF; auto; rewrite ? negb_false_iff in SF;
unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
- - destruct c0; simpl; try congruence;
+ - destruct c0; cbn; try congruence;
intros SF; auto; rewrite ? negb_false_iff in SF;
unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
- - destruct c0; simpl; try congruence;
+ - destruct c0; cbn; try congruence;
intros SF; auto; rewrite ? negb_false_iff in SF;
unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
- - destruct c0; simpl; try congruence;
+ - destruct c0; cbn; try congruence;
intros SF; auto; rewrite ? negb_false_iff in SF;
unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
Qed.
@@ -1387,19 +1387,19 @@ Lemma eval_condition_inj:
eval_condition cond vl1 m1 = Some b ->
eval_condition cond vl2 m2 = Some b.
Proof.
- intros. destruct cond; simpl in H0; FuncInv; InvInject; simpl; auto.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
+ intros. destruct cond; cbn in H0; FuncInv; InvInject; cbn; auto.
+- inv H3; inv H2; cbn in H0; inv H0; auto.
- eauto 3 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
-- inv H3; simpl in H0; inv H0; auto.
+- inv H3; cbn in H0; inv H0; auto.
- eauto 3 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
+- inv H3; inv H2; cbn in H0; inv H0; auto.
- eauto 3 using Val.cmplu_bool_inject, Mem.valid_pointer_implies.
-- inv H3; simpl in H0; inv H0; auto.
+- inv H3; cbn in H0; inv H0; auto.
- eauto 3 using Val.cmplu_bool_inject, Mem.valid_pointer_implies.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
-- inv H3; inv H2; simpl in H0; inv H0; auto.
+- inv H3; inv H2; cbn in H0; inv H0; auto.
+- inv H3; inv H2; cbn in H0; inv H0; auto.
+- inv H3; inv H2; cbn in H0; inv H0; auto.
+- inv H3; inv H2; cbn in H0; inv H0; auto.
Qed.
Lemma eval_condition0_inj:
@@ -1408,10 +1408,10 @@ Lemma eval_condition0_inj:
eval_condition0 cond v1 m1 = Some b ->
eval_condition0 cond v2 m2 = Some b.
Proof.
- intros. destruct cond; simpl in H0; FuncInv; InvInject; simpl; auto.
- - inv H; simpl in *; congruence.
+ intros. destruct cond; cbn in H0; FuncInv; InvInject; cbn; auto.
+ - inv H; cbn in *; congruence.
- eauto 3 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
- - inv H; simpl in *; congruence.
+ - inv H; cbn in *; congruence.
- eauto 3 using Val.cmplu_bool_inject, Mem.valid_pointer_implies.
Qed.
@@ -1432,248 +1432,244 @@ Lemma eval_operation_inj:
eval_operation ge1 sp1 op vl1 m1 = Some v1 ->
exists v2, eval_operation ge2 sp2 op vl2 m2 = Some v2 /\ Val.inject f v1 v2.
Proof.
- intros until v1; intros GL; intros. destruct op; simpl in H1; simpl; FuncInv; InvInject; TrivialExists.
+ intros until v1; intros GL; intros. destruct op; cbn in H1; cbn; FuncInv; InvInject; TrivialExists.
(* addrsymbol *)
- - apply GL; simpl; auto.
+ - apply GL; cbn; auto.
(* addrstack *)
- apply Val.offset_ptr_inject; auto.
(* castsigned *)
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* add, addimm *)
- apply Val.add_inject; auto.
- apply Val.add_inject; auto.
(* addx, addximm *)
- apply Val.add_inject; trivial.
- inv H4; inv H2; simpl; try destruct (Int.ltu _ _); simpl; auto.
- - inv H4; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
+ inv H4; inv H2; cbn; try destruct (Int.ltu _ _); cbn; auto.
+ - inv H4; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
(* neg, sub *)
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
- apply Val.sub_inject; auto.
(* revsubimm, revsubx, revsubximm *)
- - inv H4; simpl; trivial.
+ - inv H4; cbn; trivial.
- apply Val.sub_inject; trivial.
- inv H4; inv H2; simpl; try destruct (Int.ltu _ _); simpl; auto.
- - inv H4; simpl; try destruct (Int.ltu _ _); simpl; auto.
+ inv H4; inv H2; cbn; try destruct (Int.ltu _ _); cbn; auto.
+ - inv H4; cbn; try destruct (Int.ltu _ _); cbn; auto.
(* mul, mulimm, mulhs, mulhu *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
- - inv H4; inv H2; simpl; auto.
- - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
(* div, divu *)
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int.eq i0 Int.zero
- || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2.
+ - inv H4; inv H3; cbn in H1; inv H1. cbn.
+ destruct (_ || _); inv H2.
TrivialExists.
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ - inv H4; inv H3; cbn in H1; inv H1. cbn.
destruct (Int.eq i0 Int.zero); inv H2. TrivialExists.
(* mod, modu *)
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int.eq i0 Int.zero
- || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2.
+ - inv H4; inv H3; cbn in H1; inv H1. cbn.
+ destruct (_ || _); inv H2.
TrivialExists.
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ - inv H4; inv H3; cbn in H1; inv H1. cbn.
destruct (Int.eq i0 Int.zero); inv H2. TrivialExists.
(* and, andimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* nand, nandimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* or, orimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* nor, norimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* xor, xorimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* nxor, nxorimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* not *)
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
(* andn, andnimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* orn, ornimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* shl, shlimm *)
- - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
- - inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
+ - inv H4; inv H2; cbn; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
+ - inv H4; cbn; auto. destruct (Int.ltu n Int.iwordsize); auto.
(* shr, shrimm *)
- - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
- - inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
+ - inv H4; inv H2; cbn; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
+ - inv H4; cbn; auto. destruct (Int.ltu n Int.iwordsize); auto.
(* shru, shruimm *)
- - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
- - inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
+ - inv H4; inv H2; cbn; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
+ - inv H4; cbn; auto. destruct (Int.ltu n Int.iwordsize); auto.
(* shrx *)
- - inv H4; simpl; auto.
- destruct (Int.ltu n (Int.repr 31)); inv H; simpl; auto.
+ - inv H4; cbn; auto.
+ destruct (Int.ltu n (Int.repr 31)); inv H; cbn; auto.
(* rorimm *)
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
(* madd, maddim *)
- - inv H2; inv H3; inv H4; simpl; auto.
- - inv H2; inv H4; simpl; auto.
+ - inv H2; inv H3; inv H4; cbn; auto.
+ - inv H2; inv H4; cbn; auto.
(* msub *)
- apply Val.sub_inject; auto.
- inv H3; inv H2; simpl; auto.
+ inv H3; inv H2; cbn; auto.
(* makelong, highlong, lowlong *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* cast32 *)
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* addl, addlimm *)
- apply Val.addl_inject; auto.
- apply Val.addl_inject; auto.
(* addxl, addxlimm *)
- apply Val.addl_inject; auto.
- inv H4; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
- - inv H4; simpl; trivial.
- destruct (Int.ltu _ _); simpl; trivial.
+ inv H4; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
+ - inv H4; cbn; trivial.
+ destruct (Int.ltu _ _); cbn; trivial.
(* negl, subl *)
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
- apply Val.subl_inject; auto.
- inv H4; inv H2; simpl; trivial;
- destruct (Int.ltu _ _); simpl; trivial.
- - inv H4; simpl; trivial;
- destruct (Int.ltu _ _); simpl; trivial.
- - inv H4; simpl; auto.
+ inv H4; inv H2; cbn; trivial;
+ destruct (Int.ltu _ _); cbn; trivial.
+ - inv H4; cbn; trivial;
+ destruct (Int.ltu _ _); cbn; trivial.
+ - inv H4; cbn; auto.
- apply Val.subl_inject; auto.
(* mull, mullhs, mullhu *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
- - inv H4; inv H2; simpl; auto.
- - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
(* divl, divlu *)
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int64.eq i0 Int64.zero
- || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2.
+ - inv H4; inv H3; cbn in H1; inv H1. cbn.
+ destruct (_ || _); inv H2.
TrivialExists.
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ - inv H4; inv H3; cbn in H1; inv H1. cbn.
destruct (Int64.eq i0 Int64.zero); inv H2. TrivialExists.
(* modl, modlu *)
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
- destruct (Int64.eq i0 Int64.zero
- || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2.
+ - inv H4; inv H3; cbn in H1; inv H1. cbn.
+ destruct (_ || _); inv H2.
TrivialExists.
- - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ - inv H4; inv H3; cbn in H1; inv H1. cbn.
destruct (Int64.eq i0 Int64.zero); inv H2. TrivialExists.
(* andl, andlimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* nandl, nandlimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* orl, orlimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* norl, norlimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* xorl, xorlimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* nxorl, nxorlimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* notl *)
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
(* andnl, andnlimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* ornl, ornlimm *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; cbn; auto.
(* shll, shllimm *)
- - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
- - inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
+ - inv H4; inv H2; cbn; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
+ - inv H4; cbn; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
(* shr, shrimm *)
- - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
- - inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
+ - inv H4; inv H2; cbn; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
+ - inv H4; cbn; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
(* shru, shruimm *)
- - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
- - inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
+ - inv H4; inv H2; cbn; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
+ - inv H4; cbn; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
(* shrx *)
- - inv H4; simpl; auto.
- destruct (Int.ltu n (Int.repr 63)); simpl; auto.
+ - inv H4; cbn; auto.
+ destruct (Int.ltu n (Int.repr 63)); cbn; auto.
(* maddl, maddlimm *)
- apply Val.addl_inject; auto.
- inv H2; inv H3; inv H4; simpl; auto.
+ inv H2; inv H3; inv H4; cbn; auto.
- apply Val.addl_inject; auto.
- inv H4; inv H2; simpl; auto.
+ inv H4; inv H2; cbn; auto.
(* msubl, msublimm *)
- apply Val.subl_inject; auto.
- inv H2; inv H3; inv H4; simpl; auto.
+ inv H2; inv H3; inv H4; cbn; auto.
(* negf, absf *)
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* addf, subf *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
(* mulf, divf *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
(* minf, maxf *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
(* fmaddf, fmsubf *)
- - inv H4; inv H3; inv H2; simpl; auto.
- - inv H4; inv H3; inv H2; simpl; auto.
+ - inv H4; inv H3; inv H2; cbn; auto.
+ - inv H4; inv H3; inv H2; cbn; auto.
(* negfs, absfs *)
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* addfs, subfs *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
(* mulfs, divfs *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
(* minfs, maxfs *)
- - inv H4; inv H2; simpl; auto.
- - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; cbn; auto.
+ - inv H4; inv H2; cbn; auto.
(* invfs *)
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
(* fmaddfs, fmsubfs *)
- - inv H4; inv H3; inv H2; simpl; auto.
- - inv H4; inv H3; inv H2; simpl; auto.
+ - inv H4; inv H3; inv H2; cbn; auto.
+ - inv H4; inv H3; inv H2; cbn; auto.
(* singleoffloat, floatofsingle *)
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* intoffloat, intuoffloat *)
- - inv H4; simpl; auto. destruct (Float.to_int f0); simpl; auto.
- - inv H4; simpl; auto. destruct (Float.to_intu f0); simpl; auto.
+ - inv H4; cbn; auto. destruct (Float.to_int f0); cbn; auto.
+ - inv H4; cbn; auto. destruct (Float.to_intu f0); cbn; auto.
(* intofsingle, intuofsingle *)
- - inv H4; simpl; auto. destruct (Float32.to_int f0); simpl; auto.
- - inv H4; simpl; auto. destruct (Float32.to_intu f0); simpl; auto.
+ - inv H4; cbn; auto. destruct (Float32.to_int f0); cbn; auto.
+ - inv H4; cbn; auto. destruct (Float32.to_intu f0); cbn; auto.
(* singleofint, singleofintu *)
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* longoffloat, longuoffloat *)
- - inv H4; simpl; auto. destruct (Float.to_long f0); simpl; auto.
- - inv H4; simpl; auto. destruct (Float.to_longu f0); simpl; auto.
+ - inv H4; cbn; auto. destruct (Float.to_long f0); cbn; auto.
+ - inv H4; cbn; auto. destruct (Float.to_longu f0); cbn; auto.
(* floatoflong, floatoflongu *)
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* longofsingle, longuofsingle *)
- - inv H4; simpl; auto. destruct (Float32.to_long f0); simpl; auto.
- - inv H4; simpl; auto. destruct (Float32.to_longu f0); simpl; auto.
+ - inv H4; cbn; auto. destruct (Float32.to_long f0); cbn; auto.
+ - inv H4; cbn; auto. destruct (Float32.to_longu f0); cbn; auto.
(* singleoflong, singleoflongu *)
- - inv H4; simpl; auto.
- - inv H4; simpl; auto.
+ - inv H4; cbn; auto.
+ - inv H4; cbn; auto.
(* cmp *)
- subst v1. destruct (eval_condition cond vl1 m1) eqn:?.
exploit eval_condition_inj; eauto. intros EQ; rewrite EQ.
- destruct b; simpl; constructor.
- simpl; constructor.
+ destruct b; cbn; constructor.
+ cbn; constructor.
(* extfz *)
- unfold extfz.
@@ -1703,16 +1699,16 @@ Proof.
- unfold insf.
destruct (is_bitfield _ _).
+ inv H4; inv H2; trivial.
- simpl. destruct (Int.ltu _ _); trivial.
- simpl. trivial.
+ cbn. destruct (Int.ltu _ _); trivial.
+ cbn. trivial.
+ trivial.
(* insfl *)
- unfold insfl.
destruct (is_bitfieldl _ _).
+ inv H4; inv H2; trivial.
- simpl. destruct (Int.ltu _ _); trivial.
- simpl. trivial.
+ cbn. destruct (Int.ltu _ _); trivial.
+ cbn. trivial.
+ trivial.
(* Osel *)
@@ -1750,13 +1746,13 @@ Lemma eval_addressing_inj:
eval_addressing ge1 sp1 addr vl1 = Some v1 ->
exists v2, eval_addressing ge2 sp2 addr vl2 = Some v2 /\ Val.inject f v1 v2.
Proof.
- intros. destruct addr; simpl in H2; simpl; FuncInv; InvInject; TrivialExists.
+ intros. destruct addr; cbn in H2; cbn; FuncInv; InvInject; TrivialExists.
- apply Val.addl_inject; trivial.
- destruct v0; destruct v'0; simpl; trivial; destruct (Int.ltu _ _); simpl; trivial; inv H3.
+ destruct v0; destruct v'0; cbn; trivial; destruct (Int.ltu _ _); cbn; trivial; inv H3.
apply Val.inject_long.
- apply Val.addl_inject; auto.
- apply Val.offset_ptr_inject; auto.
- - apply H; simpl; auto.
+ - apply H; cbn; auto.
- apply Val.offset_ptr_inject; auto.
Qed.
@@ -1771,7 +1767,7 @@ Lemma eval_addressing_inj_none:
eval_addressing ge2 sp2 addr vl2 = None.
Proof.
intros until vl2. intros Hglobal Hinjsp Hinjvl.
- destruct addr; simpl in *.
+ destruct addr; cbn in *.
1,2: inv Hinjvl; trivial;
inv H0; trivial;
inv H2; trivial;
@@ -1895,7 +1891,7 @@ Lemma eval_addressing_lessdef_none:
eval_addressing genv sp addr vl2 = None.
Proof.
intros until vl2. intros Hlessdef Heval1.
- destruct addr; simpl in *.
+ destruct addr; cbn in *.
1, 2, 4, 5: inv Hlessdef; trivial;
inv H0; trivial;
inv H2; trivial;
@@ -1980,7 +1976,7 @@ Lemma eval_operation_inject:
/\ Val.inject f v1 v2.
Proof.
intros.
- rewrite eval_shift_stack_operation. simpl.
+ rewrite eval_shift_stack_operation. cbn.
eapply eval_operation_inj with (sp1 := Vptr sp1 Ptrofs.zero) (m1 := m1); eauto.
intros; eapply Mem.valid_pointer_inject_val; eauto.
intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
diff --git a/kvx/OpWeights.ml b/kvx/OpWeights.ml
index 4c3c40d0..9614fd92 100644
--- a/kvx/OpWeights.ml
+++ b/kvx/OpWeights.ml
@@ -1,7 +1,7 @@
open Op;;
open PostpassSchedulingOracle;;
let resource_bounds = PostpassSchedulingOracle.resource_bounds;;
-
+let nr_non_pipelined_units = 0;;
let rec nlist_rec x l = function
| 0 -> l
@@ -66,6 +66,8 @@ let resources_of_op (op : operation) (nargs : int) =
let insn = insn_of_op op nargs in
let record = basic_rec insn in
rec_to_usage record;;
+
+let non_pipelined_resources_of_op (op : operation) (nargs : int) = [| |]
let resources_of_cond (cond : condition) (nargs : int) =
let insn = insn_of_cond cond nargs in
diff --git a/kvx/Peephole.v b/kvx/Peephole.v
index 35f4bbd9..5adb823b 100644
--- a/kvx/Peephole.v
+++ b/kvx/Peephole.v
@@ -153,6 +153,6 @@ Program Definition optimize_bblock (bb : bblock) :=
exit := exit bb |}.
Next Obligation.
destruct (wf_bblockb (optimize_body (body bb))) eqn:Rwf.
- - rewrite Rwf. simpl. trivial.
+ - rewrite Rwf. cbn. trivial.
- exact (correct bb).
Qed.
diff --git a/kvx/PostpassSchedulingOracle.ml b/kvx/PostpassSchedulingOracle.ml
index 67e3f80f..2107ce22 100644
--- a/kvx/PostpassSchedulingOracle.ml
+++ b/kvx/PostpassSchedulingOracle.ml
@@ -504,8 +504,7 @@ let alu_lite_y : int array = let resmap = fun r -> match r with
| Rissue -> 3 | Rtiny -> 1 | Rlite -> 1 | _ -> 0
in Array.of_list (List.map resmap resource_names)
-let alu_nop : int array = let resmap = fun r -> match r with
- | Rissue -> 1 | Rnop -> 1 | _ -> 0
+let alu_nop : int array = let resmap = fun r -> 0
in Array.of_list (List.map resmap resource_names)
let alu_tiny : int array = let resmap = fun r -> match r with
@@ -627,16 +626,16 @@ let rec_to_usage r =
| Some U27L5 | Some U27L10 -> alu_tiny_x
| Some E27U27L10 -> alu_tiny_y
| _ -> raise InvalidEncoding)
- | Maddw -> (match encoding with None -> mau_auxr
+ | Maddw | Msbfw -> (match encoding with None -> mau_auxr
| Some U6 | Some S10 | Some U27L5 -> mau_auxr_x
| _ -> raise InvalidEncoding)
- | Maddd -> (match encoding with None | Some U6 | Some S10 -> mau_auxr
+ | Maddd | Msbfd -> (match encoding with None | Some U6 | Some S10 -> mau_auxr
| Some U27L5 | Some U27L10 -> mau_auxr_x
| Some E27U27L10 -> mau_auxr_y)
- | Mulw| Msbfw -> (match encoding with None -> mau
+ | Mulw -> (match encoding with None -> mau
| Some U6 | Some S10 | Some U27L5 -> mau_x
| _ -> raise InvalidEncoding)
- | Muld | Msbfd -> (match encoding with None | Some U6 | Some S10 -> mau
+ | Muld -> (match encoding with None | Some U6 | Some S10 -> mau
| Some U27L5 | Some U27L10 -> mau_x
| Some E27U27L10 -> mau_y)
| Nop -> alu_nop
@@ -914,12 +913,20 @@ let print_bb oc bb =
let asm_instructions = Asm.unfold_bblock bb
in List.iter (print_inst oc) asm_instructions
+let print_schedule sched =
+ print_string "[ ";
+ Array.iter (fun x -> Printf.printf "%d; " x) sched;
+ print_endline "]";;
+
let do_schedule bb =
- let problem = build_problem bb
- in let solution = scheduler_by_name (!Clflags.option_fpostpass_sched) problem
+ let problem = build_problem bb in
+ (if debug then print_problem stdout problem);
+ let solution = scheduler_by_name (!Clflags.option_fpostpass_sched) problem
in match solution with
| None -> failwith "Could not find a valid schedule"
- | Some sol -> let bundles = bundlize_solution bb sol in
+ | Some sol ->
+ ((if debug then print_schedule sol);
+ let bundles = bundlize_solution bb sol in
(if debug then
begin
Printf.eprintf "Scheduling the following group of instructions:\n";
@@ -928,7 +935,7 @@ let do_schedule bb =
List.iter (print_bb stderr) bundles;
Printf.eprintf "--------------------------------\n"
end;
- bundles)
+ bundles))
(**
* Dumb schedule if the above doesn't work
diff --git a/kvx/PrepassSchedulingOracle.ml b/kvx/PrepassSchedulingOracle.ml
new file mode 120000
index 00000000..912e9ffa
--- /dev/null
+++ b/kvx/PrepassSchedulingOracle.ml
@@ -0,0 +1 @@
+../aarch64/PrepassSchedulingOracle.ml \ No newline at end of file
diff --git a/kvx/Stacklayout.v b/kvx/Stacklayout.v
index 46202e03..81ffcebb 100644
--- a/kvx/Stacklayout.v
+++ b/kvx/Stacklayout.v
@@ -63,7 +63,7 @@ Lemma frame_env_separated:
** P.
Proof.
Local Opaque Z.add Z.mul sepconj range.
- intros; simpl.
+ intros; cbn.
set (w := if Archi.ptr64 then 8 else 4).
set (olink := align (4 * b.(bound_outgoing)) w).
set (oretaddr := olink + w).
@@ -105,7 +105,7 @@ Lemma frame_env_range:
let fe := make_env b in
0 <= fe_stack_data fe /\ fe_stack_data fe + bound_stack_data b <= fe_size fe.
Proof.
- intros; simpl.
+ intros; cbn.
set (w := if Archi.ptr64 then 8 else 4).
set (olink := align (4 * b.(bound_outgoing)) w).
set (oretaddr := olink + w).
@@ -133,7 +133,7 @@ Lemma frame_env_aligned:
/\ (align_chunk Mptr | fe_ofs_link fe)
/\ (align_chunk Mptr | fe_ofs_retaddr fe).
Proof.
- intros; simpl.
+ intros; cbn.
set (w := if Archi.ptr64 then 8 else 4).
set (olink := align (4 * b.(bound_outgoing)) w).
set (oretaddr := olink + w).
diff --git a/kvx/ValueAOp.v b/kvx/ValueAOp.v
index e634fdc0..87554258 100644
--- a/kvx/ValueAOp.v
+++ b/kvx/ValueAOp.v
@@ -16,87 +16,6 @@
Require Import Coqlib Compopts.
Require Import AST Integers Floats Values Memory Globalenvs.
Require Import Op ExtValues ExtFloats RTL ValueDomain.
-
-Definition intoffloat_total (x: aval) :=
- match x with
- | F f =>
- match Float.to_int f with
- | Some i => I i
- | None => ntop
- end
- | _ => ntop1 x
- end.
-
-Definition intuoffloat_total (x: aval) :=
- match x with
- | F f =>
- match Float.to_intu f with
- | Some i => I i
- | None => ntop
- end
- | _ => ntop1 x
- end.
-
-Definition intofsingle_total (x: aval) :=
- match x with
- | FS f =>
- match Float32.to_int f with
- | Some i => I i
- | None => ntop
- end
- | _ => ntop1 x
- end.
-
-Definition intuofsingle_total (x: aval) :=
- match x with
- | FS f =>
- match Float32.to_intu f with
- | Some i => I i
- | None => ntop
- end
- | _ => ntop1 x
- end.
-
-Definition longoffloat_total (x: aval) :=
- match x with
- | F f =>
- match Float.to_long f with
- | Some i => L i
- | None => ntop
- end
- | _ => ntop1 x
- end.
-
-Definition longuoffloat_total (x: aval) :=
- match x with
- | F f =>
- match Float.to_longu f with
- | Some i => L i
- | None => ntop
- end
- | _ => ntop1 x
- end.
-
-Definition longofsingle_total (x: aval) :=
- match x with
- | FS f =>
- match Float32.to_long f with
- | Some i => L i
- | None => ntop
- end
- | _ => ntop1 x
- end.
-
-Definition longuofsingle_total (x: aval) :=
- match x with
- | FS f =>
- match Float32.to_longu f with
- | Some i => L i
- | None => ntop
- end
- | _ => ntop1 x
- end.
-
Definition minf := binop_float ExtFloat.min.
Definition maxf := binop_float ExtFloat.max.
Definition minfs := binop_single ExtFloat32.min.
@@ -400,196 +319,6 @@ Hypothesis GENV: genv_match bc ge.
Variable sp: block.
Hypothesis STACK: bc sp = BCstack.
-Lemma intoffloat_total_sound:
- forall v x
- (MATCH : vmatch bc v x),
- vmatch bc (Val.maketotal (Val.intoffloat v)) (intoffloat_total x).
-Proof.
- unfold Val.intoffloat, intoffloat_total. intros.
- inv MATCH; simpl in *; try constructor.
- all: destruct (Float.to_int f) as [i|] eqn:E; simpl; [auto with va | constructor].
- unfold ntop1, provenance.
- destruct (va_strict tt); constructor.
-Qed.
-
-Hint Resolve intoffloat_total_sound : va.
-
-Lemma intuoffloat_total_sound:
- forall v x
- (MATCH : vmatch bc v x),
- vmatch bc (Val.maketotal (Val.intuoffloat v)) (intuoffloat_total x).
-Proof.
- unfold Val.intoffloat, intoffloat_total. intros.
- inv MATCH; simpl in *; try constructor.
- all: destruct (Float.to_intu f) as [i|] eqn:E; simpl; [auto with va | constructor].
- unfold ntop1, provenance.
- destruct (va_strict tt); constructor.
-Qed.
-
-Hint Resolve intuoffloat_total_sound : va.
-
-Lemma intofsingle_total_sound:
- forall v x
- (MATCH : vmatch bc v x),
- vmatch bc (Val.maketotal (Val.intofsingle v)) (intofsingle_total x).
-Proof.
- unfold Val.intofsingle, intofsingle_total. intros.
- inv MATCH; simpl in *; try constructor.
- all: destruct (Float32.to_int f) as [i|] eqn:E; simpl; [auto with va | constructor].
- unfold ntop1, provenance.
- destruct (va_strict tt); constructor.
-Qed.
-
-Hint Resolve intofsingle_total_sound : va.
-
-Lemma intuofsingle_total_sound:
- forall v x
- (MATCH : vmatch bc v x),
- vmatch bc (Val.maketotal (Val.intuofsingle v)) (intuofsingle_total x).
-Proof.
- unfold Val.intofsingle, intofsingle_total. intros.
- inv MATCH; simpl in *; try constructor.
- all: destruct (Float32.to_intu f) as [i|] eqn:E; simpl; [auto with va | constructor].
- unfold ntop1, provenance.
- destruct (va_strict tt); constructor.
-Qed.
-
-Hint Resolve intuofsingle_total_sound : va.
-
-Lemma singleofint_total_sound:
- forall v x, vmatch bc v x ->
- vmatch bc (Val.maketotal (Val.singleofint v)) (singleofint x).
-Proof.
- unfold Val.singleofint, singleofint; intros.
- inv H; simpl.
- all: auto with va.
- all: unfold ntop1, provenance.
- all: try constructor.
-Qed.
-
-Hint Resolve singleofint_total_sound : va.
-
-Lemma singleofintu_total_sound:
- forall v x, vmatch bc v x ->
- vmatch bc (Val.maketotal (Val.singleofintu v)) (singleofintu x).
-Proof.
- unfold Val.singleofintu, singleofintu; intros.
- inv H; simpl.
- all: auto with va.
- all: unfold ntop1, provenance.
- all: try constructor.
-Qed.
-
-Hint Resolve singleofintu_total_sound : va.
-
-Lemma longoffloat_total_sound:
- forall v x
- (MATCH : vmatch bc v x),
- vmatch bc (Val.maketotal (Val.longoffloat v)) (longoffloat_total x).
-Proof.
- unfold Val.longoffloat, longoffloat_total. intros.
- inv MATCH; simpl in *; try constructor.
- all: destruct (Float.to_long f) as [i|] eqn:E; simpl; [auto with va | constructor].
- unfold ntop1, provenance.
- destruct (va_strict tt); constructor.
-Qed.
-
-Hint Resolve longoffloat_total_sound : va.
-
-Lemma longuoffloat_total_sound:
- forall v x
- (MATCH : vmatch bc v x),
- vmatch bc (Val.maketotal (Val.longuoffloat v)) (longuoffloat_total x).
-Proof.
- unfold Val.longoffloat, longoffloat_total. intros.
- inv MATCH; simpl in *; try constructor.
- all: destruct (Float.to_longu f) as [i|] eqn:E; simpl; [auto with va | constructor].
- unfold ntop1, provenance.
- destruct (va_strict tt); constructor.
-Qed.
-
-Hint Resolve longuoffloat_total_sound : va.
-
-Lemma longofsingle_total_sound:
- forall v x
- (MATCH : vmatch bc v x),
- vmatch bc (Val.maketotal (Val.longofsingle v)) (longofsingle_total x).
-Proof.
- unfold Val.longofsingle, longofsingle_total. intros.
- inv MATCH; simpl in *; try constructor.
- all: destruct (Float32.to_long f) as [i|] eqn:E; simpl; [auto with va | constructor].
- unfold ntop1, provenance.
- destruct (va_strict tt); constructor.
-Qed.
-
-Hint Resolve longofsingle_total_sound : va.
-
-Lemma longuofsingle_total_sound:
- forall v x
- (MATCH : vmatch bc v x),
- vmatch bc (Val.maketotal (Val.longuofsingle v)) (longuofsingle_total x).
-Proof.
- unfold Val.longofsingle, longofsingle_total. intros.
- inv MATCH; simpl in *; try constructor.
- all: destruct (Float32.to_longu f) as [i|] eqn:E; simpl; [auto with va | constructor].
- unfold ntop1, provenance.
- destruct (va_strict tt); constructor.
-Qed.
-
-Hint Resolve longuofsingle_total_sound : va.
-
-Lemma singleoflong_total_sound:
- forall v x, vmatch bc v x ->
- vmatch bc (Val.maketotal (Val.singleoflong v)) (singleoflong x).
-Proof.
- unfold Val.singleoflong, singleoflong; intros.
- inv H; simpl.
- all: auto with va.
- all: unfold ntop1, provenance.
- all: try constructor.
-Qed.
-
-Hint Resolve singleoflong_total_sound : va.
-
-Lemma singleoflongu_total_sound:
- forall v x, vmatch bc v x ->
- vmatch bc (Val.maketotal (Val.singleoflongu v)) (singleoflongu x).
-Proof.
- unfold Val.singleoflongu, singleoflongu; intros.
- inv H; simpl.
- all: auto with va.
- all: unfold ntop1, provenance.
- all: try constructor.
-Qed.
-
-Hint Resolve singleoflongu_total_sound : va.
-
-Lemma floatoflong_total_sound:
- forall v x, vmatch bc v x ->
- vmatch bc (Val.maketotal (Val.floatoflong v)) (floatoflong x).
-Proof.
- unfold Val.floatoflong, floatoflong; intros.
- inv H; simpl.
- all: auto with va.
- all: unfold ntop1, provenance.
- all: try constructor.
-Qed.
-
-Hint Resolve floatoflong_total_sound : va.
-
-Lemma floatoflongu_total_sound:
- forall v x, vmatch bc v x ->
- vmatch bc (Val.maketotal (Val.floatoflongu v)) (floatoflongu x).
-Proof.
- unfold Val.floatoflongu, floatoflongu; intros.
- inv H; simpl.
- all: auto with va.
- all: unfold ntop1, provenance.
- all: try constructor.
-Qed.
-
-Hint Resolve floatoflongu_total_sound : va.
-
Lemma minf_sound:
forall v x w y, vmatch bc v x -> vmatch bc w y -> vmatch bc (ExtValues.minf v w) (minf x y).
Proof.
@@ -620,7 +349,7 @@ Proof.
intros v x;
intro MATCH;
inversion MATCH;
- simpl;
+ cbn;
constructor.
Qed.
@@ -632,9 +361,9 @@ Lemma triple_op_float_sound:
Proof.
intros until z.
intros Hax Hby Hcz.
- inv Hax; simpl; try constructor;
- inv Hby; simpl; try constructor;
- inv Hcz; simpl; try constructor.
+ inv Hax; cbn; try constructor;
+ inv Hby; cbn; try constructor;
+ inv Hcz; cbn; try constructor.
Qed.
Lemma triple_op_single_sound:
@@ -645,9 +374,9 @@ Lemma triple_op_single_sound:
Proof.
intros until z.
intros Hax Hby Hcz.
- inv Hax; simpl; try constructor;
- inv Hby; simpl; try constructor;
- inv Hcz; simpl; try constructor.
+ inv Hax; cbn; try constructor;
+ inv Hby; cbn; try constructor;
+ inv Hcz; cbn; try constructor.
Qed.
Lemma fmaddf_sound :
@@ -691,9 +420,9 @@ Proof.
intros until aargs; intros VM. inv VM.
destruct cond; auto with va.
inv H0.
- destruct cond; simpl; eauto with va.
+ destruct cond; cbn; eauto with va.
inv H2.
- destruct cond; simpl; eauto with va.
+ destruct cond; cbn; eauto with va.
destruct cond; auto with va.
Qed.
@@ -703,7 +432,7 @@ Theorem eval_static_condition0_sound:
cmatch (eval_condition0 cond varg m) (eval_static_condition0 cond aarg).
Proof.
intros until aarg; intro VM.
- destruct cond; simpl; eauto with va.
+ destruct cond; cbn; eauto with va.
Qed.
Lemma symbol_address_sound:
@@ -812,19 +541,9 @@ Proof.
+ eauto with va.
+ destruct n; destruct shift; reflexivity.
- (* shrx *)
- inv H1; simpl; try constructor.
- all: destruct Int.ltu; [simpl | constructor; fail].
+ inv H1; cbn; try constructor.
+ all: destruct Int.ltu; [cbn | constructor; fail].
all: auto with va.
- - replace (match Val.shll a1 (Vint (int_of_shift1_4 shift)) with
- | Vlong n2 => Vlong (Int64.add n n2)
- | Vptr b2 ofs2 =>
- if Archi.ptr64
- then Vptr b2 (Ptrofs.add ofs2 (Ptrofs.of_int64 n))
- else Vundef
- | _ => Vundef
- end) with (Val.addl (Vlong n) (Val.shll a1 (Vint (int_of_shift1_4 shift)))).
- + eauto with va.
- + destruct a1; destruct shift; reflexivity.
- inv H1; constructor.
- replace (match Val.shll a1 (Vint (int_of_shift1_4 shift)) with
| Vlong n2 => Vlong (Int64.sub n n2)
@@ -832,10 +551,6 @@ Proof.
end) with (Val.subl (Vlong n) (Val.shll a1 (Vint (int_of_shift1_4 shift)))).
+ eauto with va.
+ destruct a1; destruct shift; reflexivity.
- - (* shrxl *)
- inv H1; simpl; try constructor.
- all: destruct Int.ltu; [simpl | constructor; fail].
- all: auto with va.
- apply of_optbool_sound. eapply eval_static_condition_sound; eauto.
(* extfz *)
@@ -865,12 +580,12 @@ Proof.
(* insf *)
- unfold insf, eval_static_insf.
destruct (is_bitfield _ _).
- + inv H1; inv H0; simpl; try constructor; destruct (Int.ltu _ _); simpl; constructor.
+ + inv H1; inv H0; cbn; try constructor; destruct (Int.ltu _ _); cbn; constructor.
+ constructor.
(* insfl *)
- unfold insfl, eval_static_insfl.
destruct (is_bitfieldl _ _).
- + inv H1; inv H0; simpl; try constructor; destruct (Int.ltu _ _); simpl; constructor.
+ + inv H1; inv H0; cbn; try constructor; destruct (Int.ltu _ _); cbn; constructor.
+ constructor.
(* select *)
- apply select_sound; auto. eapply eval_static_condition0_sound; eauto.
diff --git a/kvx/abstractbb/AbstractBasicBlocksDef.v b/kvx/abstractbb/AbstractBasicBlocksDef.v
index 0b1c502d..6960f363 100644
--- a/kvx/abstractbb/AbstractBasicBlocksDef.v
+++ b/kvx/abstractbb/AbstractBasicBlocksDef.v
@@ -45,7 +45,7 @@ End LangParam.
-(** * Syntax and (sequential) semantics of "basic blocks" *)
+(** * Syntax and (sequential) semantics of "abstract basic blocks" *)
Module MkSeqLanguage(P: LangParam).
Export P.
@@ -62,12 +62,12 @@ Definition assign (m: mem) (x:R.t) (v: value): mem
:= fun y => if R.eq_dec x y then v else m y.
-(** expressions *)
+(** Expressions *)
Inductive exp :=
- | PReg (x:R.t)
- | Op (o:op) (le: list_exp)
- | Old (e: 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)
@@ -95,7 +95,8 @@ with list_exp_eval (le: list_exp) (m old: mem): option (list value) :=
| LOld le => list_exp_eval le old old
end.
-Definition inst := list (R.t * exp). (* = a sequence of assignments *)
+(** 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
@@ -107,6 +108,7 @@ Fixpoint inst_run (i: inst) (m old: mem): option mem :=
end
end.
+(** A basic block is a sequence of instructions. *)
Definition bblock := list inst.
Fixpoint run (p: bblock) (m: mem): option mem :=
@@ -168,7 +170,7 @@ Lemma exp_equiv e old1 old2:
(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); simpl; try congruence; auto.
+ 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.
@@ -181,38 +183,38 @@ Lemma inst_equiv_refl i old1 old2:
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]]; simpl; eauto.
+ intro H; induction i as [ | [x e]]; cbn; eauto.
intros m1 m2 H1. erewrite exp_equiv; eauto.
- destruct (exp_eval e m2 old2); simpl; auto.
+ 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']; simpl; eauto.
+ 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); simpl.
- - intros [m3 [H1 H2]]; rewrite H1; simpl; auto.
- - intros H1; rewrite H1; simpl; auto.
+ 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; simpl.
- - intros [m2 [H1 H2]]; subst; simpl. eauto.
- - intros; subst; simpl; eauto.
+ 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; simpl.
- - intros [m2 [H1 H2]]; subst; simpl.
- intros [m3 [H3 H4]]; subst; simpl.
+ 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; simpl; 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)).
@@ -230,8 +232,8 @@ Lemma run_app p1: forall m1 p2,
| None => None
end.
Proof.
- induction p1; simpl; try congruence.
- intros; destruct (inst_run _ _ _); simpl; auto.
+ induction p1; cbn; try congruence.
+ intros; destruct (inst_run _ _ _); cbn; auto.
Qed.
Lemma run_app_None p1 m1 p2:
@@ -250,12 +252,16 @@ Qed.
End SEQLANG.
-Module Terms.
-(** terms in the symbolic evaluation
-NB: such a term represents the successive computations in one given pseudo-register
+(** * 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)
@@ -328,17 +334,27 @@ Fixpoint allvalid ge (l: list term) m : Prop :=
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]; simpl; try (tauto).
+ 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.
@@ -346,26 +362,24 @@ Proof.
symmetry; eauto.
Qed.
-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 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; simpl; intros; intuition congruence.
+ 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] t.
Lemma identity_fail_correct (t: term): match_pt t (identity_fail t).
Proof.
- eapply intro_fail_correct; simpl; tauto.
+ 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 ([])%list t
@@ -376,15 +390,16 @@ Definition nofail (is_constant: op -> bool) (t: term):=
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; simpl.
- + intros; eapply intro_fail_correct; simpl; intuition congruence.
- + intros; destruct l; simpl; auto with wlp.
- destruct (is_constant o) eqn:Heqo; simpl; intuition eauto with wlp.
- eapply intro_fail_correct; simpl; intuition eauto with wlp.
+ 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.
@@ -401,6 +416,7 @@ Proof.
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) |}.
@@ -409,7 +425,7 @@ Lemma app_fail_allvalid_correct l pt t1 t2: forall
(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; simpl. clear V1 V2.
+ 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.
@@ -431,6 +447,8 @@ 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;
diff --git a/kvx/abstractbb/ImpSimuTest.v b/kvx/abstractbb/ImpSimuTest.v
index c914eee1..b1a3b985 100644
--- a/kvx/abstractbb/ImpSimuTest.v
+++ b/kvx/abstractbb/ImpSimuTest.v
@@ -10,13 +10,16 @@
(* *)
(* *************************************************************)
-(** Implementation of a symbolic execution of sequential semantics of Abstract Basic Blocks
+(** 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.
-with imperative hash-consing, and rewriting.
*)
-Require Export Impure.ImpHCons.
+Require Export Impure.ImpHCons. (**r Import the Impure library. See https://github.com/boulme/ImpureDemo *)
Export Notations.
Import HConsing.
@@ -32,6 +35,7 @@ Import ListNotations.
Local Open Scope list_scope.
+(** * Interface of (impure) equality tests for operators *)
Module Type ImpParam.
Include LangParam.
@@ -54,6 +58,8 @@ Include MkSeqLanguage LP.
End ISeqLanguage.
+(** * A generic dictinary on PseudoRegisters with an impure equality test *)
+
Module Type ImpDict.
Declare Module R: PseudoRegisters.
@@ -91,26 +97,27 @@ Parameter eq_test_correct: forall A (d1 d2: t A),
(* NB: we could also take an eq_test on R.t (but not really useful with "pure" dictionaries *)
-
-(* only for debugging *)
+(** 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) ->
@@ -127,6 +134,7 @@ Parameter verb_bblock_simu_test_correct:
End ImpSimuInterface.
+(** * Implementation of the provided tests *)
Module ImpSimu (L: ISeqLanguage) (Dict: ImpDict with Module R:=L.LP.R): ImpSimuInterface with Module CoreL := L.
@@ -152,13 +160,13 @@ Definition list_term_set_hid (l: list_term) (hid: hashcode): list_term :=
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; simpl; auto.
+ 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; simpl; auto.
+ destruct l; cbn; auto.
Qed.
(* Local nickname *)
@@ -168,7 +176,7 @@ Section SimuWithReduce.
Variable reduce: reduction.
-Section CanonBuilding.
+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.
@@ -307,7 +315,7 @@ Proof.
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 * |- *; simpl.
+ rewrite !allvalid_extensionality in * |- *; cbn.
intuition (subst; eauto).
+ eapply smem_valid_set_proof; eauto.
erewrite <- EQT; eauto.
@@ -315,11 +323,11 @@ Proof.
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 * |- *; simpl.
+ - 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; simpl; eauto.
- + intros; rewrite !Dict.set_spec_diff; simpl; eauto.
+ + 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.
@@ -396,10 +404,10 @@ 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']; simpl; wlp_xsimplify ltac:(eauto with wlp).
+ 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.
- simpl; intuition (subst; eauto with wlp localhint).
+ cbn; intuition (subst; eauto with wlp localhint).
Qed.
(*Local Hint Resolve hterm_append_correct: wlp.*)
Global Opaque hterm_append.
@@ -423,8 +431,8 @@ Lemma smart_set_correct hd x ht:
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; simpl. case (R.eq_dec x0 y).
- - intros; subst. rewrite Dict.set_spec_eq, Dict.rem_spec_eq. simpl; congruence.
+ 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.*)
@@ -448,17 +456,17 @@ Proof.
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; simpl.
+ eapply equiv_hsmem_models; eauto; unfold equiv_hsmem; cbn.
destruct H as (VALID & EFFECT); split.
- intros; rewrite APPEND, <- VALID.
- rewrite !allvalid_extensionality in * |- *; simpl; intuition (subst; eauto).
+ 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; simpl. rewrite !Dict.set_spec_eq.
+ + intros; subst. unfold hsmem_post_eval; cbn. rewrite !Dict.set_spec_eq.
erewrite LIFT, EFFECT; eauto.
- + intros; unfold hsmem_post_eval; simpl. rewrite !Dict.set_spec_diff; auto.
- * rewrite allvalid_extensionality in ALLVALID; destruct (ALLVALID ht); simpl; auto.
+ + 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.
@@ -473,7 +481,7 @@ 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 * |- * ; simpl; intuition eauto.
+ unfold smem_model in * |- * ; cbn; intuition eauto.
- erewrite IHe; eauto.
- erewrite IHe0, IHe; eauto.
Qed.
@@ -508,10 +516,10 @@ Lemma exp_hterm_correct_x ge e hod od:
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 * |- * ; simpl; wlp_simplify.
+ unfold smem_model, hsmem_post_eval in * |- * ; cbn; wlp_simplify.
- rewrite H1, <- H4; auto.
- - rewrite H4, <- H0; simpl; auto.
- - rewrite H5, <- H0, <- H4; simpl; auto.
+ - rewrite H4, <- H0; cbn; auto.
+ - rewrite H5, <- H0, <- H4; cbn; auto.
Qed.
Global Opaque exp_hterm.
@@ -536,7 +544,7 @@ Lemma hinst_smem_correct i: forall hd hod,
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; simpl; wlp_simplify; eauto 15 with wlp.
+ induction i; cbn; wlp_simplify; eauto 15 with wlp.
Qed.
Global Opaque hinst_smem.
Local Hint Resolve hinst_smem_correct: wlp.
@@ -556,7 +564,7 @@ Fixpoint bblock_hsmem_rec (p: bblock) (d: hsmem): ?? hsmem :=
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; simpl; wlp_simplify.
+ induction p; cbn; wlp_simplify.
Qed.
Global Opaque bblock_hsmem_rec.
Local Hint Resolve bblock_hsmem_rec_correct: wlp.
@@ -565,8 +573,8 @@ 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; simpl; intuition try congruence.
- rewrite !Dict.empty_spec; simpl; auto.
+ unfold smem_model, smem_valid, hsmem_post_eval; cbn; intuition try congruence.
+ rewrite !Dict.empty_spec; cbn; auto.
Qed.
Definition bblock_hsmem: bblock -> ?? hsmem
@@ -714,7 +722,7 @@ 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; simpl in * |- *; auto.
+ destruct exta0; cbn in * |- *; auto.
Qed.
Global Opaque g_bblock_simu_test.
@@ -1117,9 +1125,9 @@ Extraction Inline lift.
End ImpSimu.
-Require Import FMapPositive.
-
+(** * Implementation of the Dictionary (based on PositiveMap) *)
+Require Import FMapPositive.
Require Import PArith.
Require Import FMapPositive.
@@ -1201,12 +1209,12 @@ 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]; simpl;
- wlp_simplify; (discriminate || (subst; destruct x; simpl; auto)).
+ 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 *)
+(** 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"
@@ -1219,7 +1227,7 @@ Fixpoint pick {A} (d: t A): ?? R.t :=
RET (xO p)
end.
-(* ONLY FOR DEBUGGING INFO: find one variable on which d1 and d2 differs *)
+(** 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
diff --git a/kvx/abstractbb/Impure/ImpConfig.v b/kvx/abstractbb/Impure/ImpConfig.v
deleted file mode 100644
index dd9785b5..00000000
--- a/kvx/abstractbb/Impure/ImpConfig.v
+++ /dev/null
@@ -1,85 +0,0 @@
-(** Impure Config for UNTRUSTED backend !!! *)
-
-Require Import ImpMonads.
-Require Extraction.
-(** Pure computations (used for extraction !)
-
-We keep module [Impure] opaque in order to check that Coq proof do not depend on
-the implementation of [Impure].
-
-*)
-
-Module Type ImpureView.
-
- Include MayReturnMonad.
-
-(* WARNING: THIS IS REALLY UNSAFE TO DECOMMENT THE "UnsafeImpure" module !
-
- unsafe_coerce coerces an impure computation into a pure one !
-
-*)
-
-(* START COMMENT *)
- Module UnsafeImpure.
-
- Parameter unsafe_coerce: forall {A}, t A -> option A.
-
- Parameter unsafe_coerce_not_really_correct: forall A (k: t A) (x:A), (unsafe_coerce k)=Some x -> mayRet k x.
-
- Extraction Inline unsafe_coerce.
-
- End UnsafeImpure.
-(* END COMMENT *)
-
-
-End ImpureView.
-
-
-Module Impure: ImpureView.
-
- Include IdentityMonad.
-
- Module UnsafeImpure.
-
- Definition unsafe_coerce {A} (x:t A) := Some x.
-
- Lemma unsafe_coerce_not_really_correct: forall A (k: t A) x, (unsafe_coerce k)=Some x -> mayRet k x.
- Proof.
- unfold unsafe_coerce, mayRet; congruence.
- Qed.
-
- End UnsafeImpure.
-
-End Impure.
-
-
-(** Comment the above code and decomment this to test that coq proofs still work with an impure monad !
-
-- this should fail only on extraction or if unsafe_coerce is used !
-
-*)
-(*
-Module Impure: MayReturnMonad := PowerSetMonad.
-*)
-
-Export Impure.
-
-Extraction Inline ret mk_annot.
-
-
-(* WARNING. The following directive is unsound.
-
- Extraction Inline bind
-
-For example, it may lead to extract the following code as "true" (instead of an error raising code)
- failwith "foo";;true
-
-*)
-
-Extract Inlined Constant bind => "(|>)".
-
-
-Extract Constant t "" => "". (* This weird directive extracts [t] as "'a" instead of "'a t" *)
-Extraction Inline t.
-
-Global Opaque t.
diff --git a/kvx/abstractbb/Impure/ImpCore.v b/kvx/abstractbb/Impure/ImpCore.v
deleted file mode 100644
index 508b3f19..00000000
--- a/kvx/abstractbb/Impure/ImpCore.v
+++ /dev/null
@@ -1,196 +0,0 @@
-(** Impure monad for interface with impure code
-
-*)
-
-Require Export Program.
-Require Export ImpConfig.
-
-(* Theory: bind + embed => dbind
-
-Program Definition dbind {A B} (k1: t A) (k2: forall (a:A), (mayRet k1 a) -> t B) : t B
- := bind (mk_annot k1) (fun a => k2 a _).
-
-Lemma mayRet_dbind: forall (A B:Type) k1 k2 (b:B),
- mayRet (dbind k1 k2) b -> exists a:A, exists H: (mayRet k1 a), mayRet (k2 a H) b.
-Proof.
- intros A B k1 k2 b H; decompose [ex and] (mayRet_bind _ _ _ _ _ H).
- eapply ex_intro.
- eapply ex_intro.
- eauto.
-Qed.
-
-*)
-
-Definition wlp {A:Type} (k: t A) (P: A -> Prop): Prop
- := forall a, mayRet k a -> P a.
-
-(* Notations *)
-
-(* Print Grammar constr. *)
-
-Module Notations.
-
- Bind Scope impure_scope with t.
- Delimit Scope impure_scope with impure.
-
- Notation "?? A" := (t A) (at level 0, A at level 95): impure_scope.
-
- Notation "k '~~>' a" := (mayRet k a) (at level 75, no associativity): impure_scope.
-
- Notation "'RET' a" := (ret a) (at level 0): impure_scope.
-
- Notation "'DO' x '<~' k1 ';;' k2" := (bind k1 (fun x => k2))
- (at level 55, k1 at level 53, x at level 99, right associativity): impure_scope.
-
- Notation "k1 ';;' k2" := (bind k1 (fun _ => k2))
- (at level 55, right associativity): impure_scope.
-
- Notation "'WHEN' k '~>' a 'THEN' R" := (wlp k (fun a => R))
- (at level 73, R at level 100, right associativity): impure_scope.
-
- Notation "'ASSERT' P" := (ret (A:=P) _) (at level 0, only parsing): impure_scope.
-
-End Notations.
-
-Import Notations.
-Local Open Scope impure.
-
-Goal ((?? list nat * ??nat -> nat) = ((?? ((list nat) * ?? nat) -> nat)))%type.
-Proof.
- apply refl_equal.
-Qed.
-
-
-(* wlp lemmas for tactics *)
-
-Lemma wlp_unfold A (k:??A)(P: A -> Prop):
- (forall a, k ~~> a -> P a)
- -> wlp k P.
-Proof.
- auto.
-Qed.
-
-Lemma wlp_monotone A (k:?? A) (P1 P2: A -> Prop):
- wlp k P1
- -> (forall a, k ~~> a -> P1 a -> P2 a)
- -> wlp k P2.
-Proof.
- unfold wlp; eauto.
-Qed.
-
-Lemma wlp_forall A B (k:?? A) (P: B -> A -> Prop):
- (forall x, wlp k (P x))
- -> wlp k (fun a => forall x, P x a).
-Proof.
- unfold wlp; auto.
-Qed.
-
-Lemma wlp_ret A (P: A -> Prop) a:
- P a -> wlp (ret a) P.
-Proof.
- unfold wlp.
- intros H b H0.
- rewrite <- (mayRet_ret _ a b H0).
- auto.
-Qed.
-
-Lemma wlp_bind A B (k1:??A) (k2: A -> ??B) (P: B -> Prop):
- wlp k1 (fun a => wlp (k2 a) P) -> wlp (bind k1 k2) P.
-Proof.
- unfold wlp.
- intros H a H0.
- case (mayRet_bind _ _ _ _ _ H0); clear H0.
- intuition eauto.
-Qed.
-
-Lemma wlp_ifbool A (cond: bool) (k1 k2: ?? A) (P: A -> Prop):
- (cond=true -> wlp k1 P) -> (cond=false -> wlp k2 P) -> wlp (if cond then k1 else k2) P.
-Proof.
- destruct cond; auto.
-Qed.
-
-Lemma wlp_letprod (A B C: Type) (p: A*B) (k: A -> B -> ??C) (P: C -> Prop):
- (wlp (k (fst p) (snd p)) P)
- -> (wlp (let (x,y):=p in (k x y)) P).
-Proof.
- destruct p; simpl; auto.
-Qed.
-
-Lemma wlp_sum (A B C: Type) (x: A+B) (k1: A -> ??C) (k2: B -> ??C) (P: C -> Prop):
- (forall a, x=inl a -> wlp (k1 a) P) ->
- (forall b, x=inr b -> wlp (k2 b) P) ->
- (wlp (match x with inl a => k1 a | inr b => k2 b end) P).
-Proof.
- destruct x; simpl; auto.
-Qed.
-
-Lemma wlp_sumbool (A B:Prop) (C: Type) (x: {A}+{B}) (k1: A -> ??C) (k2: B -> ??C) (P: C -> Prop):
- (forall a, x=left a -> wlp (k1 a) P) ->
- (forall b, x=right b -> wlp (k2 b) P) ->
- (wlp (match x with left a => k1 a | right b => k2 b end) P).
-Proof.
- destruct x; simpl; auto.
-Qed.
-
-Lemma wlp_option (A B: Type) (x: option A) (k1: A -> ??B) (k2: ??B) (P: B -> Prop):
- (forall a, x=Some a -> wlp (k1 a) P) ->
- (x=None -> wlp k2 P) ->
- (wlp (match x with Some a => k1 a | None => k2 end) P).
-Proof.
- destruct x; simpl; auto.
-Qed.
-
-(* Tactics
-
-MAIN tactics:
- - xtsimplify "base": simplification using from hints in "base" database (in particular "wlp" lemmas).
- - xtstep "base": only one step of simplification.
-
-For good performance, it is recommanded to have several databases.
-
-*)
-
-Ltac introcomp :=
- let a:= fresh "exta" in
- let H:= fresh "Hexta" in
- intros a H.
-
-(* decompose the current wlp goal using "introduction" rules *)
-Ltac wlp_decompose :=
- apply wlp_ret
- || apply wlp_bind
- || apply wlp_ifbool
- || apply wlp_letprod
- || apply wlp_sum
- || apply wlp_sumbool
- || apply wlp_option
- .
-
-(* this tactic simplifies the current "wlp" goal using any hint found via tactic "hint". *)
-Ltac apply_wlp_hint hint :=
- eapply wlp_monotone;
- [ hint; fail | idtac ] ;
- simpl; introcomp.
-
-(* one step of wlp_xsimplify
-*)
-Ltac wlp_step hint :=
- match goal with
- | |- (wlp _ _) =>
- wlp_decompose
- || apply_wlp_hint hint
- || (apply wlp_unfold; introcomp)
- end.
-
-(* main general tactic
-WARNING: for the good behavior of "wlp_xsimplify", "hint" must at least perform a "eauto".
-
-Example of use:
- wlp_xsimplify (intuition eauto with base).
-*)
-Ltac wlp_xsimplify hint :=
- repeat (intros; subst; wlp_step hint; simpl; (tauto || hint)).
-
-Create HintDb wlp discriminated.
-
-Ltac wlp_simplify := wlp_xsimplify ltac:(intuition eauto with wlp).
diff --git a/kvx/abstractbb/Impure/ImpExtern.v b/kvx/abstractbb/Impure/ImpExtern.v
deleted file mode 100644
index 8fb3cf3b..00000000
--- a/kvx/abstractbb/Impure/ImpExtern.v
+++ /dev/null
@@ -1,7 +0,0 @@
-(** Exporting Extern functions
-*)
-
-Require Export ImpPrelude.
-Require Export ImpIO.
-Require Export ImpLoops.
-Require Export ImpHCons.
diff --git a/kvx/abstractbb/Impure/ImpHCons.v b/kvx/abstractbb/Impure/ImpHCons.v
deleted file mode 100644
index 637116cc..00000000
--- a/kvx/abstractbb/Impure/ImpHCons.v
+++ /dev/null
@@ -1,199 +0,0 @@
-Require Export ImpIO.
-
-Import Notations.
-Local Open Scope impure.
-
-
-Axiom string_of_hashcode: hashcode -> ?? caml_string.
-Extract Constant string_of_hashcode => "string_of_int".
-
-Axiom hash: forall {A}, A -> ?? hashcode.
-Extract Constant hash => "Hashtbl.hash".
-
-(**************************)
-(* (Weak) Sets *)
-
-
-Import Dict.
-
-Axiom make_dict: forall {A B}, (hash_params A) -> ?? Dict.t A B.
-Extract Constant make_dict => "ImpHConsOracles.make_dict".
-
-
-Module Sets.
-
-Definition t {A} (mod: A -> Prop) := Dict.t A {x | mod x}.
-
-Definition empty {A} (hp: hash_params A) {mod:A -> Prop}: ?? t mod :=
- make_dict hp.
-
-Program Fixpoint add {A} (l: list A) {mod: A -> Prop} (d: t mod): forall {H:forall x, List.In x l -> mod x}, ?? unit :=
- match l with
- | nil => fun H => RET ()
- | x::l' => fun H =>
- d.(set)(x,x);;
- add l' d
- end.
-
-Program Definition create {A} (hp: hash_params A) (l:list A): ?? t (fun x => List.In x l) :=
- DO d <~ empty hp (mod:=fun x => List.In x l);;
- add l (mod:=fun x => List.In x l) d (H:=_);;
- RET d.
-Global Opaque create.
-
-Definition is_present {A} (hp: hash_params A) (x:A) {mod} (d:t mod): ?? bool :=
- DO oy <~ (d.(get)) x;;
- match oy with
- | Some y => hp.(test_eq) x (`y)
- | None => RET false
- end.
-
-Local Hint Resolve test_eq_correct: wlp.
-
-Lemma is_present_correct A (hp: hash_params A) x mod (d:t mod):
- WHEN is_present hp x d ~> b THEN b=true -> mod x.
-Proof.
- wlp_simplify; subst; eauto.
- - apply proj2_sig.
- - discriminate.
-Qed.
-Hint Resolve is_present_correct: wlp.
-Global Opaque is_present.
-
-Definition msg_assert_incl: pstring := "Sets.assert_incl".
-
-Fixpoint assert_incl {A} (hp: hash_params A) (l: list A) {mod} (d:t mod): ?? unit :=
- match l with
- | nil => RET ()
- | x::l' =>
- DO b <~ is_present hp x d;;
- if b then
- assert_incl hp l' d
- else (
- hp.(log) x;;
- FAILWITH msg_assert_incl
- )
- end.
-
-Lemma assert_incl_correct A (hp: hash_params A) l mod (d:t mod):
- WHEN assert_incl hp l d ~> _ THEN forall x, List.In x l -> mod x.
-Proof.
- induction l; wlp_simplify; subst; eauto.
-Qed.
-Hint Resolve assert_incl_correct: wlp.
-Global Opaque assert_incl.
-
-Definition assert_list_incl {A} (hp: hash_params A) (l1 l2: list A): ?? unit :=
- (* println "";;print("dict_create ");;*)
- DO d <~ create hp l2;;
- (*print("assert_incl ");;*)
- assert_incl hp l1 d.
-
-Lemma assert_list_incl_correct A (hp: hash_params A) l1 l2:
- WHEN assert_list_incl hp l1 l2 ~> _ THEN List.incl l1 l2.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque assert_list_incl.
-Hint Resolve assert_list_incl_correct: wlp.
-
-End Sets.
-
-
-
-
-(********************************)
-(* (Weak) HConsing *)
-
-Module HConsing.
-
-Export HConsingDefs.
-
-(* NB: this axiom is NOT intended to be called directly, but only through [hCons...] functions below. *)
-Axiom xhCons: forall {A}, (hashP A) -> ?? hashConsing A.
-Extract Constant xhCons => "ImpHConsOracles.xhCons".
-
-Definition hCons_eq_msg: pstring := "xhCons: hash eq differs".
-
-Definition hCons {A} (hp: hashP A): ?? (hashConsing A) :=
- DO hco <~ xhCons hp ;;
- RET {|
- hC := (fun x =>
- DO x' <~ hC hco x ;;
- DO b0 <~ hash_eq hp x.(hdata) x' ;;
- assert_b b0 hCons_eq_msg;;
- RET x');
- next_hid := hco.(next_hid);
- next_log := hco.(next_log);
- export := hco.(export);
- remove := hco.(remove)
- |}.
-
-
-Lemma hCons_correct A (hp: hashP A):
- WHEN hCons hp ~> hco THEN
- (forall x y, WHEN hp.(hash_eq) x y ~> b THEN b=true -> (ignore_hid hp x)=(ignore_hid hp y)) ->
- forall x, WHEN hco.(hC) x ~> x' THEN ignore_hid hp x.(hdata)=ignore_hid hp x'.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque hCons.
-Hint Resolve hCons_correct: wlp.
-
-
-
-(* hashV: extending a given type with hash-consing *)
-Record hashV {A:Type}:= {
- data: A;
- hid: hashcode
-}.
-Arguments hashV: clear implicits.
-
-Definition hashV_C {A} (test_eq: A -> A -> ?? bool) : hashP (hashV A) := {|
- hash_eq := fun v1 v2 => test_eq v1.(data) v2.(data);
- get_hid := hid;
- set_hid := fun v id => {| data := v.(data); hid := id |}
-|}.
-
-Definition liftHV (x:nat) := {| data := x; hid := unknown_hid |}.
-
-Definition hConsV {A} (hasheq: A -> A -> ?? bool): ?? (hashConsing (hashV A)) :=
- hCons (hashV_C hasheq).
-
-Lemma hConsV_correct A (hasheq: A -> A -> ?? bool):
- WHEN hConsV hasheq ~> hco THEN
- (forall x y, WHEN hasheq x y ~> b THEN b=true -> x=y) ->
- forall x, WHEN hco.(hC) x ~> x' THEN x.(hdata).(data)=x'.(data).
-Proof.
- Local Hint Resolve f_equal2: core.
- wlp_simplify.
- exploit H; eauto.
- + wlp_simplify.
- + intros; congruence.
-Qed.
-Global Opaque hConsV.
-Hint Resolve hConsV_correct: wlp.
-
-Definition hC_known {A} (hco:hashConsing (hashV A)) (unknownHash_msg: hashinfo (hashV A) -> ?? pstring) (x:hashinfo (hashV A)): ?? hashV A :=
- DO clock <~ hco.(next_hid)();;
- DO x' <~ hco.(hC) x;;
- DO ok <~ hash_older x'.(hid) clock;;
- if ok
- then RET x'
- else
- hco.(remove) x;;
- DO msg <~ unknownHash_msg x;;
- FAILWITH msg.
-
-Lemma hC_known_correct A (hco:hashConsing (hashV A)) msg x:
- WHEN hC_known hco msg x ~> x' THEN
- (forall x, WHEN hco.(hC) x ~> x' THEN x.(hdata).(data)=x'.(data)) ->
- x.(hdata).(data)=x'.(data).
-Proof.
- wlp_simplify.
- unfold wlp in * |- ; eauto.
-Qed.
-Global Opaque hC_known.
-Hint Resolve hC_known_correct: wlp.
-
-End HConsing.
diff --git a/kvx/abstractbb/Impure/ImpIO.v b/kvx/abstractbb/Impure/ImpIO.v
deleted file mode 100644
index 6c02c395..00000000
--- a/kvx/abstractbb/Impure/ImpIO.v
+++ /dev/null
@@ -1,159 +0,0 @@
-(** Extension of Coq language with some IO and exception-handling operators.
-
-TODO: integration with http://coq.io/ ?
-
-*)
-
-Require Export ImpPrelude.
-
-Import Notations.
-Local Open Scope impure.
-
-(** Printing functions *)
-
-Axiom print: pstring -> ?? unit.
-Extract Constant print => "ImpIOOracles.print".
-
-Axiom println: pstring -> ?? unit.
-Extract Constant println => "ImpIOOracles.println".
-
-Axiom read_line: unit -> ?? pstring.
-Extract Constant read_line => "ImpIOOracles.read_line".
-
-Require Import ZArith.
-Axiom string_of_Z: Z -> ?? pstring.
-Extract Constant string_of_Z => "ImpIOOracles.string_of_Z".
-
-(** timer *)
-
-Axiom timer: forall {A B}, (A -> ?? B)*A -> ?? B.
-Extract Constant timer => "ImpIOOracles.timer".
-
-(** Exception Handling *)
-
-Axiom exit_observer: Type.
-Extract Constant exit_observer => "((unit -> unit) ref)".
-
-Axiom new_exit_observer: (unit -> ??unit) -> ??exit_observer.
-Extract Constant new_exit_observer => "ImpIOOracles.new_exit_observer".
-
-Axiom set_exit_observer: exit_observer * (unit -> ??unit) -> ??unit.
-Extract Constant set_exit_observer => "ImpIOOracles.set_exit_observer".
-
-Axiom exn: Type.
-Extract Inlined Constant exn => "exn".
-
-Axiom raise: forall {A}, exn -> ?? A.
-Extract Constant raise => "raise".
-
-Axiom exn2string: exn -> ?? pstring.
-Extract Constant exn2string => "ImpIOOracles.exn2string".
-
-Axiom fail: forall {A}, pstring -> ?? A.
-Extract Constant fail => "ImpIOOracles.fail".
-
-Axiom try_with_fail: forall {A}, (unit -> ?? A) * (pstring -> exn -> ??A) -> ??A.
-Extract Constant try_with_fail => "ImpIOOracles.try_with_fail".
-
-Axiom try_with_any: forall {A}, (unit -> ?? A) * (exn -> ??A) -> ??A.
-Extract Constant try_with_any => "ImpIOOracles.try_with_any".
-
-Notation "'RAISE' e" := (DO r <~ raise (A:=False) e ;; RET (match r with end)) (at level 0): impure_scope.
-Notation "'FAILWITH' msg" := (DO r <~ fail (A:=False) msg ;; RET (match r with end)) (at level 0): impure_scope.
-
-Definition _FAILWITH {A:Type} msg: ?? A := FAILWITH msg.
-
-Example _FAILWITH_correct A msg (P: A -> Prop):
- WHEN _FAILWITH msg ~> r THEN P r.
-Proof.
- wlp_simplify.
-Qed.
-
-Notation "'TRY' k1 'WITH_FAIL' s ',' e '=>' k2" := (try_with_fail (fun _ => k1, fun s e => k2))
- (at level 55, k1 at level 53, right associativity): impure_scope.
-
-Notation "'TRY' k1 'WITH_ANY' e '=>' k2" := (try_with_any (fun _ => k1, fun e => k2))
- (at level 55, k1 at level 53, right associativity): impure_scope.
-
-
-Program Definition assert_b (b: bool) (msg: pstring): ?? b=true :=
- match b with
- | true => RET _
- | false => FAILWITH msg
- end.
-
-Lemma assert_wlp_true msg b: WHEN assert_b b msg ~> _ THEN b=true.
-Proof.
- wlp_simplify.
-Qed.
-
-Lemma assert_false_wlp msg (P: Prop): WHEN assert_b false msg ~> _ THEN P.
-Proof.
- simpl; wlp_simplify.
-Qed.
-
-Program Definition try_catch_fail_ensure {A} (k1: unit -> ?? A) (k2: pstring -> exn -> ??A) (P: A -> Prop | wlp (k1 tt) P /\ (forall s e, wlp (k2 s e) P)): ?? { r | P r }
- := TRY
- DO r <~ mk_annot (k1 tt);;
- RET (exist P r _)
- WITH_FAIL s, e =>
- DO r <~ mk_annot (k2 s e);;
- RET (exist P r _).
-Obligation 2.
- unfold wlp in * |- *; eauto.
-Qed.
-
-Notation "'TRY' k1 'CATCH_FAIL' s ',' e '=>' k2 'ENSURE' P" := (try_catch_fail_ensure (fun _ => k1) (fun s e => k2) (exist _ P _))
- (at level 55, k1 at level 53, right associativity): impure_scope.
-
-Definition is_try_post {A} (P: A -> Prop) k1 k2 : Prop :=
- wlp (k1 ()) P /\ forall (e:exn), wlp (k2 e) P.
-
-Program Definition try_catch_ensure {A} k1 k2 (P:A->Prop|is_try_post P k1 k2): ?? { r | P r }
- := TRY
- DO r <~ mk_annot (k1 ());;
- RET (exist P r _)
- WITH_ANY e =>
- DO r <~ mk_annot (k2 e);;
- RET (exist P r _).
-Obligation 1.
- unfold is_try_post, wlp in * |- *; intuition eauto.
-Qed.
-Obligation 2.
- unfold is_try_post, wlp in * |- *; intuition eauto.
-Qed.
-
-Notation "'TRY' k1 'CATCH' e '=>' k2 'ENSURE' P" := (try_catch_ensure (fun _ => k1) (fun e => k2) (exist _ P _))
- (at level 55, k1 at level 53, right associativity): impure_scope.
-
-
-Program Example tryex {A} (x y:A) :=
- TRY (RET x)
- CATCH _ => (RET y)
- ENSURE (fun r => r = x \/ r = y).
-Obligation 1.
- split; wlp_simplify.
-Qed.
-
-Program Example tryex_test {A} (x y:A):
- WHEN tryex x y ~> r THEN `r <> x -> `r = y.
-Proof.
- wlp_simplify. destruct exta as [r [X|X]]; intuition.
-Qed.
-
-
-Program Example try_branch1 {A} (x:A): ?? { r | r = x} :=
- TRY (RET x)
- CATCH e => (FAILWITH "!")
- ENSURE _.
-Obligation 1.
- split; wlp_simplify.
-Qed.
-
-Program Example try_branch2 {A} (x:A): ?? { r | r = x} :=
- TRY (FAILWITH "!")
- CATCH e => (RET x)
- ENSURE _.
-Obligation 1.
- split; wlp_simplify.
-Qed.
diff --git a/kvx/abstractbb/Impure/ImpLoops.v b/kvx/abstractbb/Impure/ImpLoops.v
deleted file mode 100644
index 33376c19..00000000
--- a/kvx/abstractbb/Impure/ImpLoops.v
+++ /dev/null
@@ -1,123 +0,0 @@
-(** Extension of Coq language with generic loops. *)
-
-Require Export ImpIO.
-
-Import Notations.
-Local Open Scope impure.
-
-
-(** While-loop iterations *)
-
-Axiom loop: forall {A B}, A * (A -> ?? (A+B)) -> ?? B.
-Extract Constant loop => "ImpLoopOracles.loop".
-
-
-Section While_Loop.
-
-(** Local Definition of "while-loop-invariant" *)
-Let wli {S} cond body (I: S -> Prop) := forall s, I s -> cond s = true -> WHEN (body s) ~> s' THEN I s'.
-
-Program Definition while {S} cond body (I: S -> Prop | wli cond body I) s0: ?? {s | (I s0 -> I s) /\ cond s = false}
- := loop (A:={s | I s0 -> I s})
- (s0,
- fun s =>
- match (cond s) with
- | true =>
- DO s' <~ mk_annot (body s) ;;
- RET (inl (A:={s | I s0 -> I s }) s')
- | false =>
- RET (inr (B:={s | (I s0 -> I s) /\ cond s = false}) s)
- end).
-Obligation 2.
- unfold wli, wlp in * |-; eauto.
-Qed.
-Extraction Inline while.
-
-End While_Loop.
-
-
-Section Loop_Until_None.
-(** useful to demonstrate a UNSAT property *)
-
-(** Local Definition of "loop-until-None-invariant" *)
-Let luni {S} (body: S -> ?? (option S)) (I: S -> Prop) := forall s, I s -> WHEN (body s) ~> s' THEN match s' with Some s1 => I s1 | None => False end.
-
-Program Definition loop_until_None {S} body (I: S -> Prop | luni body I) s0: ?? ~(I s0)
- := loop (A:={s | I s0 -> I s})
- (s0,
- fun s =>
- DO s' <~ mk_annot (body s) ;;
- match s' with
- | Some s1 => RET (inl (A:={s | I s0 -> I s }) s1)
- | None => RET (inr (B:=~(I s0)) _)
- end).
-Obligation 2.
- refine (H2 s _ _ H0). auto.
-Qed.
-Obligation 3.
- intros X; refine (H1 s _ _ H). auto.
-Qed.
-Extraction Inline loop_until_None.
-
-End Loop_Until_None.
-
-
-(*********************************************)
-(* A generic fixpoint from an equality test *)
-
-Record answ {A B: Type} {R: A -> B -> Prop} := {
- input: A ;
- output: B ;
- correct: R input output
-}.
-Arguments answ {A B}.
-
-Definition msg: pstring := "wapply fails".
-
-Definition beq_correct {A} (beq: A -> A -> ?? bool) :=
- forall x y, WHEN beq x y ~> b THEN b=true -> x=y.
-
-Definition wapply {A B} {R: A -> B -> Prop} (beq: A -> A -> ?? bool) (k: A -> ?? answ R) (x:A): ?? B :=
- DO a <~ k x;;
- DO b <~ beq x (input a) ;;
- assert_b b msg;;
- RET (output a).
-
-Lemma wapply_correct A B (R: A -> B -> Prop) (beq: A -> A -> ?? bool) (k: A -> ?? answ R) x:
- beq_correct beq
- -> WHEN wapply beq k x ~> y THEN R x y.
-Proof.
- unfold beq_correct; wlp_simplify.
- destruct exta; simpl; auto.
-Qed.
-Local Hint Resolve wapply_correct: wlp.
-Global Opaque wapply.
-
-Axiom xrec_set_option: recMode -> ?? unit.
-Extract Constant xrec_set_option => "ImpLoopOracles.xrec_set_option".
-
-(* TODO: generalizaton to get beq (and a Hash function ?) in parameters ? *)
-Axiom xrec: forall {A B}, ((A -> ?? B) -> A -> ?? B) -> ?? (A -> ?? B).
-Extract Constant xrec => "ImpLoopOracles.xrec".
-
-Definition rec_preserv {A B} (recF: (A -> ?? B) -> A -> ?? B) (R: A -> B -> Prop) :=
- forall f x, WHEN recF f x ~> z THEN (forall x', WHEN f x' ~> y THEN R x' y) -> R x z.
-
-
-Program Definition rec {A B} beq recF (R: A -> B -> Prop) (H1: rec_preserv recF R) (H2: beq_correct beq): ?? (A -> ?? B) :=
- DO f <~ xrec (B:=answ R) (fun f x =>
- DO y <~ mk_annot (recF (wapply beq f) x) ;;
- RET {| input := x; output := `y |});;
- RET (wapply beq f).
-Obligation 1.
- eapply H1; eauto. clear H H1.
- wlp_simplify.
-Qed.
-
-Lemma rec_correct A B beq recF (R: A -> B -> Prop) (H1: rec_preserv recF R) (H2: beq_correct beq):
- WHEN rec beq recF R H1 H2 ~> f THEN forall x, WHEN f x ~> y THEN R x y.
-Proof.
- wlp_simplify.
-Qed.
-Hint Resolve rec_correct: wlp.
-Global Opaque rec.
diff --git a/kvx/abstractbb/Impure/ImpMonads.v b/kvx/abstractbb/Impure/ImpMonads.v
deleted file mode 100644
index f01a2755..00000000
--- a/kvx/abstractbb/Impure/ImpMonads.v
+++ /dev/null
@@ -1,148 +0,0 @@
-(** Impure monad for interface with impure code
-*)
-
-
-Require Import Program.
-
-
-Module Type MayReturnMonad.
-
- Axiom t: Type -> Type.
-
- Axiom mayRet: forall {A:Type}, t A -> A -> Prop.
-
- Axiom ret: forall {A}, A -> t A.
-
- Axiom bind: forall {A B}, (t A) -> (A -> t B) -> t B.
-
- Axiom mk_annot: forall {A} (k: t A), t { a: A | mayRet k a }.
-
- Axiom mayRet_ret: forall A (a b:A),
- mayRet (ret a) b -> a=b.
-
- Axiom mayRet_bind: forall A B k1 k2 (b:B),
- mayRet (bind k1 k2) b -> exists a:A, mayRet k1 a /\ mayRet (k2 a) b.
-
-End MayReturnMonad.
-
-
-
-(** Model of impure computation as predicate *)
-Module PowerSetMonad<: MayReturnMonad.
-
- Definition t (A:Type) := A -> Prop.
-
- Definition mayRet {A:Type} (k: t A) a: Prop := k a.
-
- Definition ret {A:Type} (a:A) := eq a.
-
- Definition bind {A B:Type} (k1: t A) (k2: A -> t B) :=
- fun b => exists a, k1 a /\ k2 a b.
-
- Definition mk_annot {A} (k: t A) : t { a | mayRet k a } := fun _ => True.
-
- Lemma mayRet_ret A (a b:A): mayRet (ret a) b -> a=b.
- Proof.
- unfold mayRet, ret. firstorder.
- Qed.
-
- Lemma mayRet_bind A B k1 k2 (b:B):
- mayRet (bind k1 k2) b -> exists (a:A), mayRet k1 a /\ mayRet (k2 a) b.
- Proof.
- unfold mayRet, bind.
- firstorder.
- Qed.
-
-End PowerSetMonad.
-
-
-(** The identity interpretation *)
-Module IdentityMonad<: MayReturnMonad.
-
- Definition t (A:Type) := A.
-
- (* may-return semantics of computations *)
- Definition mayRet {A:Type} (a b:A): Prop := a=b.
-
- Definition ret {A:Type} (a:A) := a.
-
- Definition bind {A B:Type} (k1: A) (k2: A -> B) := k2 k1.
-
- Definition mk_annot {A} (k: t A) : t { a: A | mayRet k a }
- := exist _ k (eq_refl k) .
-
- Lemma mayRet_ret (A:Type) (a b:A): mayRet (ret a) b -> a=b.
- Proof.
- intuition.
- Qed.
-
- Lemma mayRet_bind (A B:Type) (k1:t A) k2 (b:B):
- mayRet (bind k1 k2) b -> exists (a:A), mayRet k1 a /\ mayRet (k2 a) b.
- Proof.
- firstorder.
- Qed.
-
-End IdentityMonad.
-
-
-(** Model of impure computation as state-transformers *)
-Module StateMonad<: MayReturnMonad.
-
- Parameter St: Type. (* A global state *)
-
- Definition t (A:Type) := St -> A * St.
-
- Definition mayRet {A:Type} (k: t A) a: Prop :=
- exists s, fst (k s)=a.
-
- Definition ret {A:Type} (a:A) := fun (s:St) => (a,s).
-
- Definition bind {A B:Type} (k1: t A) (k2: A -> t B) :=
- fun s0 => let r := k1 s0 in k2 (fst r) (snd r).
-
- Program Definition mk_annot {A} (k: t A) : t { a | mayRet k a } :=
- fun s0 => let r := k s0 in (exist _ (fst r) _, snd r).
- Obligation 1.
- unfold mayRet; eauto.
- Qed.
-
- Lemma mayRet_ret {A:Type} (a b:A): mayRet (ret a) b -> a=b.
- Proof.
- unfold mayRet, ret. firstorder.
- Qed.
-
- Lemma mayRet_bind {A B:Type} k1 k2 (b:B):
- mayRet (bind k1 k2) b -> exists (a:A), mayRet k1 a /\ mayRet (k2 a) b.
- Proof.
- unfold mayRet, bind. firstorder eauto.
- Qed.
-
-End StateMonad.
-
-(** The deferred interpretation *)
-Module DeferredMonad<: MayReturnMonad.
-
- Definition t (A:Type) := unit -> A.
-
- (* may-return semantics of computations *)
- Definition mayRet {A:Type} (a: t A) (b:A): Prop := a tt=b.
-
- Definition ret {A:Type} (a:A) : t A := fun _ => a.
-
- Definition bind {A B:Type} (k1: t A) (k2: A -> t B) : t B := fun _ => k2 (k1 tt) tt.
-
- Definition mk_annot {A} (k: t A) : t { a: A | mayRet k a }
- := fun _ => exist _ (k tt) (eq_refl (k tt)).
-
- Lemma mayRet_ret (A:Type) (a b: A): mayRet (ret a) b -> a=b.
- Proof.
- intuition.
- Qed.
-
- Lemma mayRet_bind (A B:Type) (k1:t A) k2 (b:B):
- mayRet (bind k1 k2) b -> exists (a:A), mayRet k1 a /\ mayRet (k2 a) b.
- Proof.
- firstorder.
- Qed.
-
-End DeferredMonad.
diff --git a/kvx/abstractbb/Impure/ImpPrelude.v b/kvx/abstractbb/Impure/ImpPrelude.v
deleted file mode 100644
index de4c7973..00000000
--- a/kvx/abstractbb/Impure/ImpPrelude.v
+++ /dev/null
@@ -1,206 +0,0 @@
-Require Export String.
-Require Export List.
-Require Extraction.
-Require Import Ascii.
-Require Import BinNums.
-Require Export ImpCore.
-Require Export PArith.
-
-
-Import Notations.
-Local Open Scope impure.
-
-(** Impure lazy andb of booleans *)
-Definition iandb (k1 k2: ??bool): ?? bool :=
- DO r1 <~ k1 ;;
- if r1 then k2 else RET false.
-
-Extraction Inline iandb. (* Juste pour l'efficacité à l'extraction ! *)
-
-(** Strings for pretty-printing *)
-
-Axiom caml_string: Type.
-Extract Constant caml_string => "string".
-
-(* New line *)
-Definition nl: string := String (ascii_of_pos 10%positive) EmptyString.
-
-Inductive pstring: Type :=
- | Str: string -> pstring
- | CamlStr: caml_string -> pstring
- | Concat: pstring -> pstring -> pstring.
-
-Coercion Str: string >-> pstring.
-Bind Scope string_scope with pstring.
-
-Notation "x +; y" := (Concat x y)
- (at level 65, left associativity): string_scope.
-
-(** Coq references *)
-
-Record cref {A} := {
- set: A -> ?? unit;
- get: unit -> ?? A
-}.
-Arguments cref: clear implicits.
-
-Axiom make_cref: forall {A}, A -> ?? cref A.
-Extract Constant make_cref => "(fun x -> let r = ref x in { set = (fun y -> r:=y); get = (fun () -> !r) })".
-
-
-(** Data-structure for a logger *)
-
-Record logger {A:Type} := {
- log_insert: A -> ?? unit;
- log_info: unit -> ?? pstring;
-}.
-Arguments logger: clear implicits.
-
-Axiom count_logger: unit -> ?? logger unit.
-Extract Constant count_logger => "(fun () -> let count = ref 0 in { log_insert = (fun () -> count := !count + 1); log_info = (fun () -> (CamlStr (string_of_int !count))) })".
-
-
-(** Axioms of Physical equality *)
-
-Axiom phys_eq: forall {A}, A -> A -> ?? bool.
-
-Axiom phys_eq_correct: forall A (x y:A), WHEN phys_eq x y ~> b THEN b=true -> x=y.
-
-
-(* We only check here that above axioms are not trivially inconsistent...
- (but this does not prove the correctness of the extraction directive below).
- *)
-Module PhysEqModel.
-
-Definition phys_eq {A} (x y: A) := ret false.
-
-Lemma phys_eq_correct: forall A (x y:A), WHEN phys_eq x y ~> b THEN b=true -> x=y.
-Proof.
- wlp_simplify. discriminate.
-Qed.
-
-End PhysEqModel.
-
-Extract Inlined Constant phys_eq => "(==)".
-Hint Resolve phys_eq_correct: wlp.
-
-
-Axiom struct_eq: forall {A}, A -> A -> ?? bool.
-Axiom struct_eq_correct: forall A (x y:A), WHEN struct_eq x y ~> b THEN if b then x=y else x<>y.
-Extract Inlined Constant struct_eq => "(=)".
-Hint Resolve struct_eq_correct: wlp.
-
-
-(** Data-structure for generic hash-consing *)
-
-Axiom hashcode: Type.
-Extract Constant hashcode => "int".
-
-(* NB: hashConsing is assumed to generate hash-code in ascending order.
- This gives a way to check that a hash-consed value is older than an other one.
-*)
-Axiom hash_older: hashcode -> hashcode -> ?? bool.
-Extract Inlined Constant hash_older => "(<)".
-
-Module Dict.
-
-Record hash_params {A:Type} := {
- test_eq: A -> A -> ??bool;
- test_eq_correct: forall x y, WHEN test_eq x y ~> r THEN r=true -> x=y;
- hashing: A -> ??hashcode;
- log: A -> ??unit (* for debugging only *)
-}.
-Arguments hash_params: clear implicits.
-
-
-Record t {A B:Type} := {
- set: A * B -> ?? unit;
- get: A -> ?? option B
-}.
-Arguments t: clear implicits.
-
-End Dict.
-
-Module HConsingDefs.
-
-Record hashinfo {A: Type} := {
- hdata: A;
- hcodes: list hashcode;
-}.
-Arguments hashinfo: clear implicits.
-
-(* for inductive types with intrinsic hash-consing *)
-Record hashP {A:Type}:= {
- hash_eq: A -> A -> ?? bool;
- get_hid: A -> hashcode;
- set_hid: A -> hashcode -> A; (* WARNING: should only be used by hash-consing machinery *)
-}.
-Arguments hashP: clear implicits.
-
-Axiom unknown_hid: hashcode.
-Extract Constant unknown_hid => "-1".
-
-Definition ignore_hid {A} (hp: hashP A) (hv:A) := set_hid hp hv unknown_hid.
-
-Record hashExport {A:Type}:= {
- get_info: hashcode -> ?? hashinfo A;
- iterall: ((list pstring) -> hashcode -> hashinfo A -> ?? unit) -> ?? unit; (* iter on all elements in the hashtbl, by order of creation *)
-}.
-Arguments hashExport: clear implicits.
-
-Record hashConsing {A:Type}:= {
- hC: hashinfo A -> ?? A;
- (**** below: debugging or internal functions ****)
- next_hid: unit -> ?? hashcode; (* should be strictly less old than ignore_hid *)
- remove: hashinfo A -> ??unit; (* SHOULD NOT BE USED ! *)
- next_log: pstring -> ?? unit; (* insert a log info (for the next introduced element) -- regiven by [iterall export] below *)
- export: unit -> ?? hashExport A ;
-}.
-Arguments hashConsing: clear implicits.
-
-End HConsingDefs.
-
-(** recMode: this is mainly for Tests ! *)
-Inductive recMode:= StdRec | MemoRec | BareRec | BuggyRec.
-
-
-(* This a copy-paste from definitions in CompCert/Lib/CoqLib.v *)
-Lemma modusponens: forall (P Q: Prop), P -> (P -> Q) -> Q.
-Proof. auto. Qed.
-
-Ltac exploit x :=
- refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _ _) _)
- || refine (modusponens _ _ (x _ _ _) _)
- || refine (modusponens _ _ (x _ _) _)
- || refine (modusponens _ _ (x _) _).
diff --git a/kvx/abstractbb/Impure/LICENSE b/kvx/abstractbb/Impure/LICENSE
deleted file mode 100644
index 65c5ca88..00000000
--- a/kvx/abstractbb/Impure/LICENSE
+++ /dev/null
@@ -1,165 +0,0 @@
- GNU LESSER GENERAL PUBLIC LICENSE
- Version 3, 29 June 2007
-
- Copyright (C) 2007 Free Software Foundation, Inc. <http://fsf.org/>
- Everyone is permitted to copy and distribute verbatim copies
- of this license document, but changing it is not allowed.
-
-
- This version of the GNU Lesser General Public License incorporates
-the terms and conditions of version 3 of the GNU General Public
-License, supplemented by the additional permissions listed below.
-
- 0. Additional Definitions.
-
- As used herein, "this License" refers to version 3 of the GNU Lesser
-General Public License, and the "GNU GPL" refers to version 3 of the GNU
-General Public License.
-
- "The Library" refers to a covered work governed by this License,
-other than an Application or a Combined Work as defined below.
-
- An "Application" is any work that makes use of an interface provided
-by the Library, but which is not otherwise based on the Library.
-Defining a subclass of a class defined by the Library is deemed a mode
-of using an interface provided by the Library.
-
- A "Combined Work" is a work produced by combining or linking an
-Application with the Library. The particular version of the Library
-with which the Combined Work was made is also called the "Linked
-Version".
-
- The "Minimal Corresponding Source" for a Combined Work means the
-Corresponding Source for the Combined Work, excluding any source code
-for portions of the Combined Work that, considered in isolation, are
-based on the Application, and not on the Linked Version.
-
- The "Corresponding Application Code" for a Combined Work means the
-object code and/or source code for the Application, including any data
-and utility programs needed for reproducing the Combined Work from the
-Application, but excluding the System Libraries of the Combined Work.
-
- 1. Exception to Section 3 of the GNU GPL.
-
- You may convey a covered work under sections 3 and 4 of this License
-without being bound by section 3 of the GNU GPL.
-
- 2. Conveying Modified Versions.
-
- If you modify a copy of the Library, and, in your modifications, a
-facility refers to a function or data to be supplied by an Application
-that uses the facility (other than as an argument passed when the
-facility is invoked), then you may convey a copy of the modified
-version:
-
- a) under this License, provided that you make a good faith effort to
- ensure that, in the event an Application does not supply the
- function or data, the facility still operates, and performs
- whatever part of its purpose remains meaningful, or
-
- b) under the GNU GPL, with none of the additional permissions of
- this License applicable to that copy.
-
- 3. Object Code Incorporating Material from Library Header Files.
-
- The object code form of an Application may incorporate material from
-a header file that is part of the Library. You may convey such object
-code under terms of your choice, provided that, if the incorporated
-material is not limited to numerical parameters, data structure
-layouts and accessors, or small macros, inline functions and templates
-(ten or fewer lines in length), you do both of the following:
-
- a) Give prominent notice with each copy of the object code that the
- Library is used in it and that the Library and its use are
- covered by this License.
-
- b) Accompany the object code with a copy of the GNU GPL and this license
- document.
-
- 4. Combined Works.
-
- You may convey a Combined Work under terms of your choice that,
-taken together, effectively do not restrict modification of the
-portions of the Library contained in the Combined Work and reverse
-engineering for debugging such modifications, if you also do each of
-the following:
-
- a) Give prominent notice with each copy of the Combined Work that
- the Library is used in it and that the Library and its use are
- covered by this License.
-
- b) Accompany the Combined Work with a copy of the GNU GPL and this license
- document.
-
- c) For a Combined Work that displays copyright notices during
- execution, include the copyright notice for the Library among
- these notices, as well as a reference directing the user to the
- copies of the GNU GPL and this license document.
-
- d) Do one of the following:
-
- 0) Convey the Minimal Corresponding Source under the terms of this
- License, and the Corresponding Application Code in a form
- suitable for, and under terms that permit, the user to
- recombine or relink the Application with a modified version of
- the Linked Version to produce a modified Combined Work, in the
- manner specified by section 6 of the GNU GPL for conveying
- Corresponding Source.
-
- 1) Use a suitable shared library mechanism for linking with the
- Library. A suitable mechanism is one that (a) uses at run time
- a copy of the Library already present on the user's computer
- system, and (b) will operate properly with a modified version
- of the Library that is interface-compatible with the Linked
- Version.
-
- e) Provide Installation Information, but only if you would otherwise
- be required to provide such information under section 6 of the
- GNU GPL, and only to the extent that such information is
- necessary to install and execute a modified version of the
- Combined Work produced by recombining or relinking the
- Application with a modified version of the Linked Version. (If
- you use option 4d0, the Installation Information must accompany
- the Minimal Corresponding Source and Corresponding Application
- Code. If you use option 4d1, you must provide the Installation
- Information in the manner specified by section 6 of the GNU GPL
- for conveying Corresponding Source.)
-
- 5. Combined Libraries.
-
- You may place library facilities that are a work based on the
-Library side by side in a single library together with other library
-facilities that are not Applications and are not covered by this
-License, and convey such a combined library under terms of your
-choice, if you do both of the following:
-
- a) Accompany the combined library with a copy of the same work based
- on the Library, uncombined with any other library facilities,
- conveyed under the terms of this License.
-
- b) Give prominent notice with the combined library that part of it
- is a work based on the Library, and explaining where to find the
- accompanying uncombined form of the same work.
-
- 6. Revised Versions of the GNU Lesser General Public License.
-
- The Free Software Foundation may publish revised and/or new versions
-of the GNU Lesser General Public License from time to time. Such new
-versions will be similar in spirit to the present version, but may
-differ in detail to address new problems or concerns.
-
- Each version is given a distinguishing version number. If the
-Library as you received it specifies that a certain numbered version
-of the GNU Lesser General Public License "or any later version"
-applies to it, you have the option of following the terms and
-conditions either of that published version or of any later version
-published by the Free Software Foundation. If the Library as you
-received it does not specify a version number of the GNU Lesser
-General Public License, you may choose any version of the GNU Lesser
-General Public License ever published by the Free Software Foundation.
-
- If the Library as you received it specifies that a proxy can decide
-whether future versions of the GNU Lesser General Public License shall
-apply, that proxy's public statement of acceptance of any version is
-permanent authorization for you to choose that version for the
-Library.
diff --git a/kvx/abstractbb/Impure/README.md b/kvx/abstractbb/Impure/README.md
deleted file mode 100644
index 2b19d14a..00000000
--- a/kvx/abstractbb/Impure/README.md
+++ /dev/null
@@ -1,31 +0,0 @@
-# `Impure`: importing OCaml functions as non-deterministic ones.
-
-The principle of this library is to encode the type `A -> B` of an
-OCaml function as a type `A -> ?? B` in Coq, where `?? B` is the type
-of an axiomatized monad that can be interpreted as `B -> Prop`. In
-other word, this encoding abstracts an OCaml function as a function
-returning a postcondition on its possible results (ie a relation between its
-parameter and its result). Side-effects are simply ignored. And
-reasoning on such a function is only possible in partial correctness.
-
-See further explanations and examples on [ImpureDemo](https://github.com/boulme/ImpureDemo).
-
-## Credits
-
-[Sylvain Boulmé](mailto:Sylvain.Boulme@univ-grenoble-alpes.fr).
-
-## Code Overview
-
-- [ImpMonads](ImpMonads.v) axioms of "impure computations" and some Coq models of these axioms.
-
-- [ImpConfig](ImpConfig.v) declares the `Impure` monad and defines its extraction.
-
-- [ImpCore](ImpCore.v) defines notations for the `Impure` monad and a `wlp_simplify` tactic (to reason about `Impure` functions in a Hoare-logic style).
-
-- [ImpPrelude](ImpPrelude.v) declares the data types exchanged with `Impure` oracles.
-
-- [ImpIO](ImpIO.v), [ImpLoops](ImpLoops.v), [ImpHCons](ImpHCons.v) declare `Impure` oracles and define operators from these oracles.
- [ImpExtern](ImpExtern.v) exports all these impure operators.
-
-- [ocaml/](ocaml/) subdirectory containing the OCaml implementations of `Impure` oracles.
-
diff --git a/kvx/abstractbb/Impure/ocaml/ImpHConsOracles.ml b/kvx/abstractbb/Impure/ocaml/ImpHConsOracles.ml
deleted file mode 100644
index 68a33a91..00000000
--- a/kvx/abstractbb/Impure/ocaml/ImpHConsOracles.ml
+++ /dev/null
@@ -1,72 +0,0 @@
-open ImpPrelude
-open HConsingDefs
-
-let make_dict (type key) (p: key Dict.hash_params) =
- let module MyHashedType = struct
- type t = key
- let equal = p.Dict.test_eq
- let hash = p.Dict.hashing
- end in
- let module MyHashtbl = Hashtbl.Make(MyHashedType) in
- let dict = MyHashtbl.create 1000 in
- {
- Dict.set = (fun (k,d) -> MyHashtbl.replace dict k d);
- Dict.get = (fun k -> MyHashtbl.find_opt dict k)
- }
-
-
-exception Stop;;
-
-let xhCons (type a) (hp:a hashP) =
- (* We use a hash-table, but a hash-set would be sufficient ! *)
- (* Thus, we could use a weak hash-set, but prefer avoid it for easier debugging *)
- (* Ideally, a parameter would allow to select between the weak or full version *)
- let module MyHashedType = struct
- type t = a hashinfo
- let equal x y = hp.hash_eq x.hdata y.hdata
- let hash x = Hashtbl.hash x.hcodes
- end in
- let module MyHashtbl = Hashtbl.Make(MyHashedType) in
- let pick t =
- let res = ref None in
- try
- MyHashtbl.iter (fun k d -> res:=Some (k,d); raise Stop) t;
- None
- with
- | Stop -> !res
- in
- let t = MyHashtbl.create 1000 in
- let logs = ref [] in
- {
- hC = (fun (k:a hashinfo) ->
- (* DEBUG:
- Printf.printf "*in %d -- look for hcodes= " (Obj.magic t);
- List.iter (fun i -> Printf.printf "%d " i) k.hcodes;
- print_newline();
- *)
- match MyHashtbl.find_opt t k with
- | Some d -> d
- | None ->
- (* DEBUG: Printf.printf "*in %d -- new hid:%d" (Obj.magic t) (MyHashtbl.length t); print_newline(); *)
- let d = hp.set_hid k.hdata (MyHashtbl.length t) in
- MyHashtbl.add t {k with hdata = d } d; d);
- next_log = (fun info -> logs := (MyHashtbl.length t, info)::(!logs));
- next_hid = (fun () -> MyHashtbl.length t);
- remove = (fun (x:a hashinfo) -> MyHashtbl.remove t x);
- export = fun () ->
- match pick t with
- | None -> { get_info = (fun _ -> raise Not_found); iterall = (fun _ -> ()) }
- | Some (k,_) ->
- (* the state is fully copied at export ! *)
- let logs = ref (List.rev_append (!logs) []) in
- let rec step_log i =
- match !logs with
- | (j, info)::l' when i>=j -> logs:=l'; info::(step_log i)
- | _ -> []
- in let a = Array.make (MyHashtbl.length t) k in
- MyHashtbl.iter (fun k d -> a.(hp.get_hid d) <- k) t;
- {
- get_info = (fun i -> a.(i));
- iterall = (fun iter_node -> Array.iteri (fun i k -> iter_node (step_log i) i k) a)
- }
- }
diff --git a/kvx/abstractbb/Impure/ocaml/ImpHConsOracles.mli b/kvx/abstractbb/Impure/ocaml/ImpHConsOracles.mli
deleted file mode 100644
index 5075d176..00000000
--- a/kvx/abstractbb/Impure/ocaml/ImpHConsOracles.mli
+++ /dev/null
@@ -1,5 +0,0 @@
-open ImpPrelude
-open HConsingDefs
-
-val make_dict : 'a Dict.hash_params -> ('a, 'b) Dict.t
-val xhCons: 'a hashP -> 'a hashConsing
diff --git a/kvx/abstractbb/Impure/ocaml/ImpIOOracles.ml b/kvx/abstractbb/Impure/ocaml/ImpIOOracles.ml
deleted file mode 100644
index 9e63c12d..00000000
--- a/kvx/abstractbb/Impure/ocaml/ImpIOOracles.ml
+++ /dev/null
@@ -1,142 +0,0 @@
-(* Warning
-
-These oracles assumes the following extraction directives:
- "Require Import ExtrOcamlString."
-
-*)
-
-open ImpPrelude
-(*
-open BinNums
-open Datatypes
-*)
-
-(* two auxiliary functions, for efficient mapping of "int" to "BinNums.positive" *)
-exception Overflow
-
-let aux_add: ('a, 'b) Hashtbl.t -> 'b Queue.t -> 'a -> 'b -> unit
- = fun t q i p ->
- if i < 1 then (* protection against wrap around *)
- raise Overflow;
- Queue.add p q;
- Hashtbl.add t i p
-
-let memo_int2pos: int -> int -> BinNums.positive
- = fun n ->
- (* init of the Hashtbl *)
- let n = max n 1 in
- let t = Hashtbl.create n in
- let q = Queue.create () in
- aux_add t q 1 BinNums.Coq_xH ;
- for i = 1 to (n-1)/2 do
- let last = Queue.take q in
- let ni = 2*i in
- aux_add t q ni (BinNums.Coq_xO last);
- aux_add t q (ni+1) (BinNums.Coq_xI last)
- done;
- if n mod 2 = 0 then (
- let last = Queue.take q in
- Hashtbl.add t n (BinNums.Coq_xO last)
- );
- (* memoized translation of i *)
- let rec find i =
- try
- (* Printf.printf "-> %d\n" i; *)
- Hashtbl.find t i
- with Not_found ->
- (* Printf.printf "<- %d\n" i; *)
- if i <= 0 then
- invalid_arg "non-positive integer"
- else
- let p = find (i/2) in
- let pi = if i mod 2 = 0 then BinNums.Coq_xO p else BinNums.Coq_xI p in
- Hashtbl.add t i pi;
- pi
- in find;;
-
-let new_exit_observer: (unit -> unit) -> (unit -> unit) ref
- = fun f ->
- let o = ref f in
- at_exit (fun () -> !o());
- o;;
-
-let set_exit_observer: (unit -> unit) ref * (unit -> unit) -> unit
- = fun (r, f) -> r := f
-
-let rec print: pstring -> unit
- = fun ps ->
- match ps with
- | Str l -> List.iter print_char l
- | CamlStr s -> print_string s
- | Concat(ps1,ps2) -> (print ps1; print ps2);;
-
-let println: pstring -> unit
- = fun l -> print l; print_newline()
-
-let read_line () =
- CamlStr (Stdlib.read_line());;
-
-exception ImpureFail of pstring;;
-
-let exn2string: exn -> pstring
- = fun e -> CamlStr (Printexc.to_string e)
-
-let fail: pstring -> 'a
- = fun s -> raise (ImpureFail s);;
-
-let try_with_fail: (unit -> 'a) * (pstring -> exn -> 'a) -> 'a
- = fun (k1, k2) ->
- try
- k1()
- with
- | (ImpureFail s) as e -> k2 s e
-
-let try_with_any: (unit -> 'a) * (exn -> 'a) -> 'a
- = fun (k1, k2) ->
- try
- k1()
- with
- | e -> k2 e
-
-(** MISC **)
-
-let rec posTr: BinNums.positive -> int
-= function
- | BinNums.Coq_xH -> 1
- | BinNums.Coq_xO p -> (posTr p)*2
- | BinNums.Coq_xI p -> (posTr p)*2+1;;
-
-let zTr: BinNums.coq_Z -> int
-= function
- | BinNums.Z0 -> 0
- | BinNums.Zpos p -> posTr p
- | BinNums.Zneg p -> - (posTr p)
-
-let ten = BinNums.Zpos (BinNums.Coq_xO (BinNums.Coq_xI (BinNums.Coq_xO BinNums.Coq_xH)))
-
-let rec string_of_pos (p:BinNums.positive) (acc: pstring): pstring
-= let (q,r) = BinInt.Z.pos_div_eucl p ten in
- let acc0 = Concat (CamlStr (string_of_int (zTr r)), acc) in
- match q with
- | BinNums.Z0 -> acc0
- | BinNums.Zpos p0 -> string_of_pos p0 acc0
- | _ -> assert false
-
-(*
-let string_of_Z_debug: BinNums.coq_Z -> pstring
-= fun p -> CamlStr (string_of_int (zTr p))
-*)
-
-let string_of_Z: BinNums.coq_Z -> pstring
-= function
- | BinNums.Z0 -> CamlStr "0"
- | BinNums.Zpos p -> string_of_pos p (CamlStr "")
- | BinNums.Zneg p -> Concat (CamlStr "-", string_of_pos p (CamlStr ""))
-
-let timer ((f:'a -> 'b), (x:'a)) : 'b =
- Gc.compact();
- let itime = (Unix.times()).Unix.tms_utime in
- let r = f x in
- let rt = (Unix.times()).Unix.tms_utime -. itime in
- Printf.printf "time = %f\n" rt;
- r
diff --git a/kvx/abstractbb/Impure/ocaml/ImpIOOracles.mli b/kvx/abstractbb/Impure/ocaml/ImpIOOracles.mli
deleted file mode 100644
index 6064286a..00000000
--- a/kvx/abstractbb/Impure/ocaml/ImpIOOracles.mli
+++ /dev/null
@@ -1,33 +0,0 @@
-open ImpPrelude
-
-
-(*
-Memoized version of translation from int -> BinNums.positive.
-The first arg is an indicative bound on the max int translated:
-it pre-computes all positives lower or equal to this bound.
-*)
-val memo_int2pos: int -> int -> BinNums.positive
-
-val read_line: unit -> pstring
-
-val print: pstring -> unit
-
-val println: pstring -> unit
-
-val string_of_Z: BinNums.coq_Z -> pstring
-
-val timer : (('a -> 'b ) * 'a) -> 'b
-
-val new_exit_observer: (unit -> unit) -> (unit -> unit) ref
-
-val set_exit_observer: (unit -> unit) ref * (unit -> unit) -> unit
-
-val exn2string: exn -> pstring
-
-val fail: pstring -> 'a
-
-exception ImpureFail of pstring;;
-
-val try_with_fail: (unit -> 'a) * (pstring -> exn -> 'a) -> 'a
-
-val try_with_any: (unit -> 'a) * (exn -> 'a) -> 'a
diff --git a/kvx/abstractbb/Impure/ocaml/ImpLoopOracles.ml b/kvx/abstractbb/Impure/ocaml/ImpLoopOracles.ml
deleted file mode 100644
index cb7625e5..00000000
--- a/kvx/abstractbb/Impure/ocaml/ImpLoopOracles.ml
+++ /dev/null
@@ -1,78 +0,0 @@
-open ImpPrelude
-open Datatypes
-
-(** GENERIC ITERATIVE LOOP **)
-
-(* a simple version of loop *)
-let simple_loop: ('a * ('a -> ('a, 'b) sum)) -> 'b
- = fun (a0, f) ->
- let rec iter: 'a -> 'b
- = fun a ->
- match f a with
- | Coq_inl a' -> iter a'
- | Coq_inr b -> b
- in
- iter a0;;
-
-(* loop from while *)
-let while_loop: ('a * ('a -> ('a, 'b) sum)) -> 'b
- = fun (a0, f) ->
- let s = ref (f a0) in
- while (match !s with Coq_inl _ -> true | _ -> false) do
- match !s with
- | Coq_inl a -> s:=f a
- | _ -> assert false
- done;
- match !s with
- | Coq_inr b -> b
- | _ -> assert false;;
-
-let loop = simple_loop
-
-
-(** GENERIC FIXPOINTS **)
-
-let std_rec (recf: ('a -> 'b ) -> 'a -> 'b): 'a -> 'b =
- let rec f x = recf f x in
- f
-
-let memo_rec (recf: ('a -> 'b ) -> 'a -> 'b): 'a -> 'b =
- let memo = Hashtbl.create 10 in
- let rec f x =
- try
- Hashtbl.find memo x
- with
- Not_found ->
- let r = recf f x in
- Hashtbl.replace memo x r;
- r
- in f
-
-let bare_rec (recf: ('a -> 'b ) -> 'a -> 'b): 'a -> 'b =
- let fix = ref (fun x -> failwith "init") in
- fix := (fun x -> recf !fix x);
- !fix;;
-
-let buggy_rec (recf: ('a -> 'b ) -> 'a -> 'b): 'a -> 'b =
- let memo = ref None in
- let rec f x =
- match !memo with
- | Some y -> y
- | None ->
- let r = recf f x in
- memo := Some r;
- r
- in f
-
-let xrec_mode = ref MemoRec
-
-let xrec_set_option : recMode -> unit
-= fun m -> xrec_mode := m
-
-let xrec : (('a -> 'b ) -> 'a -> 'b ) -> ('a -> 'b )
- = fun recf ->
- match !xrec_mode with
- | StdRec -> std_rec recf
- | MemoRec -> memo_rec recf
- | BareRec -> bare_rec recf
- | BuggyRec -> buggy_rec recf
diff --git a/kvx/abstractbb/Impure/ocaml/ImpLoopOracles.mli b/kvx/abstractbb/Impure/ocaml/ImpLoopOracles.mli
deleted file mode 100644
index 194696a1..00000000
--- a/kvx/abstractbb/Impure/ocaml/ImpLoopOracles.mli
+++ /dev/null
@@ -1,8 +0,0 @@
-open ImpPrelude
-open Datatypes
-
-val loop: ('a * ('a -> ('a, 'b) sum)) -> 'b
-
-val xrec_set_option: recMode -> unit
-
-val xrec: (('a -> 'b ) -> 'a -> 'b ) -> ('a -> 'b )
diff --git a/kvx/abstractbb/Parallelizability.v b/kvx/abstractbb/Parallelizability.v
index feebeee5..e5d21434 100644
--- a/kvx/abstractbb/Parallelizability.v
+++ b/kvx/abstractbb/Parallelizability.v
@@ -26,7 +26,7 @@ Require Import Sorting.Permutation.
Require Import Bool.
Local Open Scope lazy_bool_scope.
-
+(** * Definition of the parallel semantics *)
Module ParallelSemantics (L: SeqLanguage).
Export L.
@@ -50,8 +50,8 @@ Fixpoint inst_prun (i: inst) (m tmp old: mem): option mem :=
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']; simpl; auto.
- intros m old; destruct (exp_eval ge e m old); simpl; auto.
+ induction i as [|[y e] i']; cbn; auto.
+ intros m old; destruct (exp_eval ge e m old); cbn; auto.
Qed.
@@ -76,8 +76,8 @@ 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']; simpl; eauto.
- intros m1 m2 tmp H; destruct (exp_eval ge e tmp old); simpl; auto.
+ 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.
@@ -85,12 +85,12 @@ 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']; simpl; eauto.
+ 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); simpl.
- + intros (m3 & H3 & H4); rewrite H3; simpl; eauto.
- + intros H1; rewrite H1; simpl; auto.
+ destruct (inst_prun i m1 old old); cbn.
+ + intros (m3 & H3 & H4); rewrite H3; cbn; eauto.
+ + intros H1; rewrite H1; cbn; auto.
Qed.
@@ -101,8 +101,8 @@ Lemma prun_iw_app p1: forall m1 old p2,
| None => None
end.
Proof.
- induction p1; simpl; try congruence.
- intros; destruct (inst_prun _ _ _); simpl; auto.
+ induction p1; cbn; try congruence.
+ intros; destruct (inst_prun _ _ _); cbn; auto.
Qed.
Lemma prun_iw_app_None p1: forall m1 old p2,
@@ -132,12 +132,12 @@ Fixpoint notIn {A} (x: A) (l:list A): Prop :=
Lemma notIn_iff A (x:A) l: (~List.In x l) <-> notIn x l.
Proof.
- induction l; simpl; intuition.
+ 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; simpl.
+ induction l1; cbn.
- intuition.
- intros; rewrite IHl1. intuition.
Qed.
@@ -145,7 +145,7 @@ Qed.
Lemma In_Permutation A (l1 l2: list A): Permutation l1 l2 -> forall x, In x l1 -> In x l2.
Proof.
- induction 1; simpl; intuition.
+ induction 1; cbn; intuition.
Qed.
Lemma Permutation_incl A (l1 l2: list A): Permutation l1 l2 -> incl l1 l2.
@@ -174,7 +174,7 @@ 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. simpl; intuition subst; auto.
+ 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).
@@ -230,13 +230,13 @@ Fixpoint frame_assign m1 (f: list R.t) m2 :=
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] ; simpl; auto.
- - intros; destruct (notIn_dec x []); simpl in *; tauto.
- - intros; rewrite IHf; destruct (notIn_dec x (y::f)); simpl in *.
- + destruct (notIn_dec x f); simpl in *; intuition.
+ 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); simpl in *; intuition subst.
+ + destruct (notIn_dec x f); cbn in *; intuition subst.
rewrite assign_eq; auto.
rewrite <- notIn_iff in *; intuition.
Qed.
@@ -266,7 +266,7 @@ Lemma frame_eq_list_split f1 (f2: R.t -> Prop) 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 | ]; simpl; auto.
+ unfold frame_eq; destruct om1 as [ m1 | ]; cbn; auto.
intros (m2 & H0 & H1); subst.
intros H.
eexists; intuition eauto.
@@ -280,7 +280,7 @@ Lemma frame_eq_res_eq f om1 om2:
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; simpl; firstorder.
+ clear H H0; unfold frame_eq, res_eq. destruct om1; cbn; firstorder.
Qed.
*)
@@ -296,9 +296,9 @@ 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']; simpl.
+ 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); simpl; try congruence.
+ - intros m tmp H x (H1 & H2); destruct (exp_eval ge e tmp old); cbn; try congruence.
cutrewrite (m x = assign m y v x); eauto.
rewrite assign_diff; auto.
Qed.
@@ -306,9 +306,9 @@ 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']; simpl.
+ 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); simpl; auto.
+ - 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.
@@ -323,7 +323,7 @@ Lemma inst_prun_None i m1 m2 tmp old:
inst_prun ge i m2 tmp old = None.
Proof.
intros H; generalize (inst_prun_fequiv i old m1 m2 tmp).
- rewrite H; simpl; auto.
+ rewrite H; cbn; auto.
Qed.
Lemma inst_prun_Some i m1 m2 tmp old m1':
@@ -331,7 +331,7 @@ Lemma inst_prun_Some i m1 m2 tmp old 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; simpl.
+ rewrite H; cbn.
intros (m2' & H1 & H2).
eexists; intuition eauto.
rewrite frame_assign_def.
@@ -351,7 +351,7 @@ Local Hint Resolve Permutation_app_head Permutation_app_tail Permutation_app_com
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]; simpl; auto.
+ induction 1 as [|i p p'|i1 i2 p|p1 p2 p3]; cbn; auto.
- rewrite! app_assoc; auto.
- eapply Permutation_trans; eauto.
Qed.
@@ -361,11 +361,11 @@ 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']; simpl.
+ 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|]; simpl.
+ destruct om as [m1|]; cbn.
+ eapply eq_trans.
eapply IHp'; eauto.
eapply inst_wframe_correct; eauto.
@@ -375,12 +375,12 @@ 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']; simpl.
+ 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'|]; simpl.
- + intros (m2' & H1 & H2). rewrite H1; simpl.
+ 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. }
@@ -389,7 +389,7 @@ Proof.
lapply (bblock_wframe_correct p' m2'' old m2'); eauto.
intros Xm2' Xm1'.
rewrite Xm1', Xm2'; auto.
- + intro H; rewrite H; simpl; auto.
+ + intro H; rewrite H; cbn; auto.
Qed.
Lemma prun_iw_equiv p m1 m2 old:
@@ -418,7 +418,7 @@ Fixpoint is_det (p: bblock): Prop :=
Lemma is_det_Permutation p p':
Permutation p p' -> is_det p -> is_det p'.
Proof.
- induction 1; simpl; auto.
+ induction 1; cbn; auto.
- intros; intuition. eapply disjoint_incl_r. 2: eauto.
eapply Permutation_incl. eapply Permutation_sym.
eapply bblock_wframe_Permutation; auto.
@@ -431,20 +431,20 @@ Theorem is_det_correct 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 ]; simpl; eauto.
+ 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 | ]; simpl; auto.
+ 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 | ]; simpl; auto.
+ destruct om2 as [ m2 | ]; cbn; auto.
+ remember (inst_prun ge i1 m old old) as om1.
- destruct om1 as [ m1 | ]; simpl; auto.
- * lapply (inst_prun_Some i2 m m1 old old m2); simpl; auto.
- lapply (inst_prun_Some i1 m m2 old old m1); simpl; auto.
+ 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'; simpl.
+ rewrite Hm1', Hm2'; cbn.
eapply prun_iw_equiv.
intros x; rewrite <- Xm1', <- Xm2'. clear Xm2' Xm1' Hm1' Hm2' m1' m2'.
rewrite frame_assign_def.
@@ -455,9 +455,9 @@ Proof.
erewrite (inst_wframe_correct i1 m1 old m old); eauto.
}
rewrite frame_assign_notIn; auto.
- * erewrite inst_prun_None; eauto. simpl; auto.
+ * erewrite inst_prun_None; eauto. cbn; auto.
+ remember (inst_prun ge i1 m old old) as om1.
- destruct om1 as [ m1 | ]; simpl; auto.
+ destruct om1 as [ m1 | ]; cbn; auto.
erewrite inst_prun_None; eauto.
- intros; eapply res_eq_trans.
eapply IHPermutation1; eauto.
@@ -486,7 +486,7 @@ Lemma exp_frame_correct e old1 old2:
(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)); simpl; auto.
+ (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;
@@ -501,7 +501,7 @@ Fixpoint inst_frame (i: inst): list R.t :=
Lemma inst_wframe_frame i x: In x (inst_wframe i) -> In x (inst_frame i).
Proof.
- induction i as [ | [y e] i']; simpl; intuition.
+ induction i as [ | [y e] i']; cbn; intuition.
Qed.
@@ -511,13 +511,13 @@ Lemma inst_frame_correct i wframe old1 old2: forall m tmp1 tmp2,
(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']; simpl; auto.
+ 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.
cutrewrite (exp_eval ge e tmp1 old1 = exp_eval ge e tmp2 old2).
- destruct (exp_eval ge e tmp2 old2); auto.
eapply IHi'; eauto.
- simpl; intros x0 H0; unfold assign. destruct (R.eq_dec x x0); simpl; auto.
+ cbn; intros x0 H0; unfold assign. destruct (R.eq_dec x x0); cbn; auto.
- unfold disjoint in H2; apply exp_frame_correct.
intros;apply H6; auto.
intros;apply H7; auto.
@@ -535,8 +535,8 @@ Fixpoint pararec (p: bblock) (wframe: list R.t): Prop :=
Lemma pararec_disjoint (p: bblock): forall wframe, pararec p wframe -> disjoint (bblock_wframe p) wframe.
Proof.
- induction p as [|i p']; simpl.
- - unfold disjoint; simpl; intuition.
+ 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.
@@ -546,7 +546,7 @@ Qed.
Lemma pararec_det p: forall wframe, pararec p wframe -> is_det p.
Proof.
- induction p as [|i p']; simpl; auto.
+ induction p as [|i p']; cbn; auto.
intros wframe [H0 H1]. generalize (pararec_disjoint _ _ H1). rewrite disjoint_app_r.
intuition.
- apply disjoint_sym; auto.
@@ -558,7 +558,7 @@ Lemma pararec_correct p old: forall wframe m,
(forall x, notIn x wframe -> m x = old x) ->
run ge p m = prun_iw ge p m old.
Proof.
- elim p; clear p; simpl; auto.
+ 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.
@@ -590,17 +590,17 @@ End PARALLELI.
End ParallelizablityChecking.
-Module Type PseudoRegSet.
-
-Declare Module R: PseudoRegisters.
-
-(** We assume a datatype [t] refining (list R.t)
+(** * 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.
@@ -646,7 +646,7 @@ Fixpoint inst_wsframe(i:inst): S.t :=
Lemma inst_wsframe_correct i: S.match_frame (inst_wsframe i) (inst_wframe i).
Proof.
- induction i; simpl; auto.
+ induction i; cbn; auto.
Qed.
Fixpoint exp_sframe (e: exp): S.t :=
@@ -664,7 +664,7 @@ with list_exp_sframe (le: list_exp): S.t :=
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)); simpl; auto.
+ 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 :=
@@ -677,7 +677,7 @@ 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']; simpl; auto.
+ induction i as [|[y e] i']; cbn; auto.
Qed.
Local Hint Resolve inst_wsframe_correct inst_sframe_correct: core.
@@ -692,7 +692,7 @@ Fixpoint is_pararec (p: bblock) (wsframe: S.t): bool :=
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; simpl; auto.
+ 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.
@@ -716,6 +716,11 @@ 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.
@@ -724,12 +729,6 @@ Module PosPseudoRegSet <: PseudoRegSet with Module R:=Pos.
Module R:=Pos.
-(** We assume a datatype [t] refining (list R.t)
-
-This data-refinement is given by an abstract "invariant" match_frame below,
-preserved by the following operations.
-
-*)
Definition t:=PositiveSet.t.
@@ -740,14 +739,14 @@ Definition empty:=PositiveSet.empty.
Lemma empty_match_frame: match_frame empty nil.
Proof.
- unfold match_frame, empty, PositiveSet.In; simpl; intuition.
+ 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; simpl.
+ unfold match_frame, add; cbn.
intros s x l H y. rewrite PositiveSet.add_spec, H.
intuition.
Qed.
@@ -773,13 +772,13 @@ Fixpoint is_disjoint (s s': PositiveSet.t) : bool :=
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]; simpl; try discriminate.
- destruct s' as [|l' o' r']; simpl; try discriminate.
+ 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'); simpl in X; intuition. }
+ { 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|]; simpl; eauto.
+ 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.
diff --git a/kvx/abstractbb/SeqSimuTheory.v b/kvx/abstractbb/SeqSimuTheory.v
index 61f8f2ec..df6b9963 100644
--- a/kvx/abstractbb/SeqSimuTheory.v
+++ b/kvx/abstractbb/SeqSimuTheory.v
@@ -55,13 +55,14 @@ with list_term_eval ge (l: list_term) (m: mem) {struct l}: option (list value) :
end
end.
-(* the symbolic memory:
- - pre: pre-condition expressing that the computation has not yet abort on a None.
- - post: the post-condition for each pseudo-register
-*)
-Record smem:= {pre: genv -> mem -> Prop; post:> R.t -> term}.
-
-(** initial symbolic memory *)
+(** 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 :=
@@ -78,11 +79,12 @@ with list_exp_term (le: list_exp) (d old: smem) : list_term :=
end.
-(** assignment of the symbolic memory *)
+(** 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.
@@ -90,13 +92,13 @@ 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; simpl; case (R.eq_dec x x); try congruence.
+ 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; simpl; case (R.eq_dec x y); try congruence.
+ unfold smem_set; cbn; case (R.eq_dec x y); try congruence.
Qed.
Fixpoint inst_smem (i: inst) (d old: smem): smem :=
@@ -121,15 +123,15 @@ Definition bblock_smem: bblock -> smem
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]; simpl; auto.
+ induction i as [|[y e] i IHi]; cbn; auto.
intros d a H; generalize (IHi _ _ H); clear H IHi.
- unfold smem_set; simpl; intuition.
+ 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']; simpl; eauto.
+ induction p as [|i p' IHp']; cbn; eauto.
intros d a H; eapply inst_smem_pre_monotonic; eauto.
Qed.
@@ -144,7 +146,7 @@ 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);
- simpl; auto.
+ cbn; auto.
- intros; erewrite IHe; eauto.
- intros. erewrite IHe, IHe0; eauto.
Qed.
@@ -154,12 +156,12 @@ Lemma inst_smem_abort i m0 x old: forall (d:smem),
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]; simpl; auto.
+ induction i as [|[y e] i IHi]; cbn; auto.
intros d VALID H; erewrite IHi; eauto. clear IHi.
- unfold smem_set; simpl; destruct (R.eq_dec y x); auto.
+ unfold smem_set; cbn; destruct (R.eq_dec y x); auto.
subst;
generalize (inst_smem_pre_monotonic _ _ _ _ VALID); clear VALID.
- unfold smem_set; simpl. intuition congruence.
+ unfold smem_set; cbn. intuition congruence.
Qed.
Lemma block_smem_rec_abort p m0 x: forall d,
@@ -167,7 +169,7 @@ Lemma block_smem_rec_abort p m0 x: forall d,
term_eval ge (d x) m0 = None ->
term_eval ge (bblock_smem_rec p d x) m0 = None.
Proof.
- induction p; simpl; auto.
+ induction p; cbn; auto.
intros d VALID H; erewrite IHp; eauto. clear IHp.
eapply inst_smem_abort; eauto.
Qed.
@@ -179,13 +181,13 @@ Lemma inst_smem_Some_correct1 i m0 old (od:smem):
(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]; simpl; intros m1 m2 d H.
+ 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; simpl. destruct (R.eq_dec x x0); auto.
+ unfold assign, smem_set; cbn. destruct (R.eq_dec x x0); auto.
subst; erewrite term_eval_exp; eauto.
Qed.
@@ -195,7 +197,7 @@ Lemma bblocks_smem_rec_Some_correct1 p m0: forall (m1 m2: mem) (d: smem),
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]; simpl; intros m1 m2 d H.
+ 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.
@@ -216,15 +218,15 @@ Lemma inst_smem_None_correct i m0 old (od: smem):
(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]; simpl; intros m1 d.
+ 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; simpl. destruct (R.eq_dec x x0); auto.
+ 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); simpl.
+ constructor 1 with (x:=x); cbn.
apply inst_smem_abort; auto.
rewrite set_spec_eq.
erewrite term_eval_exp; eauto.
@@ -239,14 +241,14 @@ Lemma inst_smem_Some_correct2 i m0 old (od: smem):
res_eq (Some m2) (inst_run ge i m1 old).
Proof.
intro X.
- induction i as [|[x e] i IHi]; simpl; intros m1 m2 d VALID H0.
+ 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; simpl; destruct (R.eq_dec x x0); auto.
+ 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.
@@ -260,7 +262,7 @@ Lemma bblocks_smem_rec_Some_correct2 p m0: forall (m1 m2: mem) d,
(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]; simpl; intros m1 m2 d VALID H0.
+ 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.
@@ -291,13 +293,13 @@ Lemma inst_valid i m0 old (od:smem):
(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]; simpl; auto.
+ 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; simpl; try congruence.
+ destruct (exp_eval ge e m1 old) eqn: Heq; cbn; try congruence.
eapply IHi; eauto.
- + unfold smem_set in * |- *; simpl.
+ + unfold smem_set in * |- *; cbn.
rewrite Hm1; intuition congruence.
- + intros x0. unfold assign, smem_set; simpl; destruct (R.eq_dec x x0); auto.
+ + intros x0. unfold assign, smem_set; cbn; destruct (R.eq_dec x x0); auto.
subst; erewrite term_eval_exp; eauto.
Qed.
@@ -309,7 +311,7 @@ Lemma block_smem_rec_valid p m0: forall (m1 m2: mem) (d:smem),
pre (bblock_smem_rec p d) ge m0.
Proof.
Local Hint Resolve inst_valid: core.
- induction p as [ | i p]; simpl; intros m1 d H; auto.
+ 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.
@@ -320,7 +322,7 @@ Lemma bblock_smem_valid p m0 m1:
pre (bblock_smem p) ge m0.
Proof.
intros; eapply block_smem_rec_valid; eauto.
- unfold smem_empty; simpl. auto.
+ 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.
@@ -337,7 +339,7 @@ Theorem bblock_smem_simu 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; simpl; try congruence.
+ 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.
@@ -369,21 +371,20 @@ Lemma smem_valid_set_proof d x t m:
Proof.
unfold smem_valid; intros (PRE & VALID) PREt. split.
+ split; auto.
- + intros x0; unfold smem_set; simpl; case (R.eq_dec x x0); intros; subst; auto.
+ + intros x0; unfold smem_set; cbn; case (R.eq_dec x x0); intros; subst; auto.
Qed.
End SIMU_THEORY.
-(** REMARKS: more abstract formulation of the proof...
- but relying on functional_extensionality.
+(** 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; simpl; intros m'; split.
+ unfold smem_correct; cbn; intros m'; split.
+ intros; split.
* eapply bblock_smem_valid; eauto.
* eapply bblock_smem_Some_correct1; eauto.
diff --git a/kvx/lib/ForwardSimulationBlock.v b/kvx/lib/ForwardSimulationBlock.v
index f79814f2..61466dad 100644
--- a/kvx/lib/ForwardSimulationBlock.v
+++ b/kvx/lib/ForwardSimulationBlock.v
@@ -42,11 +42,11 @@ Lemma starN_split 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; simpl.
+ induction 1; cbn.
+ intros m k H; assert (X: m=0); try omega.
assert (X0: k=0); try omega.
subst; repeat (eapply ex_intro); intuition eauto.
- + intros m; destruct m as [| m']; simpl.
+ + 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.
@@ -61,7 +61,7 @@ Lemma starN_tailstep 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; simpl.
+ induction 1; cbn.
+ intros t t1 s0; autorewrite with trace_rewrite.
intros; subst; eapply starN_step; eauto.
autorewrite with trace_rewrite; auto.
@@ -153,8 +153,8 @@ Definition head (s: memostate): state L1 :=
Lemma head_followed (s: memostate): follows_in_block (head s) (real s).
Proof.
- destruct s as [rs ms Hs]. simpl.
- destruct ms as [ms|]; unfold head; simpl; auto.
+ destruct s as [rs ms Hs]. cbn.
+ destruct ms as [ms|]; unfold head; cbn; auto.
constructor 1.
omega.
cutrewrite ((dist_end_block rs - dist_end_block rs)%nat=O).
@@ -198,21 +198,21 @@ Definition memoL1 := {|
Lemma discr_dist_end s:
{dist_end_block s = O} + {dist_end_block s <> O}.
Proof.
- destruct (dist_end_block s); simpl; intuition.
+ 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. simpl in H2; subst.
+ 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 |} _); simpl.
+ + refine (ex_intro _ {|real:=s1'; memorized:=None |} _); cbn.
intuition.
+ destruct ms2 as [s|].
- - refine (ex_intro _ {|real:=s1'; memorized:=Some s |} _); simpl.
+ - refine (ex_intro _ {|real:=s1'; memorized:=Some s |} _); cbn.
intuition.
- - refine (ex_intro _ {|real:=s1'; memorized:=Some rs2 |} _); simpl.
+ - refine (ex_intro _ {|real:=s1'; memorized:=Some rs2 |} _); cbn.
intuition.
Unshelve.
* intros; discriminate.
@@ -228,7 +228,7 @@ 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 |}); simpl.
+ + intros s1 H; eapply ex_intro with (x:={|real:=s1; memorized:=None |}); cbn.
intuition.
+ intros; subst; auto.
+ intros; exploit memo_simulation_step; eauto.
@@ -239,8 +239,8 @@ Qed.
Lemma forward_memo_simulation_2: forward_simulation memoL1 L2.
Proof.
- unfold memoL1; simpl.
- apply forward_simulation_opt with (measure:=fun s => dist_end_block (real s)) (match_states:=fun s1 s2 => match_states (head s1) s2); simpl; auto.
+ 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.
@@ -254,14 +254,14 @@ Proof.
- (* 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); simpl; auto. tauto.
+ 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].
cutrewrite (dist_end_block (head s1) - dist_end_block (real s1) = dist_end_block (head s1)) in H3; try omega.
eapply starN_tailstep; eauto.
- * unfold head; rewrite H2; simpl. intuition eauto.
+ * unfold head; rewrite H2; cbn. intuition eauto.
Qed.
Lemma forward_simulation_block_rel: forward_simulation L1 L2.
diff --git a/kvx/lib/Machblock.v b/kvx/lib/Machblock.v
index edae0ed4..404c2a96 100644
--- a/kvx/lib/Machblock.v
+++ b/kvx/lib/Machblock.v
@@ -70,7 +70,7 @@ Lemma bblock_eq:
b1 = b2.
Proof.
intros. destruct b1. destruct b2.
- simpl in *. subst. auto.
+ cbn in *. subst. auto.
Qed.
Definition length_opt {A} (o: option A) : nat :=
@@ -85,15 +85,15 @@ Lemma size_null b:
size b = 0%nat ->
header b = nil /\ body b = nil /\ exit b = None.
Proof.
- destruct b as [h b e]. simpl. unfold size. simpl.
+ destruct b as [h b e]. cbn. unfold size. cbn.
intros H.
assert (length h = 0%nat) as Hh; [ omega |].
assert (length b = 0%nat) as Hb; [ omega |].
assert (length_opt e = 0%nat) as He; [ omega|].
repeat split.
- destruct h; try (simpl in Hh; discriminate); auto.
- destruct b; try (simpl in Hb; discriminate); auto.
- destruct e; try (simpl in He; discriminate); auto.
+ 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 *)
@@ -127,13 +127,13 @@ Definition is_label (lbl: label) (bb: bblock) : bool :=
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)); simpl; intuition.
+ 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)); simpl; intuition.
+ unfold is_label; destruct (in_dec lbl (header bb)); cbn; intuition.
Qed.
diff --git a/kvx/lib/Machblockgen.v b/kvx/lib/Machblockgen.v
index ab186083..3d5d7b2c 100644
--- a/kvx/lib/Machblockgen.v
+++ b/kvx/lib/Machblockgen.v
@@ -148,11 +148,11 @@ 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]; simpl.
+ 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; simpl in * |- *; congruence.
+ + 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.
* cutrewrite (add_basic bi empty_bblock = add_to_new_bblock (MB_basic bi)); auto.
@@ -170,7 +170,7 @@ 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]; simpl; auto.
+ induction c1 as [| i c1]; cbn; auto.
Qed.
Lemma trans_code_is_trans_code c: is_trans_code c (trans_code c).
@@ -186,17 +186,17 @@ 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; simpl; auto.
+ + (* case Tr_end_block *) inversion H3; subst; cbn; auto.
* destruct (header bh); congruence.
- * destruct bl0; simpl; congruence.
- + (* case Tr_add_basic *) rewrite H3. simpl. 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]; simpl; eauto.
+ 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.
diff --git a/kvx/lib/Machblockgenproof.v b/kvx/lib/Machblockgenproof.v
index dfb97bfe..fc722887 100644
--- a/kvx/lib/Machblockgenproof.v
+++ b/kvx/lib/Machblockgenproof.v
@@ -146,16 +146,16 @@ 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; simpl; auto.
- unfold trans_stackframe. destruct s; simpl; auto.
+ 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; simpl; auto.
- unfold trans_stackframe. destruct s; simpl; auto.
+ unfold parent_ra. unfold Mach.parent_ra. destruct s; cbn; auto.
+ unfold trans_stackframe. destruct s; cbn; auto.
Qed.
Lemma external_call_preserved:
@@ -175,11 +175,11 @@ Proof.
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. simpl. rewrite peq_true.
+ 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. simpl. rewrite peq_false; 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:
@@ -192,24 +192,24 @@ Proof.
remember (is_label l _) as b.
cutrewrite (b = false); auto.
subst; unfold is_label.
- destruct i; simpl in * |- *; try (destruct (in_dec l nil); intuition).
+ destruct i; cbn in * |- *; try (destruct (in_dec l nil); intuition).
inversion H.
destruct (in_dec l (l0::nil)) as [H6|H6]; auto.
- simpl in H6; intuition try congruence.
+ 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; simpl; auto.
+ 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; simpl; auto.
+ unfold find_label; rewrite is_label_correct_false; intro H; rewrite H; cbn; auto.
Qed.
Definition concat (h: list label) (c: code): code :=
@@ -227,18 +227,18 @@ Proof.
rewrite <- is_trans_code_inv in * |-.
induction Heqbl.
+ (* Tr_nil *)
- intros; exists (l::nil); simpl in * |- *; intuition.
+ 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; simpl; auto.
+ - subst. rewrite find_label_at_begin; cbn; auto.
inversion H as [mbi H1 H2| | ].
subst.
inversion Heqbl.
subst.
- exists (l :: nil); simpl; eauto.
+ exists (l :: nil); cbn; eauto.
- exploit IHHeqbl; eauto.
destruct 1 as (h & H3 & H4).
exists h.
@@ -251,21 +251,21 @@ Proof.
- subst.
inversion H0 as [H1].
clear H0.
- erewrite find_label_at_begin; simpl; eauto.
+ erewrite find_label_at_begin; cbn; eauto.
subst_is_trans_code Heqbl.
- exists (l :: nil); simpl; eauto.
+ exists (l :: nil); cbn; eauto.
- subst; assert (H: l0 <> l); try congruence; clear H0.
exploit IHHeqbl; eauto.
clear IHHeqbl Heqbl.
intros (h & H3 & H4).
- simpl; unfold is_label, add_label; simpl.
- destruct (in_dec l (l0::header bh)) as [H5|H5]; simpl in H5.
+ 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); simpl; intuition.
+ 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'); simpl in * |- *;
- inversion H5; subst; simpl; auto.
+ destruct (trans_code c'); cbn in * |- *;
+ inversion H5; subst; cbn; auto.
* exists h. intuition.
erewrite <- find_label_add_label_diff; eauto.
+ (* Tr_add_basic *)
@@ -318,12 +318,12 @@ Local Hint Resolve exec_MBgetstack exec_MBsetstack exec_MBgetparam exec_MBop exe
Lemma size_add_label l bh: size (add_label l bh) = size bh + 1.
Proof.
- unfold add_label, size; simpl; omega.
+ unfold add_label, size; cbn; omega.
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; simpl. omega.
+ intro H. unfold add_basic, size; rewrite H; cbn. omega.
Qed.
@@ -418,8 +418,8 @@ Proof.
+ exists lbl.
unfold trans_inst in H1.
destruct i; congruence.
- + unfold add_basic in H; simpl in H; congruence.
- + unfold cfi_bblock in H; simpl in H; 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.
@@ -433,11 +433,11 @@ Proof.
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 * |- *); simpl in * |- *;
+ 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; simpl in Heqti; congruence. }
+ assert (X:i=Mlabel lbl). { destruct i; cbn in Heqti; congruence. }
subst. repeat econstructor; eauto.
+ (* label at mid block *)
exploit IHc; eauto.
@@ -451,12 +451,12 @@ Proof.
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; simpl; try congruence.
+ repeat econstructor; eauto; inversion H0; subst; repeat econstructor; cbn; try congruence.
* exploit (add_to_new_block_is_label i0); eauto.
- intros (l & H8); subst; simpl; congruence.
+ intros (l & H8); subst; cbn; congruence.
* exploit H3; eauto.
* exploit (add_to_new_block_is_label i0); eauto.
- intros (l & H8); subst; simpl; congruence.
+ intros (l & H8); subst; cbn; congruence.
+ (* basic at mid block *)
inversion H1; subst.
exploit IHc; eauto.
@@ -476,7 +476,7 @@ Lemma step_simu_header st f sp rs m s c h c' t:
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; simpl; intros hs; try (inversion hs; tauto).
+ 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.
@@ -487,21 +487,21 @@ Lemma step_simu_basic_step (i: Mach.instruction) (bi: basic_inst) (c: Mach.code)
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; simpl in * |-;
+ 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: simpl; rewrite <- parent_sp_preserved; auto.
- - eapply exec_MBop; eauto. rewrite <- H. destruct o; simpl; auto. destruct (rs ## l); simpl; auto.
+ 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; simpl; auto; destruct (rs ## l); simpl; 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; simpl; auto; destruct (rs ## l); simpl; 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; simpl; auto; destruct (rs ## l); simpl; 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; simpl; auto; destruct (rs ## l); simpl; 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.
@@ -511,7 +511,7 @@ Lemma star_step_simu_body_step s f sp c bdy c':
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; simpl.
+ induction 1; cbn.
+ intros. inversion H. exists rs. exists m. auto.
+ intros. inversion H0. exists rs. exists m. auto.
+ intros. inversion H1; subst.
@@ -531,15 +531,15 @@ Local Hint Resolve eval_builtin_args_preserved external_call_symbols_preserved f
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; simpl.
+ 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.
- simpl; constructor 1;
- intros X; inversion X as [d1 d2 d3 d4 d5 d6 d7 rs' m' d10 d11 X1 X2| | | ]; subst; simpl in * |- *;
- eapply exec_bblock; eauto; simpl;
+ 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.
@@ -551,7 +551,7 @@ Lemma step_simu_cfi_step (i: Mach.instruction) (cfi: control_flow_inst) (c: Mach
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; simpl in * |-;
+ 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.
@@ -561,8 +561,8 @@ Proof.
intuition auto.
eapply exec_MBtailcall;eauto.
- rewrite <-H; exploit (find_function_ptr_same); eauto.
- - simpl; rewrite <- parent_sp_preserved; auto.
- - simpl; rewrite <- parent_ra_preserved; auto.
+ - cbn; rewrite <- parent_sp_preserved; auto.
+ - cbn; rewrite <- parent_ra_preserved; auto.
* eapply ex_intro.
intuition auto.
eapply exec_MBbuiltin ;eauto.
@@ -605,7 +605,7 @@ Proof.
inversion H1; subst.
exploit (step_simu_cfi_step); eauto.
intros [s2 [Hcfi1 Hcfi3]].
- inversion H4. subst; simpl.
+ inversion H4. subst; cbn.
autorewrite with trace_rewrite.
exists s2.
split;eauto.
@@ -616,7 +616,7 @@ Lemma simu_end_block:
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; simpl.
+ destruct s1; cbn.
+ (* State *)
remember (trans_code _) as tc.
rewrite <- is_trans_code_inv in Heqtc.
@@ -624,7 +624,7 @@ Proof.
destruct tc as [|b bl].
{ (* nil => absurd *)
inversion Heqtc. subst.
- unfold dist_end_block_code; simpl.
+ unfold dist_end_block_code; cbn.
inversion_clear H;
inversion_clear H0.
}
@@ -659,7 +659,7 @@ Proof.
intros t s1' H; inversion_clear H.
eapply ex_intro; constructor 1; eauto.
inversion H1; subst; clear H1.
- inversion_clear H0; simpl.
+ inversion_clear H0; cbn.
- (* function_internal*)
cutrewrite (trans_code (Mach.fn_code f0) = fn_code (transf_function f0)); eauto.
eapply exec_function_internal; eauto.
@@ -674,7 +674,7 @@ Proof.
intros t s1' H; inversion_clear H.
eapply ex_intro; constructor 1; eauto.
inversion H1; subst; clear H1.
- inversion_clear H0; simpl.
+ inversion_clear H0; cbn.
eapply exec_return.
Qed.
@@ -685,10 +685,10 @@ dist_end_block_code (i :: c) = 0.
Proof.
unfold dist_end_block_code.
intro H. destruct H as [cfi H].
- destruct i;simpl in H;try(congruence); (
+ destruct i;cbn in H;try(congruence); (
remember (trans_code _) as bl;
rewrite <- is_trans_code_inv in Heqbl;
- inversion Heqbl; subst; simpl in * |- *; try (congruence)).
+ inversion Heqbl; subst; cbn in * |- *; try (congruence)).
Qed.
Theorem transf_program_correct:
@@ -697,23 +697,23 @@ 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; simpl in * |- *; omega || (intuition auto);
- destruct H2; eapply cfi_dist_end_block; simpl; eauto.
+ destruct H1; cbn in * |- *; omega || (intuition auto);
+ destruct H2; eapply cfi_dist_end_block; cbn; eauto.
(* public_preserved *)
- apply senv_preserved.
(* match_initial_states *)
- - intros. simpl.
+ - 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. simpl. destruct H. split with (r := r); auto.
+ - intros. cbn. destruct H. split with (r := r); auto.
(* final_states_end_block *)
- - intros. simpl in H0.
+ - intros. cbn in H0.
inversion H0.
- inversion H; simpl; auto.
+ inversion H; cbn; auto.
all: try (subst; discriminate).
apply cfi_dist_end_block; exists MBreturn; eauto.
(* simu_end_block *)
@@ -733,8 +733,8 @@ Proof.
intro H; destruct c as [|i' c]. { inversion H. }
remember (trans_inst i) as ti.
destruct ti as [lbl|bi|cfi].
- - (*i=lbl *) cutrewrite (i = Mlabel lbl). 2: ( destruct i; simpl in * |- *; try congruence ).
- exists nil; simpl; eexists. eapply Tr_add_label; eauto.
+ - (*i=lbl *) cutrewrite (i = 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.
@@ -742,11 +742,11 @@ Proof.
rewrite Heqti.
eapply Tr_end_block; eauto.
rewrite <-Heqti.
- eapply End_basic. inversion H; try(simpl; congruence).
- simpl in H5; congruence. }
- all: try(exists nil; simpl; eexists; eapply Tr_add_basic; eauto; inversion H; try(eauto || congruence)).
+ 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(simpl in Heqti; congruence).
+ destruct i; try(cbn in Heqti; congruence).
all: exists (add_to_new_bblock (MB_cfi cfi)::nil); exists b;
cutrewrite ((add_to_new_bblock (MB_cfi cfi) :: nil) ++ (b::l)=(add_to_new_bblock (MB_cfi cfi) :: (b::l)));eauto;
rewrite Heqti;
@@ -768,13 +768,13 @@ Qed.
(* FIXME: these two lemma should go into [Coqlib.v] *)
Lemma is_tail_app A (l1: list A): forall l2, is_tail l2 (l1 ++ l2).
Proof.
- induction l1; simpl; auto with coqlib.
+ induction l1; cbn; auto with coqlib.
Qed.
Hint Resolve is_tail_app: coqlib.
Lemma is_tail_app_inv A (l1: list A): forall l2 l3, is_tail (l1 ++ l2) l3 -> is_tail l2 l3.
Proof.
- induction l1; simpl; auto with coqlib.
+ induction l1; cbn; auto with coqlib.
intros l2 l3 H; inversion H; eauto with coqlib.
Qed.
Hint Resolve is_tail_app_inv: coqlib.
@@ -787,17 +787,17 @@ Proof.
- intros; subst.
remember (trans_code (Mcall _ _::c)) as tc2.
rewrite <- is_trans_code_inv in Heqtc2.
- inversion Heqtc2; simpl in * |- *; subst; try congruence.
+ 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'; simpl; eauto with coqlib.
+ exists b'; cbn; eauto with coqlib.
* exploit (trans_code_monotonic i c2); eauto.
intros (l' & b' & Hl'); rewrite Hl'.
- simpl; eapply ex_intro.
+ cbn; eapply ex_intro.
eapply is_tail_trans; eauto with coqlib.
Qed.
diff --git a/kvx/lib/PrepassSchedulingOracle.ml b/kvx/lib/PrepassSchedulingOracle.ml
deleted file mode 100644
index 78961310..00000000
--- a/kvx/lib/PrepassSchedulingOracle.ml
+++ /dev/null
@@ -1,432 +0,0 @@
-open AST
-open RTL
-open Maps
-open InstructionScheduler
-open OpWeights
-open Registers
-
-let use_alias_analysis () = false
-
-let length_of_chunk = function
-| Mint8signed
-| Mint8unsigned -> 1
-| Mint16signed
-| Mint16unsigned -> 2
-| Mint32
-| Mfloat32
-| Many32 -> 4
-| Mint64
-| Mfloat64
-| Many64 -> 8;;
-
-let get_simple_dependencies (seqa : (instruction*Regset.t) array) =
- let last_reg_reads : int list PTree.t ref = ref PTree.empty
- and last_reg_write : (int*int) PTree.t ref = ref PTree.empty
- and last_mem_reads : int list ref = ref []
- and last_mem_write : int option ref = ref None
- and last_branch : int option ref = ref None
- and latency_constraints : latency_constraint list ref = ref [] in
- let add_constraint instr_from instr_to latency =
- assert (instr_from <= instr_to);
- assert (latency >= 0);
- if instr_from = instr_to
- then (if latency = 0
- then ()
- else failwith "PrepassSchedulingOracle.get_dependencies: negative self-loop")
- else
- latency_constraints :=
- { instr_from = instr_from;
- instr_to = instr_to;
- latency = latency
- }:: !latency_constraints
- and get_last_reads reg =
- match PTree.get reg !last_reg_reads
- with Some l -> l
- | None -> [] in
- let add_input_mem i =
- if not (use_alias_analysis ())
- then
- begin
- begin
- (* Read after write *)
- match !last_mem_write with
- | None -> ()
- | Some j -> add_constraint j i 1
- end;
- last_mem_reads := i :: !last_mem_reads
- end
- and add_output_mem i =
- if not (use_alias_analysis ())
- then
- begin
- begin
- (* Write after write *)
- match !last_mem_write with
- | None -> ()
- | Some j -> add_constraint j i 1
- end;
- (* Write after read *)
- List.iter (fun j -> add_constraint j i 0) !last_mem_reads;
- last_mem_write := Some i;
- last_mem_reads := []
- end
- and add_input_reg i reg =
- begin
- (* Read after write *)
- match PTree.get reg !last_reg_write with
- | None -> ()
- | Some (j, latency) -> add_constraint j i latency
- end;
- last_reg_reads := PTree.set reg
- (i :: get_last_reads reg)
- !last_reg_reads
- and add_output_reg i latency reg =
- begin
- (* Write after write *)
- match PTree.get reg !last_reg_write with
- | None -> ()
- | Some (j, _) -> add_constraint j i 1
- end;
- begin
- (* Write after read *)
- List.iter (fun j -> add_constraint j i 0) (get_last_reads reg)
- end;
- last_reg_write := PTree.set reg (i, latency) !last_reg_write;
- last_reg_reads := PTree.remove reg !last_reg_reads
- in
- let add_input_regs i regs = List.iter (add_input_reg i) regs in
- let rec add_builtin_res i (res : reg builtin_res) =
- match res with
- | BR r -> add_output_reg i 10 r
- | BR_none -> ()
- | BR_splitlong (hi, lo) -> add_builtin_res i hi;
- add_builtin_res i lo in
- let rec add_builtin_arg i (ba : reg builtin_arg) =
- match ba with
- | BA r -> add_input_reg i r
- | BA_int _ | BA_long _ | BA_float _ | BA_single _ -> ()
- | BA_loadstack(_,_) -> add_input_mem i
- | BA_addrstack _ -> ()
- | BA_loadglobal(_, _, _) -> add_input_mem i
- | BA_addrglobal _ -> ()
- | BA_splitlong(hi, lo) -> add_builtin_arg i hi;
- add_builtin_arg i lo
- | BA_addptr(a1, a2) -> add_builtin_arg i a1;
- add_builtin_arg i a2 in
- let irreversible_action i =
- match !last_branch with
- | None -> ()
- | Some j -> add_constraint j i 1 in
- let set_branch i =
- irreversible_action i;
- last_branch := Some i
- in
- Array.iteri
- begin
- fun i (insn, other_uses) ->
- List.iter (fun use ->
- add_input_reg i use)
- (Regset.elements other_uses);
-
- match insn with
- | Inop _ -> ()
- | Iop(op, inputs, output, _) ->
- (if Op.is_trapping_op op then irreversible_action i);
- add_input_regs i inputs;
- add_output_reg i (latency_of_op op (List.length inputs)) output
- | Iload(trap, chunk, addressing, addr_regs, output, _) ->
- (if trap=TRAP then irreversible_action i);
- add_input_mem i;
- add_input_regs i addr_regs;
- add_output_reg i (latency_of_load trap chunk addressing (List.length addr_regs)) output
- | Istore(chunk, addressing, addr_regs, input, _) ->
- irreversible_action i;
- add_input_regs i addr_regs;
- add_input_reg i input;
- add_output_mem i
- | Icall(signature, ef, inputs, output, _) ->
- set_branch i;
- (match ef with
- | Datatypes.Coq_inl r -> add_input_reg i r
- | Datatypes.Coq_inr symbol -> ()
- );
- add_input_mem i;
- add_input_regs i inputs;
- add_output_reg i (latency_of_call signature ef) output;
- add_output_mem i;
- failwith "Icall"
- | Itailcall(signature, ef, inputs) ->
- set_branch i;
- (match ef with
- | Datatypes.Coq_inl r -> add_input_reg i r
- | Datatypes.Coq_inr symbol -> ()
- );
- add_input_mem i;
- add_input_regs i inputs;
- failwith "Itailcall"
- | Ibuiltin(ef, builtin_inputs, builtin_output, _) ->
- set_branch i;
- add_input_mem i;
- List.iter (add_builtin_arg i) builtin_inputs;
- add_builtin_res i builtin_output;
- add_output_mem i;
- failwith "Ibuiltin"
- | Icond(cond, inputs, _, _, _) ->
- set_branch i;
- add_input_mem i;
- add_input_regs i inputs
- | Ijumptable(input, _) ->
- set_branch i;
- add_input_reg i input;
- failwith "Ijumptable"
- | Ireturn(Some input) ->
- set_branch i;
- add_input_reg i input;
- failwith "Ireturn"
- | Ireturn(None) ->
- set_branch i;
- failwith "Ireturn none"
- end seqa;
- !latency_constraints;;
-
-let resources_of_instruction = function
- | Inop _ -> Array.map (fun _ -> 0) resource_bounds
- | Iop(op, inputs, output, _) -> resources_of_op op (List.length inputs)
- | Iload(trap, chunk, addressing, addr_regs, output, _) ->
- resources_of_load trap chunk addressing (List.length addr_regs)
- | Istore(chunk, addressing, addr_regs, input, _) ->
- resources_of_store chunk addressing (List.length addr_regs)
- | Icall(signature, ef, inputs, output, _) ->
- resources_of_call signature ef
- | Ibuiltin(ef, builtin_inputs, builtin_output, _) ->
- resources_of_builtin ef
- | Icond(cond, args, _, _ , _) ->
- resources_of_cond cond (List.length args)
- | Itailcall _ | Ijumptable _ | Ireturn _ -> resource_bounds
-
-let print_sequence pp (seqa : instruction array) =
- Array.iteri (
- fun i (insn : instruction) ->
- PrintRTL.print_instruction pp (i, insn)) seqa;;
-
-type unique_id = int
-
-type 'a symbolic_term_node =
- | STop of Op.operation * 'a list
- | STinitial_reg of int
- | STother of int;;
-
-type symbolic_term = {
- hash_id : unique_id;
- hash_ct : symbolic_term symbolic_term_node
- };;
-
-let rec print_term channel term =
- match term.hash_ct with
- | STop(op, args) ->
- PrintOp.print_operation print_term channel (op, args)
- | STinitial_reg n -> Printf.fprintf channel "x%d" n
- | STother n -> Printf.fprintf channel "y%d" n;;
-
-type symbolic_term_table = {
- st_table : (unique_id symbolic_term_node, symbolic_term) Hashtbl.t;
- mutable st_next_id : unique_id };;
-
-let hash_init () = {
- st_table = Hashtbl.create 20;
- st_next_id = 0
- };;
-
-let ground_to_id = function
- | STop(op, l) -> STop(op, List.map (fun t -> t.hash_id) l)
- | STinitial_reg r -> STinitial_reg r
- | STother i -> STother i;;
-
-let hash_node (table : symbolic_term_table) (term : symbolic_term symbolic_term_node) : symbolic_term =
- let grounded = ground_to_id term in
- match Hashtbl.find_opt table.st_table grounded with
- | Some x -> x
- | None ->
- let term' = { hash_id = table.st_next_id;
- hash_ct = term } in
- (if table.st_next_id = max_int then failwith "hash: max_int");
- table.st_next_id <- table.st_next_id + 1;
- Hashtbl.add table.st_table grounded term';
- term';;
-
-type access = {
- base : symbolic_term;
- offset : int64;
- length : int
- };;
-
-let term_equal a b = (a.hash_id = b.hash_id);;
-
-let access_of_addressing get_reg chunk addressing args =
- match addressing, args with
- | (Op.Aindexed ofs), [reg] -> Some
- { base = get_reg reg;
- offset = Camlcoq.camlint64_of_ptrofs ofs;
- length = length_of_chunk chunk
- }
- | _, _ -> None ;;
-(* TODO: global *)
-
-let symbolic_execution (seqa : instruction array) =
- let regs = ref PTree.empty
- and table = hash_init() in
- let assign reg term = regs := PTree.set reg term !regs
- and hash term = hash_node table term in
- let get_reg reg =
- match PTree.get reg !regs with
- | None -> hash (STinitial_reg (Camlcoq.P.to_int reg))
- | Some x -> x in
- let targets = Array.make (Array.length seqa) None in
- Array.iteri
- begin
- fun i insn ->
- match insn with
- | Iop(Op.Omove, [input], output, _) ->
- assign output (get_reg input)
- | Iop(op, inputs, output, _) ->
- assign output (hash (STop(op, List.map get_reg inputs)))
-
- | Iload(trap, chunk, addressing, args, output, _) ->
- let access = access_of_addressing get_reg chunk addressing args in
- targets.(i) <- access;
- assign output (hash (STother(i)))
-
- | Icall(_, _, _, output, _)
- | Ibuiltin(_, _, BR output, _) ->
- assign output (hash (STother(i)))
-
- | Istore(chunk, addressing, args, va, _) ->
- let access = access_of_addressing get_reg chunk addressing args in
- targets.(i) <- access
-
- | Inop _ -> ()
- | Ibuiltin(_, _, BR_none, _) -> ()
- | Ibuiltin(_, _, BR_splitlong _, _) -> failwith "BR_splitlong"
-
- | Itailcall (_, _, _)
- |Icond (_, _, _, _, _)
- |Ijumptable (_, _)
- |Ireturn _ -> ()
- end seqa;
- targets;;
-
-let print_access channel = function
- | None -> Printf.fprintf channel "any"
- | Some x -> Printf.fprintf channel "%a + %Ld" print_term x.base x.offset;;
-
-let print_targets channel seqa =
- let targets = symbolic_execution seqa in
- Array.iteri
- (fun i insn ->
- match insn with
- | Iload _ -> Printf.fprintf channel "%d: load %a\n"
- i print_access targets.(i)
- | Istore _ -> Printf.fprintf channel "%d: store %a\n"
- i print_access targets.(i)
- | _ -> ()
- ) seqa;;
-
-let may_overlap a0 b0 =
- match a0, b0 with
- | (None, _) | (_ , None) -> true
- | (Some a), (Some b) ->
- if term_equal a.base b.base
- then (max a.offset b.offset) <
- (min (Int64.add (Int64.of_int a.length) a.offset)
- (Int64.add (Int64.of_int b.length) b.offset))
- else match a.base.hash_ct, b.base.hash_ct with
- | STop(Op.Oaddrsymbol(ida, ofsa),[]),
- STop(Op.Oaddrsymbol(idb, ofsb),[]) ->
- (ida=idb) &&
- let ao = Int64.add a.offset (Camlcoq.camlint64_of_ptrofs ofsa)
- and bo = Int64.add b.offset (Camlcoq.camlint64_of_ptrofs ofsb) in
- (max ao bo) <
- (min (Int64.add (Int64.of_int a.length) ao)
- (Int64.add (Int64.of_int b.length) bo))
- | STop(Op.Oaddrstack _, []),
- STop(Op.Oaddrsymbol _, [])
- | STop(Op.Oaddrsymbol _, []),
- STop(Op.Oaddrstack _, []) -> false
- | STop(Op.Oaddrstack(ofsa),[]),
- STop(Op.Oaddrstack(ofsb),[]) ->
- let ao = Int64.add a.offset (Camlcoq.camlint64_of_ptrofs ofsa)
- and bo = Int64.add b.offset (Camlcoq.camlint64_of_ptrofs ofsb) in
- (max ao bo) <
- (min (Int64.add (Int64.of_int a.length) ao)
- (Int64.add (Int64.of_int b.length) bo))
- | _ -> true;;
-
-(*
-(* TODO suboptimal quadratic algorithm *)
-let get_alias_dependencies seqa =
- let targets = symbolic_execution seqa
- and deps = ref [] in
- let add_constraint instr_from instr_to latency =
- deps := { instr_from = instr_from;
- instr_to = instr_to;
- latency = latency
- }:: !deps in
- for i=0 to (Array.length seqa)-1
- do
- for j=0 to i-1
- do
- match seqa.(j), seqa.(i) with
- | (Istore _), ((Iload _) | (Istore _)) ->
- if may_overlap targets.(j) targets.(i)
- then add_constraint j i 1
- | (Iload _), (Istore _) ->
- if may_overlap targets.(j) targets.(i)
- then add_constraint j i 0
- | (Istore _ | Iload _), (Icall _ | Ibuiltin _)
- | (Icall _ | Ibuiltin _), (Icall _ | Ibuiltin _ | Iload _ | Istore _) ->
- add_constraint j i 1
- | (Inop _ | Iop _), _
- | _, (Inop _ | Iop _)
- | (Iload _), (Iload _) -> ()
- done
- done;
- !deps;;
- *)
-
-let define_problem seqa =
- let simple_deps = get_simple_dependencies seqa in
- { max_latency = -1;
- resource_bounds = OpWeights.resource_bounds;
- instruction_usages = Array.map resources_of_instruction (Array.map fst seqa);
- latency_constraints =
- (* if (use_alias_analysis ())
- then (get_alias_dependencies seqa) @ simple_deps
- else *) simple_deps };;
-
-let schedule_sequence (seqa : (instruction*Regset.t) array) =
- try
- if (Array.length seqa) <= 1
- then None
- else
- begin
- let nr_instructions = Array.length seqa in
- Printf.printf "prepass scheduling length = %d\n" (Array.length seqa);
- let problem = define_problem seqa in
- print_sequence stdout (Array.map fst seqa);
- print_problem stdout problem;
- match scheduler_by_name (!Clflags.option_fprepass_sched) problem with
- | None -> Printf.printf "no solution in prepass scheduling\n";
- None
- | Some solution ->
- let positions = Array.init nr_instructions (fun i -> i) in
- Array.sort (fun i j ->
- let si = solution.(i) and sj = solution.(j) in
- if si < sj then -1
- else if si > sj then 1
- else i - j) positions;
- Some positions
- end
- with (Failure s) ->
- Printf.printf "failure in prepass scheduling: %s\n" s;
- None;;
-
diff --git a/kvx/lib/RTLpath.v b/kvx/lib/RTLpath.v
deleted file mode 100644
index 35512652..00000000
--- a/kvx/lib/RTLpath.v
+++ /dev/null
@@ -1,1067 +0,0 @@
-(** We introduce a data-structure extending the RTL CFG into a control-flow graph over "traces" (in the sense of "trace-scheduling")
- Here, we use the word "path" instead of "trace" because "trace" has already a meaning in CompCert:
- a "path" is simply a list of successive nodes in the CFG (modulo some additional wellformness conditions).
-
- Actually, we extend syntactically the notion of RTL programs with a structure of "path_map":
- this gives an alternative view of the CFG -- where "nodes" are paths instead of simple instructions.
- Our wellformness condition on paths express that:
- - the CFG on paths is wellformed: any successor of a given path points to another path (possibly the same).
- - execution of a paths only emit single events.
-
- We represent each path only by a natural: the number of nodes in the path. These nodes are recovered from a static notion of "default successor".
- This notion of path is thus incomplete. For example, if a path contains a whole loop (and for example, unrools it several times),
- then this loop must be a suffix of the path.
-
- However: it is sufficient in order to represent superblocks (each superblock being represented as a path).
- A superblock decomposition of the CFG exactly corresponds to the case where each node is in at most one path.
-
- Our goal is to provide two bisimulable semantics:
- - one is simply the RTL semantics
- - the other is based on a notion of "path-step": each path is executed in a single step.
-
- Remark that all analyses on RTL programs should thus be appliable for "free" also for RTLpath programs !
-*)
-
-Require Import Coqlib Maps.
-Require Import AST Integers Values Events Memory Globalenvs Smallstep.
-Require Import Op Registers.
-Require Import RTL Linking.
-
-Declare Scope option_monad_scope.
-
-Notation "'SOME' X <- A 'IN' B" := (match A with Some X => B | None => None end)
- (at level 200, X ident, A at level 100, B at level 200)
- : option_monad_scope.
-
-Notation "'ASSERT' A 'IN' B" := (if A then B else None)
- (at level 200, A at level 100, B at level 200)
- : option_monad_scope.
-
-Local Open Scope option_monad_scope.
-
-(** * Syntax of RTLpath programs *)
-
-(** Internal instruction = instruction with a default successor in a path. *)
-
-Definition default_succ (i: instruction): option node :=
- match i with
- | Inop s => Some s
- | Iop op args res s => Some s
- | Iload _ chunk addr args dst s => Some s
- | Istore chunk addr args src s => Some s
- | Icond cond args ifso ifnot _ => Some ifnot
- | _ => None (* TODO: we could choose a successor for jumptable ? *)
- end.
-
-Definition early_exit (i: instruction): option node := (* FIXME: for jumptable, replace [node] by [list node] *)
- match i with
- | Icond cond args ifso ifnot _ => Some ifso
- | _ => None
- end.
-
-(** Our notion of path.
-
- We do not formally require that the set of path is a partition of the CFG.
- path may have intersections !
-
- Moreover, we do not formally require that path have a single entry-point (a superblock structure)
-
- But, in practice, these properties are probably necessary in order to ensure the success of dynamic verification of scheduling.
-
- Here: we only require that each exit-point of a path is the entry-point of a path
- (and that internal node of a path are internal instructions)
-*)
-
-
-(* By convention, we say that node [n] is the entry-point of a path if it is a key of the path_map.
-
- Such a path of entry [n] is defined from a natural [path] representing the [path] default-successors of [n].
-
- Remark: a path can loop several times in the CFG.
-
-*)
-
-Record path_info := {
- psize: nat; (* number minus 1 of instructions in the path *)
- input_regs: Regset.t;
- (** Registers that are used (as input_regs) by the "fallthrough successors" of the path *)
- (** This field is not used by the verificator, but is helpful for the superblock scheduler *)
- output_regs: Regset.t
-}.
-
-Definition path_map: Type := PTree.t path_info.
-
-Definition path_entry (*c:code*) (pm: path_map) (n: node): Prop
- := pm!n <> None (*/\ c!n <> None*).
-
-Inductive wellformed_path (c:code) (pm: path_map): nat -> node -> Prop :=
- | wf_last_node i pc:
- c!pc = Some i ->
- (forall n, List.In n (successors_instr i) -> path_entry (*c*) pm n) ->
- wellformed_path c pm 0 pc
- | wf_internal_node path i pc pc':
- c!pc = Some i ->
- default_succ i = Some pc' ->
- (forall n, early_exit i = Some n -> path_entry (*c*) pm n) ->
- wellformed_path c pm path pc' ->
- wellformed_path c pm (S path) pc.
-
-(* all paths defined from the path_map are wellformed *)
-Definition wellformed_path_map (c:code) (pm: path_map): Prop :=
- forall n path, pm!n = Some path -> wellformed_path c pm path.(psize) n.
-
-(** We "extend" the notion of RTL program with the additional structure for path.
-
- There is thus a trivial "forgetful functor" from RTLpath programs to RTL ones.
-*)
-
-Record function : Type :=
- { fn_RTL:> RTL.function;
- fn_path: path_map;
- (* condition 1 below: the entry-point of the code is an entry-point of a path *)
- fn_entry_point_wf: path_entry (*fn_RTL.(fn_code)*) fn_path fn_RTL.(fn_entrypoint);
- (* condition 2 below: the path_map is well-formed *)
- fn_path_wf: wellformed_path_map fn_RTL.(fn_code) fn_path
- }.
-
-Definition fundef := AST.fundef function.
-Definition program := AST.program fundef unit.
-Definition genv := Genv.t fundef unit.
-
-Definition fundef_RTL (fu: fundef) : RTL.fundef :=
- match fu with
- | Internal f => Internal f.(fn_RTL)
- | External ef => External ef
- end.
-Coercion fundef_RTL: fundef >-> RTL.fundef.
-
-Definition transf_program (p: program) : RTL.program := transform_program fundef_RTL p.
-Coercion transf_program: program >-> RTL.program.
-
-(** * Path-step semantics of RTLpath programs *)
-
-(* Semantics of internal instructions (mimicking RTL semantics) *)
-
-Record istate := mk_istate { icontinue: bool; ipc: node; irs: regset; imem: mem }.
-
-(* FIXME - prediction *)
-(* Internal step through the path *)
-Definition istep (ge: RTL.genv) (i: instruction) (sp: val) (rs: regset) (m: mem): option istate :=
- match i with
- | Inop pc' => Some (mk_istate true pc' rs m)
- | Iop op args res pc' =>
- SOME v <- eval_operation ge sp op rs##args m IN
- Some (mk_istate true pc' (rs#res <- v) m)
- | Iload TRAP chunk addr args dst pc' =>
- SOME a <- eval_addressing ge sp addr rs##args IN
- SOME v <- Mem.loadv chunk m a IN
- Some (mk_istate true pc' (rs#dst <- v) m)
- | Iload NOTRAP chunk addr args dst pc' =>
- let default_state := mk_istate true pc' rs#dst <- (default_notrap_load_value chunk) m in
- match (eval_addressing ge sp addr rs##args) with
- | None => Some default_state
- | Some a => match (Mem.loadv chunk m a) with
- | None => Some default_state
- | Some v => Some (mk_istate true pc' (rs#dst <- v) m)
- end
- end
- | Istore chunk addr args src pc' =>
- SOME a <- eval_addressing ge sp addr rs##args IN
- SOME m' <- Mem.storev chunk m a rs#src IN
- Some (mk_istate true pc' rs m')
- | Icond cond args ifso ifnot _ =>
- SOME b <- eval_condition cond rs##args m IN
- Some (mk_istate (negb b) (if b then ifso else ifnot) rs m)
- | _ => None (* TODO jumptable ? *)
- end.
-
-(** Execution of a path in a single step *)
-
-(* Executes until a state [st] is reached where st.(continue) is false *)
-Fixpoint isteps ge (path:nat) (f: function) sp rs m pc: option istate :=
- match path with
- | O => Some (mk_istate true pc rs m)
- | S p =>
- SOME i <- (fn_code f)!pc IN
- SOME st <- istep ge i sp rs m IN
- if (icontinue st) then
- isteps ge p f sp (irs st) (imem st) (ipc st)
- else
- Some st
- end.
-
-Definition find_function (pge: genv) (ros: reg + ident) (rs: regset) : option fundef :=
- match ros with
- | inl r => Genv.find_funct pge rs#r
- | inr symb =>
- match Genv.find_symbol pge symb with
- | None => None
- | Some b => Genv.find_funct_ptr pge b
- end
- end.
-
-Inductive stackframe : Type :=
- | Stackframe
- (res: reg) (**r where to store the result *)
- (f: function) (**r calling function *)
- (sp: val) (**r stack pointer in calling function *)
- (pc: node) (**r program point in calling function *)
- (rs: regset) (**r register state in calling function *)
- .
-
-Definition stf_RTL (st: stackframe): RTL.stackframe :=
- match st with
- | Stackframe res f sp pc rs => RTL.Stackframe res f sp pc rs
- end.
-
-Fixpoint stack_RTL (stack: list stackframe): list RTL.stackframe :=
- match stack with
- | nil => nil
- | cons stf stack' => cons (stf_RTL stf) (stack_RTL stack')
- end.
-
-Inductive state : Type :=
- | State
- (stack: list stackframe) (**r call stack *)
- (f: function) (**r current function *)
- (sp: val) (**r stack pointer *)
- (pc: node) (**r current program point in [c] *)
- (rs: regset) (**r register state *)
- (m: mem) (**r memory state *)
- | Callstate
- (stack: list stackframe) (**r call stack *)
- (f: fundef) (**r function to call *)
- (args: list val) (**r arguments to the call *)
- (m: mem) (**r memory state *)
- | Returnstate
- (stack: list stackframe) (**r call stack *)
- (v: val) (**r return value for the call *)
- (m: mem) (**r memory state *)
- .
-
-Definition state_RTL (s: state): RTL.state :=
- match s with
- | State stack f sp pc rs m => RTL.State (stack_RTL stack) f sp pc rs m
- | Callstate stack f args m => RTL.Callstate (stack_RTL stack) f args m
- | Returnstate stack v m => RTL.Returnstate (stack_RTL stack) v m
- end.
-Coercion state_RTL: state >-> RTL.state.
-
-(* Used to execute the last instruction of a path (isteps is only in charge of executing the instructions before the last) *)
-Inductive path_last_step ge pge stack (f: function): val -> node -> regset -> mem -> trace -> state -> Prop :=
- | exec_istate i sp pc rs m st:
- (fn_code f)!pc = Some i ->
- istep ge i sp rs m = Some st ->
- path_last_step ge pge stack f sp pc rs m
- E0 (State stack f sp (ipc st) (irs st) (imem st))
- | exec_Icall sp pc rs m sig ros args res pc' fd:
- (fn_code f)!pc = Some(Icall sig ros args res pc') ->
- find_function pge ros rs = Some fd ->
- funsig fd = sig ->
- path_last_step ge pge stack f sp pc rs m
- E0 (Callstate (Stackframe res f sp pc' rs :: stack) fd rs##args m)
- | exec_Itailcall stk pc rs m sig ros args fd m':
- (fn_code f)!pc = Some(Itailcall sig ros args) ->
- find_function pge ros rs = Some fd ->
- funsig fd = sig ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m
- E0 (Callstate stack fd rs##args m')
- | exec_Ibuiltin sp pc rs m ef args res pc' vargs t vres m':
- (fn_code f)!pc = Some(Ibuiltin ef args res pc') ->
- eval_builtin_args ge (fun r => rs#r) sp m args vargs ->
- external_call ef ge vargs m t vres m' ->
- path_last_step ge pge stack f sp pc rs m
- t (State stack f sp pc' (regmap_setres res vres rs) m')
- | exec_Ijumptable sp pc rs m arg tbl n pc': (* TODO remove jumptable from here ? *)
- (fn_code f)!pc = Some(Ijumptable arg tbl) ->
- rs#arg = Vint n ->
- list_nth_z tbl (Int.unsigned n) = Some pc' ->
- path_last_step ge pge stack f sp pc rs m
- E0 (State stack f sp pc' rs m)
- | exec_Ireturn stk pc rs m or m':
- (fn_code f)!pc = Some(Ireturn or) ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- path_last_step ge pge stack f (Vptr stk Ptrofs.zero) pc rs m
- E0 (Returnstate stack (regmap_optget or Vundef rs) m').
-
-(* Executes an entire path *)
-Inductive path_step ge pge (path:nat) stack f sp rs m pc: trace -> state -> Prop :=
- | exec_early_exit st:
- isteps ge path f sp rs m pc = Some st ->
- (icontinue st) = false ->
- path_step ge pge path stack f sp rs m pc E0 (State stack f sp (ipc st) (irs st) (imem st))
- | exec_normal_exit st t s:
- isteps ge path f sp rs m pc = Some st ->
- (icontinue st) = true ->
- path_last_step ge pge stack f sp (ipc st) (irs st) (imem st) t s ->
- path_step ge pge path stack f sp rs m pc t s.
-
-(* Either internal path execution, or the usual exec_function / exec_return borrowed from RTL *)
-Inductive step ge pge: state -> trace -> state -> Prop :=
- | exec_path path stack f sp rs m pc t s:
- (fn_path f)!pc = Some path ->
- path_step ge pge path.(psize) stack f sp rs m pc t s ->
- step ge pge (State stack f sp pc rs m) t s
- | exec_function_internal s f args m m' stk:
- Mem.alloc m 0 (fn_RTL f).(fn_stacksize) = (m', stk) ->
- step ge pge (Callstate s (Internal f) args m)
- E0 (State s
- f
- (Vptr stk Ptrofs.zero)
- f.(fn_entrypoint)
- (init_regs args f.(fn_params))
- m')
- | exec_function_external s ef args res t m m':
- external_call ef ge args m t res m' ->
- step ge pge (Callstate s (External ef) args m)
- t (Returnstate s res m')
- | exec_return res f sp pc rs s vres m:
- step ge pge (Returnstate (Stackframe res f sp pc rs :: s) vres m)
- E0 (State s f sp pc (rs#res <- vres) m).
-
-Inductive initial_state (p:program) : state -> Prop :=
- initial_state_intro (b : block) (f : fundef) (m0 : mem):
- Genv.init_mem p = Some m0 ->
- Genv.find_symbol (Genv.globalenv p) (prog_main p) = Some b ->
- Genv.find_funct_ptr (Genv.globalenv p) b = Some f ->
- funsig f = signature_main -> initial_state p (Callstate nil f nil m0).
-
-Definition final_state (st: state) (i:int): Prop
- := RTL.final_state st i.
-
-Definition semantics (p: program) :=
- Semantics (step (Genv.globalenv (transf_program p))) (initial_state p) final_state (Genv.globalenv p).
-
-(** * Proving the bisimulation between (semantics p) and (RTL.semantics p). *)
-
-(** ** Preliminaries: simple tactics for option-monad *)
-
-Lemma destruct_SOME A B (P: option B -> Prop) (e: option A) (f: A -> option B):
- (forall x, e = Some x -> P (f x)) -> (e = None -> P None) -> (P (SOME x <- e IN f x)).
-Proof.
- intros; destruct e; simpl; auto.
-Qed.
-
-Lemma destruct_ASSERT B (P: option B -> Prop) (e: bool) (x: option B):
- (e = true -> P x) -> (e = false -> P None) -> (P (ASSERT e IN x)).
-Proof.
- intros; destruct e; simpl; auto.
-Qed.
-
-Ltac inversion_SOME x :=
- try (eapply destruct_SOME; [ let x := fresh x in intro x | simpl; try congruence ]).
-
-Ltac inversion_ASSERT :=
- try (eapply destruct_ASSERT; [ idtac | simpl; try congruence ]).
-
-Ltac simplify_someHyp :=
- match goal with
- | H: None = Some _ |- _ => inversion H; clear H; subst
- | H: Some _ = None |- _ => inversion H; clear H; subst
- | H: ?t = ?t |- _ => clear H
- | H: Some _ = Some _ |- _ => inversion H; clear H; subst
- | H: Some _ <> None |- _ => clear H
- | H: None <> Some _ |- _ => clear H
- | H: _ = Some _ |- _ => (try rewrite !H in * |- *); generalize H; clear H
- | H: _ = None |- _ => (try rewrite !H in * |- *); generalize H; clear H
- end.
-
-Ltac explore_destruct :=
- repeat (match goal with
- | [H: ?expr = ?val |- context[match ?expr with | _ => _ end]] => rewrite H
- | [H: match ?var with | _ => _ end |- _] => destruct var
- | [ |- context[match ?m with | _ => _ end] ] => destruct m
- | _ => discriminate
- end).
-
-Ltac simplify_someHyps :=
- repeat (simplify_someHyp; simpl in * |- *).
-
-Ltac try_simplify_someHyps :=
- try (intros; simplify_someHyps; eauto).
-
-(* TODO: try to improve this tactic with a better control over names and inversion *)
-Ltac simplify_SOME x :=
- (repeat inversion_SOME x); try_simplify_someHyps.
-
-(** ** The easy way: Forward simulation of RTLpath by RTL
-
-This way can be viewed as a correctness property: all transitions in RTLpath are valid RTL transitions !
-
-*)
-
-Local Hint Resolve RTL.exec_Inop RTL.exec_Iop RTL.exec_Iload RTL.exec_Istore RTL.exec_Icond RTL.exec_Iload_notrap1 RTL.exec_Iload_notrap2: core.
-
-(* istep reflects RTL.step *)
-Lemma istep_correct ge i stack (f:function) sp rs m st :
- istep ge i sp rs m = Some st ->
- forall pc, (fn_code f)!pc = Some i ->
- RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(ipc) st.(irs) st.(imem)).
-Proof.
- destruct i; simpl; try congruence; simplify_SOME x.
- 1-3: explore_destruct; simplify_SOME x.
-Qed.
-
-Local Hint Resolve star_refl: core.
-
-(* isteps reflects a star relation on RTL.step *)
-Lemma isteps_correct ge path stack f sp: forall rs m pc st,
- isteps ge path f sp rs m pc = Some st ->
- star RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(ipc) st.(irs) st.(imem)).
-Proof.
- induction path; simpl; try_simplify_someHyps.
- inversion_SOME i; intros Hi.
- inversion_SOME st0; intros Hst0.
- destruct (icontinue st0) eqn:cont.
- + intros; eapply star_step.
- - eapply istep_correct; eauto.
- - simpl; eauto.
- - auto.
- + intros; simplify_someHyp; eapply star_step.
- - eapply istep_correct; eauto.
- - simpl; eauto.
- - auto.
-Qed.
-
-Lemma isteps_correct_early_exit ge path stack f sp: forall rs m pc st,
- isteps ge path f sp rs m pc = Some st ->
- st.(icontinue) = false ->
- plus RTL.step ge (State stack f sp pc rs m) E0 (State stack f sp st.(ipc) st.(irs) st.(imem)).
-Proof.
- destruct path; simpl; try_simplify_someHyps; try congruence.
- inversion_SOME i; intros Hi.
- inversion_SOME st0; intros Hst0.
- destruct (icontinue st0) eqn:cont.
- + intros; eapply plus_left.
- - eapply istep_correct; eauto.
- - eapply isteps_correct; eauto.
- - auto.
- + intros X; inversion X; subst.
- eapply plus_one.
- eapply istep_correct; eauto.
-Qed.
-
-Local Hint Resolve list_forall2_nil match_globdef_fun linkorder_refl match_globvar_intro: core.
-
-Section CORRECTNESS.
-
-Variable p: program.
-
-Lemma match_prog_RTL: match_program (fun _ f tf => tf = fundef_RTL f) eq p (transf_program p).
-Proof.
- eapply match_transform_program; eauto.
-Qed.
-
-Let pge := Genv.globalenv p.
-Let ge := Genv.globalenv (transf_program p).
-
-Lemma senv_preserved: Senv.equiv pge ge.
-Proof (Genv.senv_match match_prog_RTL).
-
-Lemma symbols_preserved s: Genv.find_symbol ge s = Genv.find_symbol pge s.
-Proof (Genv.find_symbol_match match_prog_RTL s).
-
-Lemma find_function_RTL_match ros rs fd:
- find_function pge ros rs = Some fd -> RTL.find_function ge ros rs = Some (fundef_RTL fd).
-Proof.
- destruct ros; simpl.
- + intro; exploit (Genv.find_funct_match match_prog_RTL); eauto.
- intros (cuint & tf & H1 & H2 & H3); subst; auto.
- + rewrite symbols_preserved.
- destruct (Genv.find_symbol pge i); simpl; try congruence.
- intro; exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto.
- intros (cuint & tf & H1 & H2 & H3); subst; auto.
-Qed.
-
-Local Hint Resolve istep_correct RTL.exec_Ibuiltin RTL.exec_Ijumptable RTL.exec_Ireturn RTL.exec_Icall RTL.exec_Itailcall find_function_RTL_match: core.
-
-Lemma path_last_step_correct stack f sp pc rs m t s:
- path_last_step ge pge stack f sp pc rs m t s ->
- RTL.step ge (State stack f sp pc rs m) t s.
-Proof.
- destruct 1; try (eapply istep_correct); simpl; eauto.
-Qed.
-
-Lemma path_step_correct path stack f sp pc rs m t s:
- path_step ge pge path stack f sp rs m pc t s ->
- plus RTL.step ge (State stack f sp pc rs m) t s.
-Proof.
- destruct 1.
- + eapply isteps_correct_early_exit; eauto.
- + eapply plus_right.
- eapply isteps_correct; eauto.
- eapply path_last_step_correct; eauto.
- auto.
-Qed.
-
-Local Hint Resolve plus_one RTL.exec_function_internal RTL.exec_function_external RTL.exec_return: core.
-
-Lemma step_correct s t s': step ge pge s t s' -> plus RTL.step ge s t s'.
-Proof.
- destruct 1; try (eapply path_step_correct); simpl; eauto.
-Qed.
-
-Theorem RTLpath_correct: forward_simulation (semantics p) (RTL.semantics p).
-Proof.
- eapply forward_simulation_plus with (match_states := fun s1 s2 => s2 = state_RTL s1); simpl; auto.
- - apply senv_preserved.
- - destruct 1; intros; eexists; intuition eauto. econstructor; eauto.
- + apply (Genv.init_mem_match match_prog_RTL); auto.
- + rewrite (Genv.find_symbol_match match_prog_RTL).
- rewrite (match_program_main match_prog_RTL); eauto.
- + exploit (Genv.find_funct_ptr_match match_prog_RTL); eauto.
- intros (cunit & tf0 & XX); intuition subst; eauto.
- - unfold final_state; intros; subst; eauto.
- - intros; subst. eexists; intuition.
- eapply step_correct; eauto.
-Qed.
-
-End CORRECTNESS.
-
-Lemma program_equals {A B: Type} : forall (p1 p2: AST.program A B),
- prog_defs p1 = prog_defs p2 ->
- prog_public p1 = prog_public p2 ->
- prog_main p1 = prog_main p2 ->
- p1 = p2.
-Proof.
- intros. destruct p1. destruct p2. simpl in *. subst. auto.
-Qed.
-
-Lemma cons_extract {A: Type} : forall (l: list A) a b, a = b -> a::l = b::l.
-Proof.
- intros. congruence.
-Qed.
-
-(* Definition transf_program : RTLpath.program -> RTL.program := transform_program fundef_RTL.
-
-Lemma transf_program_proj: forall p, transf_program (transf_program p) = p.
-Proof.
- intros p. destruct p as [defs pub main]. unfold program_proj. simpl.
- apply program_equals; simpl; auto.
- induction defs.
- - simpl; auto.
- - simpl. rewrite IHdefs.
- destruct a as [id gd]; simpl.
- destruct gd as [f|v]; simpl; auto.
- rewrite transf_fundef_proj. auto.
-Qed. *)
-
-
-(** The hard way: Forward simulation of RTL by RTLpath
-
-This way can be viewed as a completeness property: all transitions in RTL can be represented as RTLpath transitions !
-
-*)
-
-(* This lemma is probably needed to compose a pass from RTL -> RTLpath with other passes.*)
-Lemma match_RTL_prog {LA: Linker fundef} {LV: Linker unit} p: match_program (fun _ f tf => f = fundef_RTL tf) eq (transf_program p) p.
-Proof.
- unfold match_program, match_program_gen; intuition.
- unfold transf_program at 2; simpl.
- generalize (prog_defs p).
- induction l as [|a l]; simpl; eauto.
- destruct a; simpl.
- intros; eapply list_forall2_cons; eauto.
- unfold match_ident_globdef; simpl; intuition; destruct g as [f|v]; simpl; eauto.
- eapply match_globdef_var. destruct v; eauto.
-Qed.
-
-(* Theory of wellformed paths *)
-
-Fixpoint nth_default_succ (c: code) (path:nat) (pc: node): option node :=
- match path with
- | O => Some pc
- | S path' =>
- SOME i <- c!pc IN
- SOME pc' <- default_succ i IN
- nth_default_succ c path' pc'
- end.
-
-Lemma wellformed_suffix_path c pm path path':
- (path' <= path)%nat ->
- forall pc, wellformed_path c pm path pc ->
- exists pc', nth_default_succ c (path-path') pc = Some pc' /\ wellformed_path c pm path' pc'.
-Proof.
- induction 1 as [|m].
- + intros. enough (path'-path'=0)%nat as ->; [simpl;eauto|omega].
- + intros pc WF; enough (S m-path'=S (m-path'))%nat as ->; [simpl;eauto|omega].
- inversion WF; subst; clear WF; intros; simplify_someHyps.
- intros; simplify_someHyps; eauto.
-Qed.
-
-Definition nth_default_succ_inst (c: code) (path:nat) pc: option instruction :=
- SOME pc <- nth_default_succ c path pc IN
- c!pc.
-
-Lemma final_node_path f path pc:
- (fn_path f)!pc = Some path ->
- exists i, nth_default_succ_inst (fn_code f) path.(psize) pc = Some i
- /\ (forall n, List.In n (successors_instr i) -> path_entry (*fn_code f*) (fn_path f) n).
-Proof.
- intros; exploit fn_path_wf; eauto.
- intro WF.
- set (ps:=path.(psize)).
- exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps O); omega || eauto.
- destruct 1 as (pc' & NTH_SUCC & WF'); auto.
- assert (ps - 0 = ps)%nat as HH by omega. rewrite HH in NTH_SUCC. clear HH.
- unfold nth_default_succ_inst.
- inversion WF'; clear WF'; subst. simplify_someHyps; eauto.
-Qed.
-
-Lemma internal_node_path path f path0 pc:
- (fn_path f)!pc = (Some path0) ->
- (path < path0.(psize))%nat ->
- exists i pc',
- nth_default_succ_inst (fn_code f) path pc = Some i /\
- default_succ i = Some pc' /\
- (forall n, early_exit i = Some n -> path_entry (*fn_code f*) (fn_path f) n).
-Proof.
- intros; exploit fn_path_wf; eauto.
- set (ps:=path0.(psize)).
- intro WF; exploit (wellformed_suffix_path (fn_code f) (fn_path f) ps (ps-path)); eauto. { omega. }
- destruct 1 as (pc' & NTH_SUCC & WF').
- assert (ps - (ps - path) = path)%nat as HH by omega. rewrite HH in NTH_SUCC. clear HH.
- unfold nth_default_succ_inst.
- inversion WF'; clear WF'; subst. { omega. }
- simplify_someHyps; eauto.
-Qed.
-
-Lemma initialize_path (*c*) pm n: path_entry (*c*) pm n -> exists path, pm!n = Some path.
-Proof.
- unfold path_entry; destruct pm!n; eauto. intuition congruence.
-Qed.
-Local Hint Resolve fn_entry_point_wf: core.
-Local Opaque path_entry.
-
-Lemma istep_successors ge i sp rs m st:
- istep ge i sp rs m = Some st ->
- In (ipc st) (successors_instr i).
-Proof.
- destruct i; simpl; try congruence; simplify_SOME x.
- all: explore_destruct; simplify_SOME x.
-Qed.
-
-Lemma istep_normal_exit ge i sp rs m st:
- istep ge i sp rs m = Some st ->
- st.(icontinue) = true ->
- default_succ i = Some st.(ipc).
-Proof.
- destruct i; simpl; try congruence; simplify_SOME x.
- all: explore_destruct; simplify_SOME x.
-Qed.
-
-Lemma isteps_normal_exit ge path f sp: forall rs m pc st,
- st.(icontinue) = true ->
- isteps ge path f sp rs m pc = Some st ->
- nth_default_succ (fn_code f) path pc = Some st.(ipc).
-Proof.
- induction path; simpl. { try_simplify_someHyps. }
- intros rs m pc st CONT; try_simplify_someHyps.
- inversion_SOME i; intros Hi.
- inversion_SOME st0; intros Hst0.
- destruct (icontinue st0) eqn:X; try congruence.
- try_simplify_someHyps.
- intros; erewrite istep_normal_exit; eauto.
-Qed.
-
-
-(* TODO: the three following lemmas could maybe simplified by introducing an auxiliary
- left-recursive definition equivalent to isteps ?
-*)
-Lemma isteps_step_right ge path f sp: forall rs m pc st i,
- isteps ge path f sp rs m pc = Some st ->
- st.(icontinue) = true ->
- (fn_code f)!(st.(ipc)) = Some i ->
- istep ge i sp st.(irs) st.(imem) = isteps ge (S path) f sp rs m pc.
-Proof.
- induction path.
- + simpl; intros; try_simplify_someHyps. simplify_SOME st.
- destruct st as [b]; destruct b; simpl; auto.
- + intros rs m pc st i H.
- simpl in H.
- generalize H; clear H; simplify_SOME xx.
- destruct (icontinue xx0) eqn: CONTxx0.
- * intros; erewrite IHpath; eauto.
- * intros; congruence.
-Qed.
-
-Lemma isteps_inversion_early ge path f sp: forall rs m pc st,
- isteps ge path f sp rs m pc = Some st ->
- (icontinue st)=false ->
- exists st0 i path0,
- (path > path0)%nat /\
- isteps ge path0 f sp rs m pc = Some st0 /\
- st0.(icontinue) = true /\
- (fn_code f)!(st0.(ipc)) = Some i /\
- istep ge i sp st0.(irs) st0.(imem) = Some st.
-Proof.
- induction path as [|path]; simpl.
- - intros; try_simplify_someHyps; try congruence.
- - intros rs m pc st; inversion_SOME i; inversion_SOME st0.
- destruct (icontinue st0) eqn: CONT.
- + intros STEP PC STEPS CONT0. exploit IHpath; eauto.
- clear STEPS.
- intros (st1 & i0 & path0 & BOUND & STEP1 & CONT1 & X1 & X2); auto.
- exists st1. exists i0. exists (S path0). intuition.
- simpl; try_simplify_someHyps.
- rewrite CONT. auto.
- + intros; try_simplify_someHyps; try congruence.
- eexists. exists i. exists O; simpl. intuition eauto.
- omega.
-Qed.
-
-Lemma isteps_resize ge path0 path1 f sp rs m pc st:
- (path0 <= path1)%nat ->
- isteps ge path0 f sp rs m pc = Some st ->
- (icontinue st)=false ->
- isteps ge path1 f sp rs m pc = Some st.
-Proof.
- induction 1 as [|path1]; simpl; auto.
- intros PSTEP CONT. exploit IHle; auto. clear PSTEP IHle H path0.
- generalize rs m pc st CONT; clear rs m pc st CONT.
- induction path1 as [|path]; simpl; auto.
- - intros; try_simplify_someHyps; try congruence.
- - intros rs m pc st; inversion_SOME i; inversion_SOME st0; intros; try_simplify_someHyps.
- destruct (icontinue st0) eqn: CONT0; eauto.
-Qed.
-
-(* FIXME - add prediction *)
-Inductive is_early_exit pc: instruction -> Prop :=
- | Icond_early_exit cond args ifnot predict:
- is_early_exit pc (Icond cond args pc ifnot predict)
- . (* TODO add jumptable here ? *)
-
-Lemma istep_early_exit ge i sp rs m st :
- istep ge i sp rs m = Some st ->
- st.(icontinue) = false ->
- st.(irs) = rs /\ st.(imem) = m /\ is_early_exit st.(ipc) i.
-Proof.
- Local Hint Resolve Icond_early_exit: core.
- destruct i; simpl; try congruence; simplify_SOME b; simpl; try congruence.
- all: explore_destruct; simplify_SOME b; try discriminate.
-Qed.
-
-Section COMPLETENESS.
-
-Variable p: program.
-
-Let pge := Genv.globalenv p.
-Let ge := Genv.globalenv (transf_program p).
-
-Lemma find_funct_ptr_RTL_preserv b f:
- Genv.find_funct_ptr ge b = Some f -> (exists f0, Genv.find_funct_ptr pge b = Some f0 /\ f = f0).
-Proof.
- intros; exploit (Genv.find_funct_ptr_match (match_RTL_prog p)); eauto.
- destruct 1 as (cunit & tf & X & Y & Z); subst.
- eauto.
-Qed.
-
-Lemma find_RTL_function_match ros rs fd:
- RTL.find_function ge ros rs = Some fd -> exists fd', fd = fundef_RTL fd' /\ find_function pge ros rs = Some fd'.
-Proof.
- destruct ros; simpl.
- + intro; exploit (Genv.find_funct_match (match_RTL_prog p)); eauto.
- intros (cuint & tf & H1 & H2 & H3); subst; eauto.
- + rewrite (symbols_preserved p); unfold pge.
- destruct (Genv.find_symbol (Genv.globalenv p) i); simpl; try congruence.
- intro; exploit find_funct_ptr_RTL_preserv; eauto.
- intros (tf & H1 & H2); subst; eauto.
-Qed.
-
-
-(** *** Definition of well-formed stacks and of match_states *)
-Definition wf_stf (st: stackframe): Prop :=
- match st with
- | Stackframe res f sp pc rs => path_entry (*f.(fn_code)*) f.(fn_path) pc
- end.
-
-Definition wf_stackframe (stack: list stackframe): Prop :=
- forall st, List.In st stack -> wf_stf st.
-
-Lemma wf_stackframe_nil: wf_stackframe nil.
-Proof.
- unfold wf_stackframe; simpl. tauto.
-Qed.
-Local Hint Resolve wf_stackframe_nil: core.
-
-Lemma wf_stackframe_cons st stack:
- wf_stackframe (st::stack) <-> (wf_stf st) /\ wf_stackframe stack.
-Proof.
- unfold wf_stackframe; simpl; intuition (subst; auto).
-Qed.
-
-Definition stack_of (s: state): list stackframe :=
- match s with
- | State stack f sp pc rs m => stack
- | Callstate stack f args m => stack
- | Returnstate stack v m => stack
- end.
-
-Definition is_inst (s: RTL.state): bool :=
- match s with
- | RTL.State stack f sp pc rs m => true
- | _ => false
- end.
-
-Inductive match_inst_states_goal (idx: nat) (s1:RTL.state): state -> Prop :=
- | State_match path stack f sp pc rs m s2:
- (fn_path f)!pc = Some path ->
- (idx <= path.(psize))%nat ->
- isteps ge (path.(psize)-idx) f sp rs m pc = Some s2 ->
- s1 = State stack f sp s2.(ipc) s2.(irs) s2.(imem) ->
- match_inst_states_goal idx s1 (State stack f sp pc rs m).
-
-Definition match_inst_states (idx: nat) (s1:RTL.state) (s2:state): Prop :=
- if is_inst s1 then match_inst_states_goal idx s1 s2 else s1 = state_RTL s2.
-
-Definition match_states (idx: nat) (s1:RTL.state) (s2:state): Prop :=
- match_inst_states idx s1 s2
- /\ wf_stackframe (stack_of s2).
-
-(** *** Auxiliary lemmas of completeness *)
-Lemma istep_complete t i stack f sp rs m pc s':
- RTL.step ge (State stack f sp pc rs m) t s' ->
- (fn_code f)!pc = Some i ->
- default_succ i <> None ->
- t = E0 /\ exists st, istep ge i sp rs m = Some st /\ s'=(State stack f sp st.(ipc) st.(irs) st.(imem)).
-Proof.
- intros H X; inversion H; simpl; subst; try rewrite X in * |-; clear X; simplify_someHyps; try congruence;
- (split; auto); simplify_someHyps; eexists; split; simplify_someHyps; eauto.
- all: explore_destruct; simplify_SOME a.
-Qed.
-
-Lemma stuttering path idx stack f sp rs m pc st t s1':
- isteps ge (path.(psize)-(S idx)) f sp rs m pc = Some st ->
- (fn_path f)!pc = Some path ->
- (S idx <= path.(psize))%nat ->
- st.(icontinue) = true ->
- RTL.step ge (State stack f sp st.(ipc) st.(irs) st.(imem)) t s1' ->
- t = E0 /\ match_inst_states idx s1' (State stack f sp pc rs m).
-Proof.
- intros PSTEP PATH BOUND CONT RSTEP; exploit (internal_node_path (path.(psize)-(S idx))); omega || eauto.
- intros (i & pc' & Hi & Hpc & DUM).
- unfold nth_default_succ_inst in Hi.
- erewrite isteps_normal_exit in Hi; eauto.
- exploit istep_complete; congruence || eauto.
- intros (SILENT & st0 & STEP0 & EQ).
- intuition; subst; unfold match_inst_states; simpl.
- intros; refine (State_match _ _ path stack f sp pc rs m _ PATH _ _ _); simpl; omega || eauto.
- set (ps:=path.(psize)). enough (ps - idx = S (ps - (S idx)))%nat as ->; try omega.
- erewrite <- isteps_step_right; eauto.
-Qed.
-
-Lemma normal_exit path stack f sp rs m pc st t s1':
- isteps ge path.(psize) f sp rs m pc = Some st ->
- (fn_path f)!pc = Some path ->
- st.(icontinue) = true ->
- RTL.step ge (State stack f sp st.(ipc) st.(irs) st.(imem)) t s1' ->
- wf_stackframe stack ->
- exists s2',
- (path_last_step ge pge stack f sp st.(ipc) st.(irs) st.(imem)) t s2'
- /\ (exists idx', match_states idx' s1' s2').
-Proof.
- Local Hint Resolve istep_successors list_nth_z_in: core. (* Hint for path_entry proofs *)
- intros PSTEP PATH CONT RSTEP WF; exploit (final_node_path f path); eauto.
- intros (i & Hi & SUCCS).
- unfold nth_default_succ_inst in Hi.
- erewrite isteps_normal_exit in Hi; eauto.
- destruct (default_succ i) eqn:Hn0.
- + (* exec_istate *)
- exploit istep_complete; congruence || eauto.
- intros (SILENT & st0 & STEP0 & EQ); subst.
- exploit (exec_istate ge pge); eauto.
- eexists; intuition eauto.
- unfold match_states, match_inst_states; simpl.
- destruct (initialize_path (*fn_code f*) (fn_path f) (ipc st0)) as (path0 & Hpath0); eauto.
- exists (path0.(psize)); intuition eauto.
- econstructor; eauto.
- * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega.
- * simpl; eauto.
- + generalize Hi; inversion RSTEP; clear RSTEP; subst; (repeat (simplify_someHyp; simpl in * |- * )); try congruence; eauto.
- - (* Icall *)
- intros; exploit find_RTL_function_match; eauto.
- intros (fd' & MATCHfd & Hfd'); subst.
- exploit (exec_Icall ge pge); eauto.
- eexists; intuition eauto.
- eexists O; unfold match_states, match_inst_states; simpl; intuition eauto.
- rewrite wf_stackframe_cons; intuition simpl; eauto.
- - (* Itailcall *)
- intros; exploit find_RTL_function_match; eauto.
- intros (fd' & MATCHfd & Hfd'); subst.
- exploit (exec_Itailcall ge pge); eauto.
- eexists; intuition eauto.
- eexists O; unfold match_states, match_inst_states; simpl; intuition eauto.
- - (* Ibuiltin *)
- intros; exploit exec_Ibuiltin; eauto.
- eexists; intuition eauto.
- unfold match_states, match_inst_states; simpl.
- destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto.
- exists path0.(psize); intuition eauto.
- econstructor; eauto.
- * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega.
- * simpl; eauto.
- - (* Ijumptable *)
- intros; exploit exec_Ijumptable; eauto.
- eexists; intuition eauto.
- unfold match_states, match_inst_states; simpl.
- destruct (initialize_path (*fn_code f*) (fn_path f) pc') as (path0 & Hpath0); eauto.
- exists path0.(psize); intuition eauto.
- econstructor; eauto.
- * enough (path0.(psize)-path0.(psize)=0)%nat as ->; simpl; eauto || omega.
- * simpl; eauto.
- - (* Ireturn *)
- intros; exploit exec_Ireturn; eauto.
- eexists; intuition eauto.
- eexists O; unfold match_states, match_inst_states; simpl; intuition eauto.
-Qed.
-
-Lemma path_step_complete stack f sp rs m pc t s1' idx path st:
- isteps ge (path.(psize)-idx) f sp rs m pc = Some st ->
- (fn_path f)!pc = Some path ->
- (idx <= path.(psize))%nat ->
- RTL.step ge (State stack f sp st.(ipc) st.(irs) st.(imem)) t s1' ->
- wf_stackframe stack ->
- exists idx' s2',
- (path_step ge pge path.(psize) stack f sp rs m pc t s2'
- \/ (t = E0 /\ s2'=(State stack f sp pc rs m) /\ (idx' < idx)%nat)
- \/ (exists path', path_step ge pge path.(psize) stack f sp rs m pc E0 (State stack f sp st.(ipc) st.(irs) st.(imem))
- /\ (fn_path f)!(ipc st) = Some path' /\ path'.(psize) = O
- /\ path_step ge pge path'.(psize) stack f sp st.(irs) st.(imem) st.(ipc) t s2')
- )
- /\ match_states idx' s1' s2'.
-Proof.
- Local Hint Resolve exec_early_exit exec_normal_exit: core.
- intros PSTEP PATH BOUND RSTEP WF; destruct (st.(icontinue)) eqn: CONT.
- destruct idx as [ | idx].
- + (* path_step on normal_exit *)
- assert (path.(psize)-0=path.(psize))%nat as HH by omega. rewrite HH in PSTEP. clear HH.
- exploit normal_exit; eauto.
- intros (s2' & LSTEP & (idx' & MATCH)).
- exists idx'; exists s2'; intuition eauto.
- + (* stuttering step *)
- exploit stuttering; eauto.
- unfold match_states; exists idx; exists (State stack f sp pc rs m);
- intuition.
- + (* one or two path_step on early_exit *)
- exploit (isteps_resize ge (path.(psize) - idx)%nat path.(psize)); eauto; try omega.
- clear PSTEP; intros PSTEP.
- (* TODO for clarification: move the assert below into a separate lemma *)
- assert (HPATH0: exists path0, (fn_path f)!(ipc st) = Some path0).
- { clear RSTEP.
- exploit isteps_inversion_early; eauto.
- intros (st0 & i & path0 & BOUND0 & PSTEP0 & CONT0 & PC0 & STEP0).
- exploit istep_early_exit; eauto.
- intros (X1 & X2 & EARLY_EXIT).
- destruct st as [cont pc0 rs0 m0]; simpl in * |- *; intuition subst.
- exploit (internal_node_path path0); omega || eauto.
- intros (i' & pc' & Hi' & Hpc' & ENTRY).
- unfold nth_default_succ_inst in Hi'.
- erewrite isteps_normal_exit in Hi'; eauto.
- clear pc' Hpc' STEP0 PSTEP0 BOUND0; try_simplify_someHyps; intros.
- destruct EARLY_EXIT as [cond args ifnot]; simpl in ENTRY;
- destruct (initialize_path (*fn_code f*) (fn_path f) pc0); eauto.
- }
- destruct HPATH0 as (path1 & Hpath1).
- destruct (path1.(psize)) as [|ps] eqn:Hpath1size.
- * (* two step case *)
- exploit (normal_exit path1); try rewrite Hpath1size; simpl; eauto.
- simpl; intros (s2' & LSTEP & (idx' & MATCH)).
- exists idx'. exists s2'. constructor; auto.
- right. right. eexists; intuition eauto.
- (* now, prove the last step *)
- rewrite Hpath1size; exploit exec_normal_exit. 4:{ eauto. }
- - simpl; eauto.
- - simpl; eauto.
- - simpl; eauto.
- * (* single step case *)
- exploit (stuttering path1 ps stack f sp (irs st) (imem st) (ipc st)); simpl; auto.
- - { rewrite Hpath1size; enough (S ps-S ps=0)%nat as ->; try omega. simpl; eauto. }
- - omega.
- - simpl; eauto.
- - simpl; eauto.
- - intuition subst.
- repeat eexists; intuition eauto.
-Qed.
-
-Lemma step_noninst_complete s1 t s1' s2:
- is_inst s1 = false ->
- s1 = state_RTL s2 ->
- RTL.step ge s1 t s1' ->
- wf_stackframe (stack_of s2) ->
- exists s2', step ge pge s2 t s2' /\ exists idx, match_states idx s1' s2'.
-Proof.
- intros H0 H1 H2 WFSTACK; destruct s2; subst; simpl in * |- *; try congruence;
- inversion H2; clear H2; subst; try_simplify_someHyps; try congruence.
- + (* exec_function_internal *)
- destruct f; simpl in H3; inversion H3; subst; clear H3.
- eexists; constructor 1.
- * eapply exec_function_internal; eauto.
- * unfold match_states, match_inst_states; simpl.
- destruct (initialize_path (*fn_code f*) (fn_path f) (fn_entrypoint (fn_RTL f))) as (path & Hpath); eauto.
- exists path.(psize). constructor; auto.
- econstructor; eauto.
- - set (ps:=path.(psize)). enough (ps-ps=O)%nat as ->; simpl; eauto.
- omega.
- - simpl; auto.
- + (* exec_function_external *)
- destruct f; simpl in H3 |-; inversion H3; subst; clear H3.
- eexists; constructor 1.
- * apply exec_function_external; eauto.
- * unfold match_states, match_inst_states; simpl. exists O; auto.
- + (* exec_return *)
- destruct stack eqn: Hstack; simpl in H1; inversion H1; clear H1; subst.
- destruct s0 eqn: Hs0; simpl in H0; inversion H0; clear H0; subst.
- eexists; constructor 1.
- * apply exec_return.
- * unfold match_states, match_inst_states; simpl.
- rewrite wf_stackframe_cons in WFSTACK.
- destruct WFSTACK as (H0 & H1); simpl in H0.
- destruct (initialize_path (*fn_code f0*) (fn_path f0) pc0) as (path & Hpath); eauto.
- exists path.(psize). constructor; auto.
- econstructor; eauto.
- - set (ps:=path.(psize)). enough (ps-ps=O)%nat as ->; simpl; eauto.
- omega.
- - simpl; auto.
-Qed.
-
-(** *** The main completeness lemma and the simulation theorem...*)
-Lemma step_complete s1 t s1' idx s2:
- match_states idx s1 s2 ->
- RTL.step ge s1 t s1' ->
- exists idx' s2', (plus (step ge) pge s2 t s2' \/ (t = E0 /\ s2=s2' /\ (idx' < idx)%nat)) /\ match_states idx' s1' s2'.
-Proof.
- Local Hint Resolve plus_one plus_two exec_path: core.
- unfold match_states at 1, match_inst_states. intros (IS_INST & WFSTACK). destruct (is_inst s1) eqn: His1.
- - clear His1; destruct IS_INST as [path stack f sp pc rs m s2 X X0 X1 X2]; auto; subst; simpl in * |- *.
- intros STEP; exploit path_step_complete; eauto.
- intros (idx' & s2' & H0 & H1).
- eexists; eexists; eauto.
- destruct H0 as [H0|[H0|(path'&H0)]]; intuition subst; eauto.
- - intros; exploit step_noninst_complete; eauto.
- intros (s2' & STEP & (idx0 & MATCH)).
- exists idx0; exists s2'; intuition auto.
-Qed.
-
-Theorem RTLpath_complete: forward_simulation (RTL.semantics p) (semantics p).
-Proof.
- eapply (Forward_simulation (L1:=RTL.semantics p) (L2:=semantics p) lt match_states).
- constructor 1; simpl.
- - apply lt_wf.
- - unfold match_states, match_inst_states. destruct 1; simpl; exists O.
- destruct (find_funct_ptr_RTL_preserv b f) as (f0 & X1 & X2); subst; eauto.
- exists (Callstate nil f0 nil m0). simpl; split; try econstructor; eauto.
- + apply (Genv.init_mem_match (match_RTL_prog p)); auto.
- + rewrite (Genv.find_symbol_match (match_RTL_prog p)).
- rewrite (match_program_main (match_RTL_prog p)); eauto.
- - unfold final_state, match_states, match_inst_states. intros i s1 s2 r (H0 & H1) H2; destruct H2.
- destruct s2; simpl in * |- *; inversion H0; subst.
- constructor.
- - Local Hint Resolve star_refl: core.
- intros; exploit step_complete; eauto.
- destruct 1 as (idx' & s2' & X).
- exists idx'. exists s2'. intuition (subst; eauto).
- - intros id; destruct (senv_preserved p); simpl in * |-. intuition.
-Qed.
-
-End COMPLETENESS.
diff --git a/kvx/lib/RTLpathLivegen.v b/kvx/lib/RTLpathLivegen.v
deleted file mode 100644
index 1f0ebe3c..00000000
--- a/kvx/lib/RTLpathLivegen.v
+++ /dev/null
@@ -1,290 +0,0 @@
-(** Building a RTLpath program with liveness annotation.
-*)
-
-
-Require Import Coqlib.
-Require Import Maps.
-Require Import Lattice.
-Require Import AST.
-Require Import Op.
-Require Import Registers.
-Require Import Globalenvs Smallstep RTL RTLpath.
-Require Import Bool Errors.
-Require Import Program.
-
-Local Open Scope lazy_bool_scope.
-
-Local Open Scope option_monad_scope.
-
-Axiom build_path_map: RTL.function -> path_map.
-
-Extract Constant build_path_map => "RTLpathLivegenaux.build_path_map".
-
-Fixpoint list_mem (rl: list reg) (alive: Regset.t) {struct rl}: bool :=
- match rl with
- | nil => true
- | r1 :: rs => Regset.mem r1 alive &&& list_mem rs alive
- end.
-
-Definition exit_checker {A} (pm: path_map) (alive: Regset.t) (pc: node) (v:A): option A :=
- SOME path <- pm!pc IN
- ASSERT Regset.subset path.(input_regs) alive IN
- Some v.
-
-Lemma exit_checker_path_entry A (pm: path_map) (alive: Regset.t) (pc: node) (v:A) res:
- exit_checker pm alive pc v = Some res -> path_entry pm pc.
-Proof.
- unfold exit_checker, path_entry.
- inversion_SOME path; simpl; congruence.
-Qed.
-
-Lemma exit_checker_res A (pm: path_map) (alive: Regset.t) (pc: node) (v:A) res:
- exit_checker pm alive pc v = Some res -> v=res.
-Proof.
- unfold exit_checker, path_entry.
- inversion_SOME path; try_simplify_someHyps.
- inversion_ASSERT; try_simplify_someHyps.
-Qed.
-
-(* FIXME - what about trap? *)
-Definition iinst_checker (pm: path_map) (alive: Regset.t) (i: instruction): option (Regset.t * node) :=
- match i with
- | Inop pc' => Some (alive, pc')
- | Iop op args dst pc' =>
- ASSERT list_mem args alive IN
- Some (Regset.add dst alive, pc')
- | Iload _ chunk addr args dst pc' =>
- ASSERT list_mem args alive IN
- Some (Regset.add dst alive, pc')
- | Istore chunk addr args src pc' =>
- ASSERT Regset.mem src alive IN
- ASSERT list_mem args alive IN
- Some (alive, pc')
- | Icond cond args ifso ifnot _ =>
- ASSERT list_mem args alive IN
- exit_checker pm alive ifso (alive, ifnot)
- | _ => None (* TODO jumptable ? *)
- end.
-
-
-Local Hint Resolve exit_checker_path_entry: core.
-
-Lemma iinst_checker_path_entry (pm: path_map) (alive: Regset.t) (i: instruction) res pc:
- iinst_checker pm alive i = Some res ->
- early_exit i = Some pc -> path_entry pm pc.
-Proof.
- destruct i; simpl; try_simplify_someHyps; subst.
- inversion_ASSERT; try_simplify_someHyps.
-Qed.
-
-Lemma iinst_checker_default_succ (pm: path_map) (alive: Regset.t) (i: instruction) res pc:
- iinst_checker pm alive i = Some res ->
- pc = snd res ->
- default_succ i = Some pc.
-Proof.
- destruct i; simpl; try_simplify_someHyps; subst;
- repeat (inversion_ASSERT); try_simplify_someHyps.
- intros; exploit exit_checker_res; eauto.
- intros; subst. simpl; auto.
-Qed.
-
-Fixpoint ipath_checker (ps:nat) (f: RTL.function) (pm: path_map) (alive: Regset.t) (pc:node): option (Regset.t * node) :=
- match ps with
- | O => Some (alive, pc)
- | S p =>
- SOME i <- f.(fn_code)!pc IN
- SOME res <- iinst_checker pm alive i IN
- ipath_checker p f pm (fst res) (snd res)
- end.
-
-Lemma ipath_checker_wellformed f pm ps: forall alive pc res,
- ipath_checker ps f pm alive pc = Some res ->
- wellformed_path f.(fn_code) pm 0 (snd res) ->
- wellformed_path f.(fn_code) pm ps pc.
-Proof.
- induction ps; simpl; try_simplify_someHyps.
- inversion_SOME i; inversion_SOME res'.
- intros. eapply wf_internal_node; eauto.
- * eapply iinst_checker_default_succ; eauto.
- * intros; eapply iinst_checker_path_entry; eauto.
-Qed.
-
-Definition reg_option_mem (or: option reg) (alive: Regset.t) :=
- match or with None => true | Some r => Regset.mem r alive end.
-
-Definition reg_sum_mem (ros: reg + ident) (alive: Regset.t) :=
- match ros with inl r => Regset.mem r alive | inr s => true end.
-
-(* NB: definition following [regmap_setres] in [RTL.step] semantics *)
-Definition reg_builtin_res (res: builtin_res reg) (alive: Regset.t): Regset.t :=
- match res with
- | BR r => Regset.add r alive
- | _ => alive
- end.
-
-Fixpoint exit_list_checker (pm: path_map) (alive: Regset.t) (l: list node): bool :=
- match l with
- | nil => true
- | pc::l' => exit_checker pm alive pc tt &&& exit_list_checker pm alive l'
- end.
-
-Lemma lazy_and_Some_true A (o: option A) (b: bool): o &&& b = true <-> (exists v, o = Some v) /\ b = true.
-Proof.
- destruct o; simpl; intuition.
- - eauto.
- - firstorder. try_simplify_someHyps.
-Qed.
-
-Lemma lazy_and_Some_tt_true (o: option unit) (b: bool): o &&& b = true <-> o = Some tt /\ b = true.
-Proof.
- intros; rewrite lazy_and_Some_true; firstorder.
- destruct x; auto.
-Qed.
-
-
-Lemma exit_list_checker_correct pm alive l pc:
- exit_list_checker pm alive l = true -> List.In pc l -> exit_checker pm alive pc tt = Some tt.
-Proof.
- intros EXIT PC; induction l; intuition.
- simpl in * |-. rewrite lazy_and_Some_tt_true in EXIT.
- firstorder (subst; eauto).
-Qed.
-
-Local Hint Resolve exit_list_checker_correct: core.
-
-Definition inst_checker (pm: path_map) (alive: Regset.t) (i: instruction): option unit :=
- match i with
- | Icall sig ros args res pc' =>
- ASSERT list_mem args alive IN
- ASSERT reg_sum_mem ros alive IN
- exit_checker pm (Regset.add res alive) pc' tt
- | Itailcall sig ros args =>
- ASSERT list_mem args alive IN
- ASSERT reg_sum_mem ros alive IN
- Some tt
- | Ibuiltin ef args res pc' =>
- ASSERT list_mem (params_of_builtin_args args) alive IN
- exit_checker pm (reg_builtin_res res alive) pc' tt
- | Ijumptable arg tbl =>
- ASSERT Regset.mem arg alive IN
- ASSERT exit_list_checker pm alive tbl IN
- Some tt
- | Ireturn optarg =>
- ASSERT (reg_option_mem optarg) alive IN
- Some tt
- | _ =>
- SOME res <- iinst_checker pm alive i IN
- exit_checker pm (fst res) (snd res) tt
- end.
-
-Lemma inst_checker_wellformed (c:code) pc (pm: path_map) (alive: Regset.t) (i: instruction):
- inst_checker pm alive i = Some tt ->
- c!pc = Some i -> wellformed_path c pm 0 pc.
-Proof.
- intros CHECK PC. eapply wf_last_node; eauto.
- clear c pc PC. intros pc PC.
- destruct i; simpl in * |- *; intuition (subst; eauto);
- try (generalize CHECK; clear CHECK; try (inversion_SOME path); repeat inversion_ASSERT; try_simplify_someHyps).
- intros X; exploit exit_checker_res; eauto.
- clear X. intros; subst; eauto.
-Qed.
-
-Definition path_checker (f: RTL.function) pm (pc: node) (path:path_info): option unit :=
- SOME res <- ipath_checker (path.(psize)) f pm (path.(input_regs)) pc IN
- SOME i <- f.(fn_code)!(snd res) IN
- inst_checker pm (fst res) i.
-
-Lemma path_checker_wellformed f pm pc path:
- path_checker f pm pc path = Some tt -> wellformed_path (f.(fn_code)) pm (path.(psize)) pc.
-Proof.
- unfold path_checker.
- inversion_SOME res.
- inversion_SOME i.
- intros; eapply ipath_checker_wellformed; eauto.
- eapply inst_checker_wellformed; eauto.
-Qed.
-
-Fixpoint list_path_checker f pm (l:list (node*path_info)): bool :=
- match l with
- | nil => true
- | (pc, path)::l' =>
- path_checker f pm pc path &&& list_path_checker f pm l'
- end.
-
-Lemma list_path_checker_correct f pm l:
- list_path_checker f pm l = true -> forall e, List.In e l -> path_checker f pm (fst e) (snd e) = Some tt.
-Proof.
- intros CHECKER e H; induction l as [|(pc & path) l]; intuition.
- simpl in * |- *. rewrite lazy_and_Some_tt_true in CHECKER. intuition (subst; auto).
-Qed.
-
-Definition function_checker (f: RTL.function) pm: bool :=
- pm!(f.(fn_entrypoint)) &&& list_path_checker f pm (PTree.elements pm).
-
-Lemma function_checker_correct f pm pc path:
- function_checker f pm = true ->
- pm!pc = Some path ->
- path_checker f pm pc path = Some tt.
-Proof.
- unfold function_checker; rewrite lazy_and_Some_true.
- intros (ENTRY & PATH) PC.
- exploit list_path_checker_correct; eauto.
- - eapply PTree.elements_correct; eauto.
- - simpl; auto.
-Qed.
-
-Lemma function_checker_wellformed_path_map f pm:
- function_checker f pm = true -> wellformed_path_map f.(fn_code) pm.
-Proof.
- unfold wellformed_path_map.
- intros; eapply path_checker_wellformed; eauto.
- intros; eapply function_checker_correct; eauto.
-Qed.
-
-Lemma function_checker_path_entry f pm:
- function_checker f pm = true -> path_entry pm (f.(fn_entrypoint)).
-Proof.
- unfold function_checker; rewrite lazy_and_Some_true;
- unfold path_entry. firstorder congruence.
-Qed.
-
-Definition liveness_ok_function (f: function): Prop :=
- forall pc path, f.(fn_path)!pc = Some path -> path_checker f f.(fn_path) pc path = Some tt.
-
-Program Definition transf_function (f: RTL.function): { r: res function | forall f', r = OK f' -> liveness_ok_function f' /\ f'.(fn_RTL) = f } :=
- let pm := build_path_map f in
- match function_checker f pm with
- | true => OK {| fn_RTL := f; fn_path := pm |}
- | false => Error(msg "RTLpathGen: function_checker failed")
- end.
-Obligation 1.
- apply function_checker_path_entry; auto.
-Qed.
-Obligation 2.
- apply function_checker_wellformed_path_map; auto.
-Qed.
-Obligation 3.
- unfold liveness_ok_function; simpl; intros; intuition.
- apply function_checker_correct; auto.
-Qed.
-
-Definition transf_fundef (f: RTL.fundef) : res fundef :=
- transf_partial_fundef (fun f => ` (transf_function f)) f.
-
-Inductive liveness_ok_fundef: fundef -> Prop :=
- | liveness_ok_Internal f: liveness_ok_function f -> liveness_ok_fundef (Internal f)
- | liveness_ok_External ef: liveness_ok_fundef (External ef).
-
-Lemma transf_fundef_correct f f':
- transf_fundef f = OK f' -> (liveness_ok_fundef f') /\ fundef_RTL f' = f.
-Proof.
- intros TRANSF; destruct f; simpl; monadInv TRANSF.
- - destruct (transf_function f) as [res H]; simpl in * |- *; auto.
- destruct (H _ EQ).
- intuition subst; auto. apply liveness_ok_Internal; auto.
- - intuition. apply liveness_ok_External; auto.
-Qed.
-
-Definition transf_program (p: RTL.program) : res program :=
- transform_partial_program transf_fundef p.
-
diff --git a/kvx/lib/RTLpathLivegenaux.ml b/kvx/lib/RTLpathLivegenaux.ml
deleted file mode 100644
index dd971db8..00000000
--- a/kvx/lib/RTLpathLivegenaux.ml
+++ /dev/null
@@ -1,309 +0,0 @@
-open RTL
-open RTLpath
-open Registers
-open Maps
-open Camlcoq
-open Datatypes
-open Kildall
-open Lattice
-
-let debug_flag = ref false
-
-let dprintf fmt = let open Printf in
- match !debug_flag with
- | true -> printf fmt
- | false -> ifprintf stdout fmt
-
-let get_some = function
-| None -> failwith "Got None instead of Some _"
-| Some thing -> thing
-
-let successors_inst = function
-| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n]
-| Icond (_,_,n1,n2,_) -> [n1; n2]
-| Ijumptable (_,l) -> l
-| Itailcall _ | Ireturn _ -> []
-
-let predicted_successor = function
-| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> Some n
-| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> None
-| Icond (_,_,n1,n2,p) -> (
- match p with
- | Some true -> Some n1
- | Some false -> Some n2
- | None -> None )
-| Ijumptable _ | Itailcall _ | Ireturn _ -> None
-
-let non_predicted_successors i =
- match predicted_successor i with
- | None -> successors_inst i
- | Some n -> List.filter (fun n' -> n != n') (successors_inst i)
-
-let rec list_to_regset = function
- | [] -> Regset.empty
- | r::l -> Regset.add r (list_to_regset l)
-
-let get_input_regs i =
- let empty = Regset.empty in
- match i with
- | Inop _ -> empty
- | Iop (_,lr,_,_) | Iload (_,_,_,lr,_,_) | Icond (_,lr,_,_,_) -> list_to_regset lr
- | Istore (_,_,lr,r,_) -> Regset.add r (list_to_regset lr)
- | Icall (_, ri, lr, _, _) | Itailcall (_, ri, lr) -> begin
- let rs = list_to_regset lr in
- match ri with
- | Coq_inr _ -> rs
- | Coq_inl r -> Regset.add r rs
- end
- | Ibuiltin (_, lbr, _, _) -> list_to_regset @@ AST.params_of_builtin_args lbr
- | Ijumptable (r, _) -> Regset.add r empty
- | Ireturn opr -> (match opr with Some r -> Regset.add r empty | None -> empty)
-
-let get_output_reg i =
- match i with
- | Inop _ | Istore _ | Icond _ | Itailcall _ | Ijumptable _ | Ireturn _ -> None
- | Iop (_, _, r, _) | Iload (_, _, _, _, r, _) | Icall (_, _, _, r, _) -> Some r
- | Ibuiltin (_, _, brr, _) -> (match brr with AST.BR r -> Some r | _ -> None)
-
-(* adapted from Linearizeaux.get_join_points *)
-let get_join_points code entry =
- let reached = ref (PTree.map (fun n i -> false) code) in
- let reached_twice = ref (PTree.map (fun n i -> false) code) in
- let rec traverse pc =
- if get_some @@ PTree.get pc !reached then begin
- if not (get_some @@ PTree.get pc !reached_twice) then
- reached_twice := PTree.set pc true !reached_twice
- end else begin
- reached := PTree.set pc true !reached;
- traverse_succs (successors_inst @@ get_some @@ PTree.get pc code)
- end
- and traverse_succs = function
- | [] -> ()
- | [pc] -> traverse pc
- | pc :: l -> traverse pc; traverse_succs l
- in traverse entry; !reached_twice
-
-(* Does not set the input_regs and liveouts field *)
-let get_path_map code entry join_points =
- let visited = ref (PTree.map (fun n i -> false) code) in
- let path_map = ref PTree.empty in
- let rec dig_path e =
- let psize = ref (-1) in
- let path_successors = ref [] in
- let rec dig_path_rec n : (path_info * node list) option =
- if not (get_some @@ PTree.get n !visited) then
- let inst = get_some @@ PTree.get n code in
- begin
- visited := PTree.set n true !visited;
- psize := !psize + 1;
- let successor = match predicted_successor inst with
- | None -> None
- | Some n' -> if get_some @@ PTree.get n' join_points then None else Some n'
- in match successor with
- | Some n' -> begin
- path_successors := !path_successors @ non_predicted_successors inst;
- dig_path_rec n'
- end
- | None -> Some ({ psize = (Camlcoq.Nat.of_int !psize);
- input_regs = Regset.empty; output_regs = Regset.empty },
- !path_successors @ successors_inst inst)
- end
- else None
- in match dig_path_rec e with
- | None -> ()
- | Some ret ->
- let (path_info, succs) = ret in
- begin
- path_map := PTree.set e path_info !path_map;
- List.iter dig_path succs
- end
- in begin
- dig_path entry;
- !path_map
- end
-
-let print_regset rs = begin
- dprintf "[";
- List.iter (fun n -> dprintf "%d " (P.to_int n)) (Regset.elements rs);
- dprintf "]"
-end
-
-let print_ptree_regset pt = begin
- dprintf "[";
- List.iter (fun (n, rs) ->
- dprintf "\n\t";
- dprintf "%d: " (P.to_int n);
- print_regset rs
- ) (PTree.elements pt);
- dprintf "]"
-end
-
-let transfer f pc after = let open Liveness in
- match PTree.get pc f.fn_code with
- | Some i ->
- (match i with
- | Inop _ -> after
- | Iop (_, args, res, _) ->
- reg_list_live args (Regset.remove res after)
- | Iload (_, _, _, args, dst, _) ->
- reg_list_live args (Regset.remove dst after)
- | Istore (_, _, args, src, _) ->
- reg_list_live args (Regset.add src after)
- | Icall (_, ros, args, res, _) ->
- reg_list_live args (reg_sum_live ros (Regset.remove res after))
- | Itailcall (_, ros, args) ->
- reg_list_live args (reg_sum_live ros Regset.empty)
- | Ibuiltin (_, args, res, _) ->
- reg_list_live (AST.params_of_builtin_args args)
- (reg_list_dead (AST.params_of_builtin_res res) after)
- | Icond (_, args, _, _, _) ->
- reg_list_live args after
- | Ijumptable (arg, _) ->
- Regset.add arg after
- | Ireturn optarg ->
- reg_option_live optarg Regset.empty)
- | None -> Regset.empty
-
-module RegsetLat = LFSet(Regset)
-
-module DS = Backward_Dataflow_Solver(RegsetLat)(NodeSetBackward)
-
-let analyze f =
- let liveouts = get_some @@ DS.fixpoint f.fn_code successors_instr (transfer f) in
- PTree.map (fun n _ -> let lo = PMap.get n liveouts in transfer f n lo) f.fn_code
-
-(** OLD CODE - If needed to have our own kildall
-
-let transfer after = let open Liveness in function
- | Inop _ -> after
- | Iop (_, args, res, _) ->
- reg_list_live args (Regset.remove res after)
- | Iload (_, _, _, args, dst, _) ->
- reg_list_live args (Regset.remove dst after)
- | Istore (_, _, args, src, _) ->
- reg_list_live args (Regset.add src after)
- | Icall (_, ros, args, res, _) ->
- reg_list_live args (reg_sum_live ros (Regset.remove res after))
- | Itailcall (_, ros, args) ->
- reg_list_live args (reg_sum_live ros Regset.empty)
- | Ibuiltin (_, args, res, _) ->
- reg_list_live (AST.params_of_builtin_args args)
- (reg_list_dead (AST.params_of_builtin_res res) after)
- | Icond (_, args, _, _, _) ->
- reg_list_live args after
- | Ijumptable (arg, _) ->
- Regset.add arg after
- | Ireturn optarg ->
- reg_option_live optarg Regset.empty
-
-let get_last_nodes f =
- let visited = ref (PTree.map (fun n i -> false) f.fn_code) in
- let rec step n =
- let inst = get_some @@ PTree.get n f.fn_code in
- let successors = successors_inst inst in
- if get_some @@ PTree.get n !visited then []
- else begin
-
-let analyze f =
- let liveness = ref (PTree.map (fun n i -> None) f.fn_code) in
- let predecessors = Duplicateaux.get_predecessors_rtl f.fn_code in
- let last_nodes = get_last_nodes f in
- let rec step liveout n = (* liveout is the input_regs from the successor *)
- let inst = get_some @@ PTree.get n f.fn_code in
- let continue = ref true in
- let alive = match get_some @@ PTree.get n !liveness with
- | None -> transfer liveout inst
- | Some pre_alive -> begin
- let union = Regset.union pre_alive liveout in
- let new_alive = transfer union inst in
- (if Regset.equal pre_alive new_alive then continue := false);
- new_alive
- end
- in begin
- liveness := PTree.set n (Some alive) !liveness;
- if !continue then
- let preds = get_some @@ PTree.get n predecessors in
- List.iter (step alive) preds
- end
- in begin
- List.iter (step Regset.empty) last_nodes;
- let liveness_noopt = PTree.map (fun n i -> get_some i) !liveness in
- begin
- debug_flag := true;
- dprintf "Liveness: "; print_ptree_regset liveness_noopt; dprintf "\n";
- debug_flag := false;
- liveness_noopt
- end
- end
-*)
-
-let rec traverse code n size =
- let inst = get_some @@ PTree.get n code in
- if (size == 0) then inst
- else
- let n' = get_some @@ predicted_successor inst in
- traverse code n' (size-1)
-
-let get_outputs liveness code n pi =
- let last_instruction = traverse code n (Camlcoq.Nat.to_int pi.psize) in
- let path_last_successors = successors_inst last_instruction in
- let list_input_regs = List.map (
- fun n -> get_some @@ PTree.get n liveness
- ) path_last_successors in
- List.fold_left Regset.union Regset.empty list_input_regs
-
-let set_pathmap_liveness f pm =
- let liveness = analyze f in
- let new_pm = ref PTree.empty in
- let code = f.fn_code in
- begin
- dprintf "Liveness: "; print_ptree_regset liveness; dprintf "\n";
- List.iter (fun (n, pi) ->
- let inputs = get_some @@ PTree.get n liveness in
- let outputs = get_outputs liveness code n pi in
- new_pm := PTree.set n
- {psize=pi.psize; input_regs=inputs; output_regs=outputs} !new_pm
- ) (PTree.elements pm);
- !new_pm
- end
-
-let print_true_nodes booltree = begin
- dprintf "[";
- List.iter (fun (n,b) ->
- if b then dprintf "%d " (P.to_int n)
- ) (PTree.elements booltree);
- dprintf "]";
-end
-
-let print_path_info pi = begin
- dprintf "(psize=%d; " (Camlcoq.Nat.to_int pi.psize);
- dprintf "input_regs=";
- print_regset pi.input_regs;
- dprintf "; output_regs=";
- print_regset pi.output_regs;
- dprintf ")"
-end
-
-let print_path_map path_map = begin
- dprintf "[";
- List.iter (fun (n,pi) ->
- dprintf "\n\t";
- dprintf "%d: " (P.to_int n);
- print_path_info pi
- ) (PTree.elements path_map);
- dprintf "]"
-end
-
-let build_path_map f =
- let code = f.fn_code in
- let entry = f.fn_entrypoint in
- let join_points = get_join_points code entry in
- let path_map = set_pathmap_liveness f @@ get_path_map code entry join_points in
- begin
- dprintf "Join points: ";
- print_true_nodes join_points;
- dprintf "\nPath map: ";
- print_path_map path_map;
- dprintf "\n";
- path_map
- end
diff --git a/kvx/lib/RTLpathLivegenproof.v b/kvx/lib/RTLpathLivegenproof.v
deleted file mode 100644
index 1bf1af72..00000000
--- a/kvx/lib/RTLpathLivegenproof.v
+++ /dev/null
@@ -1,736 +0,0 @@
-(** Proofs of the liveness properties from the liveness checker of RTLpathLivengen.
-*)
-
-
-Require Import Coqlib.
-Require Import Maps.
-Require Import Lattice.
-Require Import AST.
-Require Import Op.
-Require Import Registers.
-Require Import Globalenvs Smallstep RTL RTLpath RTLpathLivegen.
-Require Import Bool Errors Linking Values Events.
-Require Import Program.
-
-Definition match_prog (p: RTL.program) (tp: program) :=
- match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
-
-Lemma transf_program_match:
- forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog.
-Proof.
- intros. eapply match_transform_partial_program_contextual; eauto.
-Qed.
-
-Section PRESERVATION.
-
-Variables prog: RTL.program.
-Variables tprog: program.
-Hypothesis TRANSL: match_prog prog tprog.
-Let ge := Genv.globalenv prog.
-Let tpge := Genv.globalenv tprog.
-Let tge := Genv.globalenv (RTLpath.transf_program tprog).
-
-Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof.
- rewrite <- (Genv.find_symbol_match TRANSL).
- apply (Genv.find_symbol_match (match_prog_RTL tprog)).
-Qed.
-
-Lemma senv_transitivity x y z: Senv.equiv x y -> Senv.equiv y z -> Senv.equiv x z.
-Proof.
- unfold Senv.equiv. intuition congruence.
-Qed.
-
-Lemma senv_preserved: Senv.equiv ge tge.
-Proof.
- eapply senv_transitivity. { eapply (Genv.senv_match TRANSL). }
- eapply RTLpath.senv_preserved.
-Qed.
-
-Lemma function_ptr_preserved v f: Genv.find_funct_ptr ge v = Some f ->
- exists tf, Genv.find_funct_ptr tpge v = Some tf /\ transf_fundef f = OK tf.
-Proof.
- intros; apply (Genv.find_funct_ptr_transf_partial TRANSL); eauto.
-Qed.
-
-
-Lemma function_ptr_RTL_preserved v f: Genv.find_funct_ptr ge v = Some f -> Genv.find_funct_ptr tge v = Some f.
-Proof.
- intros; exploit function_ptr_preserved; eauto.
- intros (tf & Htf & TRANS).
- exploit (Genv.find_funct_ptr_match (match_prog_RTL tprog)); eauto.
- intros (cunit & tf0 & X & Y & DUM); subst.
- unfold tge. rewrite X.
- exploit transf_fundef_correct; eauto.
- intuition subst; auto.
-Qed.
-
-Lemma find_function_preserved ros rs fd:
- RTL.find_function ge ros rs = Some fd -> RTL.find_function tge ros rs = Some fd.
-Proof.
- intros H; assert (X: exists tfd, find_function tpge ros rs = Some tfd /\ fd = fundef_RTL tfd).
- * destruct ros; simpl in * |- *.
- + intros; exploit (Genv.find_funct_match TRANSL); eauto.
- intros (cuint & tf & H1 & H2 & H3); subst; repeat econstructor; eauto.
- exploit transf_fundef_correct; eauto.
- intuition auto.
- + rewrite <- (Genv.find_symbol_match TRANSL) in H.
- unfold tpge. destruct (Genv.find_symbol _ i); simpl; try congruence.
- exploit function_ptr_preserved; eauto.
- intros (tf & H1 & H2); subst; repeat econstructor; eauto.
- exploit transf_fundef_correct; eauto.
- intuition auto.
- * destruct X as (tf & X1 & X2); subst.
- eapply find_function_RTL_match; eauto.
-Qed.
-
-
-Local Hint Resolve symbols_preserved senv_preserved: core.
-
-Lemma transf_program_RTL_correct:
- forward_simulation (RTL.semantics prog) (RTL.semantics (RTLpath.transf_program tprog)).
-Proof.
- eapply forward_simulation_step with (match_states:=fun (s1 s2:RTL.state) => s1=s2); simpl; eauto.
- - eapply senv_preserved.
- - (* initial states *)
- intros s1 INIT. destruct INIT as [b f m0 ge0 INIT SYMB PTR SIG]. eexists; intuition eauto.
- econstructor; eauto.
- + intros; eapply (Genv.init_mem_match (match_prog_RTL tprog)). apply (Genv.init_mem_match TRANSL); auto.
- + rewrite symbols_preserved.
- replace (prog_main (RTLpath.transf_program tprog)) with (prog_main prog).
- * eapply SYMB.
- * erewrite (match_program_main (match_prog_RTL tprog)). erewrite (match_program_main TRANSL); auto.
- + exploit function_ptr_RTL_preserved; eauto.
- - intros; subst; auto.
- - intros s t s2 STEP s1 H; subst.
- eexists; intuition.
- destruct STEP.
- + (* Inop *) eapply exec_Inop; eauto.
- + (* Iop *) eapply exec_Iop; eauto.
- erewrite eval_operation_preserved; eauto.
- + (* Iload *) eapply exec_Iload; eauto.
- erewrite eval_addressing_preserved; eauto.
- + (* Iload notrap1 *) eapply exec_Iload_notrap1; eauto.
- erewrite eval_addressing_preserved; eauto.
- + (* Iload notrap2 *) eapply exec_Iload_notrap2; eauto.
- erewrite eval_addressing_preserved; eauto.
- + (* Istore *) eapply exec_Istore; eauto.
- erewrite eval_addressing_preserved; eauto.
- + (* Icall *)
- eapply RTL.exec_Icall; eauto.
- eapply find_function_preserved; eauto.
- + (* Itailcall *)
- eapply RTL.exec_Itailcall; eauto.
- eapply find_function_preserved; eauto.
- + (* Ibuiltin *)
- eapply RTL.exec_Ibuiltin; eauto.
- * eapply eval_builtin_args_preserved; eauto.
- * eapply external_call_symbols_preserved; eauto.
- + (* Icond *)
- eapply exec_Icond; eauto.
- + (* Ijumptable *)
- eapply RTL.exec_Ijumptable; eauto.
- + (* Ireturn *)
- eapply RTL.exec_Ireturn; eauto.
- + (* exec_function_internal *)
- eapply RTL.exec_function_internal; eauto.
- + (* exec_function_external *)
- eapply RTL.exec_function_external; eauto.
- eapply external_call_symbols_preserved; eauto.
- + (* exec_return *)
- eapply RTL.exec_return; eauto.
-Qed.
-
-Theorem transf_program_correct:
- forward_simulation (RTL.semantics prog) (RTLpath.semantics tprog).
-Proof.
- eapply compose_forward_simulations.
- + eapply transf_program_RTL_correct.
- + eapply RTLpath_complete.
-Qed.
-
-
-(* Properties used in hypothesis of [RTLpathLiveproofs.step_eqlive] theorem *)
-Theorem all_fundef_liveness_ok b f:
- Genv.find_funct_ptr tpge b = Some f -> liveness_ok_fundef f.
-Proof.
- unfold match_prog, match_program in TRANSL.
- unfold Genv.find_funct_ptr, tpge; simpl; intro X.
- destruct (Genv.find_def_match_2 TRANSL b) as [|f0 y H]; try congruence.
- destruct y as [tf0|]; try congruence.
- inversion X as [H1]. subst. clear X.
- remember (@Gfun fundef unit f) as f2.
- destruct H as [ctx' f1 f2 H0|]; try congruence.
- inversion Heqf2 as [H2]. subst; clear Heqf2.
- exploit transf_fundef_correct; eauto.
- intuition.
-Qed.
-
-End PRESERVATION.
-
-Local Open Scope lazy_bool_scope.
-Local Open Scope option_monad_scope.
-
-Local Notation ext alive := (fun r => Regset.In r alive).
-
-Lemma regset_add_spec live r1 r2: Regset.In r1 (Regset.add r2 live) <-> (r1 = r2 \/ Regset.In r1 live).
-Proof.
- destruct (Pos.eq_dec r1 r2).
- - subst. intuition; eapply Regset.add_1; auto.
- - intuition.
- * right. eapply Regset.add_3; eauto.
- * eapply Regset.add_2; auto.
-Qed.
-
-Definition eqlive_reg (alive: Regset.elt -> Prop) (rs1 rs2: regset): Prop :=
- forall r, (alive r) -> rs1#r = rs2#r.
-
-Lemma eqlive_reg_refl alive rs: eqlive_reg alive rs rs.
-Proof.
- unfold eqlive_reg; auto.
-Qed.
-
-Lemma eqlive_reg_symmetry alive rs1 rs2: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs1.
-Proof.
- unfold eqlive_reg; intros; symmetry; auto.
-Qed.
-
-Lemma eqlive_reg_trans alive rs1 rs2 rs3: eqlive_reg alive rs1 rs2 -> eqlive_reg alive rs2 rs3 -> eqlive_reg alive rs1 rs3.
-Proof.
- unfold eqlive_reg; intros H0 H1 r H. rewrite H0; eauto.
-Qed.
-
-Lemma eqlive_reg_update (alive: Regset.elt -> Prop) rs1 rs2 r v: eqlive_reg (fun r1 => r1 <> r /\ alive r1) rs1 rs2 -> eqlive_reg alive (rs1 # r <- v) (rs2 # r <- v).
-Proof.
- unfold eqlive_reg; intros EQLIVE r0 ALIVE.
- destruct (Pos.eq_dec r r0) as [H|H].
- - subst. rewrite! Regmap.gss. auto.
- - rewrite! Regmap.gso; auto.
-Qed.
-
-Lemma eqlive_reg_monotonic (alive1 alive2: Regset.elt -> Prop) rs1 rs2: eqlive_reg alive2 rs1 rs2 -> (forall r, alive1 r -> alive2 r) -> eqlive_reg alive1 rs1 rs2.
-Proof.
- unfold eqlive_reg; intuition.
-Qed.
-
-Lemma eqlive_reg_triv rs1 rs2: (forall r, rs1#r = rs2#r) <-> eqlive_reg (fun _ => True) rs1 rs2.
-Proof.
- unfold eqlive_reg; intuition.
-Qed.
-
-Lemma eqlive_reg_triv_trans alive rs1 rs2 rs3: eqlive_reg alive rs1 rs2 -> (forall r, rs2#r = rs3#r) -> eqlive_reg alive rs1 rs3.
-Proof.
- rewrite eqlive_reg_triv; intros; eapply eqlive_reg_trans; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; eauto.
-Qed.
-
-Local Hint Resolve Regset.mem_2 Regset.subset_2: core.
-
-Lemma lazy_and_true (b1 b2: bool): b1 &&& b2 = true <-> b1 = true /\ b2 = true.
-Proof.
- destruct b1; simpl; intuition.
-Qed.
-
-Lemma list_mem_correct (rl: list reg) (alive: Regset.t):
- list_mem rl alive = true -> forall r, List.In r rl -> ext alive r.
-Proof.
- induction rl; simpl; try rewrite lazy_and_true; intuition subst; auto.
-Qed.
-
-Lemma eqlive_reg_list (alive: Regset.elt -> Prop) args rs1 rs2: eqlive_reg alive rs1 rs2 -> (forall r, List.In r args -> (alive r)) -> rs1##args = rs2##args.
-Proof.
- induction args; simpl; auto.
- intros EQLIVE ALIVE; rewrite IHargs; auto.
- unfold eqlive_reg in EQLIVE.
- rewrite EQLIVE; auto.
-Qed.
-
-Lemma eqlive_reg_listmem (alive: Regset.t) args rs1 rs2: eqlive_reg (ext alive) rs1 rs2 -> list_mem args alive = true -> rs1##args = rs2##args.
-Proof.
- intros; eapply eqlive_reg_list; eauto.
- intros; eapply list_mem_correct; eauto.
-Qed.
-
-Record eqlive_istate alive (st1 st2: istate): Prop :=
- { eqlive_continue: icontinue st1 = icontinue st2;
- eqlive_ipc: ipc st1 = ipc st2;
- eqlive_irs: eqlive_reg alive (irs st1) (irs st2);
- eqlive_imem: (imem st1) = (imem st2) }.
-
-Lemma iinst_checker_eqlive ge sp pm alive i res rs1 rs2 m st1:
- eqlive_reg (ext alive) rs1 rs2 ->
- iinst_checker pm alive i = Some res ->
- istep ge i sp rs1 m = Some st1 ->
- exists st2, istep ge i sp rs2 m = Some st2 /\ eqlive_istate (ext (fst res)) st1 st2.
-Proof.
- intros EQLIVE.
- destruct i; simpl; try_simplify_someHyps.
- - (* Inop *)
- repeat (econstructor; eauto).
- - (* Iop *)
- inversion_ASSERT; try_simplify_someHyps.
- inversion_SOME v. intros EVAL.
- erewrite <- eqlive_reg_listmem; eauto.
- try_simplify_someHyps.
- repeat (econstructor; simpl; eauto).
- eapply eqlive_reg_update.
- eapply eqlive_reg_monotonic; eauto.
- intros r0; rewrite regset_add_spec.
- intuition.
- - (* Iload *)
- inversion_ASSERT; try_simplify_someHyps.
- destruct t. (* TODO - simplify that proof ? *)
- + inversion_SOME a0. intros EVAL.
- erewrite <- eqlive_reg_listmem; eauto.
- try_simplify_someHyps.
- inversion_SOME v; try_simplify_someHyps.
- repeat (econstructor; simpl; eauto).
- eapply eqlive_reg_update.
- eapply eqlive_reg_monotonic; eauto.
- intros r0; rewrite regset_add_spec.
- intuition.
- + erewrite <- (eqlive_reg_listmem _ _ rs1 rs2); eauto.
- destruct (eval_addressing _ _ _ _).
- * destruct (Memory.Mem.loadv _ _ _).
- ** intros. inv H1. repeat (econstructor; simpl; eauto).
- eapply eqlive_reg_update.
- eapply eqlive_reg_monotonic; eauto.
- intros r0; rewrite regset_add_spec.
- intuition.
- ** intros. inv H1. repeat (econstructor; simpl; eauto).
- eapply eqlive_reg_update.
- eapply eqlive_reg_monotonic; eauto.
- intros r0; rewrite regset_add_spec.
- intuition.
- * intros. inv H1. repeat (econstructor; simpl; eauto).
- eapply eqlive_reg_update.
- eapply eqlive_reg_monotonic; eauto.
- intros r0; rewrite regset_add_spec.
- intuition.
- - (* Istore *)
- (repeat inversion_ASSERT); try_simplify_someHyps.
- inversion_SOME a0. intros EVAL.
- erewrite <- eqlive_reg_listmem; eauto.
- rewrite <- (EQLIVE r); auto.
- inversion_SOME v; try_simplify_someHyps.
- try_simplify_someHyps.
- repeat (econstructor; simpl; eauto).
- - (* Icond *)
- inversion_ASSERT.
- inversion_SOME b. intros EVAL.
- intros ARGS; erewrite <- eqlive_reg_listmem; eauto.
- try_simplify_someHyps.
- repeat (econstructor; simpl; eauto).
- exploit exit_checker_res; eauto.
- intro; subst; simpl. auto.
-Qed.
-
-Lemma iinst_checker_istep_continue ge sp pm alive i res rs m st:
- iinst_checker pm alive i = Some res ->
- istep ge i sp rs m = Some st ->
- icontinue st = true ->
- (snd res)=(ipc st).
-Proof.
- intros; exploit iinst_checker_default_succ; eauto.
- erewrite istep_normal_exit; eauto.
- congruence.
-Qed.
-
-Lemma exit_checker_eqlive A (pm: path_map) (alive: Regset.t) (pc: node) (v:A) res rs1 rs2:
- exit_checker pm alive pc v = Some res ->
- eqlive_reg (ext alive) rs1 rs2 ->
- exists path, pm!pc = Some path /\ eqlive_reg (ext path.(input_regs)) rs1 rs2.
-Proof.
- unfold exit_checker.
- inversion_SOME path.
- inversion_ASSERT. try_simplify_someHyps.
- repeat (econstructor; eauto).
- intros; eapply eqlive_reg_monotonic; eauto.
- intros; exploit Regset.subset_2; eauto.
-Qed.
-
-Lemma iinst_checker_eqlive_stopped ge sp pm alive i res rs1 rs2 m st1:
- eqlive_reg (ext alive) rs1 rs2 ->
- istep ge i sp rs1 m = Some st1 ->
- iinst_checker pm alive i = Some res ->
- icontinue st1 = false ->
- exists path st2, pm!(ipc st1) = Some path /\ istep ge i sp rs2 m = Some st2 /\ eqlive_istate (ext path.(input_regs)) st1 st2.
-Proof.
- intros EQLIVE.
- set (tmp := istep ge i sp rs2).
- destruct i; simpl; try_simplify_someHyps; repeat (inversion_ASSERT || inversion_SOME b); try_simplify_someHyps; try congruence.
- 1-3: explore_destruct; simpl; try_simplify_someHyps; repeat (inversion_ASSERT || inversion_SOME b); try_simplify_someHyps; try congruence.
- (* Icond *)
- unfold tmp; clear tmp; simpl.
- intros EVAL; erewrite <- eqlive_reg_listmem; eauto.
- try_simplify_someHyps.
- destruct b eqn:EQb; simpl in * |-; try congruence.
- intros; exploit exit_checker_eqlive; eauto.
- intros (path & PATH & EQLIVE2).
- repeat (econstructor; simpl; eauto).
-Qed.
-
-Lemma ipath_checker_eqlive_normal ge ps (f:function) sp pm: forall alive pc res rs1 rs2 m st1,
- eqlive_reg (ext alive) rs1 rs2 ->
- ipath_checker ps f pm alive pc = Some res ->
- isteps ge ps f sp rs1 m pc = Some st1 ->
- icontinue st1 = true ->
- exists st2, isteps ge ps f sp rs2 m pc = Some st2 /\ eqlive_istate (ext (fst res)) st1 st2.
-Proof.
- induction ps as [|ps]; simpl; try_simplify_someHyps.
- - repeat (econstructor; simpl; eauto).
- - inversion_SOME i; try_simplify_someHyps.
- inversion_SOME res0.
- inversion_SOME st0.
- intros.
- exploit iinst_checker_eqlive; eauto.
- destruct 1 as (st2 & ISTEP & [CONT PC RS MEM]).
- try_simplify_someHyps.
- rewrite <- CONT, <- MEM, <- PC.
- destruct (icontinue st0) eqn:CONT'.
- * intros; exploit iinst_checker_istep_continue; eauto.
- rewrite <- PC; intros X; rewrite X in * |-. eauto.
- * try_simplify_someHyps.
- congruence.
-Qed.
-
-Lemma ipath_checker_isteps_continue ge ps (f:function) sp pm: forall alive pc res rs m st,
- ipath_checker ps f pm alive pc = Some res ->
- isteps ge ps f sp rs m pc = Some st ->
- icontinue st = true ->
- (snd res)=(ipc st).
-Proof.
- induction ps as [|ps]; simpl; try_simplify_someHyps.
- inversion_SOME i; try_simplify_someHyps.
- inversion_SOME res0.
- inversion_SOME st0.
- destruct (icontinue st0) eqn:CONT'.
- - intros; exploit iinst_checker_istep_continue; eauto.
- intros EQ; rewrite EQ in * |-; clear EQ; eauto.
- - try_simplify_someHyps; congruence.
-Qed.
-
-Lemma ipath_checker_eqlive_stopped ge ps (f:function) sp pm: forall alive pc res rs1 rs2 m st1,
- eqlive_reg (ext alive) rs1 rs2 ->
- ipath_checker ps f pm alive pc = Some res ->
- isteps ge ps f sp rs1 m pc = Some st1 ->
- icontinue st1 = false ->
- exists path st2, pm!(ipc st1) = Some path /\ isteps ge ps f sp rs2 m pc = Some st2 /\ eqlive_istate (ext path.(input_regs)) st1 st2.
-Proof.
- induction ps as [|ps]; simpl; try_simplify_someHyps; try congruence.
- inversion_SOME i; try_simplify_someHyps.
- inversion_SOME res0.
- inversion_SOME st0.
- intros.
- destruct (icontinue st0) eqn:CONT'; try_simplify_someHyps; intros.
- * intros; exploit iinst_checker_eqlive; eauto.
- destruct 1 as (st2 & ISTEP & [CONT PC RS MEM]).
- exploit iinst_checker_istep_continue; eauto.
- intros PC'.
- try_simplify_someHyps.
- rewrite PC', <- CONT, <- MEM, <- PC, CONT'.
- eauto.
- * intros; exploit iinst_checker_eqlive_stopped; eauto.
- intros EQLIVE; generalize EQLIVE; destruct 1 as (path & st2 & PATH & ISTEP & [CONT PC RS MEM]).
- try_simplify_someHyps.
- rewrite <- CONT, <- MEM, <- PC, CONT'.
- try_simplify_someHyps.
-Qed.
-
-Inductive eqlive_stackframes: stackframe -> stackframe -> Prop :=
- | eqlive_stackframes_intro path res f sp pc rs1 rs2
- (LIVE: liveness_ok_function f)
- (PATH: f.(fn_path)!pc = Some path)
- (EQUIV: forall v, eqlive_reg (ext path.(input_regs)) (rs1 # res <- v) (rs2 # res <- v)):
- eqlive_stackframes (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2).
-
-Inductive eqlive_states: state -> state -> Prop :=
- | eqlive_states_intro
- path st1 st2 f sp pc rs1 rs2 m
- (STACKS: list_forall2 eqlive_stackframes st1 st2)
- (LIVE: liveness_ok_function f)
- (PATH: f.(fn_path)!pc = Some path)
- (EQUIV: eqlive_reg (ext path.(input_regs)) rs1 rs2):
- eqlive_states (State st1 f sp pc rs1 m) (State st2 f sp pc rs2 m)
- | eqlive_states_call st1 st2 f args m
- (LIVE: liveness_ok_fundef f)
- (STACKS: list_forall2 eqlive_stackframes st1 st2):
- eqlive_states (Callstate st1 f args m) (Callstate st2 f args m)
- | eqlive_states_return st1 st2 v m
- (STACKS: list_forall2 eqlive_stackframes st1 st2):
- eqlive_states (Returnstate st1 v m) (Returnstate st2 v m).
-
-
-Section LivenessProperties.
-
-Variable prog: program.
-
-Let pge := Genv.globalenv prog.
-Let ge := Genv.globalenv (RTLpath.transf_program prog).
-
-Hypothesis all_fundef_liveness_ok: forall b f,
- Genv.find_funct_ptr pge b = Some f ->
- liveness_ok_fundef f.
-
-Lemma find_funct_liveness_ok v fd:
- Genv.find_funct pge v = Some fd -> liveness_ok_fundef fd.
-Proof.
- unfold Genv.find_funct.
- destruct v; try congruence.
- destruct (Integers.Ptrofs.eq_dec _ _); try congruence.
- eapply all_fundef_liveness_ok; eauto.
-Qed.
-
-Lemma find_function_liveness_ok ros rs f:
- find_function pge ros rs = Some f -> liveness_ok_fundef f.
-Proof.
- destruct ros as [r|i]; simpl.
- - intros; eapply find_funct_liveness_ok; eauto.
- - destruct (Genv.find_symbol pge i); try congruence.
- eapply all_fundef_liveness_ok; eauto.
-Qed.
-
-Lemma find_function_eqlive alive ros rs1 rs2:
- eqlive_reg (ext alive) rs1 rs2 ->
- reg_sum_mem ros alive = true ->
- find_function pge ros rs1 = find_function pge ros rs2.
-Proof.
- intros EQLIVE.
- destruct ros; simpl; auto.
- intros H; erewrite (EQLIVE r); eauto.
-Qed.
-
-Lemma inst_checker_from_iinst_checker i sp rs m st pm alive:
- istep ge i sp rs m = Some st ->
- inst_checker pm alive i = (SOME res <- iinst_checker pm alive i IN exit_checker pm (fst res) (snd res) tt).
-Proof.
- destruct i; simpl; try congruence.
-Qed.
-
-Lemma exit_checker_eqlive_ext1 (pm: path_map) (alive: Regset.t) (pc: node) r rs1 rs2:
- exit_checker pm (Regset.add r alive) pc tt = Some tt ->
- eqlive_reg (ext alive) rs1 rs2 ->
- exists path, pm!pc = Some path /\ (forall v, eqlive_reg (ext path.(input_regs)) (rs1 # r <- v) (rs2 # r <- v)).
-Proof.
- unfold exit_checker.
- inversion_SOME path.
- inversion_ASSERT. try_simplify_someHyps.
- repeat (econstructor; eauto).
- intros; eapply eqlive_reg_update; eauto.
- eapply eqlive_reg_monotonic; eauto.
- intros r0 [X1 X2]; exploit Regset.subset_2; eauto.
- rewrite regset_add_spec. intuition subst.
-Qed.
-
-Local Hint Resolve in_or_app: local.
-Lemma eqlive_eval_builtin_args alive rs1 rs2 sp m args vargs:
- eqlive_reg alive rs1 rs2 ->
- Events.eval_builtin_args ge (fun r => rs1 # r) sp m args vargs ->
- (forall r, List.In r (params_of_builtin_args args) -> alive r) ->
- Events.eval_builtin_args ge (fun r => rs2 # r) sp m args vargs.
-Proof.
- unfold Events.eval_builtin_args.
- intros EQLIVE; induction 1 as [|a1 al b1 bl EVAL1 EVALL]; simpl.
- { econstructor; eauto. }
- intro X.
- assert (X1: eqlive_reg (fun r => In r (params_of_builtin_arg a1)) rs1 rs2).
- { eapply eqlive_reg_monotonic; eauto with local. }
- lapply IHEVALL; eauto with local.
- clear X IHEVALL; intro X. econstructor; eauto.
- generalize X1; clear EVALL X1 X.
- induction EVAL1; simpl; try (econstructor; eauto; fail).
- - intros X1; erewrite X1; [ econstructor; eauto | eauto ].
- - intros; econstructor.
- + eapply IHEVAL1_1; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; intros; eauto with local.
- + eapply IHEVAL1_2; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; intros; eauto with local.
- - intros; econstructor.
- + eapply IHEVAL1_1; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; intros; eauto with local.
- + eapply IHEVAL1_2; eauto.
- eapply eqlive_reg_monotonic; eauto.
- simpl; intros; eauto with local.
-Qed.
-
-Lemma exit_checker_eqlive_builtin_res (pm: path_map) (alive: Regset.t) (pc: node) rs1 rs2 (res:builtin_res reg):
- exit_checker pm (reg_builtin_res res alive) pc tt = Some tt ->
- eqlive_reg (ext alive) rs1 rs2 ->
- exists path, pm!pc = Some path /\ (forall vres, eqlive_reg (ext path.(input_regs)) (regmap_setres res vres rs1) (regmap_setres res vres rs2)).
-Proof.
- destruct res; simpl.
- - intros; exploit exit_checker_eqlive_ext1; eauto.
- - intros; exploit exit_checker_eqlive; eauto.
- intros (path & PATH & EQLIVE).
- eexists; intuition eauto.
- - intros; exploit exit_checker_eqlive; eauto.
- intros (path & PATH & EQLIVE).
- eexists; intuition eauto.
-Qed.
-
-Lemma exit_list_checker_eqlive (pm: path_map) (alive: Regset.t) (tbl: list node) rs1 rs2 pc: forall n,
- exit_list_checker pm alive tbl = true ->
- eqlive_reg (ext alive) rs1 rs2 ->
- list_nth_z tbl n = Some pc ->
- exists path, pm!pc = Some path /\ eqlive_reg (ext path.(input_regs)) rs1 rs2.
-Proof.
- induction tbl; simpl.
- - intros; try congruence.
- - intros n; rewrite lazy_and_Some_tt_true; destruct (zeq n 0) eqn: Hn.
- * try_simplify_someHyps; intuition.
- exploit exit_checker_eqlive; eauto.
- * intuition. eapply IHtbl; eauto.
-Qed.
-
-Lemma inst_checker_eqlive (f: function) sp alive pc i rs1 rs2 m stk1 stk2 t s1:
- list_forall2 eqlive_stackframes stk1 stk2 ->
- eqlive_reg (ext alive) rs1 rs2 ->
- liveness_ok_function f ->
- (fn_code f) ! pc = Some i ->
- path_last_step ge pge stk1 f sp pc rs1 m t s1 ->
- inst_checker (fn_path f) alive i = Some tt ->
- exists s2, path_last_step ge pge stk2 f sp pc rs2 m t s2 /\ eqlive_states s1 s2.
-Proof.
- intros STACKS EQLIVE LIVENESS PC;
- destruct 1 as [i' sp pc rs1 m st1|
- sp pc rs1 m sig ros args res pc' fd|
- st1 pc rs1 m sig ros args fd m'|
- sp pc rs1 m ef args res pc' vargs t vres m'|
- sp pc rs1 m arg tbl n pc' |
- st1 pc rs1 m optr m'];
- try_simplify_someHyps.
- + (* istate *)
- intros PC ISTEP. erewrite inst_checker_from_iinst_checker; eauto.
- inversion_SOME res.
- intros.
- destruct (icontinue st1) eqn: CONT.
- - (* CONT => true *)
- exploit iinst_checker_eqlive; eauto.
- destruct 1 as (st2 & ISTEP2 & [CONT' PC2 RS MEM]).
- repeat (econstructor; simpl; eauto).
- rewrite <- MEM, <- PC2.
- exploit exit_checker_eqlive; eauto.
- intros (path & PATH & EQLIVE2).
- eapply eqlive_states_intro; eauto.
- erewrite <- iinst_checker_istep_continue; eauto.
- - (* CONT => false *)
- intros; exploit iinst_checker_eqlive_stopped; eauto.
- destruct 1 as (path & st2 & PATH & ISTEP2 & [CONT2 PC2 RS MEM]).
- repeat (econstructor; simpl; eauto).
- rewrite <- MEM, <- PC2.
- eapply eqlive_states_intro; eauto.
- + (* Icall *)
- repeat inversion_ASSERT. intros.
- exploit exit_checker_eqlive_ext1; eauto.
- intros (path & PATH & EQLIVE2).
- eexists; split.
- - eapply exec_Icall; eauto.
- erewrite <- find_function_eqlive; eauto.
- - erewrite eqlive_reg_listmem; eauto.
- eapply eqlive_states_call; eauto.
- eapply find_function_liveness_ok; eauto.
- repeat (econstructor; eauto).
- + (* Itailcall *)
- repeat inversion_ASSERT. intros.
- eexists; split.
- - eapply exec_Itailcall; eauto.
- erewrite <- find_function_eqlive; eauto.
- - erewrite eqlive_reg_listmem; eauto.
- eapply eqlive_states_call; eauto.
- eapply find_function_liveness_ok; eauto.
- + (* Ibuiltin *)
- repeat inversion_ASSERT. intros.
- exploit exit_checker_eqlive_builtin_res; eauto.
- intros (path & PATH & EQLIVE2).
- eexists; split.
- - eapply exec_Ibuiltin; eauto.
- eapply eqlive_eval_builtin_args; eauto.
- intros; eapply list_mem_correct; eauto.
- - repeat (econstructor; simpl; eauto).
- + (* Ijumptable *)
- repeat inversion_ASSERT. intros.
- exploit exit_list_checker_eqlive; eauto.
- intros (path & PATH & EQLIVE2).
- eexists; split.
- - eapply exec_Ijumptable; eauto.
- erewrite <- EQLIVE; eauto.
- - repeat (econstructor; simpl; eauto).
- + (* Ireturn *)
- repeat inversion_ASSERT. intros.
- eexists; split.
- - eapply exec_Ireturn; eauto.
- - destruct optr; simpl in * |- *.
- * erewrite (EQLIVE r); eauto.
- eapply eqlive_states_return; eauto.
- * eapply eqlive_states_return; eauto.
-Qed.
-
-Lemma path_step_eqlive path stk1 f sp rs1 m pc t s1 stk2 rs2:
- path_step ge pge (psize path) stk1 f sp rs1 m pc t s1 ->
- list_forall2 eqlive_stackframes stk1 stk2 ->
- eqlive_reg (ext (input_regs path)) rs1 rs2 ->
- liveness_ok_function f ->
- (fn_path f) ! pc = Some path ->
- exists s2, path_step ge pge (psize path) stk2 f sp rs2 m pc t s2 /\ eqlive_states s1 s2.
-Proof.
- intros STEP STACKS EQLIVE LIVE PC.
- unfold liveness_ok_function in LIVE.
- exploit LIVE; eauto.
- unfold path_checker.
- inversion_SOME res; (* destruct res as [alive pc']. *) intros ICHECK. (* simpl. *)
- inversion_SOME i; intros PC'.
- destruct STEP as [st ISTEPS CONT|].
- - (* early_exit *)
- intros; exploit ipath_checker_eqlive_stopped; eauto.
- destruct 1 as (path2 & st2 & PATH & ISTEP2 & [CONT2 PC2 RS MEM]).
- repeat (econstructor; simpl; eauto).
- rewrite <- MEM, <- PC2.
- eapply eqlive_states_intro; eauto.
- - (* normal_exit *)
- intros; exploit ipath_checker_eqlive_normal; eauto.
- destruct 1 as (st2 & ISTEP2 & [CONT' PC2 RS MEM]).
- exploit ipath_checker_isteps_continue; eauto.
- intros PC3; rewrite <- PC3, <- PC2 in * |-.
- exploit inst_checker_eqlive; eauto.
- intros (s2 & LAST_STEP & EQLIVE2).
- eexists; split; eauto.
- eapply exec_normal_exit; eauto.
- rewrite <- PC3, <- MEM; auto.
-Qed.
-
-Theorem step_eqlive t s1 s1' s2:
- step ge pge s1 t s1' ->
- eqlive_states s1 s2 ->
- exists s2', step ge pge s2 t s2' /\ eqlive_states s1' s2'.
-Proof.
- destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ].
- - intros EQLIVE; inv EQLIVE; simplify_someHyps.
- intro PATH.
- exploit path_step_eqlive; eauto.
- intros (s2 & STEP2 & EQUIV2).
- eexists; split; eauto.
- eapply exec_path; eauto.
- - intros EQLIVE; inv EQLIVE; inv LIVE.
- exploit initialize_path. { eapply fn_entry_point_wf. }
- intros (path & Hpath).
- eexists; split.
- * eapply exec_function_internal; eauto.
- * eapply eqlive_states_intro; eauto.
- eapply eqlive_reg_refl.
- - intros EQLIVE; inv EQLIVE.
- eexists; split.
- * eapply exec_function_external; eauto.
- * eapply eqlive_states_return; eauto.
- - intros EQLIVE; inv EQLIVE.
- inversion STACKS as [|s1 st1 s' s2 STACK STACKS']; subst; clear STACKS.
- inv STACK.
- exists (State s2 f sp pc (rs2 # res <- vres) m); split.
- * apply exec_return.
- * eapply eqlive_states_intro; eauto.
-Qed.
-
-End LivenessProperties.
diff --git a/kvx/lib/RTLpathSE_impl.v b/kvx/lib/RTLpathSE_impl.v
deleted file mode 100644
index 87a064d5..00000000
--- a/kvx/lib/RTLpathSE_impl.v
+++ /dev/null
@@ -1,1529 +0,0 @@
-(** Implementation and refinement of the symbolic execution *)
-
-Require Import Coqlib Maps Floats.
-Require Import AST Integers Values Events Memory Globalenvs Smallstep.
-Require Import Op Registers.
-Require Import RTL RTLpath.
-Require Import Errors.
-Require Import RTLpathSE_theory RTLpathLivegenproof.
-Require Import Axioms RTLpathSE_simu_specs.
-
-Local Open Scope error_monad_scope.
-Local Open Scope option_monad_scope.
-
-Require Import Impure.ImpHCons.
-Import Notations.
-Import HConsing.
-
-Local Open Scope impure.
-Local Open Scope hse.
-
-Import ListNotations.
-Local Open Scope list_scope.
-
-Definition XDEBUG {A} (x:A) (k: A -> ?? pstring): ?? unit := RET tt. (* TO REMOVE DEBUG INFO *)
-(* Definition XDEBUG {A} (x:A) (k: A -> ?? pstring): ?? unit := DO s <~ k x;; println ("DEBUG simu_check:" +; s). (* TO INSERT DEBUG INFO *) *)
-
-Definition DEBUG (s: pstring): ?? unit := XDEBUG tt (fun _ => RET s).
-
-(** * Implementation of Data-structure use in Hash-consing *)
-
-Definition hsval_get_hid (hsv: hsval): hashcode :=
- match hsv with
- | HSinput _ hid => hid
- | HSop _ _ _ hid => hid
- | HSload _ _ _ _ _ hid => hid
- end.
-
-Definition list_hsval_get_hid (lhsv: list_hsval): hashcode :=
- match lhsv with
- | HSnil hid => hid
- | HScons _ _ hid => hid
- end.
-
-Definition hsmem_get_hid (hsm: hsmem): hashcode :=
- match hsm with
- | HSinit hid => hid
- | HSstore _ _ _ _ _ hid => hid
- end.
-
-Definition hsval_set_hid (hsv: hsval) (hid: hashcode): hsval :=
- match hsv with
- | HSinput r _ => HSinput r hid
- | HSop o lhsv hsm _ => HSop o lhsv hsm hid
- | HSload hsm trap chunk addr lhsv _ => HSload hsm trap chunk addr lhsv hid
- end.
-
-Definition list_hsval_set_hid (lhsv: list_hsval) (hid: hashcode): list_hsval :=
- match lhsv with
- | HSnil _ => HSnil hid
- | HScons hsv lhsv _ => HScons hsv lhsv hid
- end.
-
-Definition hsmem_set_hid (hsm: hsmem) (hid: hashcode): hsmem :=
- match hsm with
- | HSinit _ => HSinit hid
- | HSstore hsm chunk addr lhsv srce _ => HSstore hsm chunk addr lhsv srce hid
- end.
-
-
-Lemma hsval_set_hid_correct x y ge sp rs0 m0:
- hsval_set_hid x unknown_hid = hsval_set_hid y unknown_hid ->
- seval_hsval ge sp x rs0 m0 = seval_hsval ge sp y rs0 m0.
-Proof.
- destruct x, y; intro H; inversion H; subst; simpl; auto.
-Qed.
-Local Hint Resolve hsval_set_hid_correct: core.
-
-Lemma list_hsval_set_hid_correct x y ge sp rs0 m0:
- list_hsval_set_hid x unknown_hid = list_hsval_set_hid y unknown_hid ->
- seval_list_hsval ge sp x rs0 m0 = seval_list_hsval ge sp y rs0 m0.
-Proof.
- destruct x, y; intro H; inversion H; subst; simpl; auto.
-Qed.
-Local Hint Resolve list_hsval_set_hid_correct: core.
-
-Lemma hsmem_set_hid_correct x y ge sp rs0 m0:
- hsmem_set_hid x unknown_hid = hsmem_set_hid y unknown_hid ->
- seval_hsmem ge sp x rs0 m0 = seval_hsmem ge sp y rs0 m0.
-Proof.
- destruct x, y; intro H; inversion H; subst; simpl; auto.
-Qed.
-Local Hint Resolve hsmem_set_hid_correct: core.
-
-(** Now, we build the hash-Cons value from a "hash_eq".
-
- Informal specification:
- [hash_eq] must be consistent with the "hashed" constructors defined above.
-
- We expect that hashinfo values in the code of these "hashed" constructors verify:
- (hash_eq (hdata x) (hdata y) ~> true) <-> (hcodes x)=(hcodes y)
-*)
-
-
-Definition hsval_hash_eq (sv1 sv2: hsval): ?? bool :=
- match sv1, sv2 with
- | HSinput r1 _, HSinput r2 _ => struct_eq r1 r2 (* NB: really need a struct_eq here ? *)
- | HSop op1 lsv1 sm1 _, HSop op2 lsv2 sm2 _ =>
- DO b1 <~ phys_eq lsv1 lsv2;;
- DO b2 <~ phys_eq sm1 sm2;;
- if b1 && b2
- then struct_eq op1 op2 (* NB: really need a struct_eq here ? *)
- else RET false
- | HSload sm1 trap1 chk1 addr1 lsv1 _, HSload sm2 trap2 chk2 addr2 lsv2 _ =>
- DO b1 <~ phys_eq lsv1 lsv2;;
- DO b2 <~ phys_eq sm1 sm2;;
- DO b3 <~ struct_eq trap1 trap2;;
- DO b4 <~ struct_eq chk1 chk2;;
- if b1 && b2 && b3 && b4
- then struct_eq addr1 addr2
- else RET false
- | _,_ => RET false
- end.
-
-
-Lemma and_true_split a b: a && b = true <-> a = true /\ b = true.
-Proof.
- destruct a; simpl; intuition.
-Qed.
-
-Lemma hsval_hash_eq_correct x y:
- WHEN hsval_hash_eq x y ~> b THEN
- b = true -> hsval_set_hid x unknown_hid = hsval_set_hid y unknown_hid.
-Proof.
- destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
-Qed.
-Global Opaque hsval_hash_eq.
-Local Hint Resolve hsval_hash_eq_correct: wlp.
-
-Definition list_hsval_hash_eq (lsv1 lsv2: list_hsval): ?? bool :=
- match lsv1, lsv2 with
- | HSnil _, HSnil _ => RET true
- | HScons sv1 lsv1' _, HScons sv2 lsv2' _ =>
- DO b <~ phys_eq lsv1' lsv2';;
- if b
- then phys_eq sv1 sv2
- else RET false
- | _,_ => RET false
- end.
-
-Lemma list_hsval_hash_eq_correct x y:
- WHEN list_hsval_hash_eq x y ~> b THEN
- b = true -> list_hsval_set_hid x unknown_hid = list_hsval_set_hid y unknown_hid.
-Proof.
- destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
-Qed.
-Global Opaque list_hsval_hash_eq.
-Local Hint Resolve list_hsval_hash_eq_correct: wlp.
-
-Definition hsmem_hash_eq (sm1 sm2: hsmem): ?? bool :=
- match sm1, sm2 with
- | HSinit _, HSinit _ => RET true
- | HSstore sm1 chk1 addr1 lsv1 sv1 _, HSstore sm2 chk2 addr2 lsv2 sv2 _ =>
- DO b1 <~ phys_eq lsv1 lsv2;;
- DO b2 <~ phys_eq sm1 sm2;;
- DO b3 <~ phys_eq sv1 sv2;;
- DO b4 <~ struct_eq chk1 chk2;;
- if b1 && b2 && b3 && b4
- then struct_eq addr1 addr2
- else RET false
- | _,_ => RET false
- end.
-
-Lemma hsmem_hash_eq_correct x y:
- WHEN hsmem_hash_eq x y ~> b THEN
- b = true -> hsmem_set_hid x unknown_hid = hsmem_set_hid y unknown_hid.
-Proof.
- destruct x, y; wlp_simplify; try (rewrite !and_true_split in *); intuition; subst; try congruence.
-Qed.
-Global Opaque hsmem_hash_eq.
-Local Hint Resolve hsmem_hash_eq_correct: wlp.
-
-
-Definition hSVAL: hashP hsval := {| hash_eq := hsval_hash_eq; get_hid:=hsval_get_hid; set_hid:=hsval_set_hid |}.
-Definition hLSVAL: hashP list_hsval := {| hash_eq := list_hsval_hash_eq; get_hid:= list_hsval_get_hid; set_hid:= list_hsval_set_hid |}.
-Definition hSMEM: hashP hsmem := {| hash_eq := hsmem_hash_eq; get_hid:= hsmem_get_hid; set_hid:= hsmem_set_hid |}.
-
-Program Definition mk_hash_params: Dict.hash_params hsval :=
- {|
- Dict.test_eq := phys_eq;
- Dict.hashing := fun (ht: hsval) => RET (hsval_get_hid ht);
- Dict.log := fun hv =>
- DO hv_name <~ string_of_hashcode (hsval_get_hid hv);;
- println ("unexpected undef behavior of hashcode:" +; (CamlStr hv_name)) |}.
-Obligation 1.
- wlp_simplify.
-Qed.
-
-(** ** various auxiliary (trivial lemmas) *)
-Lemma hsilocal_refines_sreg ge sp rs0 m0 hst st:
- hsilocal_refines ge sp rs0 m0 hst st -> hsok_local ge sp rs0 m0 hst -> forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (si_sreg st r) rs0 m0.
-Proof.
- unfold hsilocal_refines; intuition.
-Qed.
-Local Hint Resolve hsilocal_refines_sreg: core.
-
-Lemma hsilocal_refines_valid_pointer ge sp rs0 m0 hst st:
- hsilocal_refines ge sp rs0 m0 hst st -> forall m b ofs, seval_smem ge sp st.(si_smem) rs0 m0 = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs.
-Proof.
- unfold hsilocal_refines; intuition.
-Qed.
-Local Hint Resolve hsilocal_refines_valid_pointer: core.
-
-Lemma hsilocal_refines_smem_refines ge sp rs0 m0 hst st:
- hsilocal_refines ge sp rs0 m0 hst st -> hsok_local ge sp rs0 m0 hst -> smem_refines ge sp rs0 m0 (hsi_smem hst) (st.(si_smem)).
-Proof.
- unfold hsilocal_refines; intuition.
-Qed.
-Local Hint Resolve hsilocal_refines_smem_refines: core.
-
-Lemma hsistate_refines_dyn_exits ge sp rs0 m0 hst st:
- hsistate_refines_dyn ge sp rs0 m0 hst st -> hsiexits_refines_dyn ge sp rs0 m0 (hsi_exits hst) (si_exits st).
-Proof.
- unfold hsistate_refines_dyn; intuition.
-Qed.
-Local Hint Resolve hsistate_refines_dyn_exits: core.
-
-Lemma hsistate_refines_dyn_local ge sp rs0 m0 hst st:
- hsistate_refines_dyn ge sp rs0 m0 hst st -> hsilocal_refines ge sp rs0 m0 (hsi_local hst) (si_local st).
-Proof.
- unfold hsistate_refines_dyn; intuition.
-Qed.
-Local Hint Resolve hsistate_refines_dyn_local: core.
-
-Lemma hsistate_refines_dyn_nested ge sp rs0 m0 hst st:
- hsistate_refines_dyn ge sp rs0 m0 hst st -> nested_sok ge sp rs0 m0 (si_local st) (si_exits st).
-Proof.
- unfold hsistate_refines_dyn; intuition.
-Qed.
-Local Hint Resolve hsistate_refines_dyn_nested: core.
-
-(** * Implementation of symbolic execution *)
-Section CanonBuilding.
-
-Variable hC_hsval: hashinfo hsval -> ?? hsval.
-
-Hypothesis hC_hsval_correct: forall hs,
- WHEN hC_hsval hs ~> hs' THEN forall ge sp rs0 m0,
- seval_hsval ge sp (hdata hs) rs0 m0 = seval_hsval ge sp hs' rs0 m0.
-
-Variable hC_list_hsval: hashinfo list_hsval -> ?? list_hsval.
-Hypothesis hC_list_hsval_correct: forall lh,
- WHEN hC_list_hsval lh ~> lh' THEN forall ge sp rs0 m0,
- seval_list_hsval ge sp (hdata lh) rs0 m0 = seval_list_hsval ge sp lh' rs0 m0.
-
-Variable hC_hsmem: hashinfo hsmem -> ?? hsmem.
-Hypothesis hC_hsmem_correct: forall hm,
- WHEN hC_hsmem hm ~> hm' THEN forall ge sp rs0 m0,
- seval_hsmem ge sp (hdata hm) rs0 m0 = seval_hsmem ge sp hm' rs0 m0.
-
-(* First, we wrap constructors for hashed values !*)
-
-Definition reg_hcode := 1.
-Definition op_hcode := 2.
-Definition load_hcode := 3.
-
-Definition hSinput_hcodes (r: reg) :=
- DO hc <~ hash reg_hcode;;
- DO hv <~ hash r;;
- RET [hc;hv].
-Extraction Inline hSinput_hcodes.
-
-Definition hSinput (r:reg): ?? hsval :=
- DO hv <~ hSinput_hcodes r;;
- hC_hsval {| hdata:=HSinput r unknown_hid; hcodes :=hv; |}.
-
-Lemma hSinput_correct r:
- WHEN hSinput r ~> hv THEN forall ge sp rs0 m0,
- sval_refines ge sp rs0 m0 hv (Sinput r).
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque hSinput.
-Local Hint Resolve hSinput_correct: wlp.
-
-Definition hSop_hcodes (op:operation) (lhsv: list_hsval) (hsm: hsmem) :=
- DO hc <~ hash op_hcode;;
- DO hv <~ hash op;;
- RET [hc;hv;list_hsval_get_hid lhsv; hsmem_get_hid hsm].
-Extraction Inline hSop_hcodes.
-
-Definition hSop (op:operation) (lhsv: list_hsval) (hsm: hsmem): ?? hsval :=
- DO hv <~ hSop_hcodes op lhsv hsm;;
- hC_hsval {| hdata:=HSop op lhsv hsm unknown_hid; hcodes :=hv |}.
-
-Lemma hSop_correct op lhsv hsm:
- WHEN hSop op lhsv hsm ~> hv THEN forall ge sp rs0 m0 lsv sm
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv)
- (MR: smem_refines ge sp rs0 m0 hsm sm),
- sval_refines ge sp rs0 m0 hv (Sop op lsv sm).
-Proof.
- wlp_simplify.
- rewrite <- LR, <- MR.
- auto.
-Qed.
-Global Opaque hSop.
-Local Hint Resolve hSop_correct: wlp.
-
-Definition hSload_hcodes (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval):=
- DO hc <~ hash load_hcode;;
- DO hv1 <~ hash trap;;
- DO hv2 <~ hash chunk;;
- DO hv3 <~ hash addr;;
- RET [hc; hsmem_get_hid hsm; hv1; hv2; hv3; list_hsval_get_hid lhsv].
-Extraction Inline hSload_hcodes.
-
-Definition hSload (hsm: hsmem) (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval): ?? hsval :=
- DO hv <~ hSload_hcodes hsm trap chunk addr lhsv;;
- hC_hsval {| hdata := HSload hsm trap chunk addr lhsv unknown_hid; hcodes := hv |}.
-
-Lemma hSload_correct hsm trap chunk addr lhsv:
- WHEN hSload hsm trap chunk addr lhsv ~> hv THEN forall ge sp rs0 m0 lsv sm
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv)
- (MR: smem_refines ge sp rs0 m0 hsm sm),
- sval_refines ge sp rs0 m0 hv (Sload sm trap chunk addr lsv).
-Proof.
- wlp_simplify.
- rewrite <- LR, <- MR.
- auto.
-Qed.
-Global Opaque hSload.
-Local Hint Resolve hSload_correct: wlp.
-
-Definition hSnil (_: unit): ?? list_hsval :=
- hC_list_hsval {| hdata := HSnil unknown_hid; hcodes := nil |}.
-
-Lemma hSnil_correct:
- WHEN hSnil() ~> hv THEN forall ge sp rs0 m0,
- list_sval_refines ge sp rs0 m0 hv Snil.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque hSnil.
-Local Hint Resolve hSnil_correct: wlp.
-
-Definition hScons (hsv: hsval) (lhsv: list_hsval): ?? list_hsval :=
- hC_list_hsval {| hdata := HScons hsv lhsv unknown_hid; hcodes := [hsval_get_hid hsv; list_hsval_get_hid lhsv] |}.
-
-Lemma hScons_correct hsv lhsv:
- WHEN hScons hsv lhsv ~> lhsv' THEN forall ge sp rs0 m0 sv lsv
- (VR: sval_refines ge sp rs0 m0 hsv sv)
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv),
- list_sval_refines ge sp rs0 m0 lhsv' (Scons sv lsv).
-Proof.
- wlp_simplify.
- rewrite <- VR, <- LR.
- auto.
-Qed.
-Global Opaque hScons.
-Local Hint Resolve hScons_correct: wlp.
-
-Definition hSinit (_: unit): ?? hsmem :=
- hC_hsmem {| hdata := HSinit unknown_hid; hcodes := nil |}.
-
-Lemma hSinit_correct:
- WHEN hSinit() ~> hm THEN forall ge sp rs0 m0,
- smem_refines ge sp rs0 m0 hm Sinit.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque hSinit.
-Local Hint Resolve hSinit_correct: wlp.
-
-Definition hSstore_hcodes (hsm: hsmem) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval) (srce: hsval):=
- DO hv1 <~ hash chunk;;
- DO hv2 <~ hash addr;;
- RET [hsmem_get_hid hsm; hv1; hv2; list_hsval_get_hid lhsv; hsval_get_hid srce].
-Extraction Inline hSstore_hcodes.
-
-Definition hSstore (hsm: hsmem) (chunk: memory_chunk) (addr: addressing) (lhsv: list_hsval) (srce: hsval): ?? hsmem :=
- DO hv <~ hSstore_hcodes hsm chunk addr lhsv srce;;
- hC_hsmem {| hdata := HSstore hsm chunk addr lhsv srce unknown_hid; hcodes := hv |}.
-
-Lemma hSstore_correct hsm chunk addr lhsv hsv:
- WHEN hSstore hsm chunk addr lhsv hsv ~> hsm' THEN forall ge sp rs0 m0 lsv sm sv
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv)
- (MR: smem_refines ge sp rs0 m0 hsm sm)
- (VR: sval_refines ge sp rs0 m0 hsv sv),
- smem_refines ge sp rs0 m0 hsm' (Sstore sm chunk addr lsv sv).
-Proof.
- wlp_simplify.
- rewrite <- LR, <- MR, <- VR.
- auto.
-Qed.
-Global Opaque hSstore.
-Local Hint Resolve hSstore_correct: wlp.
-
-Definition hsi_sreg_get (hst: PTree.t hsval) r: ?? hsval :=
- match PTree.get r hst with
- | None => hSinput r
- | Some sv => RET sv
- end.
-
-Lemma hsi_sreg_get_correct hst r:
- WHEN hsi_sreg_get hst r ~> hsv THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- sval_refines ge sp rs0 m0 hsv (f r).
-Proof.
- unfold hsi_sreg_eval, hsi_sreg_proj; wlp_simplify; rewrite <- RR; try_simplify_someHyps.
-Qed.
-Global Opaque hsi_sreg_get.
-Local Hint Resolve hsi_sreg_get_correct: wlp.
-
-Fixpoint hlist_args (hst: PTree.t hsval) (l: list reg): ?? list_hsval :=
- match l with
- | nil => hSnil()
- | r::l =>
- DO v <~ hsi_sreg_get hst r;;
- DO lhsv <~ hlist_args hst l;;
- hScons v lhsv
- end.
-
-Lemma hlist_args_correct hst l:
- WHEN hlist_args hst l ~> lhsv THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- list_sval_refines ge sp rs0 m0 lhsv (list_sval_inj (List.map f l)).
-Proof.
- induction l; wlp_simplify.
-Qed.
-Global Opaque hlist_args.
-Local Hint Resolve hlist_args_correct: wlp.
-
-(** ** Assignment of memory *)
-Definition hslocal_set_smem (hst:hsistate_local) hm :=
- {| hsi_smem := hm;
- hsi_ok_lsval := hsi_ok_lsval hst;
- hsi_sreg:= hsi_sreg hst
- |}.
-
-Lemma sok_local_set_mem ge sp rs0 m0 st sm:
- sok_local ge sp rs0 m0 (slocal_set_smem st sm)
- <-> (sok_local ge sp rs0 m0 st /\ seval_smem ge sp sm rs0 m0 <> None).
-Proof.
- unfold slocal_set_smem, sok_local; simpl; intuition (subst; eauto).
-Qed.
-
-Lemma hsok_local_set_mem ge sp rs0 m0 hst hsm:
- (seval_hsmem ge sp (hsi_smem hst) rs0 m0 = None -> seval_hsmem ge sp hsm rs0 m0 = None) ->
- hsok_local ge sp rs0 m0 (hslocal_set_smem hst hsm)
- <-> (hsok_local ge sp rs0 m0 hst /\ seval_hsmem ge sp hsm rs0 m0 <> None).
-Proof.
- unfold hslocal_set_smem, hsok_local; simpl; intuition.
-Qed.
-
-Lemma hslocal_set_mem_correct ge sp rs0 m0 hst st hsm sm:
- (seval_hsmem ge sp (hsi_smem hst) rs0 m0 = None -> seval_hsmem ge sp hsm rs0 m0 = None) ->
- (forall m b ofs, seval_smem ge sp sm rs0 m0 = Some m -> Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs) ->
- hsilocal_refines ge sp rs0 m0 hst st ->
- (hsok_local ge sp rs0 m0 hst -> smem_refines ge sp rs0 m0 hsm sm) ->
- hsilocal_refines ge sp rs0 m0 (hslocal_set_smem hst hsm) (slocal_set_smem st sm).
-Proof.
- intros PRESERV SMVALID (OKEQ & SMEMEQ' & REGEQ & MVALID) SMEMEQ.
- split; rewrite! hsok_local_set_mem; simpl; eauto; try tauto.
- rewrite sok_local_set_mem.
- intuition congruence.
-Qed.
-
-Definition hslocal_store (hst: hsistate_local) chunk addr args src: ?? hsistate_local :=
- let pt := hst.(hsi_sreg) in
- DO hargs <~ hlist_args pt args;;
- DO hsrc <~ hsi_sreg_get pt src;;
- DO hm <~ hSstore hst chunk addr hargs hsrc;;
- RET (hslocal_set_smem hst hm).
-
-Lemma hslocal_store_correct hst chunk addr args src:
- WHEN hslocal_store hst chunk addr args src ~> hst' THEN forall ge sp rs0 m0 st
- (REF: hsilocal_refines ge sp rs0 m0 hst st),
- hsilocal_refines ge sp rs0 m0 hst' (slocal_store st chunk addr args src).
-Proof.
- wlp_simplify.
- eapply hslocal_set_mem_correct; simpl; eauto.
- + intros X; erewrite H1; eauto.
- rewrite X. simplify_SOME z.
- + unfold hsilocal_refines in *;
- simplify_SOME z; intuition.
- erewrite <- Mem.storev_preserv_valid; [| eauto].
- eauto.
- + unfold hsilocal_refines in *; intuition eauto.
-Qed.
-Global Opaque hslocal_store.
-Local Hint Resolve hslocal_store_correct: wlp.
-
-(** ** Assignment of local state *)
-
-Definition hsist_set_local (hst: hsistate) (pc: node) (hnxt: hsistate_local): hsistate :=
- {| hsi_pc := pc; hsi_exits := hst.(hsi_exits); hsi_local:= hnxt |}.
-
-Lemma hsist_set_local_correct_stat hst st pc hnxt nxt:
- hsistate_refines_stat hst st ->
- hsistate_refines_stat (hsist_set_local hst pc hnxt) (sist_set_local st pc nxt).
-Proof.
- unfold hsistate_refines_stat; simpl; intuition.
-Qed.
-
-Lemma hsist_set_local_correct_dyn ge sp rs0 m0 hst st pc hnxt nxt:
- hsistate_refines_dyn ge sp rs0 m0 hst st ->
- hsilocal_refines ge sp rs0 m0 hnxt nxt ->
- (sok_local ge sp rs0 m0 nxt -> sok_local ge sp rs0 m0 (si_local st)) ->
- hsistate_refines_dyn ge sp rs0 m0 (hsist_set_local hst pc hnxt) (sist_set_local st pc nxt).
-Proof.
- unfold hsistate_refines_dyn; simpl.
- intros (EREF & LREF & NESTED) LREFN SOK; intuition.
- destruct NESTED as [|st0 se lse TOP NEST]; econstructor; simpl; auto.
-Qed.
-
-(** ** Assignment of registers *)
-
-(** locally new symbolic values during symbolic execution *)
-Inductive root_sval: Type :=
-| Rop (op: operation)
-| Rload (trap: trapping_mode) (chunk: memory_chunk) (addr: addressing)
-.
-
-Definition root_apply (rsv: root_sval) (lr: list reg) (st: sistate_local): sval :=
- let lsv := list_sval_inj (List.map (si_sreg st) lr) in
- let sm := si_smem st in
- match rsv with
- | Rop op => Sop op lsv sm
- | Rload trap chunk addr => Sload sm trap chunk addr lsv
- end.
-Coercion root_apply: root_sval >-> Funclass.
-
-Definition hSop_hSinit (op:operation) (lhsv: list_hsval): ?? hsval :=
- DO hsi <~ hSinit ();;
- hSop op lhsv hsi (** magically remove the dependency on sm ! *)
- .
-
-Lemma hSop_hSinit_correct op lhsv:
- WHEN hSop_hSinit op lhsv ~> hv THEN forall ge sp rs0 m0 lsv sm m
- (MEM: seval_smem ge sp sm rs0 m0 = Some m)
- (MVALID: forall b ofs, Mem.valid_pointer m b ofs = Mem.valid_pointer m0 b ofs)
- (LR: list_sval_refines ge sp rs0 m0 lhsv lsv),
- sval_refines ge sp rs0 m0 hv (Sop op lsv sm).
-Proof.
- wlp_simplify.
- erewrite H0; [ idtac | eauto | eauto ].
- rewrite H, MEM.
- destruct (seval_list_sval _ _ lsv _); try congruence.
- eapply op_valid_pointer_eq; eauto.
-Qed.
-Global Opaque hSop_hSinit.
-Local Hint Resolve hSop_hSinit_correct: wlp.
-
-Definition root_happly (rsv: root_sval) (lr: list reg) (hst: hsistate_local) : ?? hsval :=
- DO lhsv <~ hlist_args hst lr;;
- match rsv with
- | Rop op => hSop_hSinit op lhsv
- | Rload trap chunk addr => hSload hst trap chunk addr lhsv
- end.
-
-Lemma root_happly_correct (rsv: root_sval) lr hst:
- WHEN root_happly rsv lr hst ~> hv' THEN forall ge sp rs0 m0 st
- (REF:hsilocal_refines ge sp rs0 m0 hst st)
- (OK:hsok_local ge sp rs0 m0 hst),
- sval_refines ge sp rs0 m0 hv' (rsv lr st).
-Proof.
- unfold hsilocal_refines, root_apply, root_happly; destruct rsv; wlp_simplify.
- unfold sok_local in *.
- generalize (H0 ge sp rs0 m0 (list_sval_inj (map (si_sreg st) lr)) (si_smem st)); clear H0.
- destruct (seval_smem ge sp (si_smem st) rs0 m0) as [m|] eqn:X; eauto.
- intuition congruence.
-Qed.
-Global Opaque root_happly.
-Hint Resolve root_happly_correct: wlp.
-
-Local Open Scope lazy_bool_scope.
-
-(* NB: return [false] if the rsv cannot fail *)
-Definition may_trap (rsv: root_sval) (lr: list reg): bool :=
- match rsv with
- | Rop op => is_trapping_op op ||| negb (Nat.eqb (length lr) (args_of_operation op)) (* cf. lemma is_trapping_op_sound *)
- | Rload TRAP _ _ => true
- | _ => false
- end.
-
-Lemma lazy_orb_negb_false (b1 b2:bool):
- (b1 ||| negb b2) = false <-> (b1 = false /\ b2 = true).
-Proof.
- unfold negb; explore; simpl; intuition (try congruence).
-Qed.
-
-Lemma seval_list_sval_length ge sp rs0 m0 (f: reg -> sval) (l:list reg):
- forall l', seval_list_sval ge sp (list_sval_inj (List.map f l)) rs0 m0 = Some l' ->
- Datatypes.length l = Datatypes.length l'.
-Proof.
- induction l.
- - simpl. intros. inv H. reflexivity.
- - simpl. intros. destruct (seval_sval _ _ _ _ _); [|discriminate].
- destruct (seval_list_sval _ _ _ _ _) eqn:SLS; [|discriminate]. inv H. simpl.
- erewrite IHl; eauto.
-Qed.
-
-Lemma may_trap_correct (ge: RTL.genv) (sp:val) (rsv: root_sval) (rs0: regset) (m0: mem) (lr: list reg) st:
- may_trap rsv lr = false ->
- seval_list_sval ge sp (list_sval_inj (List.map (si_sreg st) lr)) rs0 m0 <> None ->
- seval_smem ge sp (si_smem st) rs0 m0 <> None ->
- seval_sval ge sp (rsv lr st) rs0 m0 <> None.
-Proof.
- destruct rsv; simpl; try congruence.
- - rewrite lazy_orb_negb_false. intros (TRAP1 & TRAP2) OK1 OK2.
- explore; try congruence.
- eapply is_trapping_op_sound; eauto.
- erewrite <- seval_list_sval_length; eauto.
- apply Nat.eqb_eq in TRAP2.
- assumption.
- - intros X OK1 OK2.
- explore; try congruence.
-Qed.
-
-(** simplify a symbolic value before assignment to a register *)
-Definition simplify (rsv: root_sval) (lr: list reg) (hst: hsistate_local): ?? hsval :=
- match rsv with
- | Rop op =>
- match is_move_operation op lr with
- | Some arg => hsi_sreg_get hst arg (** optimization of Omove *)
- | None =>
- DO lhsv <~ hlist_args hst lr;;
- hSop_hSinit op lhsv
- end
- | Rload _ chunk addr =>
- DO lhsv <~ hlist_args hst lr;;
- hSload hst NOTRAP chunk addr lhsv
- end.
-
-Lemma simplify_correct rsv lr hst:
- WHEN simplify rsv lr hst ~> hv THEN forall ge sp rs0 m0 st
- (REF: hsilocal_refines ge sp rs0 m0 hst st)
- (OK0: hsok_local ge sp rs0 m0 hst)
- (OK1: seval_sval ge sp (rsv lr st) rs0 m0 <> None),
- sval_refines ge sp rs0 m0 hv (rsv lr st).
-Proof.
- destruct rsv; simpl; auto.
- - (* Rop *)
- destruct (is_move_operation _ _) eqn: Hmove; wlp_simplify.
- + exploit is_move_operation_correct; eauto.
- intros (Hop & Hlsv); subst; simpl in *.
- simplify_SOME z.
- * erewrite H; eauto.
- * try_simplify_someHyps; congruence.
- * congruence.
- + clear Hmove.
- generalize (H0 ge sp rs0 m0 (list_sval_inj (map (si_sreg st) lr)) (si_smem st)); clear H0.
- destruct (seval_smem ge sp (si_smem st) rs0 m0) as [m|] eqn:X; eauto.
- intro H0; clear H0; simplify_SOME z; congruence. (* absurd case *)
- - (* Rload *)
- destruct trap; wlp_simplify.
- erewrite H0; eauto.
- erewrite H; eauto.
- erewrite hsilocal_refines_smem_refines; eauto.
- destruct (seval_list_sval _ _ _ _) as [args|] eqn: Hargs; try congruence.
- destruct (eval_addressing _ _ _ _) as [a|] eqn: Ha; try congruence.
- destruct (seval_smem _ _ _ _) as [m|] eqn: Hm; try congruence.
- destruct (Mem.loadv _ _ _); try congruence.
-Qed.
-Global Opaque simplify.
-Local Hint Resolve simplify_correct: wlp.
-
-Definition red_PTree_set (r: reg) (hsv: hsval) (hst: PTree.t hsval): PTree.t hsval :=
- match hsv with
- | HSinput r' _ =>
- if Pos.eq_dec r r'
- then PTree.remove r' hst
- else PTree.set r hsv hst
- | _ => PTree.set r hsv hst
- end.
-
-Lemma red_PTree_set_correct (r r0:reg) hsv hst ge sp rs0 m0:
- hsi_sreg_eval ge sp (red_PTree_set r hsv hst) r0 rs0 m0 = hsi_sreg_eval ge sp (PTree.set r hsv hst) r0 rs0 m0.
-Proof.
- destruct hsv; simpl; auto.
- destruct (Pos.eq_dec r r1); auto.
- subst; unfold hsi_sreg_eval, hsi_sreg_proj.
- destruct (Pos.eq_dec r0 r1); auto.
- - subst; rewrite PTree.grs, PTree.gss; simpl; auto.
- - rewrite PTree.gro, PTree.gso; simpl; auto.
-Qed.
-
-Lemma red_PTree_set_refines (r r0:reg) hsv hst sv st ge sp rs0 m0:
- hsilocal_refines ge sp rs0 m0 hst st ->
- sval_refines ge sp rs0 m0 hsv sv ->
- hsok_local ge sp rs0 m0 hst ->
- hsi_sreg_eval ge sp (red_PTree_set r hsv hst) r0 rs0 m0 = seval_sval ge sp (if Pos.eq_dec r r0 then sv else si_sreg st r0) rs0 m0.
-Proof.
- intros; rewrite red_PTree_set_correct.
- exploit hsilocal_refines_sreg; eauto.
- unfold hsi_sreg_eval, hsi_sreg_proj.
- destruct (Pos.eq_dec r r0); auto.
- - subst. rewrite PTree.gss; simpl; auto.
- - rewrite PTree.gso; simpl; eauto.
-Qed.
-
-Lemma sok_local_set_sreg (rsv:root_sval) ge sp rs0 m0 st r lr:
- sok_local ge sp rs0 m0 (slocal_set_sreg st r (rsv lr st))
- <-> (sok_local ge sp rs0 m0 st /\ seval_sval ge sp (rsv lr st) rs0 m0 <> None).
-Proof.
- unfold slocal_set_sreg, sok_local; simpl; split.
- + intros ((SVAL0 & PRE) & SMEM & SVAL).
- repeat (split; try tauto).
- - intros r0; generalize (SVAL r0); clear SVAL; destruct (Pos.eq_dec r r0); try congruence.
- - generalize (SVAL r); clear SVAL; destruct (Pos.eq_dec r r); try congruence.
- + intros ((PRE & SMEM & SVAL0) & SVAL).
- repeat (split; try tauto; eauto).
- intros r0; destruct (Pos.eq_dec r r0); try congruence.
-Qed.
-
-Definition hslocal_set_sreg (hst: hsistate_local) (r: reg) (rsv: root_sval) (lr: list reg): ?? hsistate_local :=
- DO ok_lhsv <~
- (if may_trap rsv lr
- then DO hv <~ root_happly rsv lr hst;;
- XDEBUG hv (fun hv => DO hv_name <~ string_of_hashcode (hsval_get_hid hv);; RET ("-- insert undef behavior of hashcode:" +; (CamlStr hv_name))%string);;
- RET (hv::(hsi_ok_lsval hst))
- else RET (hsi_ok_lsval hst));;
- DO simp <~ simplify rsv lr hst;;
- RET {| hsi_smem := hst;
- hsi_ok_lsval := ok_lhsv;
- hsi_sreg := red_PTree_set r simp (hsi_sreg hst) |}.
-
-Lemma hslocal_set_sreg_correct hst r rsv lr:
- WHEN hslocal_set_sreg hst r rsv lr ~> hst' THEN forall ge sp rs0 m0 st
- (REF: hsilocal_refines ge sp rs0 m0 hst st),
- hsilocal_refines ge sp rs0 m0 hst' (slocal_set_sreg st r (rsv lr st)).
-Proof.
- wlp_simplify.
- + (* may_trap ~> true *)
- assert (X: sok_local ge sp rs0 m0 (slocal_set_sreg st r (rsv lr st)) <->
- hsok_local ge sp rs0 m0 {| hsi_smem := hst; hsi_ok_lsval := exta :: hsi_ok_lsval hst; hsi_sreg := red_PTree_set r exta0 hst |}).
- { rewrite sok_local_set_sreg; generalize REF.
- intros (OKeq & MEM & REG & MVALID); rewrite OKeq; clear OKeq.
- unfold hsok_local; simpl; intuition (subst; eauto);
- erewrite <- H0 in *; eauto; unfold hsok_local; simpl; intuition eauto.
- }
- unfold hsilocal_refines; simpl; split; auto.
- rewrite <- X, sok_local_set_sreg. intuition eauto.
- - destruct REF; intuition eauto.
- - generalize REF; intros (OKEQ & _). rewrite OKEQ in * |-; erewrite red_PTree_set_refines; eauto.
- + (* may_trap ~> false *)
- assert (X: sok_local ge sp rs0 m0 (slocal_set_sreg st r (rsv lr st)) <->
- hsok_local ge sp rs0 m0 {| hsi_smem := hst; hsi_ok_lsval := hsi_ok_lsval hst; hsi_sreg := red_PTree_set r exta hst |}).
- {
- rewrite sok_local_set_sreg; generalize REF.
- intros (OKeq & MEM & REG & MVALID); rewrite OKeq.
- unfold hsok_local; simpl; intuition (subst; eauto).
- assert (X0:hsok_local ge sp rs0 m0 hst). { unfold hsok_local; intuition. }
- exploit may_trap_correct; eauto.
- * intro X1; eapply seval_list_sval_inj_not_none; eauto.
- assert (X2: sok_local ge sp rs0 m0 st). { intuition. }
- unfold sok_local in X2; intuition eauto.
- * rewrite <- MEM; eauto.
- }
- unfold hsilocal_refines; simpl; split; auto.
- rewrite <- X, sok_local_set_sreg. intuition eauto.
- - destruct REF; intuition eauto.
- - generalize REF; intros (OKEQ & _). rewrite OKEQ in * |-; erewrite red_PTree_set_refines; eauto.
-Qed.
-Global Opaque hslocal_set_sreg.
-Local Hint Resolve hslocal_set_sreg_correct: wlp.
-
-(** ** Execution of one instruction *)
-
-Definition hsiexec_inst (i: instruction) (hst: hsistate): ?? (option hsistate) :=
- match i with
- | Inop pc' =>
- RET (Some (hsist_set_local hst pc' hst.(hsi_local)))
- | Iop op args dst pc' =>
- DO next <~ hslocal_set_sreg hst.(hsi_local) dst (Rop op) args;;
- RET (Some (hsist_set_local hst pc' next))
- | Iload trap chunk addr args dst pc' =>
- DO next <~ hslocal_set_sreg hst.(hsi_local) dst (Rload trap chunk addr) args;;
- RET (Some (hsist_set_local hst pc' next))
- | Istore chunk addr args src pc' =>
- DO next <~ hslocal_store hst.(hsi_local) chunk addr args src;;
- RET (Some (hsist_set_local hst pc' next))
- | Icond cond args ifso ifnot _ =>
- let prev := hst.(hsi_local) in
- DO vargs <~ hlist_args prev args ;;
- let ex := {| hsi_cond:=cond; hsi_scondargs:=vargs; hsi_elocal := prev; hsi_ifso := ifso |} in
- RET (Some {| hsi_pc := ifnot; hsi_exits := ex::hst.(hsi_exits); hsi_local := prev |})
- | _ => RET None (* TODO jumptable ? *)
- end.
-
-
-Remark hsiexec_inst_None_correct i hst:
- WHEN hsiexec_inst i hst ~> o THEN forall st, o = None -> siexec_inst i st = None.
-Proof.
- destruct i; wlp_simplify; congruence.
-Qed.
-
-Lemma seval_condition_refines hst st ge sp cond hargs args rs m:
- hsok_local ge sp rs m hst ->
- hsilocal_refines ge sp rs m hst st ->
- list_sval_refines ge sp rs m hargs args ->
- hseval_condition ge sp cond hargs (hsi_smem hst) rs m
- = seval_condition ge sp cond args (si_smem st) rs m.
- Proof.
- intros HOK (_ & MEMEQ & _) LR. unfold hseval_condition, seval_condition.
- rewrite LR, <- MEMEQ; auto.
-Qed.
-
-Lemma sok_local_set_sreg_simp (rsv:root_sval) ge sp rs0 m0 st r lr:
- sok_local ge sp rs0 m0 (slocal_set_sreg st r (rsv lr st))
- -> sok_local ge sp rs0 m0 st.
-Proof.
- rewrite sok_local_set_sreg; intuition.
-Qed.
-
-Local Hint Resolve hsist_set_local_correct_stat: core.
-
-Lemma hsiexec_inst_correct i hst:
- WHEN hsiexec_inst i hst ~> o THEN forall hst' st,
- o = Some hst' ->
- exists st', siexec_inst i st = Some st'
- /\ (forall (REF:hsistate_refines_stat hst st), hsistate_refines_stat hst' st')
- /\ (forall ge sp rs0 m0 (REF:hsistate_refines_dyn ge sp rs0 m0 hst st), hsistate_refines_dyn ge sp rs0 m0 hst' st').
-Proof.
- destruct i; simpl; wlp_simplify; try_simplify_someHyps; eexists; intuition eauto.
- - (* refines_dyn Iop *)
- eapply hsist_set_local_correct_dyn; eauto.
- generalize (sok_local_set_sreg_simp (Rop o)); simpl; eauto.
- - (* refines_dyn Iload *)
- eapply hsist_set_local_correct_dyn; eauto.
- generalize (sok_local_set_sreg_simp (Rload t0 m a)); simpl; eauto.
- - (* refines_dyn Istore *)
- eapply hsist_set_local_correct_dyn; eauto.
- unfold sok_local; simpl; intuition.
- - (* refines_stat Icond *)
- unfold hsistate_refines_stat, hsiexits_refines_stat in *; simpl; intuition.
- constructor; simpl; eauto.
- constructor.
- - (* refines_dyn Icond *)
- destruct REF as (EXREF & LREF & NEST).
- split.
- + constructor; simpl; auto.
- constructor; simpl; auto.
- intros; erewrite seval_condition_refines; eauto.
- + split; simpl; auto.
- destruct NEST as [|st0 se lse TOP NEST];
- econstructor; simpl; auto; constructor; auto.
-Qed.
-Global Opaque hsiexec_inst.
-Local Hint Resolve hsiexec_inst_correct: wlp.
-
-
-Definition some_or_fail {A} (o: option A) (msg: pstring): ?? A :=
- match o with
- | Some x => RET x
- | None => FAILWITH msg
- end.
-
-Fixpoint hsiexec_path (path:nat) (f: function) (hst: hsistate): ?? hsistate :=
- match path with
- | O => RET hst
- | S p =>
- let pc := hst.(hsi_pc) in
- XDEBUG pc (fun pc => DO name_pc <~ string_of_Z (Zpos pc);; RET ("- sym exec node: " +; name_pc)%string);;
- DO i <~ some_or_fail ((fn_code f)!pc) "hsiexec_path.internal_error.1";;
- DO ohst1 <~ hsiexec_inst i hst;;
- DO hst1 <~ some_or_fail ohst1 "hsiexec_path.internal_error.2";;
- hsiexec_path p f hst1
- end.
-
-Lemma hsiexec_path_correct path f: forall hst,
- WHEN hsiexec_path path f hst ~> hst' THEN forall st
- (RSTAT:hsistate_refines_stat hst st),
- exists st', siexec_path path f st = Some st'
- /\ hsistate_refines_stat hst' st'
- /\ (forall ge sp rs0 m0 (REF:hsistate_refines_dyn ge sp rs0 m0 hst st), hsistate_refines_dyn ge sp rs0 m0 hst' st').
-Proof.
- induction path; wlp_simplify; try_simplify_someHyps. clear IHpath.
- generalize RSTAT; intros (PCEQ & _) INSTEQ.
- rewrite <- PCEQ, INSTEQ; simpl.
- exploit H0; eauto. clear H0.
- intros (st0 & SINST & ISTAT & IDYN); erewrite SINST.
- exploit H1; eauto. clear H1.
- intros (st' & SPATH & PSTAT & PDYN).
- eexists; intuition eauto.
-Qed.
-Global Opaque hsiexec_path.
-Local Hint Resolve hsiexec_path_correct: wlp.
-
-Fixpoint hbuiltin_arg (hst: PTree.t hsval) (arg : builtin_arg reg): ?? builtin_arg hsval :=
- match arg with
- | BA r =>
- DO v <~ hsi_sreg_get hst r;;
- RET (BA v)
- | BA_int n => RET (BA_int n)
- | BA_long n => RET (BA_long n)
- | BA_float f0 => RET (BA_float f0)
- | BA_single s => RET (BA_single s)
- | BA_loadstack chunk ptr => RET (BA_loadstack chunk ptr)
- | BA_addrstack ptr => RET (BA_addrstack ptr)
- | BA_loadglobal chunk id ptr => RET (BA_loadglobal chunk id ptr)
- | BA_addrglobal id ptr => RET (BA_addrglobal id ptr)
- | BA_splitlong ba1 ba2 =>
- DO v1 <~ hbuiltin_arg hst ba1;;
- DO v2 <~ hbuiltin_arg hst ba2;;
- RET (BA_splitlong v1 v2)
- | BA_addptr ba1 ba2 =>
- DO v1 <~ hbuiltin_arg hst ba1;;
- DO v2 <~ hbuiltin_arg hst ba2;;
- RET (BA_addptr v1 v2)
- end.
-
-Lemma hbuiltin_arg_correct hst arg:
- WHEN hbuiltin_arg hst arg ~> hargs THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- seval_builtin_sval ge sp (builtin_arg_map hsval_proj hargs) rs0 m0 = seval_builtin_sval ge sp (builtin_arg_map f arg) rs0 m0.
-Proof.
- induction arg; wlp_simplify.
- + erewrite H; eauto.
- + erewrite H; eauto.
- erewrite H0; eauto.
- + erewrite H; eauto.
- erewrite H0; eauto.
-Qed.
-Global Opaque hbuiltin_arg.
-Local Hint Resolve hbuiltin_arg_correct: wlp.
-
-Fixpoint hbuiltin_args (hst: PTree.t hsval) (args: list (builtin_arg reg)): ?? list (builtin_arg hsval) :=
- match args with
- | nil => RET nil
- | a::l =>
- DO ha <~ hbuiltin_arg hst a;;
- DO hl <~ hbuiltin_args hst l;;
- RET (ha::hl)
- end.
-
-Lemma hbuiltin_args_correct hst args:
- WHEN hbuiltin_args hst args ~> hargs THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- bargs_refines ge sp rs0 m0 hargs (List.map (builtin_arg_map f) args).
-Proof.
- unfold bargs_refines, seval_builtin_args; induction args; wlp_simplify.
- erewrite H; eauto.
- erewrite H0; eauto.
-Qed.
-Global Opaque hbuiltin_args.
-Local Hint Resolve hbuiltin_args_correct: wlp.
-
-Definition hsum_left (hst: PTree.t hsval) (ros: reg + ident): ?? (hsval + ident) :=
- match ros with
- | inl r => DO hr <~ hsi_sreg_get hst r;; RET (inl hr)
- | inr s => RET (inr s)
- end.
-
-Lemma hsum_left_correct hst ros:
- WHEN hsum_left hst ros ~> hsi THEN forall ge sp rs0 m0 (f: reg -> sval)
- (RR: forall r, hsi_sreg_eval ge sp hst r rs0 m0 = seval_sval ge sp (f r) rs0 m0),
- sum_refines ge sp rs0 m0 hsi (sum_left_map f ros).
-Proof.
- unfold sum_refines; destruct ros; wlp_simplify.
-Qed.
-Global Opaque hsum_left.
-Local Hint Resolve hsum_left_correct: wlp.
-
-Definition hsexec_final (i: instruction) (hst: PTree.t hsval): ?? hsfval :=
- match i with
- | Icall sig ros args res pc =>
- DO svos <~ hsum_left hst ros;;
- DO sargs <~ hlist_args hst args;;
- RET (HScall sig svos sargs res pc)
- | Itailcall sig ros args =>
- DO svos <~ hsum_left hst ros;;
- DO sargs <~ hlist_args hst args;;
- RET (HStailcall sig svos sargs)
- | Ibuiltin ef args res pc =>
- DO sargs <~ hbuiltin_args hst args;;
- RET (HSbuiltin ef sargs res pc)
- | Ijumptable reg tbl =>
- DO sv <~ hsi_sreg_get hst reg;;
- RET (HSjumptable sv tbl)
- | Ireturn or =>
- match or with
- | Some r => DO hr <~ hsi_sreg_get hst r;; RET (HSreturn (Some hr))
- | None => RET (HSreturn None)
- end
- | _ => RET (HSnone)
- end.
-
-Lemma hsexec_final_correct (hsl: hsistate_local) i:
- WHEN hsexec_final i hsl ~> hsf THEN forall ge sp rs0 m0 sl
- (OK: hsok_local ge sp rs0 m0 hsl)
- (REF: hsilocal_refines ge sp rs0 m0 hsl sl),
- hfinal_refines ge sp rs0 m0 hsf (sexec_final i sl).
-Proof.
- destruct i; wlp_simplify; try econstructor; simpl; eauto.
-Qed.
-Global Opaque hsexec_final.
-Local Hint Resolve hsexec_final_correct: wlp.
-
-Definition init_hsistate_local (_:unit): ?? hsistate_local
- := DO hm <~ hSinit ();;
- RET {| hsi_smem := hm; hsi_ok_lsval := nil; hsi_sreg := PTree.empty hsval |}.
-
-Lemma init_hsistate_local_correct:
- WHEN init_hsistate_local () ~> hsl THEN forall ge sp rs0 m0,
- hsilocal_refines ge sp rs0 m0 hsl init_sistate_local.
-Proof.
- unfold hsilocal_refines; wlp_simplify.
- - unfold hsok_local; simpl; intuition. erewrite H in *; congruence.
- - unfold hsok_local, sok_local; simpl in *; intuition; try congruence.
- - unfold hsi_sreg_eval, hsi_sreg_proj. rewrite PTree.gempty. reflexivity.
- - try_simplify_someHyps.
-Qed.
-Global Opaque init_hsistate_local.
-Local Hint Resolve init_hsistate_local_correct: wlp.
-
-Definition init_hsistate pc: ?? hsistate
- := DO hst <~ init_hsistate_local ();;
- RET {| hsi_pc := pc; hsi_exits := nil; hsi_local := hst |}.
-
-Lemma init_hsistate_correct pc:
- WHEN init_hsistate pc ~> hst THEN
- hsistate_refines_stat hst (init_sistate pc)
- /\ forall ge sp rs0 m0, hsistate_refines_dyn ge sp rs0 m0 hst (init_sistate pc).
-Proof.
- unfold hsistate_refines_stat, hsistate_refines_dyn, hsiexits_refines_dyn; wlp_simplify; constructor.
-Qed.
-Global Opaque init_hsistate.
-Local Hint Resolve init_hsistate_correct: wlp.
-
-Definition hsexec (f: function) (pc:node): ?? hsstate :=
- DO path <~ some_or_fail ((fn_path f)!pc) "hsexec.internal_error.1";;
- DO hinit <~ init_hsistate pc;;
- DO hst <~ hsiexec_path path.(psize) f hinit;;
- DO i <~ some_or_fail ((fn_code f)!(hst.(hsi_pc))) "hsexec.internal_error.2";;
- DO ohst <~ hsiexec_inst i hst;;
- match ohst with
- | Some hst' => RET {| hinternal := hst'; hfinal := HSnone |}
- | None => DO hsvf <~ hsexec_final i hst.(hsi_local);;
- RET {| hinternal := hst; hfinal := hsvf |}
- end.
-
-Lemma hsexec_correct_aux f pc:
- WHEN hsexec f pc ~> hst THEN
- exists st, sexec f pc = Some st /\ hsstate_refines hst st.
-Proof.
- unfold hsstate_refines, sexec; wlp_simplify.
- - (* Some *)
- rewrite H; clear H.
- exploit H0; clear H0; eauto.
- intros (st0 & EXECPATH & SREF & DREF).
- rewrite EXECPATH; clear EXECPATH.
- generalize SREF. intros (EQPC & _).
- rewrite <- EQPC, H3; clear H3.
- exploit H4; clear H4; eauto.
- intros (st' & EXECL & SREF' & DREF').
- try_simplify_someHyps.
- eexists; intuition (simpl; eauto).
- constructor.
- - (* None *)
- rewrite H; clear H H4.
- exploit H0; clear H0; eauto.
- intros (st0 & EXECPATH & SREF & DREF).
- rewrite EXECPATH; clear EXECPATH.
- generalize SREF. intros (EQPC & _).
- rewrite <- EQPC, H3; clear H3.
- erewrite hsiexec_inst_None_correct; eauto.
- eexists; intuition (simpl; eauto).
-Qed.
-
-Global Opaque hsexec.
-
-End CanonBuilding.
-
-(** Correction of concrete symbolic execution wrt abstract symbolic execution *)
-Theorem hsexec_correct
- (hC_hsval : hashinfo hsval -> ?? hsval)
- (hC_list_hsval : hashinfo list_hsval -> ?? list_hsval)
- (hC_hsmem : hashinfo hsmem -> ?? hsmem)
- (f : function)
- (pc : node):
- WHEN hsexec hC_hsval hC_list_hsval hC_hsmem f pc ~> hst THEN forall
- (hC_hsval_correct: forall hs,
- WHEN hC_hsval hs ~> hs' THEN forall ge sp rs0 m0,
- seval_sval ge sp (hsval_proj (hdata hs)) rs0 m0 =
- seval_sval ge sp (hsval_proj hs') rs0 m0)
- (hC_list_hsval_correct: forall lh,
- WHEN hC_list_hsval lh ~> lh' THEN forall ge sp rs0 m0,
- seval_list_sval ge sp (hsval_list_proj (hdata lh)) rs0 m0 =
- seval_list_sval ge sp (hsval_list_proj lh') rs0 m0)
- (hC_hsmem_correct: forall hm,
- WHEN hC_hsmem hm ~> hm' THEN forall ge sp rs0 m0,
- seval_smem ge sp (hsmem_proj (hdata hm)) rs0 m0 =
- seval_smem ge sp (hsmem_proj hm') rs0 m0),
- exists st : sstate, sexec f pc = Some st /\ hsstate_refines hst st.
-Proof.
- wlp_simplify.
- eapply hsexec_correct_aux; eauto.
-Qed.
-Local Hint Resolve hsexec_correct: wlp.
-
-(** * Implementing the simulation test with concrete hash-consed symbolic execution *)
-
-Definition phys_check {A} (x y:A) (msg: pstring): ?? unit :=
- DO b <~ phys_eq x y;;
- assert_b b msg;;
- RET tt.
-
-Definition struct_check {A} (x y: A) (msg: pstring): ?? unit :=
- DO b <~ struct_eq x y;;
- assert_b b msg;;
- RET tt.
-
-Lemma struct_check_correct {A} (a b: A) msg:
- WHEN struct_check a b msg ~> _ THEN
- a = b.
-Proof. wlp_simplify. Qed.
-Global Opaque struct_check.
-Hint Resolve struct_check_correct: wlp.
-
-Definition option_eq_check {A} (o1 o2: option A): ?? unit :=
- match o1, o2 with
- | Some x1, Some x2 => phys_check x1 x2 "option_eq_check: data physically differ"
- | None, None => RET tt
- | _, _ => FAILWITH "option_eq_check: structure differs"
- end.
-
-Lemma option_eq_check_correct A (o1 o2: option A): WHEN option_eq_check o1 o2 ~> _ THEN o1=o2.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque option_eq_check.
-Hint Resolve option_eq_check_correct:wlp.
-
-Import PTree.
-
-Fixpoint PTree_eq_check {A} (d1 d2: PTree.t A): ?? unit :=
- match d1, d2 with
- | Leaf, Leaf => RET tt
- | Node l1 o1 r1, Node l2 o2 r2 =>
- option_eq_check o1 o2;;
- PTree_eq_check l1 l2;;
- PTree_eq_check r1 r2
- | _, _ => FAILWITH "PTree_eq_check: some key is absent"
- end.
-
-Lemma PTree_eq_check_correct A d1: forall (d2: t A),
- WHEN PTree_eq_check d1 d2 ~> _ THEN forall x, PTree.get x d1 = PTree.get x d2.
-Proof.
- induction d1 as [|l1 Hl1 o1 r1 Hr1]; destruct d2 as [|l2 o2 r2]; simpl;
- wlp_simplify. destruct x; simpl; auto.
-Qed.
-Global Opaque PTree_eq_check.
-Local Hint Resolve PTree_eq_check_correct: wlp.
-
-Fixpoint PTree_frame_eq_check {A} (frame: list positive) (d1 d2: PTree.t A): ?? unit :=
- match frame with
- | nil => RET tt
- | k::l =>
- option_eq_check (PTree.get k d1) (PTree.get k d2);;
- PTree_frame_eq_check l d1 d2
- end.
-
-Lemma PTree_frame_eq_check_correct A l (d1 d2: t A):
- WHEN PTree_frame_eq_check l d1 d2 ~> _ THEN forall x, List.In x l -> PTree.get x d1 = PTree.get x d2.
-Proof.
- induction l as [|k l]; simpl; wlp_simplify.
- subst; auto.
-Qed.
-Global Opaque PTree_frame_eq_check.
-Local Hint Resolve PTree_frame_eq_check_correct: wlp.
-
-Definition hsilocal_simu_check hst1 hst2 : ?? unit :=
- DEBUG("? last check");;
- phys_check (hsi_smem hst2) (hsi_smem hst1) "hsilocal_simu_check: hsi_smem sets aren't equiv";;
- PTree_eq_check (hsi_sreg hst1) (hsi_sreg hst2);;
- Sets.assert_list_incl mk_hash_params (hsi_ok_lsval hst2) (hsi_ok_lsval hst1);;
- DEBUG("=> last check: OK").
-
-Lemma hsilocal_simu_check_correct hst1 hst2:
- WHEN hsilocal_simu_check hst1 hst2 ~> _ THEN
- hsilocal_simu_spec None hst1 hst2.
-Proof.
- unfold hsilocal_simu_spec; wlp_simplify.
-Qed.
-Hint Resolve hsilocal_simu_check_correct: wlp.
-Global Opaque hsilocal_simu_check.
-
-Definition hsilocal_frame_simu_check frame hst1 hst2 : ?? unit :=
- DEBUG("? frame check");;
- phys_check (hsi_smem hst2) (hsi_smem hst1) "hsilocal_frame_simu_check: hsi_smem sets aren't equiv";;
- PTree_frame_eq_check frame (hsi_sreg hst1) (hsi_sreg hst2);;
- Sets.assert_list_incl mk_hash_params (hsi_ok_lsval hst2) (hsi_ok_lsval hst1);;
- DEBUG("=> frame check: OK").
-
-Lemma setoid_in {A: Type} (a: A): forall l,
- SetoidList.InA (fun x y => x = y) a l ->
- In a l.
-Proof.
- induction l; intros; inv H.
- - constructor. reflexivity.
- - right. auto.
-Qed.
-
-Lemma regset_elements_in r rs:
- Regset.In r rs ->
- In r (Regset.elements rs).
-Proof.
- intros. exploit Regset.elements_1; eauto. intro SIN.
- apply setoid_in. assumption.
-Qed.
-Local Hint Resolve regset_elements_in: core.
-
-Lemma hsilocal_frame_simu_check_correct hst1 hst2 alive:
- WHEN hsilocal_frame_simu_check (Regset.elements alive) hst1 hst2 ~> _ THEN
- hsilocal_simu_spec (Some alive) hst1 hst2.
-Proof.
- unfold hsilocal_simu_spec; wlp_simplify. symmetry; eauto.
-Qed.
-Hint Resolve hsilocal_frame_simu_check_correct: wlp.
-Global Opaque hsilocal_frame_simu_check.
-
-Definition revmap_check_single (dm: PTree.t node) (n tn: node) : ?? unit :=
- DO res <~ some_or_fail (dm ! tn) "revmap_check_single: no mapping for tn";;
- struct_check n res "revmap_check_single: n and res are physically different".
-
-Lemma revmap_check_single_correct dm pc1 pc2:
- WHEN revmap_check_single dm pc1 pc2 ~> _ THEN
- dm ! pc2 = Some pc1.
-Proof.
- wlp_simplify. congruence.
-Qed.
-Hint Resolve revmap_check_single_correct: wlp.
-Global Opaque revmap_check_single.
-
-Definition hsiexit_simu_check (dm: PTree.t node) (f: RTLpath.function) (hse1 hse2: hsistate_exit): ?? unit :=
- struct_check (hsi_cond hse1) (hsi_cond hse2) "hsiexit_simu_check: conditions do not match";;
- phys_check (hsi_scondargs hse1) (hsi_scondargs hse2) "hsiexit_simu_check: args do not match";;
- revmap_check_single dm (hsi_ifso hse1) (hsi_ifso hse2);;
- DO path <~ some_or_fail ((fn_path f) ! (hsi_ifso hse1)) "hsiexit_simu_check: internal error";;
- hsilocal_frame_simu_check (Regset.elements path.(input_regs)) (hsi_elocal hse1) (hsi_elocal hse2).
-
-Lemma hsiexit_simu_check_correct dm f hse1 hse2:
- WHEN hsiexit_simu_check dm f hse1 hse2 ~> _ THEN
- hsiexit_simu_spec dm f hse1 hse2.
-Proof.
- unfold hsiexit_simu_spec; wlp_simplify.
-Qed.
-Hint Resolve hsiexit_simu_check_correct: wlp.
-Global Opaque hsiexit_simu_check.
-
-Fixpoint hsiexits_simu_check (dm: PTree.t node) (f: RTLpath.function) (lhse1 lhse2: list hsistate_exit) :=
- match lhse1,lhse2 with
- | nil, nil => RET tt
- | hse1 :: lhse1, hse2 :: lhse2 =>
- hsiexit_simu_check dm f hse1 hse2;;
- hsiexits_simu_check dm f lhse1 lhse2
- | _, _ => FAILWITH "siexists_simu_check: lengths do not match"
- end.
-
-Lemma hsiexits_simu_check_correct dm f: forall le1 le2,
- WHEN hsiexits_simu_check dm f le1 le2 ~> _ THEN
- hsiexits_simu_spec dm f le1 le2.
-Proof.
- unfold hsiexits_simu_spec; induction le1; simpl; destruct le2; wlp_simplify; constructor; eauto.
-Qed.
-Hint Resolve hsiexits_simu_check_correct: wlp.
-Global Opaque hsiexits_simu_check.
-
-Definition hsistate_simu_check (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsistate) :=
- hsiexits_simu_check dm f (hsi_exits hst1) (hsi_exits hst2);;
- hsilocal_simu_check (hsi_local hst1) (hsi_local hst2).
-
-Lemma hsistate_simu_check_correct dm f hst1 hst2:
- WHEN hsistate_simu_check dm f hst1 hst2 ~> _ THEN
- hsistate_simu_spec dm f hst1 hst2.
-Proof.
- unfold hsistate_simu_spec; wlp_simplify.
-Qed.
-Hint Resolve hsistate_simu_check_correct: wlp.
-Global Opaque hsistate_simu_check.
-
-
-Fixpoint revmap_check_list (dm: PTree.t node) (ln ln': list node): ?? unit :=
- match ln, ln' with
- | nil, nil => RET tt
- | n::ln, n'::ln' =>
- revmap_check_single dm n n';;
- revmap_check_list dm ln ln'
- | _, _ => FAILWITH "revmap_check_list: lists have different lengths"
- end.
-
-Lemma revmap_check_list_correct dm: forall lpc lpc',
- WHEN revmap_check_list dm lpc lpc' ~> _ THEN
- ptree_get_list dm lpc' = Some lpc.
-Proof.
- induction lpc.
- - destruct lpc'; wlp_simplify.
- - destruct lpc'; wlp_simplify. try_simplify_someHyps.
-Qed.
-Global Opaque revmap_check_list.
-Hint Resolve revmap_check_list_correct: wlp.
-
-
-Definition svos_simu_check (svos1 svos2: hsval + ident) :=
- match svos1, svos2 with
- | inl sv1, inl sv2 => phys_check sv1 sv2 "svos_simu_check: sval mismatch"
- | inr id1, inr id2 => phys_check id1 id2 "svos_simu_check: symbol mismatch"
- | _, _ => FAILWITH "svos_simu_check: type mismatch"
- end.
-
-Lemma svos_simu_check_correct svos1 svos2:
- WHEN svos_simu_check svos1 svos2 ~> _ THEN
- svos1 = svos2.
-Proof.
- destruct svos1; destruct svos2; wlp_simplify.
-Qed.
-Global Opaque svos_simu_check.
-Hint Resolve svos_simu_check_correct: wlp.
-
-
-Fixpoint builtin_arg_simu_check (bs bs': builtin_arg hsval) :=
- match bs with
- | BA sv =>
- match bs' with
- | BA sv' => phys_check sv sv' "builtin_arg_simu_check: sval mismatch"
- | _ => FAILWITH "builtin_arg_simu_check: BA mismatch"
- end
- | BA_splitlong lo hi =>
- match bs' with
- | BA_splitlong lo' hi' =>
- builtin_arg_simu_check lo lo';;
- builtin_arg_simu_check hi hi'
- | _ => FAILWITH "builtin_arg_simu_check: BA_splitlong mismatch"
- end
- | BA_addptr b1 b2 =>
- match bs' with
- | BA_addptr b1' b2' =>
- builtin_arg_simu_check b1 b1';;
- builtin_arg_simu_check b2 b2'
- | _ => FAILWITH "builtin_arg_simu_check: BA_addptr mismatch"
- end
- | bs => struct_check bs bs' "builtin_arg_simu_check: basic mismatch"
- end.
-
-Lemma builtin_arg_simu_check_correct: forall bs1 bs2,
- WHEN builtin_arg_simu_check bs1 bs2 ~> _ THEN
- builtin_arg_map hsval_proj bs1 = builtin_arg_map hsval_proj bs2.
-Proof.
- induction bs1.
- all: try (wlp_simplify; subst; reflexivity).
- all: destruct bs2; wlp_simplify; congruence.
-Qed.
-Global Opaque builtin_arg_simu_check.
-Hint Resolve builtin_arg_simu_check_correct: wlp.
-
-Fixpoint list_builtin_arg_simu_check lbs1 lbs2 :=
- match lbs1, lbs2 with
- | nil, nil => RET tt
- | bs1::lbs1, bs2::lbs2 =>
- builtin_arg_simu_check bs1 bs2;;
- list_builtin_arg_simu_check lbs1 lbs2
- | _, _ => FAILWITH "list_builtin_arg_simu_check: length mismatch"
- end.
-
-Lemma list_builtin_arg_simu_check_correct: forall lbs1 lbs2,
- WHEN list_builtin_arg_simu_check lbs1 lbs2 ~> _ THEN
- List.map (builtin_arg_map hsval_proj) lbs1 = List.map (builtin_arg_map hsval_proj) lbs2.
-Proof.
- induction lbs1; destruct lbs2; wlp_simplify. congruence.
-Qed.
-Global Opaque list_builtin_arg_simu_check.
-Hint Resolve list_builtin_arg_simu_check_correct: wlp.
-
-Definition sfval_simu_check (dm: PTree.t node) (f: RTLpath.function) (pc1 pc2: node) (fv1 fv2: hsfval) :=
- match fv1, fv2 with
- | HSnone, HSnone => revmap_check_single dm pc1 pc2
- | HScall sig1 svos1 lsv1 res1 pc1, HScall sig2 svos2 lsv2 res2 pc2 =>
- revmap_check_single dm pc1 pc2;;
- phys_check sig1 sig2 "sfval_simu_check: Scall different signatures";;
- phys_check res1 res2 "sfval_simu_check: Scall res do not match";;
- svos_simu_check svos1 svos2;;
- phys_check lsv1 lsv2 "sfval_simu_check: Scall args do not match"
- | HStailcall sig1 svos1 lsv1, HStailcall sig2 svos2 lsv2 =>
- phys_check sig1 sig2 "sfval_simu_check: Stailcall different signatures";;
- svos_simu_check svos1 svos2;;
- phys_check lsv1 lsv2 "sfval_simu_check: Stailcall args do not match"
- | HSbuiltin ef1 lbs1 br1 pc1, HSbuiltin ef2 lbs2 br2 pc2 =>
- revmap_check_single dm pc1 pc2;;
- phys_check ef1 ef2 "sfval_simu_check: builtin ef do not match";;
- phys_check br1 br2 "sfval_simu_check: builtin br do not match";;
- list_builtin_arg_simu_check lbs1 lbs2
- | HSjumptable sv ln, HSjumptable sv' ln' =>
- revmap_check_list dm ln ln';;
- phys_check sv sv' "sfval_simu_check: Sjumptable sval do not match"
- | HSreturn osv1, HSreturn osv2 =>
- option_eq_check osv1 osv2
- | _, _ => FAILWITH "sfval_simu_check: structure mismatch"
- end.
-
-Lemma sfval_simu_check_correct dm f opc1 opc2 fv1 fv2:
- WHEN sfval_simu_check dm f opc1 opc2 fv1 fv2 ~> _ THEN
- hfinal_simu_spec dm f opc1 opc2 fv1 fv2.
-Proof.
- unfold hfinal_simu_spec; destruct fv1; destruct fv2; wlp_simplify; try congruence.
-Qed.
-Hint Resolve sfval_simu_check_correct: wlp.
-Global Opaque sfval_simu_check.
-
-Definition hsstate_simu_check (dm: PTree.t node) (f: RTLpath.function) (hst1 hst2: hsstate) :=
- hsistate_simu_check dm f (hinternal hst1) (hinternal hst2);;
- sfval_simu_check dm f (hsi_pc hst1) (hsi_pc hst2) (hfinal hst1) (hfinal hst2).
-
-Lemma hsstate_simu_check_correct dm f hst1 hst2:
- WHEN hsstate_simu_check dm f hst1 hst2 ~> _ THEN
- hsstate_simu_spec dm f hst1 hst2.
-Proof.
- unfold hsstate_simu_spec; wlp_simplify.
-Qed.
-Hint Resolve hsstate_simu_check_correct: wlp.
-Global Opaque hsstate_simu_check.
-
-Definition simu_check_single (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) (m: node * node): ?? unit :=
- let (pc2, pc1) := m in
- (* creating the hash-consing tables *)
- DO hC_sval <~ hCons hSVAL;;
- DO hC_list_hsval <~ hCons hLSVAL;;
- DO hC_hsmem <~ hCons hSMEM;;
- let hsexec := hsexec hC_sval.(hC) hC_list_hsval.(hC) hC_hsmem.(hC) in
- (* performing the hash-consed executions *)
- XDEBUG pc1 (fun pc => DO name_pc <~ string_of_Z (Zpos pc);; RET ("entry-point of input superblock: " +; name_pc)%string);;
- DO hst1 <~ hsexec f pc1;;
- XDEBUG pc2 (fun pc => DO name_pc <~ string_of_Z (Zpos pc);; RET ("entry-point of output superblock: " +; name_pc)%string);;
- DO hst2 <~ hsexec tf pc2;;
- (* comparing the executions *)
- hsstate_simu_check dm f hst1 hst2.
-
-Lemma simu_check_single_correct dm tf f pc1 pc2:
- WHEN simu_check_single dm f tf (pc2, pc1) ~> _ THEN
- sexec_simu dm f tf pc1 pc2.
-Proof.
- unfold sexec_simu; wlp_simplify.
- exploit H2; clear H2. 1-3: wlp_simplify.
- intros (st2 & SEXEC2 & REF2). try_simplify_someHyps.
- exploit H3; clear H3. 1-3: wlp_simplify.
- intros (st3 & SEXEC3 & REF3). try_simplify_someHyps.
- eexists. split; eauto.
- intros ctx.
- eapply hsstate_simu_spec_correct; eauto.
-Qed.
-Global Opaque simu_check_single.
-Global Hint Resolve simu_check_single_correct: wlp.
-
-Fixpoint simu_check_rec (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) lm : ?? unit :=
- match lm with
- | nil => RET tt
- | m :: lm =>
- simu_check_single dm f tf m;;
- simu_check_rec dm f tf lm
- end.
-
-Lemma simu_check_rec_correct dm f tf lm:
- WHEN simu_check_rec dm f tf lm ~> _ THEN
- forall pc1 pc2, In (pc2, pc1) lm -> sexec_simu dm f tf pc1 pc2.
-Proof.
- induction lm; wlp_simplify.
- match goal with
- | X: (_,_) = (_,_) |- _ => inversion X; subst
- end.
- subst; eauto.
-Qed.
-Global Opaque simu_check_rec.
-Global Hint Resolve simu_check_rec_correct: wlp.
-
-Definition imp_simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function): ?? unit :=
- simu_check_rec dm f tf (PTree.elements dm);;
- DEBUG("simu_check OK!").
-
-Local Hint Resolve PTree.elements_correct: core.
-Lemma imp_simu_check_correct dm f tf:
- WHEN imp_simu_check dm f tf ~> _ THEN
- forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2.
-Proof.
- wlp_simplify.
-Qed.
-Global Opaque imp_simu_check.
-Global Hint Resolve imp_simu_check_correct: wlp.
-
-Program Definition aux_simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function): ?? bool :=
- DO r <~
- (TRY
- imp_simu_check dm f tf;;
- RET true
- CATCH_FAIL s, _ =>
- println ("simu_check_failure:" +; s);;
- RET false
- ENSURE (fun b => b=true -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2));;
- RET (`r).
-Obligation 1.
- split; wlp_simplify. discriminate.
-Qed.
-
-Lemma aux_simu_check_correct dm f tf:
- WHEN aux_simu_check dm f tf ~> b THEN
- b=true -> forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2.
-Proof.
- unfold aux_simu_check; wlp_simplify.
- destruct exta; simpl; auto.
-Qed.
-
-(* Coerce aux_simu_check into a pure function (this is a little unsafe like all oracles in CompCert). *)
-
-Import UnsafeImpure.
-
-Definition simu_check (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) : res unit :=
- match unsafe_coerce (aux_simu_check dm f tf) with
- | Some true => OK tt
- | _ => Error (msg "simu_check has failed")
- end.
-
-Lemma simu_check_correct dm f tf:
- simu_check dm f tf = OK tt ->
- forall pc1 pc2, dm ! pc2 = Some pc1 ->
- sexec_simu dm f tf pc1 pc2.
-Proof.
- unfold simu_check.
- destruct (unsafe_coerce (aux_simu_check dm f tf)) as [[|]|] eqn:Hres; simpl; try discriminate.
- intros; eapply aux_simu_check_correct; eauto.
- eapply unsafe_coerce_not_really_correct; eauto.
-Qed. \ No newline at end of file
diff --git a/kvx/lib/RTLpathSE_theory.v b/kvx/lib/RTLpathSE_theory.v
deleted file mode 100644
index 06b3a646..00000000
--- a/kvx/lib/RTLpathSE_theory.v
+++ /dev/null
@@ -1,1854 +0,0 @@
-(* A theory of symbolic execution on RTLpath
-
-NB: an efficient implementation with hash-consing will be defined in RTLpathSE_impl.v
-
-*)
-
-Require Import Coqlib Maps Floats.
-Require Import AST Integers Values Events Memory Globalenvs Smallstep.
-Require Import Op Registers.
-Require Import RTL RTLpath.
-Require Import Errors Duplicate.
-
-Local Open Scope error_monad_scope.
-
-(* Enhanced from kvx/Asmblockgenproof.v *)
-Ltac explore_hyp :=
- repeat match goal with
- | [ H : match ?var with | _ => _ end = _ |- _ ] => (let EQ1 := fresh "EQ" in (destruct var eqn:EQ1; try discriminate))
- | [ H : OK _ = OK _ |- _ ] => monadInv H
- | [ H : bind _ _ = OK _ |- _ ] => monadInv H
- | [ H : Error _ = OK _ |- _ ] => inversion H
- | [ H : Some _ = Some _ |- _ ] => inv H
- | [ x : unit |- _ ] => destruct x
- end.
-
-Ltac explore := explore_hyp;
- repeat match goal with
- | [ |- context[if ?b then _ else _] ] => (let EQ1 := fresh "IEQ" in destruct b eqn:EQ1)
- | [ |- context[match ?m with | _ => _ end] ] => (let DEQ1 := fresh "DEQ" in destruct m eqn:DEQ1)
- | [ |- context[match ?m as _ return _ with | _ => _ end]] => (let DREQ1 := fresh "DREQ" in destruct m eqn:DREQ1)
- end.
-
-(* Ltac explore :=
- repeat match goal with
- | [ H : match ?var with | _ => _ end = _ |- _ ] => (let EQ1 := fresh "EQ" in (destruct var eqn:EQ1; try discriminate))
- | [ H : OK _ = OK _ |- _ ] => monadInv H
- | [ |- context[if ?b then _ else _] ] => (let EQ1 := fresh "IEQ" in destruct b eqn:EQ1)
- | [ |- context[match ?m with | _ => _ end] ] => (let DEQ1 := fresh "DEQ" in destruct m eqn:DEQ1)
- | [ |- context[match ?m as _ return _ with | _ => _ end]] => (let DREQ1 := fresh "DREQ" in destruct m eqn:DREQ1)
- | [ H : bind _ _ = OK _ |- _ ] => monadInv H
- | [ H : Error _ = OK _ |- _ ] => inversion H
- | [ H : Some _ = Some _ |- _ ] => inv H
- | [ x : unit |- _ ] => destruct x
- end. *)
-
-(** * Syntax and semantics of symbolic values *)
-
-(* symbolic value *)
-Inductive sval :=
- | Sinput (r: reg)
- | Sop (op:operation) (lsv: list_sval) (sm: smem)
- | Sload (sm: smem) (trap: trapping_mode) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval)
-with list_sval :=
- | Snil
- | Scons (sv: sval) (lsv: list_sval)
-(* symbolic memory *)
-with smem :=
- | Sinit
- | Sstore (sm: smem) (chunk:memory_chunk) (addr:addressing) (lsv:list_sval) (srce: sval).
-
-Scheme sval_mut := Induction for sval Sort Prop
-with list_sval_mut := Induction for list_sval Sort Prop
-with smem_mut := Induction for smem Sort Prop.
-
-Fixpoint list_sval_inj (l: list sval): list_sval :=
- match l with
- | nil => Snil
- | v::l => Scons v (list_sval_inj l)
- end.
-
-Local Open Scope option_monad_scope.
-
-Fixpoint seval_sval (ge: RTL.genv) (sp:val) (sv: sval) (rs0: regset) (m0: mem): option val :=
- match sv with
- | Sinput r => Some (rs0#r)
- | Sop op l sm =>
- SOME args <- seval_list_sval ge sp l rs0 m0 IN
- SOME m <- seval_smem ge sp sm rs0 m0 IN
- eval_operation ge sp op args m
- | Sload sm trap chunk addr lsv =>
- match trap with
- | TRAP =>
- SOME args <- seval_list_sval ge sp lsv rs0 m0 IN
- SOME a <- eval_addressing ge sp addr args IN
- SOME m <- seval_smem ge sp sm rs0 m0 IN
- Mem.loadv chunk m a
- | NOTRAP =>
- SOME args <- seval_list_sval ge sp lsv rs0 m0 IN
- match (eval_addressing ge sp addr args) with
- | None => Some (default_notrap_load_value chunk)
- | Some a =>
- SOME m <- seval_smem ge sp sm rs0 m0 IN
- match (Mem.loadv chunk m a) with
- | None => Some (default_notrap_load_value chunk)
- | Some val => Some val
- end
- end
- end
- end
-with seval_list_sval (ge: RTL.genv) (sp:val) (lsv: list_sval) (rs0: regset) (m0: mem): option (list val) :=
- match lsv with
- | Snil => Some nil
- | Scons sv lsv' =>
- SOME v <- seval_sval ge sp sv rs0 m0 IN
- SOME lv <- seval_list_sval ge sp lsv' rs0 m0 IN
- Some (v::lv)
- end
-with seval_smem (ge: RTL.genv) (sp:val) (sm: smem) (rs0: regset) (m0: mem): option mem :=
- match sm with
- | Sinit => Some m0
- | Sstore sm chunk addr lsv srce =>
- SOME args <- seval_list_sval ge sp lsv rs0 m0 IN
- SOME a <- eval_addressing ge sp addr args IN
- SOME m <- seval_smem ge sp sm rs0 m0 IN
- SOME sv <- seval_sval ge sp srce rs0 m0 IN
- Mem.storev chunk m a sv
- end.
-
-(* Syntax and Semantics of local symbolic internal states *)
-(* [si_pre] is a precondition on initial ge, sp, rs0, m0 *)
-Record sistate_local := { si_pre: RTL.genv -> val -> regset -> mem -> Prop; si_sreg: reg -> sval; si_smem: smem }.
-
-(* Predicate on which (rs, m) is a possible final state after evaluating [st] on (rs0, m0) *)
-Definition ssem_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem) (rs: regset) (m: mem): Prop :=
- st.(si_pre) ge sp rs0 m0
- /\ seval_smem ge sp st.(si_smem) rs0 m0 = Some m
- /\ forall (r:reg), seval_sval ge sp (st.(si_sreg) r) rs0 m0 = Some (rs#r).
-
-Definition sabort_local (ge: RTL.genv) (sp:val) (st: sistate_local) (rs0: regset) (m0: mem): Prop :=
- ~(st.(si_pre) ge sp rs0 m0)
- \/ seval_smem ge sp st.(si_smem) rs0 m0 = None
- \/ exists (r: reg), seval_sval ge sp (st.(si_sreg) r) rs0 m0 = None.
-
-(* Syntax and semantics of symbolic exit states *)
-Record sistate_exit := mk_sistate_exit
- { si_cond: condition; si_scondargs: list_sval; si_elocal: sistate_local; si_ifso: node }.
-
-Definition seval_condition ge sp (cond: condition) (lsv: list_sval) (sm: smem) rs0 m0 : option bool :=
- SOME args <- seval_list_sval ge sp lsv rs0 m0 IN
- SOME m <- seval_smem ge sp sm rs0 m0 IN
- eval_condition cond args m.
-
-Definition all_fallthrough ge sp (lx: list sistate_exit) rs0 m0: Prop :=
- forall ext, List.In ext lx ->
- seval_condition ge sp ext.(si_cond) ext.(si_scondargs) ext.(si_elocal).(si_smem) rs0 m0 = Some false.
-
-Lemma all_fallthrough_revcons ge sp ext rs m lx:
- all_fallthrough ge sp (ext::lx) rs m ->
- seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs m = Some false
- /\ all_fallthrough ge sp lx rs m.
-Proof.
- intros ALLFU. constructor.
- - assert (In ext (ext::lx)) by (constructor; auto). apply ALLFU in H. assumption.
- - intros ext' INEXT. assert (In ext' (ext::lx)) by (apply in_cons; auto).
- apply ALLFU in H. assumption.
-Qed.
-
-(** Semantic of an exit in pseudo code:
- if si_cond (si_condargs)
- si_elocal; goto if_so
- else ()
-*)
-
-Definition ssem_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) rs' m' (pc': node) : Prop :=
- seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m = Some true
- /\ ssem_local ge sp (si_elocal ext) rs m rs' m'
- /\ (si_ifso ext) = pc'.
-
-(* Either an abort on the condition evaluation OR an abort on the sistate_local IF the condition was true *)
-Definition sabort_exit (ge: RTL.genv) (sp: val) (ext: sistate_exit) (rs: regset) (m: mem) : Prop :=
- let sev_cond := seval_condition ge sp (si_cond ext) (si_scondargs ext) ext.(si_elocal).(si_smem) rs m in
- sev_cond = None
- \/ (sev_cond = Some true /\ sabort_local ge sp ext.(si_elocal) rs m).
-
-(** * Syntax and Semantics of symbolic internal state *)
-Record sistate := { si_pc: node; si_exits: list sistate_exit; si_local: sistate_local }.
-
-Definition all_fallthrough_upto_exit ge sp ext lx' lx rs m : Prop :=
- is_tail (ext::lx') lx /\ all_fallthrough ge sp lx' rs m.
-
-(** Semantic of a sistate in pseudo code:
- si_exit1; si_exit2; ...; si_exitn;
- si_local; goto si_pc *)
-
-(* Note: in RTLpath, is.(icontinue) = false iff we took an early exit *)
-
-Definition ssem_internal (ge: RTL.genv) (sp:val) (st: sistate) (rs: regset) (m: mem) (is: istate): Prop :=
- if (is.(icontinue))
- then
- ssem_local ge sp st.(si_local) rs m is.(irs) is.(imem)
- /\ st.(si_pc) = is.(ipc)
- /\ all_fallthrough ge sp st.(si_exits) rs m
- else exists ext lx,
- ssem_exit ge sp ext rs m is.(irs) is.(imem) is.(ipc)
- /\ all_fallthrough_upto_exit ge sp ext lx st.(si_exits) rs m.
-
-Definition sabort (ge: RTL.genv) (sp: val) (st: sistate) (rs: regset) (m: mem): Prop :=
- (* No early exit was met but we aborted on the si_local *)
- (all_fallthrough ge sp st.(si_exits) rs m /\ sabort_local ge sp st.(si_local) rs m)
- (* OR we aborted on an evaluation of one of the early exits *)
- \/ (exists ext lx, all_fallthrough_upto_exit ge sp ext lx st.(si_exits) rs m /\ sabort_exit ge sp ext rs m).
-
-Definition ssem_internal_opt ge sp (st: sistate) rs0 m0 (ois: option istate): Prop :=
- match ois with
- | Some is => ssem_internal ge sp st rs0 m0 is
- | None => sabort ge sp st rs0 m0
- end.
-
-Definition ssem_internal_opt2 ge sp (ost: option sistate) rs0 m0 (ois: option istate) : Prop :=
- match ost with
- | Some st => ssem_internal_opt ge sp st rs0 m0 ois
- | None => ois=None
- end.
-
-(** * An internal state represents a parallel program !
-
- We prove below that the semantics [ssem_internal_opt] is deterministic.
-
- *)
-
-Definition istate_eq ist1 ist2 :=
- ist1.(icontinue) = ist2.(icontinue) /\
- ist1.(ipc) = ist2.(ipc) /\
- (forall r, (ist1.(irs)#r) = ist2.(irs)#r) /\
- ist1.(imem) = ist2.(imem).
-
-Lemma all_fallthrough_noexit ge sp ext lx rs0 m0 rs m pc:
- ssem_exit ge sp ext rs0 m0 rs m pc ->
- In ext lx ->
- all_fallthrough ge sp lx rs0 m0 ->
- False.
-Proof.
- Local Hint Resolve is_tail_in: core.
- intros SSEM INE ALLF.
- destruct SSEM as (SSEM & SSEM').
- unfold all_fallthrough in ALLF. rewrite ALLF in SSEM; eauto.
- discriminate.
-Qed.
-
-Lemma ssem_internal_exclude_incompatible_continue ge sp st rs m is1 is2:
- is1.(icontinue) = true ->
- is2.(icontinue) = false ->
- ssem_internal ge sp st rs m is1 ->
- ssem_internal ge sp st rs m is2 ->
- False.
-Proof.
- Local Hint Resolve all_fallthrough_noexit: core.
- unfold ssem_internal.
- intros CONT1 CONT2.
- rewrite CONT1, CONT2; simpl.
- intuition eauto.
- destruct H0 as (ext & lx & SSEME & ALLFU).
- destruct ALLFU as (ALLFU & ALLFU').
- eapply all_fallthrough_noexit; eauto.
-Qed.
-
-Lemma ssem_internal_determ_continue ge sp st rs m is1 is2:
- ssem_internal ge sp st rs m is1 ->
- ssem_internal ge sp st rs m is2 ->
- is1.(icontinue) = is2.(icontinue).
-Proof.
- Local Hint Resolve ssem_internal_exclude_incompatible_continue: core.
- destruct (Bool.bool_dec is1.(icontinue) is2.(icontinue)) as [|H]; auto.
- intros H1 H2. assert (absurd: False); intuition.
- destruct (icontinue is1) eqn: His1, (icontinue is2) eqn: His2; eauto.
-Qed.
-
-Lemma ssem_local_determ ge sp st rs0 m0 rs1 m1 rs2 m2:
- ssem_local ge sp st rs0 m0 rs1 m1 ->
- ssem_local ge sp st rs0 m0 rs2 m2 ->
- (forall r, rs1#r = rs2#r) /\ m1 = m2.
-Proof.
- unfold ssem_local. intuition try congruence.
- generalize (H5 r); rewrite H4; congruence.
-Qed.
-
-(* TODO: lemma to move in Coqlib *)
-Lemma is_tail_bounded_total {A} (l1 l2 l3: list A): is_tail l1 l3 -> is_tail l2 l3
- -> is_tail l1 l2 \/ is_tail l2 l1.
-Proof.
- Local Hint Resolve is_tail_cons: core.
- induction 1 as [|i l1 l3 T1 IND]; simpl; auto.
- intros T2; inversion T2; subst; auto.
-Qed.
-
-Lemma exit_cond_determ ge sp rs0 m0 l1 l2:
- is_tail l1 l2 -> forall ext1 lx1 ext2 lx2,
- l1=(ext1 :: lx1) ->
- l2=(ext2 :: lx2) ->
- all_fallthrough ge sp lx1 rs0 m0 ->
- seval_condition ge sp (si_cond ext1) (si_scondargs ext1) (si_smem (si_elocal ext1)) rs0 m0 = Some true ->
- all_fallthrough ge sp lx2 rs0 m0 ->
- ext1=ext2.
-Proof.
- destruct 1 as [l1|i l1 l3 T1]; intros ext1 lx1 ext2 lx2 EQ1 EQ2; subst;
- inversion EQ2; subst; auto.
- intros D1 EVAL NYE.
- Local Hint Resolve is_tail_in: core.
- unfold all_fallthrough in NYE.
- rewrite NYE in EVAL; eauto.
- try congruence.
-Qed.
-
-Lemma ssem_exit_determ ge sp ext rs0 m0 rs1 m1 pc1 rs2 m2 pc2:
- ssem_exit ge sp ext rs0 m0 rs1 m1 pc1 ->
- ssem_exit ge sp ext rs0 m0 rs2 m2 pc2 ->
- pc1 = pc2 /\ (forall r, rs1#r = rs2#r) /\ m1 = m2.
-Proof.
- Local Hint Resolve exit_cond_determ eq_sym: core.
- intros SSEM1 SSEM2. destruct SSEM1 as (SEVAL1 & SLOC1 & PCEQ1). destruct SSEM2 as (SEVAL2 & SLOC2 & PCEQ2). subst.
- destruct (ssem_local_determ ge sp (si_elocal ext) rs0 m0 rs1 m1 rs2 m2); auto.
-Qed.
-
-Remark is_tail_inv_left {A: Type} (a a': A) l l':
- is_tail (a::l) (a'::l') ->
- (a = a' /\ l = l') \/ (In a l' /\ is_tail l (a'::l')).
-Proof.
- intros. inv H.
- - left. eauto.
- - right. econstructor.
- + eapply is_tail_in; eauto.
- + eapply is_tail_cons_left; eauto.
-Qed.
-
-Lemma ssem_internal_determ ge sp st rs m is1 is2:
- ssem_internal ge sp st rs m is1 ->
- ssem_internal ge sp st rs m is2 ->
- istate_eq is1 is2.
-Proof.
- unfold istate_eq.
- intros SEM1 SEM2.
- exploit (ssem_internal_determ_continue ge sp st rs m is1 is2); eauto.
- intros CONTEQ. unfold ssem_internal in * |-. rewrite CONTEQ in * |- *.
- destruct (icontinue is2).
- - destruct (ssem_local_determ ge sp (si_local st) rs m (irs is1) (imem is1) (irs is2) (imem is2));
- intuition (try congruence).
- - destruct SEM1 as (ext1 & lx1 & SSEME1 & ALLFU1). destruct SEM2 as (ext2 & lx2 & SSEME2 & ALLFU2).
- destruct ALLFU1 as (ALLFU1 & ALLFU1'). destruct ALLFU2 as (ALLFU2 & ALLFU2').
- destruct SSEME1 as (SSEME1 & SSEME1' & SSEME1''). destruct SSEME2 as (SSEME2 & SSEME2' & SSEME2'').
- assert (X:ext1=ext2).
- { destruct (is_tail_bounded_total (ext1 :: lx1) (ext2 :: lx2) (si_exits st)) as [TAIL|TAIL]; eauto. }
- subst. destruct (ssem_local_determ ge sp (si_elocal ext2) rs m (irs is1) (imem is1) (irs is2) (imem is2)); auto.
- intuition. congruence.
-Qed.
-
-Lemma ssem_local_exclude_sabort_local ge sp loc rs m rs' m':
- ssem_local ge sp loc rs m rs' m' ->
-(* all_fallthrough ge sp (si_exits st) rs m -> *)
- sabort_local ge sp loc rs m ->
- False.
-Proof.
- intros SIML (* ALLF *) ABORT. inv SIML. destruct H0 as (H0 & H0').
- inversion ABORT as [ABORT1 | [ABORT2 | ABORT3]]; [ | | inv ABORT3]; congruence.
-Qed.
-
-(* TODO: remove this JUNK ?
-Lemma ssem_local_exclude_sabort_exit ge sp st ext lx rs m rs' m':
- ssem_local ge sp (si_local st) rs m rs' m' ->
- all_fallthrough ge sp (si_exits st) rs m ->
- is_tail (ext :: lx) (si_exits st) ->
- sabort_exit ge sp ext rs m ->
- False.
-Proof.
- intros SSEML ALLF TAIL ABORT.
- inv ABORT.
- - exploit ALLF; eauto. congruence.
- - (* FIXME Problem : if we have a ssem_local, this means we ONLY evaluated the conditions,
- but we NEVER actually evaluated the si_elocal from the sistate_exit ! So we cannot prove
- a lack of abort on the si_elocal.. We must change the definitions *)
-Abort.
-*)
-
-Lemma ssem_local_exclude_sabort ge sp st rs m rs' m':
- ssem_local ge sp (si_local st) rs m rs' m' ->
- all_fallthrough ge sp (si_exits st) rs m ->
- sabort ge sp st rs m ->
- False.
-Proof.
- intros SIML ALLF ABORT.
- inv ABORT.
- - intuition; eapply ssem_local_exclude_sabort_local; eauto.
- - destruct H as (ext & lx & ALLFU & SABORT).
- destruct ALLFU as (TAIL & _). eapply is_tail_in in TAIL.
- eapply ALLF in TAIL.
- destruct SABORT as [CONDFAIL | (CONDTRUE & ABORTL)]; congruence.
-Qed.
-
-Lemma ssem_exit_fallthrough_upto_exit ge sp ext ext' lx lx' exits rs m rs' m' pc':
- ssem_exit ge sp ext rs m rs' m' pc' ->
- all_fallthrough_upto_exit ge sp ext lx exits rs m ->
- all_fallthrough_upto_exit ge sp ext' lx' exits rs m ->
- is_tail (ext'::lx') (ext::lx).
-Proof.
- intros SSEME ALLFU ALLFU'.
- destruct ALLFU as (ISTAIL & ALLFU). destruct ALLFU' as (ISTAIL' & ALLFU').
- destruct (is_tail_bounded_total (ext::lx) (ext'::lx') exits); eauto.
- inv H.
- - econstructor; eauto.
- - eapply is_tail_in in H2. eapply ALLFU' in H2.
- destruct SSEME as (SEVAL & _). congruence.
-Qed.
-
-Lemma ssem_exit_exclude_sabort_exit ge sp ext rs m rs' m' pc':
- ssem_exit ge sp ext rs m rs' m' pc' ->
- sabort_exit ge sp ext rs m ->
- False.
-Proof.
- intros A B. destruct A as (A & A' & A''). inv B.
- - congruence.
- - destruct H as (_ & H). eapply ssem_local_exclude_sabort_local; eauto.
-Qed.
-
-Lemma ssem_exit_exclude_sabort ge sp ext st lx rs m rs' m' pc':
- ssem_exit ge sp ext rs m rs' m' pc' ->
- all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs m ->
- sabort ge sp st rs m ->
- False.
-Proof.
- intros SSEM ALLFU ABORT.
- inv ABORT.
- - destruct H as (ALLF & _). destruct ALLFU as (TAIL & _).
- eapply is_tail_in in TAIL.
- destruct SSEM as (SEVAL & _ & _).
- eapply ALLF in TAIL. congruence.
- - destruct H as (ext' & lx' & ALLFU' & ABORT).
- exploit ssem_exit_fallthrough_upto_exit; eauto. intros ITAIL.
- destruct ALLFU as (ALLFU1 & ALLFU2). destruct ALLFU' as (ALLFU1' & ALLFU2').
- exploit (is_tail_inv_left ext' ext lx' lx); eauto. intro. inv H.
- + inv H0. eapply ssem_exit_exclude_sabort_exit; eauto.
- + destruct H0 as (INE & TAIL). eapply ALLFU2 in INE. destruct ABORT as [ABORT | (ABORT & ABORT')]; congruence.
-Qed.
-
-Lemma ssem_internal_exclude_sabort ge sp st rs m is:
- sabort ge sp st rs m ->
- ssem_internal ge sp st rs m is -> False.
-Proof.
- intros ABORT SEM.
- unfold ssem_internal in SEM. destruct icontinue.
- - destruct SEM as (SEM1 & SEM2 & SEM3).
- eapply ssem_local_exclude_sabort; eauto.
- - destruct SEM as (ext & lx & SEM1 & SEM2). eapply ssem_exit_exclude_sabort; eauto.
-Qed.
-
-Definition istate_eq_opt ist1 oist :=
- exists ist2, oist = Some ist2 /\ istate_eq ist1 ist2.
-
-Lemma ssem_internal_opt_determ ge sp st rs m ois is:
- ssem_internal_opt ge sp st rs m ois ->
- ssem_internal ge sp st rs m is ->
- istate_eq_opt is ois.
-Proof.
- destruct ois as [is1|]; simpl; eauto.
- - intros; eexists; intuition; eapply ssem_internal_determ; eauto.
- - intros; exploit ssem_internal_exclude_sabort; eauto. destruct 1.
-Qed.
-
-(** * Symbolic execution of one internal step *)
-
-Definition slocal_set_sreg (st:sistate_local) (r:reg) (sv:sval) :=
- {| si_pre:=(fun ge sp rs m => seval_sval ge sp (st.(si_sreg) r) rs m <> None /\ (st.(si_pre) ge sp rs m));
- si_sreg:=fun y => if Pos.eq_dec r y then sv else st.(si_sreg) y;
- si_smem:= st.(si_smem)|}.
-
-Definition slocal_set_smem (st:sistate_local) (sm:smem) :=
- {| si_pre:=(fun ge sp rs m => seval_smem ge sp st.(si_smem) rs m <> None /\ (st.(si_pre) ge sp rs m));
- si_sreg:= st.(si_sreg);
- si_smem:= sm |}.
-
-Definition sist_set_local (st: sistate) (pc: node) (nxt: sistate_local): sistate :=
- {| si_pc := pc; si_exits := st.(si_exits); si_local:= nxt |}.
-
-Definition slocal_store st chunk addr args src : sistate_local :=
- let args := list_sval_inj (List.map (si_sreg st) args) in
- let src := si_sreg st src in
- let sm := Sstore (si_smem st) chunk addr args src
- in slocal_set_smem st sm.
-
-Definition siexec_inst (i: instruction) (st: sistate): option sistate :=
- match i with
- | Inop pc' =>
- Some (sist_set_local st pc' st.(si_local))
- | Iop op args dst pc' =>
- let prev := st.(si_local) in
- let vargs := list_sval_inj (List.map prev.(si_sreg) args) in
- let next := slocal_set_sreg prev dst (Sop op vargs prev.(si_smem)) in
- Some (sist_set_local st pc' next)
- | Iload trap chunk addr args dst pc' =>
- let prev := st.(si_local) in
- let vargs := list_sval_inj (List.map prev.(si_sreg) args) in
- let next := slocal_set_sreg prev dst (Sload prev.(si_smem) trap chunk addr vargs) in
- Some (sist_set_local st pc' next)
- | Istore chunk addr args src pc' =>
- let next := slocal_store st.(si_local) chunk addr args src in
- Some (sist_set_local st pc' next)
- | Icond cond args ifso ifnot _ =>
- let prev := st.(si_local) in
- let vargs := list_sval_inj (List.map prev.(si_sreg) args) in
- let ex := {| si_cond:=cond; si_scondargs:=vargs; si_elocal := prev; si_ifso := ifso |} in
- Some {| si_pc := ifnot; si_exits := ex::st.(si_exits); si_local := prev |}
- | _ => None (* TODO jumptable ? *)
- end.
-
-Lemma seval_list_sval_inj ge sp l rs0 m0 (sreg: reg -> sval) rs:
- (forall r : reg, seval_sval ge sp (sreg r) rs0 m0 = Some (rs # r)) ->
- seval_list_sval ge sp (list_sval_inj (map sreg l)) rs0 m0 = Some (rs ## l).
-Proof.
- intros H; induction l as [|r l]; simpl; auto.
- inversion_SOME v.
- inversion_SOME lv.
- generalize (H r).
- try_simplify_someHyps.
-Qed.
-
-Lemma slocal_set_sreg_preserves_sabort_local ge sp st rs0 m0 r sv:
- sabort_local ge sp st rs0 m0 ->
- sabort_local ge sp (slocal_set_sreg st r sv) rs0 m0.
-Proof.
- unfold sabort_local. simpl; intuition.
- destruct H as [r1 H]. destruct (Pos.eq_dec r r1) as [TEST|TEST] eqn: HTEST.
- - subst; rewrite H; intuition.
- - right. right. exists r1. rewrite HTEST. auto.
-Qed.
-
-Lemma slocal_set_smem_preserves_sabort_local ge sp st rs0 m0 m:
- sabort_local ge sp st rs0 m0 ->
- sabort_local ge sp (slocal_set_smem st m) rs0 m0.
-Proof.
- unfold sabort_local. simpl; intuition.
-Qed.
-
-Lemma all_fallthrough_upto_exit_cons ge sp ext lx ext' exits rs m:
- all_fallthrough_upto_exit ge sp ext lx exits rs m ->
- all_fallthrough_upto_exit ge sp ext lx (ext'::exits) rs m.
-Proof.
- intros. inv H. econstructor; eauto.
-Qed.
-
-Lemma all_fallthrough_cons ge sp exits rs m ext:
- all_fallthrough ge sp exits rs m ->
- seval_condition ge sp (si_cond ext) (si_scondargs ext) (si_smem (si_elocal ext)) rs m = Some false ->
- all_fallthrough ge sp (ext::exits) rs m.
-Proof.
- intros. unfold all_fallthrough in *. intros.
- inv H1; eauto.
-Qed.
-
-Lemma siexec_inst_preserves_sabort i ge sp rs m st st':
- siexec_inst i st = Some st' ->
- sabort ge sp st rs m -> sabort ge sp st' rs m.
-Proof.
- intros SISTEP ABORT.
- destruct i; simpl in SISTEP; try discriminate; inv SISTEP; unfold sabort; simpl.
- (* NOP *)
- * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - left. constructor; eauto.
- - right. exists ext0, lx0. constructor; eauto.
- (* OP *)
- * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - left. constructor; eauto. eapply slocal_set_sreg_preserves_sabort_local; eauto.
- - right. exists ext0, lx0. constructor; eauto.
- (* LOAD *)
- * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - left. constructor; eauto. eapply slocal_set_sreg_preserves_sabort_local; eauto.
- - right. exists ext0, lx0. constructor; eauto.
- (* STORE *)
- * destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - left. constructor; eauto. eapply slocal_set_smem_preserves_sabort_local; eauto.
- - right. exists ext0, lx0. constructor; eauto.
- (* COND *)
- * remember ({| si_cond := _; si_scondargs := _; si_elocal := _; si_ifso := _ |}) as ext.
- destruct ABORT as [(ALLF & ABORTL) | (ext0 & lx0 & ALLFU & ABORTE)].
- - destruct (seval_condition ge sp (si_cond ext) (si_scondargs ext)
- (si_smem (si_elocal ext)) rs m) eqn:SEVAL; [destruct b|].
- (* case true *)
- + right. exists ext, (si_exits st).
- constructor.
- ++ constructor. econstructor; eauto. eauto.
- ++ unfold sabort_exit. right. constructor; eauto.
- subst. simpl. eauto.
- (* case false *)
- + left. constructor; eauto. eapply all_fallthrough_cons; eauto.
- (* case None *)
- + right. exists ext, (si_exits st). constructor.
- ++ constructor. econstructor; eauto. eauto.
- ++ unfold sabort_exit. left. eauto.
- - right. exists ext0, lx0. constructor; eauto. eapply all_fallthrough_upto_exit_cons; eauto.
-Qed.
-
-Lemma siexec_inst_WF i st:
- siexec_inst i st = None -> default_succ i = None.
-Proof.
- destruct i; simpl; unfold sist_set_local; simpl; congruence.
-Qed.
-
-Lemma siexec_inst_default_succ i st st':
- siexec_inst i st = Some st' -> default_succ i = Some (st'.(si_pc)).
-Proof.
- destruct i; simpl; unfold sist_set_local; simpl; try congruence;
- intro H; inversion_clear H; simpl; auto.
-Qed.
-
-
-Lemma seval_list_sval_inj_not_none ge sp st rs0 m0: forall l,
- (forall r, List.In r l -> seval_sval ge sp (si_sreg st r) rs0 m0 = None -> False) ->
- seval_list_sval ge sp (list_sval_inj (map (si_sreg st) l)) rs0 m0 = None -> False.
-Proof.
- induction l.
- - intuition discriminate.
- - intros ALLR. simpl.
- inversion_SOME v.
- + intro SVAL. inversion_SOME lv; [discriminate|].
- assert (forall r : reg, In r l -> seval_sval ge sp (si_sreg st r) rs0 m0 = None -> False).
- { intros r INR. eapply ALLR. right. assumption. }
- intro SVALLIST. intro. eapply IHl; eauto.
- + intros. exploit (ALLR a); simpl; eauto.
-Qed.
-
-Lemma siexec_inst_correct ge sp i st rs0 m0 rs m:
- ssem_local ge sp st.(si_local) rs0 m0 rs m ->
- all_fallthrough ge sp st.(si_exits) rs0 m0 ->
- ssem_internal_opt2 ge sp (siexec_inst i st) rs0 m0 (istep ge i sp rs m).
-Proof.
- intros (PRE & MEM & REG) NYE.
- destruct i; simpl; auto.
- + (* Nop *)
- constructor; [|constructor]; simpl; auto.
- constructor; auto.
- + (* Op *)
- inversion_SOME v; intros OP; simpl.
- - constructor; [|constructor]; simpl; auto.
- constructor; simpl; auto.
- * constructor; auto. congruence.
- * constructor; auto.
- intro r0. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst. rewrite Regmap.gss; simpl; auto.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- - left. constructor; simpl; auto.
- unfold sabort_local. right. right.
- simpl. exists r. destruct (Pos.eq_dec r r); try congruence.
- simpl. erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- + (* LOAD *)
- inversion_SOME a0; intro ADD.
- { inversion_SOME v; intros LOAD; simpl.
- - explore_destruct; unfold ssem_internal, ssem_local; simpl; intuition.
- * unfold ssem_internal. simpl. constructor; [|constructor]; auto.
- constructor; constructor; simpl; auto. congruence. intro r0.
- destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst; rewrite Regmap.gss; simpl.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- * unfold ssem_internal. simpl. constructor; [|constructor]; auto.
- constructor; constructor; simpl; auto. congruence. intro r0.
- destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst; rewrite Regmap.gss; simpl.
- inversion_SOME args; intros ARGS.
- 2: { exploit seval_list_sval_inj_not_none; eauto; intuition congruence. }
- exploit seval_list_sval_inj; eauto. intro ARGS'. erewrite ARGS in ARGS'. inv ARGS'. rewrite ADD.
- inversion_SOME m2. intro SMEM.
- assert (m = m2) by congruence. subst. rewrite LOAD. reflexivity.
- - explore_destruct; unfold sabort, sabort_local; simpl.
- * unfold sabort. simpl. left. constructor; auto.
- right. right. exists r. simpl. destruct (Pos.eq_dec r r); try congruence.
- simpl. erewrite seval_list_sval_inj; simpl; auto.
- rewrite ADD; simpl; auto. try_simplify_someHyps.
- * unfold ssem_internal. simpl. constructor; [|constructor]; auto.
- constructor; constructor; simpl; auto. congruence. intro r0.
- destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst; rewrite Regmap.gss; simpl.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- } { rewrite ADD. destruct t.
- - simpl. left; eauto. simpl. econstructor; eauto.
- right. right. simpl. exists r. destruct (Pos.eq_dec r r); [|contradiction].
- simpl. inversion_SOME args. intro SLS.
- eapply seval_list_sval_inj in REG. rewrite REG in SLS. inv SLS.
- rewrite ADD. reflexivity.
- - simpl. constructor; [|constructor]; simpl; auto.
- constructor; simpl; constructor; auto; [congruence|].
- intro r0. destruct (Pos.eq_dec r r0); [|rewrite Regmap.gso; auto].
- subst. simpl. rewrite Regmap.gss.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- }
- + (* STORE *)
- inversion_SOME a0; intros ADD.
- { inversion_SOME m'; intros STORE; simpl.
- - unfold ssem_internal, ssem_local; simpl; intuition.
- * congruence.
- * erewrite seval_list_sval_inj; simpl; auto.
- erewrite REG.
- try_simplify_someHyps.
- - unfold sabort, sabort_local; simpl.
- left. constructor; auto. right. left.
- erewrite seval_list_sval_inj; simpl; auto.
- erewrite REG.
- try_simplify_someHyps. }
- { unfold sabort, sabort_local; simpl.
- left. constructor; auto. right. left.
- erewrite seval_list_sval_inj; simpl; auto.
- erewrite ADD; simpl; auto. }
- + (* COND *)
- Local Hint Resolve is_tail_refl: core.
- Local Hint Unfold ssem_local: core.
- inversion_SOME b; intros COND.
- { destruct b; simpl; unfold ssem_internal, ssem_local; simpl.
- - remember (mk_sistate_exit _ _ _ _) as ext. exists ext, (si_exits st).
- constructor; constructor; subst; simpl; auto.
- unfold seval_condition. subst; simpl.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps.
- - intuition. unfold all_fallthrough in * |- *. simpl.
- intuition. subst. simpl.
- unfold seval_condition.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps. }
- { unfold sabort. simpl. right.
- remember (mk_sistate_exit _ _ _ _) as ext. exists ext, (si_exits st).
- constructor; [constructor; subst; simpl; auto|].
- left. subst; simpl; auto.
- unfold seval_condition.
- erewrite seval_list_sval_inj; simpl; auto.
- try_simplify_someHyps. }
-Qed.
-
-
-Lemma siexec_inst_correct_None ge sp i st rs0 m0 rs m:
- ssem_local ge sp (st.(si_local)) rs0 m0 rs m ->
- siexec_inst i st = None ->
- istep ge i sp rs m = None.
-Proof.
- intros (PRE & MEM & REG).
- destruct i; simpl; unfold sist_set_local, ssem_internal, ssem_local; simpl; try_simplify_someHyps.
-Qed.
-
-(** * Symbolic execution of the internal steps of a path *)
-Fixpoint siexec_path (path:nat) (f: function) (st: sistate): option sistate :=
- match path with
- | O => Some st
- | S p =>
- SOME i <- (fn_code f)!(st.(si_pc)) IN
- SOME st1 <- siexec_inst i st IN
- siexec_path p f st1
- end.
-
-Lemma siexec_inst_add_exits i st st':
- siexec_inst i st = Some st' ->
- ( si_exits st' = si_exits st \/ exists ext, si_exits st' = ext :: si_exits st ).
-Proof.
- destruct i; simpl; intro SISTEP; inversion_clear SISTEP; unfold siexec_inst; simpl; (discriminate || eauto).
-Qed.
-
-Lemma siexec_inst_preserves_allfu ge sp ext lx rs0 m0 st st' i:
- all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs0 m0 ->
- siexec_inst i st = Some st' ->
- all_fallthrough_upto_exit ge sp ext lx (si_exits st') rs0 m0.
-Proof.
- intros ALLFU SISTEP. destruct ALLFU as (ISTAIL & ALLF).
- constructor; eauto.
- destruct i; simpl in SISTEP; inversion_clear SISTEP; simpl; (discriminate || eauto).
-Qed.
-
-Lemma siexec_path_correct_false ge sp f rs0 m0 st' is:
- forall path,
- is.(icontinue)=false ->
- forall st, ssem_internal ge sp st rs0 m0 is ->
- siexec_path path f st = Some st' ->
- ssem_internal ge sp st' rs0 m0 is.
-Proof.
- induction path; simpl.
- - intros. congruence.
- - intros ICF st SSEM STEQ'.
- destruct ((fn_code f) ! (si_pc st)) eqn:FIC; [|discriminate].
- destruct (siexec_inst _ _) eqn:SISTEP; [|discriminate].
- eapply IHpath. 3: eapply STEQ'. eauto.
- unfold ssem_internal in SSEM. rewrite ICF in SSEM.
- destruct SSEM as (ext & lx & SEXIT & ALLFU).
- unfold ssem_internal. rewrite ICF. exists ext, lx.
- constructor; auto. eapply siexec_inst_preserves_allfu; eauto.
-Qed.
-
-Lemma siexec_path_preserves_sabort ge sp path f rs0 m0 st': forall st,
- siexec_path path f st = Some st' ->
- sabort ge sp st rs0 m0 -> sabort ge sp st' rs0 m0.
-Proof.
- Local Hint Resolve siexec_inst_preserves_sabort: core.
- induction path; simpl.
- + unfold sist_set_local; try_simplify_someHyps.
- + intros st; inversion_SOME i.
- inversion_SOME st1; eauto.
-Qed.
-
-Lemma siexec_path_WF path f: forall st,
- siexec_path path f st = None -> nth_default_succ (fn_code f) path st.(si_pc) = None.
-Proof.
- induction path; simpl.
- + unfold sist_set_local. intuition congruence.
- + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try tauto.
- destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl.
- - intros; erewrite siexec_inst_default_succ; eauto.
- - intros; erewrite siexec_inst_WF; eauto.
-Qed.
-
-Lemma siexec_path_default_succ path f st': forall st,
- siexec_path path f st = Some st' -> nth_default_succ (fn_code f) path st.(si_pc) = Some st'.(si_pc).
-Proof.
- induction path; simpl.
- + unfold sist_set_local. intros st H. inversion_clear H; simpl; try congruence.
- + intros st; destruct ((fn_code f) ! (si_pc st)); simpl; try congruence.
- destruct (siexec_inst i st) as [st1|] eqn: Hst1; simpl; try congruence.
- intros; erewrite siexec_inst_default_succ; eauto.
-Qed.
-
-Lemma siexec_path_correct_true ge sp path (f:function) rs0 m0: forall st is,
- is.(icontinue)=true ->
- ssem_internal ge sp st rs0 m0 is ->
- nth_default_succ (fn_code f) path st.(si_pc) <> None ->
- ssem_internal_opt2 ge sp (siexec_path path f st) rs0 m0
- (isteps ge path f sp is.(irs) is.(imem) is.(ipc))
- .
-Proof.
- Local Hint Resolve siexec_path_correct_false siexec_path_preserves_sabort siexec_path_WF: core.
- induction path; simpl.
- + intros st is CONT INV WF;
- unfold ssem_internal, sist_set_local in * |- *;
- try_simplify_someHyps. simpl.
- destruct is; simpl in * |- *; subst; intuition auto.
- + intros st is CONT; unfold ssem_internal at 1; rewrite CONT.
- intros (LOCAL & PC & NYE) WF.
- rewrite <- PC.
- inversion_SOME i; intro Hi; rewrite Hi in WF |- *; simpl; auto.
- exploit siexec_inst_correct; eauto.
- inversion_SOME st1; intros Hst1; erewrite Hst1; simpl.
- - inversion_SOME is1; intros His1;rewrite His1; simpl.
- * destruct (icontinue is1) eqn:CONT1.
- (* icontinue is0 = true *)
- intros; eapply IHpath; eauto.
- destruct i; simpl in * |- *; unfold sist_set_local in * |- *; try_simplify_someHyps.
- (* icontinue is0 = false -> EARLY EXIT *)
- destruct (siexec_path path f st1) as [st2|] eqn: Hst2; simpl; eauto.
- destruct WF. erewrite siexec_inst_default_succ; eauto.
- (* try_simplify_someHyps; eauto. *)
- * destruct (siexec_path path f st1) as [st2|] eqn: Hst2; simpl; eauto.
- - intros His1;rewrite His1; simpl; auto.
-Qed.
-
-(** REM: in the following two unused lemmas *)
-
-Lemma siexec_path_right_assoc_decompose f path: forall st st',
- siexec_path (S path) f st = Some st' ->
- exists st0, siexec_path path f st = Some st0 /\ siexec_path 1%nat f st0 = Some st'.
-Proof.
- induction path; simpl; eauto.
- intros st st'.
- inversion_SOME i1.
- inversion_SOME st1.
- try_simplify_someHyps; eauto.
-Qed.
-
-Lemma siexec_path_right_assoc_compose f path: forall st st0 st',
- siexec_path path f st = Some st0 ->
- siexec_path 1%nat f st0 = Some st' ->
- siexec_path (S path) f st = Some st'.
-Proof.
- induction path.
- + intros st st0 st' H. simpl in H.
- try_simplify_someHyps; auto.
- + intros st st0 st'.
- assert (X:exists x, x=(S path)); eauto.
- destruct X as [x X].
- intros H1 H2. rewrite <- X.
- generalize H1; clear H1. simpl.
- inversion_SOME i1. intros Hi1; rewrite Hi1.
- inversion_SOME st1. intros Hst1; rewrite Hst1.
- subst; eauto.
-Qed.
-
-(** * Symbolic (final) value of a path *)
-Inductive sfval :=
- | Snone
- | Scall (sig:signature) (svos: sval + ident) (lsv:list_sval) (res:reg) (pc:node)
- (* NB: [res] the return register is hard-wired ! Is it restrictive ? *)
- | Stailcall: signature -> sval + ident -> list_sval -> sfval
- | Sbuiltin (ef:external_function) (sargs: list (builtin_arg sval)) (res: builtin_res reg) (pc:node)
- | Sjumptable (sv: sval) (tbl: list node)
- | Sreturn: option sval -> sfval
-.
-
-Definition sfind_function (pge: RTLpath.genv) (ge: RTL.genv) (sp: val) (svos : sval + ident) (rs0: regset) (m0: mem): option fundef :=
- match svos with
- | inl sv => SOME v <- seval_sval ge sp sv rs0 m0 IN Genv.find_funct pge v
- | inr symb => SOME b <- Genv.find_symbol pge symb IN Genv.find_funct_ptr pge b
- end.
-
-Section SEVAL_BUILTIN_ARG. (* adapted from Events.v *)
-
-Variable ge: RTL.genv.
-Variable sp: val.
-Variable m: mem.
-Variable rs0: regset.
-Variable m0: mem.
-
-Inductive seval_builtin_arg: builtin_arg sval -> val -> Prop :=
- | seval_BA: forall x v,
- seval_sval ge sp x rs0 m0 = Some v ->
- seval_builtin_arg (BA x) v
- | seval_BA_int: forall n,
- seval_builtin_arg (BA_int n) (Vint n)
- | seval_BA_long: forall n,
- seval_builtin_arg (BA_long n) (Vlong n)
- | seval_BA_float: forall n,
- seval_builtin_arg (BA_float n) (Vfloat n)
- | seval_BA_single: forall n,
- seval_builtin_arg (BA_single n) (Vsingle n)
- | seval_BA_loadstack: forall chunk ofs v,
- Mem.loadv chunk m (Val.offset_ptr sp ofs) = Some v ->
- seval_builtin_arg (BA_loadstack chunk ofs) v
- | seval_BA_addrstack: forall ofs,
- seval_builtin_arg (BA_addrstack ofs) (Val.offset_ptr sp ofs)
- | seval_BA_loadglobal: forall chunk id ofs v,
- Mem.loadv chunk m (Senv.symbol_address ge id ofs) = Some v ->
- seval_builtin_arg (BA_loadglobal chunk id ofs) v
- | seval_BA_addrglobal: forall id ofs,
- seval_builtin_arg (BA_addrglobal id ofs) (Senv.symbol_address ge id ofs)
- | seval_BA_splitlong: forall hi lo vhi vlo,
- seval_builtin_arg hi vhi -> seval_builtin_arg lo vlo ->
- seval_builtin_arg (BA_splitlong hi lo) (Val.longofwords vhi vlo)
- | seval_BA_addptr: forall a1 a2 v1 v2,
- seval_builtin_arg a1 v1 -> seval_builtin_arg a2 v2 ->
- seval_builtin_arg (BA_addptr a1 a2)
- (if Archi.ptr64 then Val.addl v1 v2 else Val.add v1 v2).
-
-Definition seval_builtin_args (al: list (builtin_arg sval)) (vl: list val) : Prop :=
- list_forall2 seval_builtin_arg al vl.
-
-Lemma seval_builtin_arg_determ:
- forall a v, seval_builtin_arg a v -> forall v', seval_builtin_arg a v' -> v' = v.
-Proof.
- induction 1; intros v' EV; inv EV; try congruence.
- f_equal; eauto.
- apply IHseval_builtin_arg1 in H3. apply IHseval_builtin_arg2 in H5. subst; auto.
-Qed.
-
-Lemma eval_builtin_args_determ:
- forall al vl, seval_builtin_args al vl -> forall vl', seval_builtin_args al vl' -> vl' = vl.
-Proof.
- induction 1; intros v' EV; inv EV; f_equal; eauto using seval_builtin_arg_determ.
-Qed.
-
-End SEVAL_BUILTIN_ARG.
-
-Inductive ssem_final (pge: RTLpath.genv) (ge: RTL.genv) (sp:val) (npc: node) stack (f: function) (rs0: regset) (m0: mem): sfval -> regset -> mem -> trace -> state -> Prop :=
- | exec_Snone rs m:
- ssem_final pge ge sp npc stack f rs0 m0 Snone rs m E0 (State stack f sp npc rs m)
- | exec_Scall rs m sig svos lsv args res pc fd:
- sfind_function pge ge sp svos rs0 m0 = Some fd ->
- funsig fd = sig ->
- seval_list_sval ge sp lsv rs0 m0 = Some args ->
- ssem_final pge ge sp npc stack f rs0 m0 (Scall sig svos lsv res pc) rs m
- E0 (Callstate (Stackframe res f sp pc rs :: stack) fd args m)
- | exec_Stailcall stk rs m sig svos args fd m' lsv:
- sfind_function pge ge sp svos rs0 m0 = Some fd ->
- funsig fd = sig ->
- sp = Vptr stk Ptrofs.zero ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- seval_list_sval ge sp lsv rs0 m0 = Some args ->
- ssem_final pge ge sp npc stack f rs0 m0 (Stailcall sig svos lsv) rs m
- E0 (Callstate stack fd args m')
- | exec_Sbuiltin m' rs m vres res pc t sargs ef vargs:
- seval_builtin_args ge sp m rs0 m0 sargs vargs ->
- external_call ef ge vargs m t vres m' ->
- ssem_final pge ge sp npc stack f rs0 m0 (Sbuiltin ef sargs res pc) rs m
- t (State stack f sp pc (regmap_setres res vres rs) m')
- | exec_Sjumptable sv tbl pc' n rs m:
- seval_sval ge sp sv rs0 m0 = Some (Vint n) ->
- list_nth_z tbl (Int.unsigned n) = Some pc' ->
- ssem_final pge ge sp npc stack f rs0 m0 (Sjumptable sv tbl) rs m
- E0 (State stack f sp pc' rs m)
- | exec_Sreturn stk osv rs m m' v:
- sp = (Vptr stk Ptrofs.zero) ->
- Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
- match osv with Some sv => seval_sval ge sp sv rs0 m0 | None => Some Vundef end = Some v ->
- ssem_final pge ge sp npc stack f rs0 m0 (Sreturn osv) rs m
- E0 (Returnstate stack v m')
-.
-
-Record sstate := { internal:> sistate; final: sfval }.
-
-Inductive ssem pge (ge: RTL.genv) (sp:val) (st: sstate) stack f (rs0: regset) (m0: mem): trace -> state -> Prop :=
- | ssem_early is:
- is.(icontinue) = false ->
- ssem_internal ge sp st rs0 m0 is ->
- ssem pge ge sp st stack f rs0 m0 E0 (State stack f sp is.(ipc) is.(irs) is.(imem))
- | ssem_normal is t s:
- is.(icontinue) = true ->
- ssem_internal ge sp st rs0 m0 is ->
- ssem_final pge ge sp st.(si_pc) stack f rs0 m0 st.(final) is.(irs) is.(imem) t s ->
- ssem pge ge sp st stack f rs0 m0 t s
- .
-
-(* NB: generic function that could be put into [AST] file *)
-Fixpoint builtin_arg_map {A B} (f: A -> B) (arg: builtin_arg A) : builtin_arg B :=
- match arg with
- | BA x => BA (f x)
- | BA_int n => BA_int n
- | BA_long n => BA_long n
- | BA_float f => BA_float f
- | BA_single s => BA_single s
- | BA_loadstack chunk ptr => BA_loadstack chunk ptr
- | BA_addrstack ptr => BA_addrstack ptr
- | BA_loadglobal chunk id ptr => BA_loadglobal chunk id ptr
- | BA_addrglobal id ptr => BA_addrglobal id ptr
- | BA_splitlong ba1 ba2 => BA_splitlong (builtin_arg_map f ba1) (builtin_arg_map f ba2)
- | BA_addptr ba1 ba2 => BA_addptr (builtin_arg_map f ba1) (builtin_arg_map f ba2)
- end.
-
-Lemma seval_builtin_arg_correct ge sp rs m rs0 m0 sreg: forall arg varg,
- (forall r, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) ->
- eval_builtin_arg ge (fun r => rs # r) sp m arg varg ->
- seval_builtin_arg ge sp m rs0 m0 (builtin_arg_map sreg arg) varg.
-Proof.
- induction arg.
- all: try (intros varg SEVAL BARG; inv BARG; constructor; congruence).
- - intros varg SEVAL BARG. inv BARG. simpl. constructor.
- eapply IHarg1; eauto. eapply IHarg2; eauto.
- - intros varg SEVAL BARG. inv BARG. simpl. constructor.
- eapply IHarg1; eauto. eapply IHarg2; eauto.
-Qed.
-
-Lemma seval_builtin_args_correct ge sp rs m rs0 m0 sreg args vargs:
- (forall r, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) ->
- eval_builtin_args ge (fun r => rs # r) sp m args vargs ->
- seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_map sreg) args) vargs.
-Proof.
- induction 2.
- - constructor.
- - simpl. constructor; [| assumption].
- eapply seval_builtin_arg_correct; eauto.
-Qed.
-
-Lemma seval_builtin_arg_complete ge sp rs m rs0 m0 sreg: forall arg varg,
- (forall r, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) ->
- seval_builtin_arg ge sp m rs0 m0 (builtin_arg_map sreg arg) varg ->
- eval_builtin_arg ge (fun r => rs # r) sp m arg varg.
-Proof.
- induction arg.
- all: intros varg SEVAL BARG; try (inv BARG; constructor; congruence).
- - inv BARG. rewrite SEVAL in H0. inv H0. constructor.
- - inv BARG. simpl. constructor.
- eapply IHarg1; eauto. eapply IHarg2; eauto.
- - inv BARG. simpl. constructor.
- eapply IHarg1; eauto. eapply IHarg2; eauto.
-Qed.
-
-Lemma seval_builtin_args_complete ge sp rs m rs0 m0 sreg: forall args vargs,
- (forall r, seval_sval ge sp (sreg r) rs0 m0 = Some rs # r) ->
- seval_builtin_args ge sp m rs0 m0 (map (builtin_arg_map sreg) args) vargs ->
- eval_builtin_args ge (fun r => rs # r) sp m args vargs.
-Proof.
- induction args.
- - simpl. intros. inv H0. constructor.
- - intros vargs SEVAL BARG. simpl in BARG. inv BARG.
- constructor; [| eapply IHargs; eauto].
- eapply seval_builtin_arg_complete; eauto.
-Qed.
-
-(** * Symbolic execution of final step *)
-Definition sexec_final (i: instruction) (prev: sistate_local): sfval :=
- match i with
- | Icall sig ros args res pc =>
- let svos := sum_left_map prev.(si_sreg) ros in
- let sargs := list_sval_inj (List.map prev.(si_sreg) args) in
- Scall sig svos sargs res pc
- | Itailcall sig ros args =>
- let svos := sum_left_map prev.(si_sreg) ros in
- let sargs := list_sval_inj (List.map prev.(si_sreg) args) in
- Stailcall sig svos sargs
- | Ibuiltin ef args res pc =>
- let sargs := List.map (builtin_arg_map prev.(si_sreg)) args in
- Sbuiltin ef sargs res pc
- | Ireturn or =>
- let sor := SOME r <- or IN Some (prev.(si_sreg) r) in
- Sreturn sor
- | Ijumptable reg tbl =>
- let sv := prev.(si_sreg) reg in
- Sjumptable sv tbl
- | _ => Snone
- end.
-
-Lemma sexec_final_correct pge ge sp i (f:function) pc st stack rs0 m0 t rs m s:
- (fn_code f) ! pc = Some i ->
- pc = st.(si_pc) ->
- ssem_local ge sp (si_local st) rs0 m0 rs m ->
- path_last_step ge pge stack f sp pc rs m t s ->
- siexec_inst i st = None ->
- ssem_final pge ge sp pc stack f rs0 m0 (sexec_final i (si_local st)) rs m t s.
-Proof.
- intros PC1 PC2 (PRE&MEM&REG) LAST. destruct LAST; subst; try_simplify_someHyps; simpl.
- + (* Snone *) intro Hi; destruct i; simpl in Hi |- *; unfold sist_set_local in Hi; try congruence.
- + (* Icall *) intros; eapply exec_Scall; auto.
- - destruct ros; simpl in * |- *; auto.
- rewrite REG; auto.
- - erewrite seval_list_sval_inj; simpl; auto.
- + (* Itailcall *) intros. eapply exec_Stailcall; auto.
- - destruct ros; simpl in * |- *; auto.
- rewrite REG; auto.
- - erewrite seval_list_sval_inj; simpl; auto.
- + (* Ibuiltin *) intros. eapply exec_Sbuiltin; eauto.
- eapply seval_builtin_args_correct; eauto.
- + (* Ijumptable *) intros. eapply exec_Sjumptable; eauto. congruence.
- + (* Ireturn *) intros; eapply exec_Sreturn; simpl; eauto.
- destruct or; simpl; auto.
-Qed.
-
-Lemma sexec_final_complete i (f:function) pc st ge pge sp stack rs0 m0 t rs m s:
- (fn_code f) ! pc = Some i ->
- pc = st.(si_pc) ->
- ssem_local ge sp (si_local st) rs0 m0 rs m ->
- ssem_final pge ge sp pc stack f rs0 m0 (sexec_final i (si_local st)) rs m t s ->
- siexec_inst i st = None ->
- path_last_step ge pge stack f sp pc rs m t s.
-Proof.
- intros PC1 PC2 (PRE&MEM&REG) LAST HSIS.
- destruct i as [ (* Inop *) | (* Iop *) | (* Iload *) | (* Istore *)
- | (* Icall *) sig ros args res pc'
- | (* Itailcall *) sig ros args
- | (* Ibuiltin *) ef bargs br pc'
- | (* Icond *)
- | (* Ijumptable *) jr tbl
- | (*Ireturn*) or];
- subst; try_simplify_someHyps; try (unfold sist_set_local in HSIS; try congruence);
- inversion LAST; subst; clear LAST; simpl in * |- *.
- + (* Icall *)
- erewrite seval_list_sval_inj in * |- ; simpl; try_simplify_someHyps; auto.
- intros; eapply exec_Icall; eauto.
- destruct ros; simpl in * |- *; auto.
- rewrite REG in * |- ; auto.
- + (* Itailcall *)
- intros HPC SMEM. erewrite seval_list_sval_inj in H10; auto. inv H10.
- eapply exec_Itailcall; eauto.
- destruct ros; simpl in * |- *; auto.
- rewrite REG in * |- ; auto.
- + (* Ibuiltin *) intros HPC SMEM.
- eapply exec_Ibuiltin; eauto.
- eapply seval_builtin_args_complete; eauto.
- + (* Ijumptable *) intros HPC SMEM.
- eapply exec_Ijumptable; eauto.
- congruence.
- + (* Ireturn *)
- intros; subst. enough (v=regmap_optget or Vundef rs) as ->.
- * eapply exec_Ireturn; eauto.
- * intros; destruct or; simpl; congruence.
-Qed.
-
-(** * Main function of the symbolic execution *)
-
-Definition init_sistate_local := {| si_pre:= fun _ _ _ _ => True; si_sreg:= fun r => Sinput r; si_smem:= Sinit |}.
-
-Definition init_sistate pc := {| si_pc:= pc; si_exits:=nil; si_local:= init_sistate_local |}.
-
-Lemma init_ssem_internal ge sp pc rs m: ssem_internal ge sp (init_sistate pc) rs m (mk_istate true pc rs m).
-Proof.
- unfold ssem_internal, ssem_local, all_fallthrough; simpl. intuition.
-Qed.
-
-Definition sexec (f: function) (pc:node): option sstate :=
- SOME path <- (fn_path f)!pc IN
- SOME st <- siexec_path path.(psize) f (init_sistate pc) IN
- SOME i <- (fn_code f)!(st.(si_pc)) IN
- Some (match siexec_inst i st with
- | Some st' => {| internal := st'; final := Snone |}
- | None => {| internal := st; final := sexec_final i st.(si_local) |}
- end).
-
-Lemma final_node_path_simpl f path pc:
- (fn_path f)!pc = Some path -> nth_default_succ_inst (fn_code f) path.(psize) pc <> None.
-Proof.
- intros; exploit final_node_path; eauto.
- intros (i & NTH & DUM).
- congruence.
-Qed.
-
-Lemma symb_path_last_step i st st' ge pge stack (f:function) sp pc rs m t s:
- (fn_code f) ! pc = Some i ->
- pc = st.(si_pc) ->
- siexec_inst i st = Some st' ->
- path_last_step ge pge stack f sp pc rs m t s ->
- exists mk_istate,
- istep ge i sp rs m = Some mk_istate
- /\ t = E0
- /\ s = (State stack f sp mk_istate.(ipc) mk_istate.(RTLpath.irs) mk_istate.(imem)).
-Proof.
- intros PC1 PC2 Hst' LAST; destruct LAST; subst; try_simplify_someHyps; simpl.
-Qed.
-
-(* NB: each concrete execution can be executed on the symbolic state (produced from [sexec])
-(sexec is a correct over-approximation)
-*)
-Theorem sexec_correct f pc pge ge sp path stack rs m t s:
- (fn_path f)!pc = Some path ->
- path_step ge pge path.(psize) stack f sp rs m pc t s ->
- exists st, sexec f pc = Some st /\ ssem pge ge sp st stack f rs m t s.
-Proof.
- Local Hint Resolve init_ssem_internal: core.
- intros PATH STEP; unfold sexec; rewrite PATH; simpl.
- lapply (final_node_path_simpl f path pc); eauto. intro WF.
- exploit (siexec_path_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto.
- { intros ABS. apply WF; unfold nth_default_succ_inst. rewrite ABS; auto. }
- (destruct (nth_default_succ_inst (fn_code f) path.(psize) pc) as [i|] eqn: Hi; [clear WF|congruence]).
- destruct STEP as [sti STEPS CONT|sti t s STEPS CONT LAST];
- (* intro Hst *)
- (rewrite STEPS; unfold ssem_internal_opt2; destruct (siexec_path _ _ _) as [st|] eqn: Hst; try congruence);
- (* intro SEM *)
- (simpl; unfold ssem_internal; simpl; rewrite CONT; intro SEM);
- (* intro Hi' *)
- ( assert (Hi': (fn_code f) ! (si_pc st) = Some i);
- [ unfold nth_default_succ_inst in Hi;
- exploit siexec_path_default_succ; eauto; simpl;
- intros DEF; rewrite DEF in Hi; auto
- | clear Hi; rewrite Hi' ]);
- (* eexists *)
- (eexists; constructor; eauto).
- - (* early *)
- eapply ssem_early; eauto.
- unfold ssem_internal; simpl; rewrite CONT.
- destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl; eauto.
- destruct SEM as (ext & lx & SEM & ALLFU). exists ext, lx.
- constructor; auto. eapply siexec_inst_preserves_allfu; eauto.
- - destruct SEM as (SEM & PC & HNYE).
- destruct (siexec_inst i st) as [st'|] eqn: Hst'; simpl.
- + (* normal on Snone *)
- rewrite <- PC in LAST.
- exploit symb_path_last_step; eauto; simpl.
- intros (mk_istate & ISTEP & Ht & Hs); subst.
- exploit siexec_inst_correct; eauto. simpl.
- erewrite Hst', ISTEP; simpl.
- clear LAST CONT STEPS PC SEM HNYE Hst Hi' Hst' ISTEP st sti i.
- intro SEM; destruct (mk_istate.(icontinue)) eqn: CONT.
- { (* icontinue mk_istate = true *)
- eapply ssem_normal; simpl; eauto.
- unfold ssem_internal in SEM.
- rewrite CONT in SEM.
- destruct SEM as (SEM & PC & HNYE).
- rewrite <- PC.
- eapply exec_Snone. }
- { eapply ssem_early; eauto. }
- + (* normal non-Snone instruction *)
- eapply ssem_normal; eauto.
- * unfold ssem_internal; simpl; rewrite CONT; intuition.
- * simpl. eapply sexec_final_correct; eauto.
- rewrite PC; auto.
-Qed.
-
-(* TODO: déplacer les trucs sur equiv_stackframe dans RTLpath ? *)
-Inductive equiv_stackframe: stackframe -> stackframe -> Prop :=
- | equiv_stackframe_intro res f sp pc rs1 rs2
- (EQUIV: forall r : positive, rs1 !! r = rs2 !! r):
- equiv_stackframe (Stackframe res f sp pc rs1) (Stackframe res f sp pc rs2).
-
-Inductive equiv_state: state -> state -> Prop :=
- | State_equiv stack f sp pc rs1 m rs2
- (EQUIV: forall r, rs1#r = rs2#r):
- equiv_state (State stack f sp pc rs1 m) (State stack f sp pc rs2 m)
- | Call_equiv stk stk' f args m
- (STACKS: list_forall2 equiv_stackframe stk stk'):
- equiv_state (Callstate stk f args m) (Callstate stk' f args m)
- | Return_equiv stk stk' v m
- (STACKS: list_forall2 equiv_stackframe stk stk'):
- equiv_state (Returnstate stk v m) (Returnstate stk' v m).
-
-Lemma equiv_stackframe_refl stf: equiv_stackframe stf stf.
-Proof.
- destruct stf. constructor; auto.
-Qed.
-
-Lemma equiv_stack_refl stk: list_forall2 equiv_stackframe stk stk.
-Proof.
- Local Hint Resolve equiv_stackframe_refl: core.
- induction stk; simpl; constructor; auto.
-Qed.
-
-Lemma equiv_state_refl s: equiv_state s s.
-Proof.
- Local Hint Resolve equiv_stack_refl: core.
- induction s; simpl; constructor; auto.
-Qed.
-
-(*
-Lemma equiv_stackframe_trans stf1 stf2 stf3:
- equiv_stackframe stf1 stf2 -> equiv_stackframe stf2 stf3 -> equiv_stackframe stf1 stf3.
-Proof.
- destruct 1; intros EQ; inv EQ; try econstructor; eauto.
- intros; eapply eq_trans; eauto.
-Qed.
-
-Lemma equiv_stack_trans stk1 stk2:
- list_forall2 equiv_stackframe stk1 stk2 ->
- forall stk3, list_forall2 equiv_stackframe stk2 stk3 ->
- list_forall2 equiv_stackframe stk1 stk3.
-Proof.
- Local Hint Resolve equiv_stackframe_trans.
- induction 1; intros stk3 EQ; inv EQ; econstructor; eauto.
-Qed.
-
-Lemma equiv_state_trans s1 s2 s3: equiv_state s1 s2 -> equiv_state s2 s3 -> equiv_state s1 s3.
-Proof.
- Local Hint Resolve equiv_stack_trans.
- destruct 1; intros EQ; inv EQ; econstructor; eauto.
- intros; eapply eq_trans; eauto.
-Qed.
-*)
-
-Lemma regmap_setres_eq (rs rs': regset) res vres:
- (forall r, rs # r = rs' # r) ->
- forall r, (regmap_setres res vres rs) # r = (regmap_setres res vres rs') # r.
-Proof.
- intros RSEQ r. destruct res; simpl; try congruence.
- destruct (peq x r).
- - subst. repeat (rewrite Regmap.gss). reflexivity.
- - repeat (rewrite Regmap.gso); auto.
-Qed.
-
-Lemma ssem_final_equiv pge ge sp (f:function) st sv stack rs0 m0 t rs1 rs2 m s:
- ssem_final pge ge sp st stack f rs0 m0 sv rs1 m t s ->
- (forall r, rs1#r = rs2#r) ->
- exists s', equiv_state s s' /\ ssem_final pge ge sp st stack f rs0 m0 sv rs2 m t s'.
-Proof.
- Local Hint Resolve equiv_stack_refl: core.
- destruct 1.
- - (* Snone *) intros; eexists; econstructor.
- + eapply State_equiv; eauto.
- + eapply exec_Snone.
- - (* Scall *)
- intros; eexists; econstructor.
- 2: { eapply exec_Scall; eauto. }
- apply Call_equiv; auto.
- repeat (constructor; auto).
- - (* Stailcall *)
- intros; eexists; econstructor; [| eapply exec_Stailcall; eauto].
- apply Call_equiv; auto.
- - (* Sbuiltin *)
- intros; eexists; econstructor; [| eapply exec_Sbuiltin; eauto].
- constructor. eapply regmap_setres_eq; eauto.
- - (* Sjumptable *)
- intros; eexists; econstructor; [| eapply exec_Sjumptable; eauto].
- constructor. assumption.
- - (* Sreturn *)
- intros; eexists; econstructor; [| eapply exec_Sreturn; eauto].
- eapply equiv_state_refl; eauto.
-Qed.
-
-Lemma siexec_inst_early_exit_absurd i st st' ge sp rs m rs' m' pc':
- siexec_inst i st = Some st' ->
- (exists ext lx, ssem_exit ge sp ext rs m rs' m' pc' /\
- all_fallthrough_upto_exit ge sp ext lx (si_exits st) rs m) ->
- all_fallthrough ge sp (si_exits st') rs m ->
- False.
-Proof.
- intros SIEXEC (ext & lx & SSEME & ALLFU) ALLF. destruct ALLFU as (TAIL & _).
- exploit siexec_inst_add_exits; eauto. destruct 1 as [SIEQ | (ext0 & SIEQ)].
- - rewrite SIEQ in *. eapply all_fallthrough_noexit. eauto. 2: eapply ALLF. eapply is_tail_in. eassumption.
- - rewrite SIEQ in *. eapply all_fallthrough_noexit. eauto. 2: eapply ALLF. eapply is_tail_in.
- constructor. eassumption.
-Qed.
-
-Lemma is_tail_false {A: Type}: forall (l: list A) a, is_tail (a::l) nil -> False.
-Proof.
- intros. eapply is_tail_incl in H. unfold incl in H. pose (H a).
- assert (In a (a::l)) by (constructor; auto). assert (In a nil) by auto. apply in_nil in H1.
- contradiction.
-Qed.
-
-Lemma cons_eq_false {A: Type}: forall (l: list A) a,
- a :: l = l -> False.
-Proof.
- induction l; intros.
- - discriminate.
- - inv H. apply IHl in H2. contradiction.
-Qed.
-
-Lemma app_cons_nil_eq {A: Type}: forall l' l (a:A),
- (l' ++ a :: nil) ++ l = l' ++ a::l.
-Proof.
- induction l'; intros.
- - simpl. reflexivity.
- - simpl. rewrite IHl'. reflexivity.
-Qed.
-
-Lemma app_eq_false {A: Type}: forall l (l': list A) a,
- l' ++ a :: l = l -> False.
-Proof.
- induction l; intros.
- - apply app_eq_nil in H. destruct H as (_ & H). apply cons_eq_false in H. contradiction.
- - destruct l' as [|a' l'].
- + simpl in H. apply cons_eq_false in H. contradiction.
- + rewrite <- app_comm_cons in H. inv H.
- apply (IHl (l' ++ (a0 :: nil)) a). rewrite app_cons_nil_eq. assumption.
-Qed.
-
-Lemma is_tail_false_gen {A: Type}: forall (l: list A) l' a, is_tail (l'++(a::l)) l -> False.
-Proof.
- induction l.
- - intros. destruct l' as [|a' l'].
- + simpl in H. apply is_tail_false in H. contradiction.
- + rewrite <- app_comm_cons in H. apply is_tail_false in H. contradiction.
- - intros. inv H.
- + apply app_eq_false in H2. contradiction.
- + apply (IHl (l' ++ (a0 :: nil)) a). rewrite app_cons_nil_eq. assumption.
-Qed.
-
-Lemma is_tail_eq {A: Type}: forall (l l': list A),
- is_tail l' l ->
- is_tail l l' ->
- l = l'.
-Proof.
- destruct l as [|a l]; intros l' ITAIL ITAIL'.
- - destruct l' as [|i' l']; auto. apply is_tail_false in ITAIL. contradiction.
- - inv ITAIL; auto.
- destruct l' as [|i' l']. { apply is_tail_false in ITAIL'. contradiction. }
- exploit is_tail_trans. eapply ITAIL'. eauto. intro ABSURD.
- apply (is_tail_false_gen l nil a) in ABSURD. contradiction.
-Qed.
-
-(* NB: each execution of a symbolic state (produced from [sexec]) represents a concrete execution
- (sexec is exact).
-*)
-Theorem sexec_exact f pc pge ge sp path stack st rs m t s1:
- (fn_path f)!pc = Some path ->
- sexec f pc = Some st ->
- ssem pge ge sp st stack f rs m t s1 ->
- exists s2, path_step ge pge path.(psize) stack f sp rs m pc t s2 /\
- equiv_state s1 s2.
-Proof.
- Local Hint Resolve init_ssem_internal: core.
- unfold sexec; intros PATH SSTEP SEM; rewrite PATH in SSTEP.
- lapply (final_node_path_simpl f path pc); eauto. intro WF.
- exploit (siexec_path_correct_true ge sp path.(psize) f rs m (init_sistate pc) (mk_istate true pc rs m)); simpl; eauto.
- { intros ABS. apply WF; unfold nth_default_succ_inst. rewrite ABS; auto. }
- (destruct (nth_default_succ_inst (fn_code f) path.(psize) pc) as [i|] eqn: Hi; [clear WF|congruence]).
- unfold nth_default_succ_inst in Hi.
- destruct (siexec_path path.(psize) f (init_sistate pc)) as [st0|] eqn: Hst0; simpl.
- 2:{ (* absurd case *)
- exploit siexec_path_WF; eauto.
- simpl; intros NDS; rewrite NDS in Hi; congruence. }
- exploit siexec_path_default_succ; eauto; simpl.
- intros NDS; rewrite NDS in Hi.
- rewrite Hi in SSTEP.
- intros ISTEPS. try_simplify_someHyps.
- destruct (siexec_inst i st0) as [st'|] eqn:Hst'; simpl.
- + (* exit on Snone instruction *)
- assert (SEM': t = E0 /\ exists is, ssem_internal ge sp st' rs m is
- /\ s1 = (State stack f sp (if (icontinue is) then (si_pc st') else (ipc is)) (irs is) (imem is))).
- { destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *.
- - repeat (econstructor; eauto).
- rewrite CONT; eauto.
- - inversion SEM2. repeat (econstructor; eauto).
- rewrite CONT; eauto. }
- clear SEM; subst. destruct SEM' as [X (is & SEM & X')]; subst.
- intros.
- destruct (isteps ge (psize path) f sp rs m pc) as [is0|] eqn:RISTEPS; simpl in *.
- * unfold ssem_internal in ISTEPS. destruct (icontinue is0) eqn: ICONT0.
- ** (* icontinue is0=true: path_step by normal_exit *)
- destruct ISTEPS as (SEMis0&H1&H2).
- rewrite H1 in * |-.
- exploit siexec_inst_correct; eauto.
- rewrite Hst'; simpl.
- intros; exploit ssem_internal_opt_determ; eauto.
- destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4).
- eexists. econstructor 1.
- *** eapply exec_normal_exit; eauto.
- eapply exec_istate; eauto.
- *** rewrite EQ1.
- enough ((ipc st) = (if icontinue st then si_pc st' else ipc is)) as ->.
- { rewrite EQ2, EQ4. eapply State_equiv; auto. }
- destruct (icontinue st) eqn:ICONT; auto.
- exploit siexec_inst_default_succ; eauto.
- erewrite istep_normal_exit; eauto.
- try_simplify_someHyps.
- ** (* The concrete execution has not reached "i" => early exit *)
- unfold ssem_internal in SEM.
- destruct (icontinue is) eqn:ICONT.
- { destruct SEM as (SEML & SIPC & ALLF).
- exploit siexec_inst_early_exit_absurd; eauto. contradiction. }
-
- eexists. econstructor 1.
- *** eapply exec_early_exit; eauto.
- *** destruct ISTEPS as (ext & lx & SSEME & ALLFU). destruct SEM as (ext' & lx' & SSEME' & ALLFU').
- eapply siexec_inst_preserves_allfu in ALLFU; eauto.
- exploit ssem_exit_fallthrough_upto_exit; eauto.
- exploit ssem_exit_fallthrough_upto_exit. eapply SSEME. eapply ALLFU. eapply ALLFU'.
- intros ITAIL ITAIL'. apply is_tail_eq in ITAIL; auto. clear ITAIL'.
- inv ITAIL. exploit ssem_exit_determ. eapply SSEME. eapply SSEME'. intros (IPCEQ & IRSEQ & IMEMEQ).
- rewrite <- IPCEQ. rewrite <- IMEMEQ. constructor. congruence.
- * (* The concrete execution has not reached "i" => abort case *)
- eapply siexec_inst_preserves_sabort in ISTEPS; eauto.
- exploit ssem_internal_exclude_sabort; eauto. contradiction.
- + destruct SEM as [is CONT SEM|is t s CONT SEM1 SEM2]; simpl in * |- *.
- - (* early exit *)
- intros.
- exploit ssem_internal_opt_determ; eauto.
- destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4).
- eexists. econstructor 1.
- * eapply exec_early_exit; eauto.
- * rewrite EQ2, EQ4; eapply State_equiv. auto.
- - (* normal exit non-Snone instruction *)
- intros.
- exploit ssem_internal_opt_determ; eauto.
- destruct 1 as (st & Hst & EQ1 & EQ2 & EQ3 & EQ4).
- unfold ssem_internal in SEM1.
- rewrite CONT in SEM1. destruct SEM1 as (SEM1 & PC0 & NYE0).
- exploit ssem_final_equiv; eauto.
- clear SEM2; destruct 1 as (s' & Ms' & SEM2).
- rewrite ! EQ4 in * |-; clear EQ4.
- rewrite ! EQ2 in * |-; clear EQ2.
- exists s'; intuition.
- eapply exec_normal_exit; eauto.
- eapply sexec_final_complete; eauto.
- * congruence.
- * unfold ssem_local in * |- *.
- destruct SEM1 as (A & B & C). constructor; [|constructor]; eauto.
- intro r. congruence.
- * congruence.
-Qed.
-
-(** * Simulation of RTLpath code w.r.t symbolic execution *)
-
-Section SymbValPreserved.
-
-Variable ge ge': RTL.genv.
-
-Hypothesis symbols_preserved_RTL: forall s, Genv.find_symbol ge' s = Genv.find_symbol ge s.
-
-Hypothesis senv_preserved_RTL: Senv.equiv ge ge'.
-
-Lemma senv_find_symbol_preserved id:
- Senv.find_symbol ge id = Senv.find_symbol ge' id.
-Proof.
- destruct senv_preserved_RTL as (A & B & C). congruence.
-Qed.
-
-Lemma senv_symbol_address_preserved id ofs:
- Senv.symbol_address ge id ofs = Senv.symbol_address ge' id ofs.
-Proof.
- unfold Senv.symbol_address. rewrite senv_find_symbol_preserved.
- reflexivity.
-Qed.
-
-Lemma seval_preserved sp sv rs0 m0:
- seval_sval ge sp sv rs0 m0 = seval_sval ge' sp sv rs0 m0.
-Proof.
- Local Hint Resolve symbols_preserved_RTL: core.
- induction sv using sval_mut with (P0 := fun lsv => seval_list_sval ge sp lsv rs0 m0 = seval_list_sval ge' sp lsv rs0 m0)
- (P1 := fun sm => seval_smem ge sp sm rs0 m0 = seval_smem ge' sp sm rs0 m0); simpl; auto.
- + rewrite IHsv; clear IHsv. destruct (seval_list_sval _ _ _ _); auto.
- rewrite IHsv0; clear IHsv0. destruct (seval_smem _ _ _ _); auto.
- erewrite eval_operation_preserved; eauto.
- + rewrite IHsv0; clear IHsv0. destruct (seval_list_sval _ _ _ _); auto.
- erewrite <- eval_addressing_preserved; eauto.
- destruct (eval_addressing _ sp _ _); auto.
- rewrite IHsv; auto.
- + rewrite IHsv; clear IHsv. destruct (seval_sval _ _ _ _); auto.
- rewrite IHsv0; auto.
- + rewrite IHsv0; clear IHsv0. destruct (seval_list_sval _ _ _ _); auto.
- erewrite <- eval_addressing_preserved; eauto.
- destruct (eval_addressing _ sp _ _); auto.
- rewrite IHsv; clear IHsv. destruct (seval_smem _ _ _ _); auto.
- rewrite IHsv1; auto.
-Qed.
-
-Lemma seval_builtin_arg_preserved sp m rs0 m0:
- forall bs varg,
- seval_builtin_arg ge sp m rs0 m0 bs varg ->
- seval_builtin_arg ge' sp m rs0 m0 bs varg.
-Proof.
- induction 1.
- all: try (constructor; auto).
- - rewrite <- seval_preserved. assumption.
- - rewrite <- senv_symbol_address_preserved. assumption.
- - rewrite senv_symbol_address_preserved. eapply seval_BA_addrglobal.
-Qed.
-
-Lemma seval_builtin_args_preserved sp m rs0 m0 lbs vargs:
- seval_builtin_args ge sp m rs0 m0 lbs vargs ->
- seval_builtin_args ge' sp m rs0 m0 lbs vargs.
-Proof.
- induction 1; constructor; eauto.
- eapply seval_builtin_arg_preserved; auto.
-Qed.
-
-Lemma list_sval_eval_preserved sp lsv rs0 m0:
- seval_list_sval ge sp lsv rs0 m0 = seval_list_sval ge' sp lsv rs0 m0.
-Proof.
- induction lsv; simpl; auto.
- rewrite seval_preserved. destruct (seval_sval _ _ _ _); auto.
- rewrite IHlsv; auto.
-Qed.
-
-Lemma smem_eval_preserved sp sm rs0 m0:
- seval_smem ge sp sm rs0 m0 = seval_smem ge' sp sm rs0 m0.
-Proof.
- induction sm; simpl; auto.
- rewrite list_sval_eval_preserved. destruct (seval_list_sval _ _ _ _); auto.
- erewrite <- eval_addressing_preserved; eauto.
- destruct (eval_addressing _ sp _ _); auto.
- rewrite IHsm; clear IHsm. destruct (seval_smem _ _ _ _); auto.
- rewrite seval_preserved; auto.
-Qed.
-
-Lemma seval_condition_preserved sp cond lsv sm rs0 m0:
- seval_condition ge sp cond lsv sm rs0 m0 = seval_condition ge' sp cond lsv sm rs0 m0.
-Proof.
- unfold seval_condition.
- rewrite list_sval_eval_preserved. destruct (seval_list_sval _ _ _ _); auto.
- rewrite smem_eval_preserved; auto.
-Qed.
-
-End SymbValPreserved.
-
-Require Import RTLpathLivegen RTLpathLivegenproof.
-
-(** * DEFINITION OF SIMULATION BETWEEN (ABSTRACT) SYMBOLIC EXECUTIONS
-*)
-
-Definition istate_simulive alive (srce: PTree.t node) (is1 is2: istate): Prop :=
- is1.(icontinue) = is2.(icontinue)
- /\ eqlive_reg alive is1.(irs) is2.(irs)
- /\ is1.(imem) = is2.(imem).
-
-Definition istate_simu f (srce: PTree.t node) is1 is2: Prop :=
- if is1.(icontinue) then
- (* TODO: il faudra raffiner le (fun _ => True) si on veut autoriser l'oracle à
- ajouter du "code mort" sur des registres non utilisés (loop invariant code motion à la David)
- Typiquement, pour connaître la frame des registres vivants, il faudra faire une propagation en arrière
- sur la dernière instruction du superblock. *)
- istate_simulive (fun _ => True) srce is1 is2
- else
- exists path, f.(fn_path)!(is1.(ipc)) = Some path
- /\ istate_simulive (fun r => Regset.In r path.(input_regs)) srce is1 is2
- /\ srce!(is2.(ipc)) = Some is1.(ipc).
-
-Record simu_proof_context {f1: RTLpath.function} := {
- liveness_hyps: liveness_ok_function f1;
- the_ge1: RTL.genv;
- the_ge2: RTL.genv;
- genv_match: forall s, Genv.find_symbol the_ge1 s = Genv.find_symbol the_ge2 s;
- the_sp: val;
- the_rs0: regset;
- the_m0: mem
-}.
-Arguments simu_proof_context: clear implicits.
-
-(* NOTE: a pure semantic definition on [sistate], for a total freedom in refinements *)
-Definition sistate_simu (dm: PTree.t node) (f: RTLpath.function) (st1 st2: sistate) (ctx: simu_proof_context f): Prop :=
- forall is1, ssem_internal (the_ge1 ctx) (the_sp ctx) st1 (the_rs0 ctx) (the_m0 ctx) is1 ->
- exists is2, ssem_internal (the_ge2 ctx) (the_sp ctx) st2 (the_rs0 ctx) (the_m0 ctx) is2
- /\ istate_simu f dm is1 is2.
-
-Inductive svident_simu (f: RTLpath.function) (ctx: simu_proof_context f): (sval + ident) -> (sval + ident) -> Prop :=
- | Sleft_simu sv1 sv2:
- (seval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx)) = (seval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx))
- -> svident_simu f ctx (inl sv1) (inl sv2)
- | Sright_simu id1 id2:
- id1 = id2
- -> svident_simu f ctx (inr id1) (inr id2)
- .
-
-
-Fixpoint ptree_get_list (pt: PTree.t node) (lp: list positive) : option (list positive) :=
- match lp with
- | nil => Some nil
- | p1::lp => SOME p2 <- pt!p1 IN
- SOME lp2 <- (ptree_get_list pt lp) IN
- Some (p2 :: lp2)
- end.
-
-Lemma ptree_get_list_nth dm p2: forall lp2 lp1,
- ptree_get_list dm lp2 = Some lp1 ->
- forall n, list_nth_z lp2 n = Some p2 ->
- exists p1,
- list_nth_z lp1 n = Some p1 /\ dm ! p2 = Some p1.
-Proof.
- induction lp2.
- - simpl. intros. inv H. simpl in *. discriminate.
- - intros lp1 PGL n LNZ. simpl in PGL. explore.
- inv LNZ. destruct (zeq n 0) eqn:ZEQ.
- + subst. inv H0. exists n0. simpl; constructor; auto.
- + exploit IHlp2; eauto. intros (p1 & LNZ & DMEQ).
- eexists. simpl. rewrite ZEQ.
- constructor; eauto.
-Qed.
-
-Lemma ptree_get_list_nth_rev dm p1: forall lp2 lp1,
- ptree_get_list dm lp2 = Some lp1 ->
- forall n, list_nth_z lp1 n = Some p1 ->
- exists p2,
- list_nth_z lp2 n = Some p2 /\ dm ! p2 = Some p1.
-Proof.
- induction lp2.
- - simpl. intros. inv H. simpl in *. discriminate.
- - intros lp1 PGL n LNZ. simpl in PGL. explore.
- inv LNZ. destruct (zeq n 0) eqn:ZEQ.
- + subst. inv H0. exists a. simpl; constructor; auto.
- + exploit IHlp2; eauto. intros (p2 & LNZ & DMEQ).
- eexists. simpl. rewrite ZEQ.
- constructor; eauto. congruence.
-Qed.
-
-Fixpoint seval_builtin_sval ge sp bsv rs0 m0 :=
- match bsv with
- | BA sv => SOME v <- seval_sval ge sp sv rs0 m0 IN Some (BA v)
- | BA_splitlong sv1 sv2 =>
- SOME v1 <- seval_builtin_sval ge sp sv1 rs0 m0 IN
- SOME v2 <- seval_builtin_sval ge sp sv2 rs0 m0 IN
- Some (BA_splitlong v1 v2)
- | BA_addptr sv1 sv2 =>
- SOME v1 <- seval_builtin_sval ge sp sv1 rs0 m0 IN
- SOME v2 <- seval_builtin_sval ge sp sv2 rs0 m0 IN
- Some (BA_addptr v1 v2)
- | BA_int i => Some (BA_int i)
- | BA_long l => Some (BA_long l)
- | BA_float f => Some (BA_float f)
- | BA_single s => Some (BA_single s)
- | BA_loadstack chk ptr => Some (BA_loadstack chk ptr)
- | BA_addrstack ptr => Some (BA_addrstack ptr)
- | BA_loadglobal chk id ptr => Some (BA_loadglobal chk id ptr)
- | BA_addrglobal id ptr => Some (BA_addrglobal id ptr)
- end.
-
-Fixpoint seval_list_builtin_sval ge sp lbsv rs0 m0 :=
- match lbsv with
- | nil => Some nil
- | bsv::lbsv => SOME v <- seval_builtin_sval ge sp bsv rs0 m0 IN
- SOME lv <- seval_list_builtin_sval ge sp lbsv rs0 m0 IN
- Some (v::lv)
- end.
-
-Lemma seval_list_builtin_sval_nil ge sp rs0 m0 lbs2:
- seval_list_builtin_sval ge sp lbs2 rs0 m0 = Some nil ->
- lbs2 = nil.
-Proof.
- destruct lbs2; simpl; auto.
- intros. destruct (seval_builtin_sval _ _ _ _ _);
- try destruct (seval_list_builtin_sval _ _ _ _ _); discriminate.
-Qed.
-
-Lemma seval_builtin_arg_sval ge sp m rs0 m0 v: forall bs,
- seval_builtin_arg ge sp m rs0 m0 bs v ->
- exists ba,
- seval_builtin_sval ge sp bs rs0 m0 = Some ba
- /\ eval_builtin_arg ge (fun id => id) sp m ba v.
-Proof.
- induction 1.
- all: try (eexists; constructor; [simpl; reflexivity | constructor]).
- 2-3: try assumption.
- - eexists. constructor.
- + simpl. rewrite H. reflexivity.
- + constructor.
- - destruct IHseval_builtin_arg1 as (ba1 & A1 & B1).
- destruct IHseval_builtin_arg2 as (ba2 & A2 & B2).
- eexists. constructor.
- + simpl. rewrite A1. rewrite A2. reflexivity.
- + constructor; assumption.
- - destruct IHseval_builtin_arg1 as (ba1 & A1 & B1).
- destruct IHseval_builtin_arg2 as (ba2 & A2 & B2).
- eexists. constructor.
- + simpl. rewrite A1. rewrite A2. reflexivity.
- + constructor; assumption.
-Qed.
-
-Lemma seval_builtin_args_sval ge sp m rs0 m0 lv: forall lbs,
- seval_builtin_args ge sp m rs0 m0 lbs lv ->
- exists lba,
- seval_list_builtin_sval ge sp lbs rs0 m0 = Some lba
- /\ list_forall2 (eval_builtin_arg ge (fun id => id) sp m) lba lv.
-Proof.
- induction 1.
- - eexists. constructor.
- + simpl. reflexivity.
- + constructor.
- - destruct IHlist_forall2 as (lba & A & B).
- apply seval_builtin_arg_sval in H. destruct H as (ba & A' & B').
- eexists. constructor.
- + simpl. rewrite A'. rewrite A. reflexivity.
- + constructor; assumption.
-Qed.
-
-Lemma seval_builtin_sval_correct ge sp m rs0 m0 v bs2: forall bs1,
- seval_builtin_arg ge sp m rs0 m0 bs1 v ->
- (seval_builtin_sval ge sp bs1 rs0 m0) = (seval_builtin_sval ge sp bs2 rs0 m0) ->
- seval_builtin_arg ge sp m rs0 m0 bs2 v.
-Proof.
- induction 1.
- - simpl. rewrite H. intros. destruct bs2; simpl in *; try congruence.
-Admitted.
-
-Lemma seval_list_builtin_sval_correct ge sp m rs0 m0 vargs: forall lbs1,
- seval_builtin_args ge sp m rs0 m0 lbs1 vargs ->
- forall lbs2, (seval_list_builtin_sval ge sp lbs1 rs0 m0) = (seval_list_builtin_sval ge sp lbs2 rs0 m0) ->
- seval_builtin_args ge sp m rs0 m0 lbs2 vargs.
-Proof.
- induction 1.
- - simpl. intros.
- exploit seval_list_builtin_sval_nil; eauto. intro. subst. constructor.
- - eapply seval_builtin_arg_sval in H. destruct H as (ba & A & B).
- eapply seval_builtin_args_sval in H0. destruct H0 as (lba & A' & B').
- intros. destruct lbs2.
- + simpl in H. rewrite A in H. rewrite A' in H.
- discriminate.
- + constructor. 2: apply IHlist_forall2. 2: admit.
-Admitted.
-
-(* NOTE: we need to mix semantical simulation and syntactic definition on [sfval] in order to abstract the [match_states] *)
-Inductive sfval_simu (dm: PTree.t node) (f: RTLpath.function) (opc1 opc2: node) (ctx: simu_proof_context f): sfval -> sfval -> Prop :=
- | Snone_simu:
- dm!opc2 = Some opc1 ->
- sfval_simu dm f opc1 opc2 ctx Snone Snone
- | Scall_simu sig svos1 svos2 lsv1 lsv2 res pc1 pc2:
- dm!pc2 = Some pc1 ->
- svident_simu f ctx svos1 svos2 ->
- (seval_list_sval (the_ge1 ctx) (the_sp ctx) lsv1 (the_rs0 ctx) (the_m0 ctx))
- = (seval_list_sval (the_ge2 ctx) (the_sp ctx) lsv2 (the_rs0 ctx) (the_m0 ctx)) ->
- sfval_simu dm f opc1 opc2 ctx (Scall sig svos1 lsv1 res pc1) (Scall sig svos2 lsv2 res pc2)
- | Stailcall_simu sig svos1 svos2 lsv1 lsv2:
- svident_simu f ctx svos1 svos2 ->
- (seval_list_sval (the_ge1 ctx) (the_sp ctx) lsv1 (the_rs0 ctx) (the_m0 ctx))
- = (seval_list_sval (the_ge2 ctx) (the_sp ctx) lsv2 (the_rs0 ctx) (the_m0 ctx)) ->
- sfval_simu dm f opc1 opc2 ctx (Stailcall sig svos1 lsv1) (Stailcall sig svos2 lsv2)
- | Sbuiltin_simu ef lbs1 lbs2 br pc1 pc2:
- dm!pc2 = Some pc1 ->
- (seval_list_builtin_sval (the_ge1 ctx) (the_sp ctx) lbs1 (the_rs0 ctx) (the_m0 ctx))
- = (seval_list_builtin_sval (the_ge2 ctx) (the_sp ctx) lbs2 (the_rs0 ctx) (the_m0 ctx)) ->
- sfval_simu dm f opc1 opc2 ctx (Sbuiltin ef lbs1 br pc1) (Sbuiltin ef lbs2 br pc2)
- | Sjumptable_simu sv1 sv2 lpc1 lpc2:
- ptree_get_list dm lpc2 = Some lpc1 ->
- (seval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx))
- = (seval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx)) ->
- sfval_simu dm f opc1 opc2 ctx (Sjumptable sv1 lpc1) (Sjumptable sv2 lpc2)
- | Sreturn_simu_none: sfval_simu dm f opc1 opc2 ctx (Sreturn None) (Sreturn None)
- | Sreturn_simu_some sv1 sv2:
- (seval_sval (the_ge1 ctx) (the_sp ctx) sv1 (the_rs0 ctx) (the_m0 ctx))
- = (seval_sval (the_ge2 ctx) (the_sp ctx) sv2 (the_rs0 ctx) (the_m0 ctx)) ->
- sfval_simu dm f opc1 opc2 ctx (Sreturn (Some sv1)) (Sreturn (Some sv2)).
-
-Definition sstate_simu dm f (s1 s2: sstate) (ctx: simu_proof_context f): Prop :=
- sistate_simu dm f s1.(internal) s2.(internal) ctx
- /\ forall is1,
- ssem_internal (the_ge1 ctx) (the_sp ctx) s1 (the_rs0 ctx) (the_m0 ctx) is1 ->
- is1.(icontinue) = true ->
- sfval_simu dm f s1.(si_pc) s2.(si_pc) ctx s1.(final) s2.(final).
-
-Definition sexec_simu dm (f1 f2: RTLpath.function) pc1 pc2: Prop :=
- forall st1, sexec f1 pc1 = Some st1 ->
- exists st2, sexec f2 pc2 = Some st2 /\ forall ctx, sstate_simu dm f1 st1 st2 ctx.
diff --git a/kvx/lib/RTLpathScheduler.v b/kvx/lib/RTLpathScheduler.v
deleted file mode 100644
index beab405f..00000000
--- a/kvx/lib/RTLpathScheduler.v
+++ /dev/null
@@ -1,330 +0,0 @@
-(** RTLpath Scheduling from an external oracle.
-
-This module is inspired from [Duplicate] and [Duplicateproof]
-
-*)
-
-Require Import AST Linking Values Maps Globalenvs Smallstep Registers.
-Require Import Coqlib Maps Events Errors Op.
-Require Import RTL RTLpath RTLpathLivegen RTLpathLivegenproof RTLpathSE_theory RTLpathSE_impl.
-
-
-Notation "'ASSERT' A 'WITH' MSG 'IN' B" := (if A then B else Error (msg MSG))
- (at level 200, A at level 100, B at level 200)
- : error_monad_scope.
-
-Local Open Scope error_monad_scope.
-Local Open Scope positive_scope.
-
-(** External oracle returning the new RTLpath function and a mapping of new path_entries to old path_entries
-
-NB: the new RTLpath function is generated from the fn_code, the fn_entrypoint and the fn_path
-It requires to check that the path structure is wf !
-
-*)
-
-(* Returns: new code, new entrypoint, new pathmap, revmap
- * Indeed, the entrypoint might not be the same if the entrypoint node is moved further down
- * a path ; same reasoning for the pathmap *)
-Axiom untrusted_scheduler: RTLpath.function -> code * node * path_map * (PTree.t node).
-
-Extract Constant untrusted_scheduler => "RTLpathScheduleraux.scheduler".
-
-Program Definition function_builder (tfr: RTL.function) (tpm: path_map) :
- { r : res RTLpath.function | forall f', r = OK f' -> fn_RTL f' = tfr} :=
- match RTLpathLivegen.function_checker tfr tpm with
- | false => Error (msg "In function_builder: (tfr, tpm) is not wellformed")
- | true => OK {| fn_RTL := tfr; fn_path := tpm |}
- end.
-Next Obligation.
- apply function_checker_path_entry. auto.
-Defined. Next Obligation.
- apply function_checker_wellformed_path_map. auto.
-Defined.
-
-Definition entrypoint_check (dm: PTree.t node) (fr tfr: RTL.function) : res unit :=
- match dm ! (fn_entrypoint tfr) with
- | None => Error (msg "No mapping for (entrypoint tfr)")
- | Some etp => if (Pos.eq_dec (fn_entrypoint fr) etp) then OK tt
- else Error (msg "Entrypoints do not match")
- end.
-
-Lemma entrypoint_check_correct fr tfr dm:
- entrypoint_check dm fr tfr = OK tt ->
- dm ! (fn_entrypoint tfr) = Some (fn_entrypoint fr).
-Proof.
- unfold entrypoint_check. explore; try discriminate. congruence.
-Qed.
-
-Definition path_entry_check_single (pm tpm: path_map) (m: node * node) :=
- let (pc2, pc1) := m in
- match (tpm ! pc2) with
- | None => Error (msg "pc2 isn't an entry of tpm")
- | Some _ =>
- match (pm ! pc1) with
- | None => Error (msg "pc1 isn't an entry of pm")
- | Some _ => OK tt
- end
- end.
-
-Lemma path_entry_check_single_correct pm tpm pc1 pc2:
- path_entry_check_single pm tpm (pc2, pc1) = OK tt ->
- path_entry tpm pc2 /\ path_entry pm pc1.
-Proof.
- unfold path_entry_check_single. intro. explore.
- constructor; congruence.
-Qed.
-
-(* Inspired from Duplicate.verify_mapping_rec *)
-Fixpoint path_entry_check_rec (pm tpm: path_map) lm :=
- match lm with
- | nil => OK tt
- | m :: lm => do u1 <- path_entry_check_single pm tpm m;
- do u2 <- path_entry_check_rec pm tpm lm;
- OK tt
- end.
-
-Lemma path_entry_check_rec_correct pm tpm pc1 pc2: forall lm,
- path_entry_check_rec pm tpm lm = OK tt ->
- In (pc2, pc1) lm ->
- path_entry tpm pc2 /\ path_entry pm pc1.
-Proof.
- induction lm.
- - simpl. intuition.
- - simpl. intros. explore. destruct H0.
- + subst. eapply path_entry_check_single_correct; eauto.
- + eapply IHlm; assumption.
-Qed.
-
-Definition path_entry_check (dm: PTree.t node) (pm tpm: path_map) := path_entry_check_rec pm tpm (PTree.elements dm).
-
-Lemma path_entry_check_correct dm pm tpm:
- path_entry_check dm pm tpm = OK tt ->
- forall pc1 pc2, dm ! pc2 = Some pc1 ->
- path_entry tpm pc2 /\ path_entry pm pc1.
-Proof.
- unfold path_entry_check. intros. eapply PTree.elements_correct in H0.
- eapply path_entry_check_rec_correct; eassumption.
-Qed.
-
-Definition function_equiv_checker (dm: PTree.t node) (f: RTLpath.function) (tf: RTLpath.function) : res unit :=
- let pm := fn_path f in
- let fr := fn_RTL f in
- let tpm := fn_path tf in
- let tfr := fn_RTL tf in
- do _ <- entrypoint_check dm fr tfr;
- do _ <- path_entry_check dm pm tpm;
- do _ <- simu_check dm f tf;
- OK tt.
-
-Lemma function_equiv_checker_entrypoint f tf dm:
- function_equiv_checker dm f tf = OK tt ->
- dm ! (fn_entrypoint tf) = Some (fn_entrypoint f).
-Proof.
- unfold function_equiv_checker. intros. explore.
- eapply entrypoint_check_correct; eauto.
-Qed.
-
-Lemma function_equiv_checker_pathentry1 f tf dm:
- function_equiv_checker dm f tf = OK tt ->
- forall pc1 pc2, dm ! pc2 = Some pc1 ->
- path_entry (fn_path tf) pc2.
-Proof.
- unfold function_equiv_checker. intros. explore.
- exploit path_entry_check_correct. eassumption. all: eauto. intuition.
-Qed.
-
-Lemma function_equiv_checker_pathentry2 f tf dm:
- function_equiv_checker dm f tf = OK tt ->
- forall pc1 pc2, dm ! pc2 = Some pc1 ->
- path_entry (fn_path f) pc1.
-Proof.
- unfold function_equiv_checker. intros. explore.
- exploit path_entry_check_correct. eassumption. all: eauto. intuition.
-Qed.
-
-Lemma function_equiv_checker_correct f tf dm:
- function_equiv_checker dm f tf = OK tt ->
- forall pc1 pc2, dm ! pc2 = Some pc1 ->
- sexec_simu dm f tf pc1 pc2.
-Proof.
- unfold function_equiv_checker. intros. explore.
- eapply simu_check_correct; eauto.
-Qed.
-
-Definition verified_scheduler (f: RTLpath.function) : res (RTLpath.function * (PTree.t node)) :=
- let (tctetpm, dm) := untrusted_scheduler f in
- let (tcte, tpm) := tctetpm in
- let (tc, te) := tcte in
- let tfr := mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te in
- do tf <- proj1_sig (function_builder tfr tpm);
- do tt <- function_equiv_checker dm f tf;
- OK (tf, dm).
-
-Theorem verified_scheduler_correct f tf dm:
- verified_scheduler f = OK (tf, dm) ->
- fn_sig f = fn_sig tf
- /\ fn_params f = fn_params tf
- /\ fn_stacksize f = fn_stacksize tf
- /\ dm ! (fn_entrypoint tf) = Some (fn_entrypoint f)
- /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path f) pc1)
- /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> path_entry (fn_path tf) pc2)
- /\ (forall pc1 pc2, dm ! pc2 = Some pc1 -> sexec_simu dm f tf pc1 pc2)
-.
-Proof.
- intros VERIF. unfold verified_scheduler in VERIF. explore.
- Local Hint Resolve function_equiv_checker_entrypoint
- function_equiv_checker_pathentry1 function_equiv_checker_pathentry2
- function_equiv_checker_correct: core.
- destruct (function_builder _ _) as [res H]; simpl in * |- *; auto.
- apply H in EQ2. rewrite EQ2. simpl.
- repeat (constructor; eauto).
- - exploit function_equiv_checker_entrypoint. eapply EQ4. rewrite EQ2. intuition.
-Qed.
-
-Record match_function (dupmap: PTree.t node) (f1 f2: RTLpath.function): Prop := {
- preserv_fnsig: fn_sig f1 = fn_sig f2;
- preserv_fnparams: fn_params f1 = fn_params f2;
- preserv_fnstacksize: fn_stacksize f1 = fn_stacksize f2;
- preserv_entrypoint: dupmap!(f2.(fn_entrypoint)) = Some f1.(fn_entrypoint);
- dupmap_path_entry1: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f1) pc1;
- dupmap_path_entry2: forall pc1 pc2, dupmap!pc2 = Some pc1 -> path_entry (fn_path f2) pc2;
- dupmap_correct: forall pc1 pc2, dupmap!pc2 = Some pc1 -> sexec_simu dupmap f1 f2 pc1 pc2;
-}.
-
-Program Definition transf_function (f: RTLpath.function):
- { r : res RTLpath.function | forall f', r = OK f' -> exists dm, match_function dm f f'} :=
- match (verified_scheduler f) with
- | Error e => Error e
- | OK (tf, dm) => OK tf
- end.
-Next Obligation.
- exploit verified_scheduler_correct; eauto.
- intros (A & B & C & D & E & F & G (* & H *)).
- exists dm. econstructor; eauto.
-Defined.
-
-Theorem match_function_preserves f f' dm:
- match_function dm f f' ->
- fn_sig f = fn_sig f' /\ fn_params f = fn_params f' /\ fn_stacksize f = fn_stacksize f'.
-Proof.
- intros.
- destruct H as [SIG PARAM SIZE ENTRY CORRECT].
- intuition.
-Qed.
-
-Definition transf_fundef (f: fundef) : res fundef :=
- transf_partial_fundef (fun f => proj1_sig (transf_function f)) f.
-
-Definition transf_program (p: program) : res program :=
- transform_partial_program transf_fundef p.
-
-(** * Preservation proof *)
-
-Local Notation ext alive := (fun r => Regset.In r alive).
-
-Inductive match_fundef: RTLpath.fundef -> RTLpath.fundef -> Prop :=
- | match_Internal dupmap f f': match_function dupmap f f' -> match_fundef (Internal f) (Internal f')
- | match_External ef: match_fundef (External ef) (External ef).
-
-Inductive match_stackframes: stackframe -> stackframe -> Prop :=
- | match_stackframe_intro dupmap res f sp pc rs1 rs2 f' pc' path
- (TRANSF: match_function dupmap f f')
- (DUPLIC: dupmap!pc' = Some pc)
- (LIVE: liveness_ok_function f)
- (PATH: f.(fn_path)!pc = Some path)
- (EQUIV: forall v, eqlive_reg (ext path.(input_regs)) (rs1 # res <- v) (rs2 # res <- v)):
- match_stackframes (Stackframe res f sp pc rs1) (Stackframe res f' sp pc' rs2).
-
-Inductive match_states: state -> state -> Prop :=
- | match_states_intro dupmap st f sp pc rs1 rs2 m st' f' pc' path
- (STACKS: list_forall2 match_stackframes st st')
- (TRANSF: match_function dupmap f f')
- (DUPLIC: dupmap!pc' = Some pc)
- (LIVE: liveness_ok_function f)
- (PATH: f.(fn_path)!pc = Some path)
- (EQUIV: eqlive_reg (ext path.(input_regs)) rs1 rs2):
- match_states (State st f sp pc rs1 m) (State st' f' sp pc' rs2 m)
- | match_states_call st st' f f' args m
- (STACKS: list_forall2 match_stackframes st st')
- (TRANSF: match_fundef f f')
- (LIVE: liveness_ok_fundef f):
- match_states (Callstate st f args m) (Callstate st' f' args m)
- | match_states_return st st' v m
- (STACKS: list_forall2 match_stackframes st st'):
- match_states (Returnstate st v m) (Returnstate st' v m).
-
-Lemma match_stackframes_equiv stf1 stf2 stf3:
- match_stackframes stf1 stf2 -> equiv_stackframe stf2 stf3 -> match_stackframes stf1 stf3.
-Proof.
- destruct 1; intros EQ; inv EQ; try econstructor; eauto.
- intros; eapply eqlive_reg_trans; eauto.
- rewrite eqlive_reg_triv in * |-.
- eapply eqlive_reg_update.
- eapply eqlive_reg_monotonic; eauto.
- simpl; auto.
-Qed.
-
-Lemma match_stack_equiv stk1 stk2:
- list_forall2 match_stackframes stk1 stk2 ->
- forall stk3, list_forall2 equiv_stackframe stk2 stk3 ->
- list_forall2 match_stackframes stk1 stk3.
-Proof.
- Local Hint Resolve match_stackframes_equiv: core.
- induction 1; intros stk3 EQ; inv EQ; econstructor; eauto.
-Qed.
-
-Lemma match_states_equiv s1 s2 s3: match_states s1 s2 -> equiv_state s2 s3 -> match_states s1 s3.
-Proof.
- Local Hint Resolve match_stack_equiv: core.
- destruct 1; intros EQ; inv EQ; econstructor; eauto.
- intros; eapply eqlive_reg_triv_trans; eauto.
-Qed.
-
-Lemma eqlive_match_stackframes stf1 stf2 stf3:
- eqlive_stackframes stf1 stf2 -> match_stackframes stf2 stf3 -> match_stackframes stf1 stf3.
-Proof.
- destruct 1; intros MS; inv MS; try econstructor; eauto.
- try_simplify_someHyps. intros; eapply eqlive_reg_trans; eauto.
-Qed.
-
-Lemma eqlive_match_stack stk1 stk2:
- list_forall2 eqlive_stackframes stk1 stk2 ->
- forall stk3, list_forall2 match_stackframes stk2 stk3 ->
- list_forall2 match_stackframes stk1 stk3.
-Proof.
- induction 1; intros stk3 MS; inv MS; econstructor; eauto.
- eapply eqlive_match_stackframes; eauto.
-Qed.
-
-Lemma eqlive_match_states s1 s2 s3: eqlive_states s1 s2 -> match_states s2 s3 -> match_states s1 s3.
-Proof.
- Local Hint Resolve eqlive_match_stack: core.
- destruct 1; intros MS; inv MS; try_simplify_someHyps; econstructor; eauto.
- eapply eqlive_reg_trans; eauto.
-Qed.
-
-Lemma eqlive_stackframes_refl stf1 stf2: match_stackframes stf1 stf2 -> eqlive_stackframes stf1 stf1.
-Proof.
- destruct 1; econstructor; eauto.
- intros; eapply eqlive_reg_refl; eauto.
-Qed.
-
-Lemma eqlive_stacks_refl stk1 stk2:
- list_forall2 match_stackframes stk1 stk2 -> list_forall2 eqlive_stackframes stk1 stk1.
-Proof.
- induction 1; simpl; econstructor; eauto.
- eapply eqlive_stackframes_refl; eauto.
-Qed.
-
-Lemma transf_fundef_correct f f':
- transf_fundef f = OK f' -> match_fundef f f'.
-Proof.
- intros TRANSF; destruct f; simpl; monadInv TRANSF.
- + destruct (transf_function f) as [res H]; simpl in * |- *; auto.
- destruct (H _ EQ).
- intuition subst; auto.
- eapply match_Internal; eauto.
- + eapply match_External.
-Qed.
-
diff --git a/kvx/lib/RTLpathScheduleraux.ml b/kvx/lib/RTLpathScheduleraux.ml
deleted file mode 100644
index 88f777a5..00000000
--- a/kvx/lib/RTLpathScheduleraux.ml
+++ /dev/null
@@ -1,368 +0,0 @@
-open RTLpath
-open RTL
-open Maps
-open RTLpathLivegenaux
-open Registers
-open Camlcoq
-
-type superblock = {
- instructions: P.t array; (* pointers to code instructions *)
- (* each predicted Pcb has its attached liveins *)
- (* This is indexed by the pc value *)
- liveins: Regset.t PTree.t;
- (* Union of the input_regs of the last successors *)
- output_regs: Regset.t;
- typing: RTLtyping.regenv
-}
-
-let print_instructions insts code =
- if (!debug_flag) then begin
- dprintf "[ ";
- List.iter (
- fun n -> (PrintRTL.print_instruction stdout (P.to_int n, get_some @@ PTree.get n code))
- ) insts; dprintf "]"
- end
-
-let print_superblock sb code =
- let insts = sb.instructions in
- let li = sb.liveins in
- let outs = sb.output_regs in
- begin
- dprintf "{ instructions = "; print_instructions (Array.to_list insts) code; dprintf "\n";
- dprintf " liveins = "; print_ptree_regset li; dprintf "\n";
- dprintf " output_regs = "; print_regset outs; dprintf "}"
- end
-
-let print_superblocks lsb code =
- let rec f = function
- | [] -> ()
- | sb :: lsb -> (print_superblock sb code; dprintf ",\n"; f lsb)
- in begin
- dprintf "[\n";
- f lsb;
- dprintf "]"
- end
-
-(* Adapted from backend/PrintRTL.ml: print_function *)
-let print_code code = let open PrintRTL in let open Printf in
- if (!debug_flag) then begin
- fprintf stdout "{\n";
- let instrs =
- List.sort
- (fun (pc1, _) (pc2, _) -> compare pc2 pc1)
- (List.rev_map
- (fun (pc, i) -> (P.to_int pc, i))
- (PTree.elements code)) in
- List.iter (print_instruction stdout) instrs;
- fprintf stdout "}"
- end
-
-let print_arrayp arr = begin
- dprintf "[| ";
- Array.iter (fun n -> dprintf "%d, " (P.to_int n)) arr;
- dprintf "|]"
-end
-
-let get_superblocks code entry pm typing =
- let visited = ref (PTree.map (fun n i -> false) code) in
- let rec get_superblocks_rec pc =
- let liveins = ref (PTree.empty) in
- let rec follow pc n =
- let inst = get_some @@ PTree.get pc code in
- if (n == 0) then begin
- (match (non_predicted_successors inst) with
- | [pcout] ->
- let live = (get_some @@ PTree.get pcout pm).input_regs in
- liveins := PTree.set pc live !liveins
- | _ -> ());
- ([pc], successors_inst inst)
- end else
- let nexts_from_exit = match (non_predicted_successors inst) with
- | [pcout] ->
- let live = (get_some @@ PTree.get pcout pm).input_regs in begin
- liveins := PTree.set pc live !liveins;
- [pcout]
- end
- | [] -> []
- | _ -> failwith "Having more than one non_predicted_successor is not handled"
- in match (predicted_successor inst) with
- | None -> failwith "Incorrect path"
- | Some succ ->
- let (insts, nexts) = follow succ (n-1) in (pc :: insts, nexts_from_exit @ nexts)
- in if (get_some @@ PTree.get pc !visited) then []
- else begin
- visited := PTree.set pc true !visited;
- let pi = get_some @@ PTree.get pc pm in
- let (insts, nexts) = follow pc (Camlcoq.Nat.to_int pi.psize) in
- let superblock = { instructions = Array.of_list insts; liveins = !liveins;
- output_regs = pi.output_regs; typing = typing } in
- superblock :: (List.concat @@ List.map get_superblocks_rec nexts)
- end
- in let lsb = get_superblocks_rec entry in begin
- (* debug_flag := true; *)
- dprintf "Superblocks identified:"; print_superblocks lsb code; dprintf "\n";
- (* debug_flag := false; *)
- lsb
-end
-
-(* TODO David *)
-let schedule_superblock sb code =
- if not !Clflags.option_fprepass
- then sb.instructions
- else
- let old_flag = !debug_flag in
- debug_flag := true;
- print_endline "ORIGINAL SUPERBLOCK";
- print_superblock sb code;
- debug_flag := old_flag;
- let nr_instr = Array.length sb.instructions in
- let trailer_length =
- match PTree.get (sb.instructions.(nr_instr-1)) code with
- | None -> 0
- | Some ii ->
- match predicted_successor ii with
- | Some _ -> 0
- | None -> 1 in
- match PrepassSchedulingOracle.schedule_sequence
- (Array.map (fun i ->
- (match PTree.get i code with
- | Some ii -> ii
- | None -> failwith "RTLpathScheduleraux.schedule_superblock"),
- (match PTree.get i sb.liveins with
- | Some s -> s
- | None -> Regset.empty))
- (Array.sub sb.instructions 0 (nr_instr-trailer_length))) with
- | None -> sb.instructions
- | Some order ->
- let ins' =
- Array.append
- (Array.map (fun i -> sb.instructions.(i)) order)
- (Array.sub sb.instructions (nr_instr-trailer_length) trailer_length) in
- Printf.printf "REORDERED SUPERBLOCK %d\n" (Array.length ins');
- debug_flag := true;
- print_instructions (Array.to_list ins') code;
- debug_flag := old_flag;
- flush stdout;
- assert ((Array.length sb.instructions) = (Array.length ins'));
- (*sb.instructions; *)
- ins';;
-
- (* stub2: reverse function *)
- (*
- let reversed = Array.of_list @@ List.rev @@ Array.to_list (sb.instructions) in
- let tmp = reversed.(0) in
- let last_index = Array.length reversed - 1 in
- begin
- reversed.(0) <- reversed.(last_index);
- reversed.(last_index) <- tmp;
- reversed
- end *)
- (* stub: identity function *)
-
-(**
- * Perform basic checks on the new order :
- * - must have the same length as the old order
- * - non basic instructions (call, tailcall, return, jumptable, non predicted CB) must not move
- *)
-let check_order code old_order new_order = begin
- assert ((Array.length old_order) == (Array.length new_order));
- let length = Array.length new_order in
- if length > 0 then
- let last_inst = Array.get old_order (length - 1) in
- let instr = get_some @@ PTree.get last_inst code in
- match predicted_successor instr with
- | None ->
- if (last_inst != Array.get new_order (length - 1)) then
- failwith "The last instruction of the superblock is not basic, but was moved"
- | _ -> ()
-end
-
-type sinst =
- (* Each middle instruction has a direct successor *)
- (* A Smid can be the last instruction of a superblock, but a Send cannot be moved *)
- | Smid of RTL.instruction * node
- | Send of RTL.instruction
-
-let rinst_to_sinst inst =
- match inst with
- | Inop n -> Smid(inst, n)
- | Iop (_,_,_,n) -> Smid(inst, n)
- | Iload (_,_,_,_,_,n) -> Smid(inst, n)
- | Istore (_,_,_,_,n) -> Smid(inst, n)
- | Icond (_,_,n1,n2,p) -> (
- match p with
- | Some true -> Smid(inst, n1)
- | Some false -> Smid(inst, n2)
- | None -> Send(inst)
- )
- | Icall _ | Ibuiltin _ | Ijumptable _ | Itailcall _ | Ireturn _ -> Send(inst)
-
-let change_predicted_successor s = function
- | Smid(i, n) -> Smid(i, s)
- | Send _ -> failwith "Called change_predicted_successor on Send. Are you trying to move a non-basic instruction in the middle of the block?"
-
-(* Forwards the successor changes into an RTL instruction *)
-let sinst_to_rinst = function
- | Smid(inst, s) -> (
- match inst with
- | Inop n -> Inop s
- | Iop (a,b,c,n) -> Iop (a,b,c,s)
- | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s)
- | Istore (a,b,c,d,n) -> Istore (a,b,c,d,s)
- | Icond (a,b,n1,n2,p) -> (
- match p with
- | Some true -> Icond(a, b, s, n2, p)
- | Some false -> Icond(a, b, n1, s, p)
- | None -> failwith "Non predicted Icond as a middle instruction!"
- )
- | _ -> failwith "That instruction shouldn't be a middle instruction"
- )
- | Send i -> i
-
-let apply_schedule code sb new_order =
- let tc = ref code in
- let old_order = sb.instructions in
- begin
- check_order code old_order new_order;
- Array.iteri (fun i n' ->
- let inst' = get_some @@ PTree.get n' code in
- let iend = Array.length old_order - 1 in
- let new_inst =
- if (i == iend) then
- let final_inst_node = Array.get old_order iend in
- let sinst' = rinst_to_sinst inst' in
- match sinst' with
- (* The below assert fails if a Send is in the middle of the original superblock *)
- | Send i -> (assert (final_inst_node == n'); i)
- | Smid _ ->
- let final_inst = get_some @@ PTree.get final_inst_node code in
- match rinst_to_sinst final_inst with
- | Smid (_, s') -> sinst_to_rinst @@ change_predicted_successor s' sinst'
- | Send _ -> assert(false) (* should have failed earlier *)
- else
- sinst_to_rinst
- (* this will fail if the moved instruction is a Send *)
- @@ change_predicted_successor (Array.get old_order (i+1))
- @@ rinst_to_sinst inst'
- in tc := PTree.set (Array.get old_order i) new_inst !tc
- ) new_order;
- !tc
- end
-
- (*
-let main_successors = function
-| Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> [n]
-| Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> [n]
-| Icond (_,_,n1,n2,p) -> (
- match p with
- | Some true -> [n1; n2]
- | Some false -> [n2; n1]
- | None -> [n1; n2] )
-| Ijumptable _ | Itailcall _ | Ireturn _ -> []
-
-let change_predicted_successor i s = match i with
- | Itailcall _ | Ireturn _ -> failwith "Wrong instruction (5) - Tailcalls and returns should not be moved in the middle of a superblock"
- | Ijumptable _ -> failwith "Wrong instruction (6) - Jumptables should not be moved in the middle of a superblock"
- | Inop n -> Inop s
- | Iop (a,b,c,n) -> Iop (a,b,c,s)
- | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s)
- | Istore (a,b,c,d,n) -> Istore (a,b,c,d,s)
- | Icall (a,b,c,d,n) -> Icall (a,b,c,d,s)
- | Ibuiltin (a,b,c,n) -> Ibuiltin (a,b,c,s)
- | Icond (a,b,n1,n2,p) -> (
- match p with
- | Some true -> Icond (a,b,s,n2,p)
- | Some false -> Icond (a,b,n1,s,p)
- | None -> failwith "Predicted a successor for an Icond with p=None - unpredicted CB should not be moved in the middle of the superblock"
- )
-
-let rec change_successors i = function
- | [] -> (
- match i with
- | Itailcall _ | Ireturn _ -> i
- | _ -> failwith "Wrong instruction (1)")
- | [s] -> (
- match i with
- | Inop n -> Inop s
- | Iop (a,b,c,n) -> Iop (a,b,c,s)
- | Iload (a,b,c,d,e,n) -> Iload (a,b,c,d,e,s)
- | Istore (a,b,c,d,n) -> Istore (a,b,c,d,s)
- | Icall (a,b,c,d,n) -> Icall (a,b,c,d,s)
- | Ibuiltin (a,b,c,n) -> Ibuiltin (a,b,c,s)
- | Ijumptable (a,[n]) -> Ijumptable (a,[s])
- | Icond (a,b,n1,n2,p) -> (
- match p with
- | Some true -> Icond (a,b,s,n2,p)
- | Some false -> Icond (a,b,n1,s,p)
- | None -> failwith "Icond Wrong instruction (2) ; should not happen?"
- )
- | _ -> failwith "Wrong instruction (2)")
- | [s1; s2] -> (
- match i with
- | Icond (a,b,n1,n2,p) -> Icond (a,b,s1,s2,p)
- | Ijumptable (a, [n1; n2]) -> Ijumptable (a, [s1; s2])
- | _ -> change_successors i [s1])
- | ls -> (
- match i with
- | Ijumptable (a, ln) -> begin
- assert ((List.length ln) == (List.length ls));
- Ijumptable (a, ls)
- end
- | _ -> failwith "Wrong instruction (4)")
-
-
-let apply_schedule code sb new_order =
- let tc = ref code in
- let old_order = sb.instructions in
- let last_node = Array.get old_order (Array.length old_order - 1) in
- let last_successors = main_successors
- @@ get_some @@ PTree.get last_node code in
- begin
- check_order code old_order new_order;
- Array.iteri (fun i n' ->
- let inst' = get_some @@ PTree.get n' code in
- let new_inst =
- if (i == (Array.length old_order - 1)) then
- change_successors inst' last_successors
- else
- change_predicted_successor inst' (Array.get old_order (i+1))
- in tc := PTree.set (Array.get old_order i) new_inst !tc
- ) new_order;
- !tc
- end
-*)
-
-let rec do_schedule code = function
- | [] -> code
- | sb :: lsb ->
- let schedule = schedule_superblock sb code in
- let new_code = apply_schedule code sb schedule in
- begin
- (* debug_flag := true; *)
- dprintf "Old Code: "; print_code code;
- dprintf "\nSchedule to apply: "; print_arrayp schedule;
- dprintf "\nNew Code: "; print_code new_code;
- dprintf "\n";
- (* debug_flag := false; *)
- do_schedule new_code lsb
- end
-
-let get_ok r = match r with Errors.OK x -> x | _ -> failwith "Did not get OK"
-
-let scheduler f =
- let code = f.fn_RTL.fn_code in
- let id_ptree = PTree.map (fun n i -> n) (f.fn_path) in
- let entry = f.fn_RTL.fn_entrypoint in
- let pm = f.fn_path in
- let typing = get_ok @@ RTLtyping.type_function f.fn_RTL in
- let lsb = get_superblocks code entry pm typing in
- begin
- (* debug_flag := true; *)
- dprintf "Pathmap:\n"; dprintf "\n";
- print_path_map pm;
- dprintf "Superblocks:\n";
- print_superblocks lsb code; dprintf "\n";
- (* debug_flag := false; *)
- let tc = do_schedule code lsb in
- (((tc, entry), pm), id_ptree)
- end
diff --git a/kvx/lib/RTLpathSchedulerproof.v b/kvx/lib/RTLpathSchedulerproof.v
deleted file mode 100644
index 4ba197b0..00000000
--- a/kvx/lib/RTLpathSchedulerproof.v
+++ /dev/null
@@ -1,363 +0,0 @@
-Require Import AST Linking Values Maps Globalenvs Smallstep Registers.
-Require Import Coqlib Maps Events Errors Op.
-Require Import RTL RTLpath RTLpathLivegen RTLpathLivegenproof RTLpathSE_theory.
-Require Import RTLpathScheduler.
-
-Definition match_prog (p tp: program) :=
- match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
-
-Lemma transf_program_match:
- forall prog tprog, transf_program prog = OK tprog -> match_prog prog tprog.
-Proof.
- intros. eapply match_transform_partial_program_contextual; eauto.
-Qed.
-
-Section PRESERVATION.
-
-Variable prog: program.
-Variable tprog: program.
-
-Hypothesis TRANSL: match_prog prog tprog.
-
-Let pge := Genv.globalenv prog.
-Let tpge := Genv.globalenv tprog.
-
-Hypothesis all_fundef_liveness_ok: forall b fd, Genv.find_funct_ptr pge b = Some fd -> liveness_ok_fundef fd.
-
-Lemma symbols_preserved s: Genv.find_symbol tpge s = Genv.find_symbol pge s.
-Proof.
- rewrite <- (Genv.find_symbol_match TRANSL). reflexivity.
-Qed.
-
-Lemma senv_preserved:
- Senv.equiv pge tpge.
-Proof.
- eapply (Genv.senv_match TRANSL).
-Qed.
-
-Lemma functions_preserved:
- forall (v: val) (f: fundef),
- Genv.find_funct pge v = Some f ->
- exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tpge v = Some tf /\ linkorder cunit prog.
-Proof.
- intros. exploit (Genv.find_funct_match TRANSL); eauto.
- intros (cu & tf & A & B & C).
- repeat eexists; intuition eauto.
- + unfold incl; auto.
- + eapply linkorder_refl.
-Qed.
-
-Lemma function_ptr_preserved:
- forall v f,
- Genv.find_funct_ptr pge v = Some f ->
- exists tf,
- Genv.find_funct_ptr tpge v = Some tf /\ transf_fundef f = OK tf.
-Proof.
- intros.
- exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto.
-Qed.
-
-Lemma function_sig_preserved:
- forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f.
-Proof.
- intros. destruct f.
- - simpl in H. monadInv H.
- destruct (transf_function f) as [res H]; simpl in * |- *; auto.
- destruct (H _ EQ).
- intuition subst; auto.
- symmetry.
- eapply match_function_preserves.
- eassumption.
- - simpl in H. monadInv H. reflexivity.
-Qed.
-
-Theorem transf_initial_states:
- forall s1, initial_state prog s1 ->
- exists s2, initial_state tprog s2 /\ match_states s1 s2.
-Proof.
- intros. inv H.
- exploit function_ptr_preserved; eauto. intros (tf & FIND & TRANSF).
- exists (Callstate nil tf nil m0).
- split.
- - econstructor; eauto.
- + intros; apply (Genv.init_mem_match TRANSL); assumption.
- + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto.
- symmetry. eapply match_program_main. eauto.
- + destruct f.
- * monadInv TRANSF. rewrite <- H3.
- destruct (transf_function f) as [res H]; simpl in * |- *; auto.
- destruct (H _ EQ).
- intuition subst; auto.
- symmetry; eapply match_function_preserves. eassumption.
- * monadInv TRANSF. assumption.
- - constructor; eauto.
- + constructor.
- + apply transf_fundef_correct; auto.
-(* + eapply all_fundef_liveness_ok; eauto. *)
-Qed.
-
-Theorem transf_final_states s1 s2 r:
- final_state s1 r -> match_states s1 s2 -> final_state s2 r.
-Proof.
- unfold final_state.
- intros H; inv H.
- intros H; inv H; simpl in * |- *; try congruence.
- inv H1.
- destruct st; simpl in * |- *; try congruence.
- inv STACKS. constructor.
-Qed.
-
-
-Let ge := Genv.globalenv (RTLpath.transf_program prog).
-Let tge := Genv.globalenv (RTLpath.transf_program tprog).
-
-Lemma senv_sym x y: Senv.equiv x y -> Senv.equiv y x.
-Proof.
- unfold Senv.equiv. intuition congruence.
-Qed.
-
-Lemma senv_transitivity x y z: Senv.equiv x y -> Senv.equiv y z -> Senv.equiv x z.
-Proof.
- unfold Senv.equiv. intuition congruence.
-Qed.
-
-Lemma senv_preserved_RTL:
- Senv.equiv ge tge.
-Proof.
- eapply senv_transitivity. { eapply senv_sym; eapply RTLpath.senv_preserved. }
- eapply senv_transitivity. { eapply senv_preserved. }
- eapply RTLpath.senv_preserved.
-Qed.
-
-Lemma symbols_preserved_RTL s: Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof.
- unfold tge, ge. erewrite RTLpath.symbols_preserved; eauto.
- rewrite symbols_preserved.
- erewrite RTLpath.symbols_preserved; eauto.
-Qed.
-
-Program Definition mkctx sp rs0 m0 {f1: RTLpath.function} (hyp: liveness_ok_function f1)
- : simu_proof_context f1
- := {| the_ge1:= ge; the_ge2 := tge; the_sp:=sp; the_rs0:=rs0; the_m0:=m0 |}.
-Obligation 2.
- erewrite symbols_preserved_RTL. eauto.
-Qed.
-
-Lemma s_find_function_preserved f sp svos1 svos2 rs0 m0 fd
- (LIVE: liveness_ok_function f):
- (svident_simu f (mkctx sp rs0 m0 LIVE) svos1 svos2) ->
- sfind_function pge ge sp svos1 rs0 m0 = Some fd ->
- exists fd', sfind_function tpge tge sp svos2 rs0 m0 = Some fd'
- /\ transf_fundef fd = OK fd'
- /\ liveness_ok_fundef fd.
-Proof.
- Local Hint Resolve symbols_preserved_RTL: core.
- unfold sfind_function. intros [sv1 sv2 SIMU|]; simpl in *.
- + rewrite !(seval_preserved ge tge) in *; eauto.
- destruct (seval_sval _ _ _ _); try congruence.
- erewrite <- SIMU; try congruence. clear SIMU.
- intros; exploit functions_preserved; eauto.
- intros (fd' & cunit & (X1 & X2 & X3)). eexists.
- repeat split; eauto.
- eapply find_funct_liveness_ok; eauto.
-(* intros. eapply all_fundef_liveness_ok; eauto. *)
- + subst. rewrite symbols_preserved. destruct (Genv.find_symbol _ _); try congruence.
- intros; exploit function_ptr_preserved; eauto.
- intros (fd' & X). eexists. intuition eauto.
-(* intros. eapply all_fundef_liveness_ok; eauto. *)
-Qed.
-
-Lemma sistate_simu f dupmap sp st st' rs m is
- (LIVE: liveness_ok_function f):
- ssem_internal ge sp st rs m is ->
- sistate_simu dupmap f st st' (mkctx sp rs m LIVE)->
- exists is',
- ssem_internal tge sp st' rs m is' /\ istate_simu f dupmap is is'.
-Proof.
- intros SEM X; eapply X; eauto.
-Qed.
-
-Lemma seval_builtin_sval_preserved sp rs m:
- forall bs, seval_builtin_sval ge sp bs rs m = seval_builtin_sval tge sp bs rs m.
-Proof.
- induction bs.
- all: try (simpl; try reflexivity; erewrite seval_preserved by eapply symbols_preserved_RTL; reflexivity).
- all: simpl; rewrite IHbs1; rewrite IHbs2; reflexivity.
-Qed.
-
-Lemma seval_list_builtin_sval_preserved sp rs m:
- forall lbs,
- seval_list_builtin_sval ge sp lbs rs m = seval_list_builtin_sval tge sp lbs rs m.
-Proof.
- induction lbs; [simpl; reflexivity|].
- simpl. rewrite seval_builtin_sval_preserved. rewrite IHlbs.
- reflexivity.
-Qed.
-
-Lemma ssem_final_simu dm f f' stk stk' sp st st' rs0 m0 sv sv' rs m t s
- (LIVE: liveness_ok_function f):
- match_function dm f f' ->
- list_forall2 match_stackframes stk stk' ->
- (* s_istate_simu f dupmap st st' -> *)
- sfval_simu dm f st.(si_pc) st'.(si_pc) (mkctx sp rs0 m0 LIVE) sv sv' ->
- ssem_final pge ge sp st.(si_pc) stk f rs0 m0 sv rs m t s ->
- exists s', ssem_final tpge tge sp st'.(si_pc) stk' f' rs0 m0 sv' rs m t s' /\ match_states s s'.
-Proof.
- Local Hint Resolve transf_fundef_correct: core.
- intros FUN STK (* SIS *) SFV. destruct SFV; intros SEM; inv SEM; simpl in *.
- - (* Snone *)
- exploit initialize_path. { eapply dupmap_path_entry1; eauto. }
- intros (path & PATH).
- eexists; split; econstructor; eauto.
- eapply eqlive_reg_refl.
- - (* Scall *)
- exploit s_find_function_preserved; eauto.
- intros (fd' & FIND & TRANSF & LIVE').
- erewrite <- function_sig_preserved; eauto.
- exploit initialize_path. { eapply dupmap_path_entry1; eauto. }
- intros (path & PATH).
- eexists; split; econstructor; eauto.
- + eapply eq_trans; try eassumption; auto.
- + simpl. repeat (econstructor; eauto).
- - (* Stailcall *)
- exploit s_find_function_preserved; eauto.
- intros (fd' & FIND & TRANSF & LIVE').
- erewrite <- function_sig_preserved; eauto.
- eexists; split; econstructor; eauto.
- + erewrite <- preserv_fnstacksize; eauto.
- + eapply eq_trans; try eassumption; auto.
- - (* Sbuiltin *)
- pose senv_preserved_RTL as SRTL.
- exploit initialize_path. { eapply dupmap_path_entry1; eauto. }
- intros (path & PATH).
- eexists; split; econstructor; eauto.
- + eapply seval_builtin_args_preserved; eauto.
- eapply seval_list_builtin_sval_correct; eauto.
- rewrite H0.
- erewrite seval_list_builtin_sval_preserved; eauto.
- + eapply external_call_symbols_preserved; eauto.
- + eapply eqlive_reg_refl.
- - (* Sjumptable *)
- exploit ptree_get_list_nth_rev; eauto. intros (p2 & LNZ & DM).
- exploit initialize_path. { eapply dupmap_path_entry1; eauto. }
- intros (path & PATH).
- eexists; split; econstructor; eauto.
- + eapply eq_trans; try eassumption; auto.
- + eapply eqlive_reg_refl.
- - (* Sreturn *)
- eexists; split; econstructor; eauto.
- erewrite <- preserv_fnstacksize; eauto.
- - (* Sreturn bis *)
- eexists; split; econstructor; eauto.
- + erewrite <- preserv_fnstacksize; eauto.
- + rewrite <- H. erewrite <- seval_preserved; eauto.
-Qed.
-
-(* The main theorem on simulation of symbolic states ! *)
-Theorem ssem_sstate_simu dm f f' stk stk' sp st st' rs m t s:
- match_function dm f f' ->
- liveness_ok_function f ->
- list_forall2 match_stackframes stk stk' ->
- ssem pge ge sp st stk f rs m t s ->
- (forall ctx: simu_proof_context f, sstate_simu dm f st st' ctx) ->
- exists s', ssem tpge tge sp st' stk' f' rs m t s' /\ match_states s s'.
-Proof.
- intros MFUNC LIVE STACKS SEM SIMU.
- destruct (SIMU (mkctx sp rs m LIVE)) as (SIMU1 & SIMU2); clear SIMU.
- destruct SEM as [is CONT SEM|is t s' CONT SEM1 SEM2]; simpl.
- - (* sem_early *)
- exploit sistate_simu; eauto.
- unfold istate_simu; rewrite CONT.
- intros (is' & SEM' & (path & PATH & (CONT' & RS' & M') & PC')).
- exists (State stk' f' sp (ipc is') (irs is') (imem is')).
- split.
- + eapply ssem_early; auto. congruence.
- + rewrite M'. econstructor; eauto.
- - (* sem_normal *)
- exploit sistate_simu; eauto.
- unfold istate_simu; rewrite CONT.
- intros (is' & SEM' & (CONT' & RS' & M')(* & DMEQ *)).
- rewrite <- eqlive_reg_triv in RS'.
- exploit ssem_final_simu; eauto.
- clear SEM2; intros (s0 & SEM2 & MATCH0).
- exploit ssem_final_equiv; eauto.
- clear SEM2; rewrite M'; rewrite CONT' in CONT; intros (s1 & EQ & SEM2).
- exists s1; split.
- + eapply ssem_normal; eauto.
- + eapply match_states_equiv; eauto.
-Qed.
-
-Lemma exec_path_simulation dupmap path stk stk' f f' sp rs m pc pc' t s:
- (fn_path f)!pc = Some path ->
- path_step ge pge path.(psize) stk f sp rs m pc t s ->
- list_forall2 match_stackframes stk stk' ->
- dupmap ! pc' = Some pc ->
- match_function dupmap f f' ->
- liveness_ok_function f ->
- exists path' s', (fn_path f')!pc' = Some path' /\ path_step tge tpge path'.(psize) stk' f' sp rs m pc' t s' /\ match_states s s'.
-Proof.
- intros PATH STEP STACKS DUPPC MATCHF LIVE.
- exploit initialize_path. { eapply dupmap_path_entry2; eauto. }
- intros (path' & PATH').
- exists path'.
- exploit (sexec_correct f pc pge ge sp path stk rs m t s); eauto.
- intros (st & SYMB & SEM); clear STEP.
- exploit dupmap_correct; eauto.
- clear SYMB; intros (st' & SYMB & SIMU).
- exploit ssem_sstate_simu; eauto.
- intros (s0 & SEM0 & MATCH).
- exploit sexec_exact; eauto.
- intros (s' & STEP' & EQ).
- exists s'; intuition.
- eapply match_states_equiv; eauto.
-Qed.
-
-Lemma step_simulation s1 t s1' s2:
- step ge pge s1 t s1' ->
- match_states s1 s2 ->
- exists s2',
- step tge tpge s2 t s2'
- /\ match_states s1' s2'.
-Proof.
- Local Hint Resolve eqlive_stacks_refl transf_fundef_correct: core.
- destruct 1 as [path stack f sp rs m pc t s PATH STEP | | | ]; intros MS; inv MS.
-(* exec_path *)
- - try_simplify_someHyps. intros.
- exploit path_step_eqlive; eauto. (* { intros. eapply all_fundef_liveness_ok; eauto. } *)
- clear STEP EQUIV rs; intros (s2 & STEP & EQLIVE).
- exploit exec_path_simulation; eauto.
- clear STEP; intros (path' & s' & PATH' & STEP' & MATCH').
- exists s'; split.
- + eapply exec_path; eauto.
- + eapply eqlive_match_states; eauto.
-(* exec_function_internal *)
- - inv LIVE.
- exploit initialize_path. { eapply (fn_entry_point_wf f). }
- destruct 1 as (path & PATH).
- inversion TRANSF as [f0 xf tf MATCHF|]; subst. eexists. split.
- + eapply exec_function_internal. erewrite <- preserv_fnstacksize; eauto.
- + erewrite preserv_fnparams; eauto.
- econstructor; eauto.
- { apply preserv_entrypoint; auto. }
- { apply eqlive_reg_refl. }
-(* exec_function_external *)
- - inversion TRANSF as [|]; subst. eexists. split.
- + econstructor. eapply external_call_symbols_preserved; eauto. apply senv_preserved_RTL.
- + constructor. assumption.
-(* exec_return *)
- - inv STACKS. destruct b1 as [res' f' sp' pc' rs']. eexists. split.
- + constructor.
- + inv H1. econstructor; eauto.
-Qed.
-
-Theorem transf_program_correct:
- forward_simulation (semantics prog) (semantics tprog).
-Proof.
- eapply forward_simulation_step with match_states.
- - eapply senv_preserved.
- - eapply transf_initial_states.
- - intros; eapply transf_final_states; eauto.
- - intros; eapply step_simulation; eauto.
-Qed.
-
-End PRESERVATION.
diff --git a/kvx/lib/RTLpathproof.v b/kvx/lib/RTLpathproof.v
deleted file mode 100644
index 20eded97..00000000
--- a/kvx/lib/RTLpathproof.v
+++ /dev/null
@@ -1,50 +0,0 @@
-Require Import Coqlib Maps.
-Require Import AST Integers Values Events Memory Globalenvs Smallstep.
-Require Import Op Registers.
-Require Import RTL Linking.
-Require Import RTLpath.
-
-Definition match_prog (p: RTLpath.program) (tp: RTL.program) :=
- match_program (fun ctx f tf => tf = fundef_RTL f) eq p tp.
-
-Lemma transf_program_match:
- forall p, match_prog p (transf_program p).
-Proof.
- intros. eapply match_transform_program; eauto.
-Qed.
-
-Lemma match_program_transf:
- forall p tp, match_prog p tp -> transf_program p = tp.
-Proof.
- intros p tp H. inversion_clear H. inv H1.
- destruct p as [defs pub main]. destruct tp as [tdefs tpub tmain]. simpl in *.
- subst. unfold transf_program. unfold transform_program. simpl.
- apply program_equals; simpl; auto.
- induction H0; simpl; auto.
- rewrite IHlist_forall2. apply cons_extract.
- destruct a1 as [ida gda]. destruct b1 as [idb gdb].
- simpl in *.
- inv H. inv H2.
- - simpl in *. subst. auto.
- - simpl in *. subst. inv H. auto.
-Qed.
-
-
-Section PRESERVATION.
-
-Variable prog: RTLpath.program.
-Variable tprog: RTL.program.
-Hypothesis TRANSF: match_prog prog tprog.
-Let ge := Genv.globalenv prog.
-Let tge := Genv.globalenv tprog.
-
-Theorem transf_program_correct:
- forward_simulation (RTLpath.semantics prog) (RTL.semantics tprog).
-Proof.
- pose proof (match_program_transf prog tprog TRANSF) as TR. subst.
- eapply RTLpath_correct.
-Qed.
-
-End PRESERVATION.
-
-