aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorSylvain Boulmé <sylvain.boulme@univ-grenoble-alpes.fr>2019-10-16 14:56:21 +0200
committerSylvain Boulmé <sylvain.boulme@univ-grenoble-alpes.fr>2019-10-16 14:56:21 +0200
commit7bbdda944e377af828af8727f6db910ff99c701f (patch)
tree8789c2d7b96dcaa4bd17fb825ad37dc0981a439d
parentb957e99a444fabffbdc3513e4a4f6cfc035c24b1 (diff)
parent72378d9371bc5da342266bcf14231ab568e0f919 (diff)
downloadcompcert-kvx-7bbdda944e377af828af8727f6db910ff99c701f.tar.gz
compcert-kvx-7bbdda944e377af828af8727f6db910ff99c701f.zip
Merge branch 'mppa-work' into RTLpath
strange conflict on cparser/Machine.ml (??)
-rw-r--r--Makefile1
-rw-r--r--backend/Duplicate.v240
-rw-r--r--backend/Duplicateaux.ml13
-rw-r--r--backend/Duplicateproof.v467
-rw-r--r--common/AST.v19
-rwxr-xr-xconfigure4
-rw-r--r--cparser/Machine.ml9
-rw-r--r--cparser/Machine.mli1
-rw-r--r--cparser/StructPassing.ml11
-rw-r--r--driver/Compiler.v64
-rw-r--r--lib/Floats.v63
-rw-r--r--lib/Integers.v34
-rw-r--r--mppa_k1c/Archi.v1
-rw-r--r--mppa_k1c/Asm.v1506
-rw-r--r--mppa_k1c/Asmaux.v2
-rw-r--r--mppa_k1c/Asmblockgen.v4
-rw-r--r--mppa_k1c/Asmblockgenproof.v3549
-rw-r--r--mppa_k1c/Asmexpand.ml4
-rw-r--r--mppa_k1c/Op.v6
-rw-r--r--test/c/Makefile16
-rw-r--r--test/c/Results/binarytrees11
-rw-r--r--test/c/Results/binarytrees-mppa_k1c4
-rw-r--r--test/c/Results/chomp6
-rw-r--r--test/c/Results/chomp-mppa_k1c9
-rw-r--r--test/c/Results/fannkuch62
-rw-r--r--test/c/Results/fannkuch-mppa_k1c31
-rw-r--r--test/c/Results/fft2
-rw-r--r--test/c/Results/fft-mppa_k1c1
-rw-r--r--test/c/Results/fftsp2
-rw-r--r--test/c/Results/fftsp-mppa_k1c1
-rw-r--r--test/c/Results/fftw-mppa_k1c16
-rw-r--r--test/c/Results/fib2
-rw-r--r--test/c/Results/fib-mppa_k1c1
-rw-r--r--test/c/Results/integr2
-rw-r--r--test/c/Results/integr-mppa_k1c1
-rw-r--r--test/c/Results/knucleotide27
-rw-r--r--test/c/Results/knucleotide-mppa_k1c0
-rw-r--r--test/c/Results/lists-mppa_k1c2
-rw-r--r--test/c/Results/mandelbrotbin709 -> 125013 bytes
-rw-r--r--test/c/Results/mandelbrot-mppa_k1cbin0 -> 709 bytes
-rw-r--r--test/c/Results/nbody2
-rw-r--r--test/c/Results/nbody-mppa_k1c2
-rw-r--r--test/c/Results/nsieve6
-rw-r--r--test/c/Results/nsieve-mppa_k1c3
-rw-r--r--test/c/Results/nsievebits6
-rw-r--r--test/c/Results/nsievebits-mppa_k1c3
-rw-r--r--test/c/Results/perlin2
-rw-r--r--test/c/Results/perlin-mppa_k1c1
-rw-r--r--test/c/Results/qsort-mppa_k1c1
-rw-r--r--test/c/Results/sha1-mppa_k1c2
-rw-r--r--test/c/Results/spectral2
-rw-r--r--test/c/Results/spectral-mppa_k1c1
-rw-r--r--test/c/Results/vmach4
-rw-r--r--test/c/Results/vmach-mppa_k1c2
-rwxr-xr-xtest/c/Runtest71
-rw-r--r--test/c/binarytrees.c4
-rw-r--r--test/c/chomp.c5
-rw-r--r--test/c/fannkuch.c4
-rw-r--r--test/c/fft.c4
-rw-r--r--test/c/fftsp.c4
-rw-r--r--test/c/fftw.c4
-rw-r--r--test/c/fib.c4
-rw-r--r--test/c/integr.c4
-rw-r--r--test/c/lists.c5
-rw-r--r--test/c/mandelbrot.c8
-rw-r--r--test/c/nbody.c4
-rw-r--r--test/c/nsieve.c8
-rw-r--r--test/c/nsievebits.c8
-rw-r--r--test/c/perlin.c6
-rw-r--r--test/c/qsort.c4
-rw-r--r--test/c/spectral.c4
-rw-r--r--test/c/vmach.c10
-rw-r--r--test/monniaux/README.md9
-rw-r--r--test/monniaux/benches.sh2
-rw-r--r--test/monniaux/genann/Makefile4
-rw-r--r--test/monniaux/genann/make.proto2
-rw-r--r--test/monniaux/glpk-4.65/Makefile39
-rw-r--r--test/monniaux/picosat-965/Makefile40
-rw-r--r--test/monniaux/rules.mk5
-rw-r--r--test/regression/Makefile7
-rw-r--r--test/regression/Results/varargs2-mppa_k1c11
-rw-r--r--test/regression/builtins-mppa_k1c.c72
-rw-r--r--test/regression/varargs2.c16
83 files changed, 3867 insertions, 2732 deletions
diff --git a/Makefile b/Makefile
index aa99786b..299f5ffe 100644
--- a/Makefile
+++ b/Makefile
@@ -80,6 +80,7 @@ BACKEND=\
Tailcall.v Tailcallproof.v \
Inlining.v Inliningspec.v Inliningproof.v \
Renumber.v Renumberproof.v \
+ Duplicate.v Duplicateproof.v \
RTLtyping.v \
Kildall.v Liveness.v \
ValueDomain.v ValueAOp.v ValueAnalysis.v \
diff --git a/backend/Duplicate.v b/backend/Duplicate.v
new file mode 100644
index 00000000..d1458bd4
--- /dev/null
+++ b/backend/Duplicate.v
@@ -0,0 +1,240 @@
+(** RTL node duplication using external oracle. Used to form superblock
+ structures *)
+
+Require Import AST RTL Maps Globalenvs.
+Require Import Coqlib Errors Op.
+
+Local Open Scope error_monad_scope.
+Local Open Scope positive_scope.
+
+(** External oracle returning the new RTL code (entry point unchanged),
+ along with the new entrypoint, and a mapping of new nodes to old nodes *)
+Axiom duplicate_aux: function -> code * node * (PTree.t node).
+
+Extract Constant duplicate_aux => "Duplicateaux.duplicate_aux".
+
+Record xfunction : Type :=
+ { fn_RTL: function;
+ fn_revmap: PTree.t node;
+ }.
+
+Definition xfundef := AST.fundef xfunction.
+Definition xprogram := AST.program xfundef unit.
+Definition xgenv := Genv.t xfundef unit.
+
+Definition fundef_RTL (fu: xfundef) : fundef :=
+ match fu with
+ | Internal f => Internal (fn_RTL f)
+ | External ef => External ef
+ end.
+
+(** * Verification of node duplications *)
+
+Definition verify_mapping_entrypoint (f: function) (xf: xfunction) : res unit :=
+ match ((fn_revmap xf)!(fn_entrypoint (fn_RTL xf))) with
+ | None => Error (msg "verify_mapping: No node in xf revmap for entrypoint")
+ | Some n => match (Pos.compare n (fn_entrypoint f)) with
+ | Eq => OK tt
+ | _ => Error (msg "verify_mapping_entrypoint: xf revmap for entrypoint does not correspond to the entrypoint of f")
+ end
+ end.
+
+Definition verify_is_copy revmap n n' :=
+ match revmap!n' with
+ | None => Error(msg "verify_is_copy None")
+ | Some revn => match (Pos.compare n revn) with Eq => OK tt | _ => Error(msg "verify_is_copy invalid map") end
+ end.
+
+Fixpoint verify_is_copy_list revmap ln ln' :=
+ match ln with
+ | n::ln => match ln' with
+ | n'::ln' => do u <- verify_is_copy revmap n n';
+ verify_is_copy_list revmap ln ln'
+ | nil => Error (msg "verify_is_copy_list: ln' bigger than ln") end
+ | nil => match ln' with
+ | n :: ln' => Error (msg "verify_is_copy_list: ln bigger than ln'")
+ | nil => OK tt end
+ end.
+
+Lemma product_eq {A B: Type} :
+ (forall (a b: A), {a=b} + {a<>b}) ->
+ (forall (c d: B), {c=d} + {c<>d}) ->
+ forall (x y: A+B), {x=y} + {x<>y}.
+Proof.
+ intros H H'. intros. decide equality.
+Qed.
+
+(** FIXME Ideally i would like to put this in AST.v but i get an "illegal application"
+ * error when doing so *)
+Remark builtin_arg_eq_pos: forall (a b: builtin_arg positive), {a=b} + {a<>b}.
+Proof.
+ intros.
+ apply (builtin_arg_eq Pos.eq_dec).
+Defined.
+Global Opaque builtin_arg_eq_pos.
+
+Remark builtin_res_eq_pos: forall (a b: builtin_res positive), {a=b} + {a<>b}.
+Proof. intros. apply (builtin_res_eq Pos.eq_dec). Qed.
+Global Opaque builtin_res_eq_pos.
+
+Definition verify_match_inst revmap inst tinst :=
+ match inst with
+ | Inop n => match tinst with Inop n' => do u <- verify_is_copy revmap n n'; OK tt | _ => Error(msg "verify_match_inst Inop") end
+
+ | Iop op lr r n => match tinst with
+ Iop op' lr' r' n' =>
+ do u <- verify_is_copy revmap n n';
+ if (eq_operation op op') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then
+ OK tt
+ else Error (msg "Different r in Iop")
+ else Error (msg "Different lr in Iop")
+ else Error(msg "Different operations in Iop")
+ | _ => Error(msg "verify_match_inst Inop") end
+
+ | Iload m a lr r n => match tinst with
+ | Iload m' a' lr' r' n' =>
+ do u <- verify_is_copy revmap n n';
+ if (chunk_eq m m') then
+ if (eq_addressing a a') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then OK tt
+ else Error (msg "Different r in Iload")
+ else Error (msg "Different lr in Iload")
+ else Error (msg "Different addressing in Iload")
+ else Error (msg "Different mchunk in Iload")
+ | _ => Error (msg "verify_match_inst Iload") end
+
+ | Istore m a lr r n => match tinst with
+ | Istore m' a' lr' r' n' =>
+ do u <- verify_is_copy revmap n n';
+ if (chunk_eq m m') then
+ if (eq_addressing a a') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then OK tt
+ else Error (msg "Different r in Istore")
+ else Error (msg "Different lr in Istore")
+ else Error (msg "Different addressing in Istore")
+ else Error (msg "Different mchunk in Istore")
+ | _ => Error (msg "verify_match_inst Istore") end
+
+ | Icall s ri lr r n => match tinst with
+ | Icall s' ri' lr' r' n' =>
+ do u <- verify_is_copy revmap n n';
+ if (signature_eq s s') then
+ if (product_eq Pos.eq_dec ident_eq ri ri') then
+ if (list_eq_dec Pos.eq_dec lr lr') then
+ if (Pos.eq_dec r r') then OK tt
+ else Error (msg "Different r r' in Icall")
+ else Error (msg "Different lr in Icall")
+ else Error (msg "Different ri in Icall")
+ else Error (msg "Different signatures in Icall")
+ | _ => Error (msg "verify_match_inst Icall") end
+
+ | Itailcall s ri lr => match tinst with
+ | Itailcall s' ri' lr' =>
+ if (signature_eq s s') then
+ if (product_eq Pos.eq_dec ident_eq ri ri') then
+ if (list_eq_dec Pos.eq_dec lr lr') then OK tt
+ else Error (msg "Different lr in Itailcall")
+ else Error (msg "Different ri in Itailcall")
+ else Error (msg "Different signatures in Itailcall")
+ | _ => Error (msg "verify_match_inst Itailcall") end
+
+ | Ibuiltin ef lbar brr n => match tinst with
+ | Ibuiltin ef' lbar' brr' n' =>
+ do u <- verify_is_copy revmap n n';
+ if (external_function_eq ef ef') then
+ if (list_eq_dec builtin_arg_eq_pos lbar lbar') then
+ if (builtin_res_eq_pos brr brr') then OK tt
+ else Error (msg "Different brr in Ibuiltin")
+ else Error (msg "Different lbar in Ibuiltin")
+ else Error (msg "Different ef in Ibuiltin")
+ | _ => Error (msg "verify_match_inst Ibuiltin") end
+
+ | Icond cond lr n1 n2 => match tinst with
+ | Icond cond' lr' n1' n2' =>
+ do u1 <- verify_is_copy revmap n1 n1';
+ do u2 <- verify_is_copy revmap n2 n2';
+ if (condition_eq cond cond') then
+ if (list_eq_dec Pos.eq_dec lr lr') then OK tt
+ else Error (msg "Different lr in Icond")
+ else Error (msg "Different cond in Icond")
+ | _ => Error (msg "verify_match_inst Icond") end
+
+ | Ijumptable r ln => match tinst with
+ | Ijumptable r' ln' =>
+ do u <- verify_is_copy_list revmap ln ln';
+ if (Pos.eq_dec r r') then OK tt
+ else Error (msg "Different r in Ijumptable")
+ | _ => Error (msg "verify_match_inst Ijumptable") end
+
+ | Ireturn or => match tinst with
+ | Ireturn or' =>
+ if (option_eq Pos.eq_dec or or') then OK tt
+ else Error (msg "Different or in Ireturn")
+ | _ => Error (msg "verify_match_inst Ireturn") end
+ end.
+
+Definition verify_mapping_mn f xf (m: positive*positive) :=
+ let (tn, n) := m in
+ match (fn_code f)!n with
+ | None => Error (msg "verify_mapping_mn: Could not get an instruction at (fn_code f)!n")
+ | Some inst => match (fn_code (fn_RTL xf))!tn with
+ | None => Error (msg "verify_mapping_mn: Could not get an instruction at (fn_code xf)!tn")
+ | Some tinst => verify_match_inst (fn_revmap xf) inst tinst
+ end
+ end.
+
+Fixpoint verify_mapping_mn_rec f xf lm :=
+ match lm with
+ | nil => OK tt
+ | m :: lm => do u <- verify_mapping_mn f xf m;
+ do u2 <- verify_mapping_mn_rec f xf lm;
+ OK tt
+ end.
+
+Definition verify_mapping_match_nodes (f: function) (xf: xfunction) : res unit :=
+ verify_mapping_mn_rec f xf (PTree.elements (fn_revmap xf)).
+
+(** Verifies that the [fn_revmap] of the translated function [xf] is giving correct information in regards to [f] *)
+Definition verify_mapping (f: function) (xf: xfunction) : res unit :=
+ do u <- verify_mapping_entrypoint f xf;
+ do v <- verify_mapping_match_nodes f xf; OK tt.
+(* TODO - verify the other axiom *)
+
+(** * Entry points *)
+
+Definition transf_function_aux (f: function) : res xfunction :=
+ let (tcte, mp) := duplicate_aux f in
+ let (tc, te) := tcte in
+ let xf := {| fn_RTL := (mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) tc te); fn_revmap := mp |} in
+ do u <- verify_mapping f xf;
+ OK xf.
+
+Theorem transf_function_aux_preserves:
+ forall f xf,
+ transf_function_aux f = OK xf ->
+ fn_sig f = fn_sig (fn_RTL xf) /\ fn_params f = fn_params (fn_RTL xf) /\ fn_stacksize f = fn_stacksize (fn_RTL xf).
+Proof.
+ intros. unfold transf_function_aux in H. destruct (duplicate_aux _) as (tcte & mp). destruct tcte as (tc & te). monadInv H.
+ repeat (split; try reflexivity).
+Qed.
+
+Remark transf_function_aux_fnsig: forall f xf, transf_function_aux f = OK xf -> fn_sig f = fn_sig (fn_RTL xf).
+ Proof. apply transf_function_aux_preserves. Qed.
+Remark transf_function_aux_fnparams: forall f xf, transf_function_aux f = OK xf -> fn_params f = fn_params (fn_RTL xf).
+ Proof. apply transf_function_aux_preserves. Qed.
+Remark transf_function_aux_fnstacksize: forall f xf, transf_function_aux f = OK xf -> fn_stacksize f = fn_stacksize (fn_RTL xf).
+ Proof. apply transf_function_aux_preserves. Qed.
+
+Definition transf_function (f: function) : res function :=
+ do xf <- transf_function_aux f;
+ OK (fn_RTL xf).
+
+Definition transf_fundef (f: fundef) : res fundef :=
+ transf_partial_fundef transf_function f.
+
+Definition transf_program (p: program) : res program :=
+ transform_partial_program transf_fundef p. \ No newline at end of file
diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml
new file mode 100644
index 00000000..9ff2ae55
--- /dev/null
+++ b/backend/Duplicateaux.ml
@@ -0,0 +1,13 @@
+open RTL
+open Maps
+
+let rec make_identity_ptree_rec = function
+| [] -> PTree.empty
+| m::lm -> let (n, _) = m in PTree.set n n (make_identity_ptree_rec lm)
+
+let make_identity_ptree f = make_identity_ptree_rec (PTree.elements (fn_code f))
+
+(* For now, identity function *)
+let duplicate_aux f =
+ let pTreeId = make_identity_ptree f
+ in (((fn_code f), (fn_entrypoint f)), pTreeId)
diff --git a/backend/Duplicateproof.v b/backend/Duplicateproof.v
new file mode 100644
index 00000000..54dd6196
--- /dev/null
+++ b/backend/Duplicateproof.v
@@ -0,0 +1,467 @@
+(** Correctness proof for code duplication *)
+Require Import AST Linking Errors Globalenvs Smallstep.
+Require Import Coqlib Maps Events Values.
+Require Import Op RTL Duplicate.
+
+Local Open Scope positive_scope.
+
+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.
+
+(* est-ce plus simple de prendre is_copy: node -> node, avec un noeud hors CFG à la place de None ? *)
+Inductive match_inst (is_copy: node -> option node): instruction -> instruction -> Prop :=
+ | match_inst_nop: forall n n',
+ is_copy n' = (Some n) -> match_inst is_copy (Inop n) (Inop n')
+ | match_inst_op: forall n n' op lr r,
+ is_copy n' = (Some n) -> match_inst is_copy (Iop op lr r n) (Iop op lr r n')
+ | match_inst_load: forall n n' m a lr r,
+ is_copy n' = (Some n) -> match_inst is_copy (Iload m a lr r n) (Iload m a lr r n')
+ | match_inst_store: forall n n' m a lr r,
+ is_copy n' = (Some n) -> match_inst is_copy (Istore m a lr r n) (Istore m a lr r n')
+ | match_inst_call: forall n n' s ri lr r,
+ is_copy n' = (Some n) -> match_inst is_copy (Icall s ri lr r n) (Icall s ri lr r n')
+ | match_inst_tailcall: forall s ri lr,
+ match_inst is_copy (Itailcall s ri lr) (Itailcall s ri lr)
+ | match_inst_builtin: forall n n' ef la br,
+ is_copy n' = (Some n) -> match_inst is_copy (Ibuiltin ef la br n) (Ibuiltin ef la br n')
+ | match_inst_cond: forall ifso ifso' ifnot ifnot' c lr,
+ is_copy ifso' = (Some ifso) -> is_copy ifnot' = (Some ifnot) ->
+ match_inst is_copy (Icond c lr ifso ifnot) (Icond c lr ifso' ifnot')
+ | match_inst_jumptable: forall ln ln' r,
+ list_forall2 (fun n n' => (is_copy n' = (Some n))) ln ln' ->
+ match_inst is_copy (Ijumptable r ln) (Ijumptable r ln')
+ | match_inst_return: forall or, match_inst is_copy (Ireturn or) (Ireturn or).
+
+Lemma verify_mapping_mn_rec_step:
+ forall lb b f xf,
+ In b lb ->
+ verify_mapping_mn_rec f xf lb = OK tt ->
+ verify_mapping_mn f xf b = OK tt.
+Proof.
+ induction lb; intros.
+ - monadInv H0. inversion H.
+ - inversion H.
+ + subst. monadInv H0. destruct x. assumption.
+ + monadInv H0. destruct x0. eapply IHlb; assumption.
+Qed.
+
+Lemma verify_is_copy_correct:
+ forall xf n n',
+ verify_is_copy (fn_revmap xf) n n' = OK tt ->
+ (fn_revmap xf) ! n' = Some n.
+Proof.
+ intros. unfold verify_is_copy in H. destruct (_ ! n') eqn:REVM; [|inversion H].
+ destruct (n ?= n0) eqn:NP; try (inversion H; fail).
+ eapply Pos.compare_eq in NP. subst.
+ reflexivity.
+Qed.
+
+Lemma verify_is_copy_list_correct:
+ forall xf ln ln',
+ verify_is_copy_list (fn_revmap xf) ln ln' = OK tt ->
+ list_forall2 (fun n n' => (fn_revmap xf) ! n' = Some n) ln ln'.
+Proof.
+ induction ln.
+ - intros. destruct ln'; monadInv H. constructor.
+ - intros. destruct ln'; monadInv H. destruct x. apply verify_is_copy_correct in EQ.
+ eapply IHln in EQ0. constructor; assumption.
+Qed.
+
+Lemma verify_match_inst_correct:
+ forall xf i i',
+ verify_match_inst (fn_revmap xf) i i' = OK tt ->
+ match_inst (fun nn => (fn_revmap xf) ! nn) i i'.
+Proof.
+ intros. unfold verify_match_inst in H.
+ destruct i; try (inversion H; fail).
+(* Inop *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct x. eapply verify_is_copy_correct in EQ.
+ constructor; eauto.
+(* Iop *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct x. eapply verify_is_copy_correct in EQ.
+ destruct (eq_operation _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (Pos.eq_dec _ _); try discriminate. clear EQ0. subst.
+ constructor. assumption.
+(* Iload *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct x. eapply verify_is_copy_correct in EQ.
+ destruct (chunk_eq _ _); try discriminate.
+ destruct (eq_addressing _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (Pos.eq_dec _ _); try discriminate. clear EQ0. subst.
+ constructor. assumption.
+(* Istore *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct x. eapply verify_is_copy_correct in EQ.
+ destruct (chunk_eq _ _); try discriminate.
+ destruct (eq_addressing _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (Pos.eq_dec _ _); try discriminate. clear EQ0. subst.
+ constructor. assumption.
+(* Icall *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct x. eapply verify_is_copy_correct in EQ.
+ destruct (signature_eq _ _); try discriminate.
+ destruct (product_eq _ _ _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (Pos.eq_dec _ _); try discriminate. subst.
+ constructor. assumption.
+(* Itailcall *)
+ - destruct i'; try (inversion H; fail).
+ destruct (signature_eq _ _); try discriminate.
+ destruct (product_eq _ _ _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate. subst. clear H.
+ constructor.
+(* Ibuiltin *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct x. eapply verify_is_copy_correct in EQ.
+ destruct (external_function_eq _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate.
+ destruct (builtin_res_eq_pos _ _); try discriminate. subst.
+ constructor. assumption.
+(* Icond *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct x. eapply verify_is_copy_correct in EQ.
+ destruct x0. eapply verify_is_copy_correct in EQ1.
+ destruct (condition_eq _ _); try discriminate.
+ destruct (list_eq_dec _ _ _); try discriminate. subst.
+ constructor; assumption.
+(* Ijumptable *)
+ - destruct i'; try (inversion H; fail). monadInv H.
+ destruct x. eapply verify_is_copy_list_correct in EQ.
+ destruct (Pos.eq_dec _ _); try discriminate. subst.
+ constructor. assumption.
+(* Ireturn *)
+ - destruct i'; try (inversion H; fail).
+ destruct (option_eq _ _ _); try discriminate. subst. clear H.
+ constructor.
+Qed.
+
+
+Lemma verify_mapping_mn_correct:
+ forall mp n n' i f xf tc,
+ mp ! n' = Some n ->
+ (fn_code f) ! n = Some i ->
+ (fn_code (fn_RTL xf)) = tc ->
+ fn_revmap xf = mp ->
+ verify_mapping_mn f xf (n', n) = OK tt ->
+ exists i',
+ tc ! n' = Some i'
+ /\ match_inst (fun nn => mp ! nn) i i'.
+Proof.
+ intros. unfold verify_mapping_mn in H3. rewrite H0 in H3. clear H0. rewrite H1 in H3. clear H1.
+ destruct (tc ! n') eqn:TCN; [| inversion H3].
+ exists i0. split; auto. rewrite <- H2.
+ eapply verify_match_inst_correct. assumption.
+Qed.
+
+
+Lemma verify_mapping_mn_rec_correct:
+ forall mp n n' i f xf tc,
+ mp ! n' = Some n ->
+ (fn_code f) ! n = Some i ->
+ (fn_code (fn_RTL xf)) = tc ->
+ fn_revmap xf = mp ->
+ verify_mapping_mn_rec f xf (PTree.elements mp) = OK tt ->
+ exists i',
+ tc ! n' = Some i'
+ /\ match_inst (fun nn => mp ! nn) i i'.
+Proof.
+ intros. exploit PTree.elements_correct. eapply H. intros IN.
+ eapply verify_mapping_mn_rec_step in H3; eauto.
+ eapply verify_mapping_mn_correct; eauto.
+Qed.
+
+
+Theorem revmap_correct: forall f xf n n',
+ transf_function_aux f = OK xf ->
+ (fn_revmap xf)!n' = Some n ->
+ (forall i, (fn_code f)!n = Some i -> exists i', (fn_code (fn_RTL xf))!n' = Some i' /\ match_inst (fun n' => (fn_revmap xf)!n') i i').
+Proof.
+ intros until n'. intros TRANSF REVM i FNC.
+ unfold transf_function_aux in TRANSF. destruct (duplicate_aux f) as (tcte & mp). destruct tcte as (tc & te). monadInv TRANSF.
+ simpl in *. monadInv EQ. clear EQ0.
+ unfold verify_mapping_match_nodes in EQ. simpl in EQ. destruct x1.
+ eapply verify_mapping_mn_rec_correct. 5: eapply EQ. all: eauto.
+Qed.
+
+
+Theorem revmap_entrypoint:
+ forall f xf, transf_function_aux f = OK xf -> (fn_revmap xf)!(fn_entrypoint (fn_RTL xf)) = Some (fn_entrypoint f).
+Proof.
+ intros. unfold transf_function_aux in H. destruct (duplicate_aux _) as (tcte & mp). destruct tcte as (tc & te).
+ monadInv H. simpl. monadInv EQ. unfold verify_mapping_entrypoint in EQ0. simpl in EQ0.
+ destruct (mp ! te) eqn:PT; try discriminate.
+ destruct (n ?= fn_entrypoint f) eqn:EQQ; try discriminate. inv EQ0.
+ apply Pos.compare_eq in EQQ. congruence.
+Qed.
+
+Section PRESERVATION.
+
+Variable prog: program.
+Variable tprog: program.
+
+Hypothesis TRANSL: match_prog prog tprog.
+
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Lemma symbols_preserved s: Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof.
+ rewrite <- (Genv.find_symbol_match TRANSL). reflexivity.
+Qed.
+
+Lemma senv_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 (Genv.senv_match TRANSL).
+Qed.
+
+Lemma functions_translated:
+ forall (v: val) (f: fundef),
+ Genv.find_funct ge v = Some f ->
+ exists tf cunit, transf_fundef f = OK tf /\ Genv.find_funct tge v = Some tf /\ linkorder cunit prog.
+Proof.
+ intros. exploit (Genv.find_funct_match TRANSL); eauto.
+ intros (cu & tf & A & B & C).
+ repeat eexists; intuition eauto.
+ + unfold incl; auto.
+ + eapply linkorder_refl.
+Qed.
+
+Lemma function_ptr_translated:
+ forall v f,
+ Genv.find_funct_ptr ge v = Some f ->
+ exists tf,
+ Genv.find_funct_ptr tge v = Some tf /\ transf_fundef f = OK tf.
+Proof.
+ intros.
+ exploit (Genv.find_funct_ptr_transf_partial TRANSL); eauto.
+Qed.
+
+Lemma function_sig_translated:
+ forall f tf, transf_fundef f = OK tf -> funsig tf = funsig f.
+Proof.
+ intros. destruct f.
+ - simpl in H. monadInv H. simpl. symmetry. monadInv EQ. apply transf_function_aux_preserves. assumption.
+ - simpl in H. monadInv H. (* monadInv EQ. *) reflexivity.
+Qed.
+
+Lemma sig_preserved:
+ forall f tf,
+ transf_fundef f = OK tf ->
+ funsig tf = funsig f.
+Proof.
+ unfold transf_fundef, transf_partial_fundef; intros.
+ destruct f. monadInv H. simpl. symmetry. monadInv EQ. apply transf_function_aux_preserves. assumption.
+ inv H. reflexivity.
+Qed.
+
+Lemma list_nth_z_revmap:
+ forall ln f ln' (pc pc': node) val,
+ list_nth_z ln val = Some pc ->
+ list_forall2 (fun n n' => (fn_revmap f)!n' = Some n) ln ln' ->
+ exists pc',
+ list_nth_z ln' val = Some pc'
+ /\ (fn_revmap f)!pc' = Some pc.
+Proof.
+ induction ln; intros until val; intros LNZ LFA.
+ - inv LNZ.
+ - inv LNZ. destruct (zeq val 0) eqn:ZEQ.
+ + inv H0. destruct ln'; inv LFA.
+ simpl. exists p. split; auto.
+ + inv LFA. simpl. rewrite ZEQ. exploit IHln. 2: eapply H0. all: eauto.
+ intros (pc'1 & LNZ & REV). exists pc'1. split; auto. congruence.
+Qed.
+
+Inductive match_stackframes: stackframe -> stackframe -> Prop :=
+ | match_stackframe_intro:
+ forall res f sp pc rs xf pc'
+ (TRANSF: transf_function_aux f = OK xf)
+ (DUPLIC: (fn_revmap xf)!pc' = Some pc),
+ match_stackframes (Stackframe res f sp pc rs) (Stackframe res (fn_RTL xf) sp pc' rs).
+
+Inductive match_states: state -> state -> Prop :=
+ | match_states_intro:
+ forall st f sp pc rs m st' xf pc'
+ (STACKS: list_forall2 match_stackframes st st')
+ (TRANSF: transf_function_aux f = OK xf)
+ (DUPLIC: (fn_revmap xf)!pc' = Some pc),
+ match_states (State st f sp pc rs m) (State st' (fn_RTL xf) sp pc' rs m)
+ | match_states_call:
+ forall st st' f f' args m
+ (STACKS: list_forall2 match_stackframes st st')
+ (TRANSF: transf_fundef f = OK f'),
+ match_states (Callstate st f args m) (Callstate st' f' args m)
+ | match_states_return:
+ forall st st' v m
+ (STACKS: list_forall2 match_stackframes st st'),
+ match_states (Returnstate st v m) (Returnstate st' v m).
+
+Theorem transf_initial_states:
+ forall s1, initial_state prog s1 ->
+ exists s2, initial_state tprog s2 /\ match_states s1 s2.
+Proof.
+ intros. inv H.
+ exploit function_ptr_translated; eauto. intros (tf & FIND & TRANSF).
+ eexists. split.
+ - econstructor.
+ + eapply (Genv.init_mem_transf_partial TRANSL); eauto.
+ + replace (prog_main tprog) with (prog_main prog). rewrite symbols_preserved. eauto.
+ symmetry. eapply match_program_main. eauto.
+ + exploit function_ptr_translated; eauto.
+ + destruct f.
+ * monadInv TRANSF. monadInv EQ. rewrite <- H3. symmetry; eapply transf_function_aux_preserves. assumption.
+ * monadInv TRANSF. (* monadInv EQ. *) assumption.
+ - constructor; eauto. constructor.
+Qed.
+
+Theorem transf_final_states:
+ forall s1 s2 r,
+ match_states s1 s2 -> final_state s1 r -> final_state s2 r.
+Proof.
+ intros. inv H0. inv H. inv STACKS. constructor.
+Qed.
+
+Theorem step_simulation:
+ forall s1 t s1', step ge s1 t s1' ->
+ forall s2 (MS: match_states s1 s2),
+ exists s2',
+ step tge s2 t s2'
+ /\ match_states s1' s2'.
+Proof.
+ induction 1; intros; inv MS.
+(* Inop *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3).
+ inv H3.
+ eexists. split.
+ + eapply exec_Inop; eauto.
+ + constructor; eauto.
+(* Iop *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ + eapply exec_Iop; eauto. erewrite eval_operation_preserved; eauto.
+ + constructor; eauto.
+(* Iload *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ + eapply exec_Iload; eauto. erewrite eval_addressing_preserved; eauto.
+ + constructor; auto.
+(* Istore *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ + eapply exec_Istore; eauto. erewrite eval_addressing_preserved; eauto.
+ + constructor; auto.
+(* Icall *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ destruct ros.
+ * simpl in H0. apply functions_translated in H0.
+ destruct H0 as (tf & cunit & TFUN & GFIND & LO).
+ eexists. split.
+ + eapply exec_Icall. eassumption. simpl. eassumption.
+ apply function_sig_translated. assumption.
+ + repeat (constructor; auto).
+ * simpl in H0. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate.
+ apply function_ptr_translated in H0. destruct H0 as (tf & GFF & TF).
+ eexists. split.
+ + eapply exec_Icall. eassumption. simpl. rewrite symbols_preserved. rewrite GFS.
+ eassumption. apply function_sig_translated. assumption.
+ + repeat (constructor; auto).
+(* Itailcall *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H10 & H11). inv H11.
+ pose symbols_preserved as SYMPRES.
+ destruct ros.
+ * simpl in H0. apply functions_translated in H0.
+ destruct H0 as (tf & cunit & TFUN & GFIND & LO).
+ eexists. split.
+ + eapply exec_Itailcall. eassumption. simpl. eassumption.
+ apply function_sig_translated. assumption.
+ erewrite <- transf_function_aux_fnstacksize; eauto.
+ + repeat (constructor; auto).
+ * simpl in H0. destruct (Genv.find_symbol _ _) eqn:GFS; try discriminate.
+ apply function_ptr_translated in H0. destruct H0 as (tf & GFF & TF).
+ eexists. split.
+ + eapply exec_Itailcall. eassumption. simpl. rewrite symbols_preserved. rewrite GFS.
+ eassumption. apply function_sig_translated. assumption.
+ erewrite <- transf_function_aux_fnstacksize; eauto.
+ + repeat (constructor; auto).
+(* Ibuiltin *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ + eapply exec_Ibuiltin; eauto. eapply eval_builtin_args_preserved; eauto.
+ eapply external_call_symbols_preserved; eauto. eapply senv_preserved.
+ + constructor; auto.
+(* Icond *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ + eapply exec_Icond; eauto.
+ + constructor; auto. destruct b; auto.
+(* Ijumptable *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ exploit list_nth_z_revmap; eauto. intros (pc'1 & LNZ & REVM).
+ eexists. split.
+ + eapply exec_Ijumptable; eauto.
+ + constructor; auto.
+(* Ireturn *)
+ - eapply revmap_correct in DUPLIC; eauto.
+ destruct DUPLIC as (i' & H2 & H3). inv H3.
+ pose symbols_preserved as SYMPRES.
+ eexists. split.
+ + eapply exec_Ireturn; eauto. erewrite <- transf_function_aux_fnstacksize; eauto.
+ + constructor; auto.
+(* exec_function_internal *)
+ - monadInv TRANSF. monadInv EQ. eexists. split.
+ + eapply exec_function_internal. erewrite <- transf_function_aux_fnstacksize; eauto.
+ + erewrite transf_function_aux_fnparams; eauto.
+ econstructor; eauto. apply revmap_entrypoint. assumption.
+(* exec_function_external *)
+ - monadInv TRANSF. (* monadInv EQ. *) eexists. split.
+ + econstructor. eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ + constructor. assumption.
+(* exec_return *)
+ - inv STACKS. destruct b1 as [res' f' sp' pc' rs']. eexists. split.
+ + constructor.
+ + inv H1. constructor; assumption.
+Qed.
+
+Theorem transf_program_correct:
+ forward_simulation (semantics prog) (semantics tprog).
+Proof.
+ eapply forward_simulation_step with match_states.
+ - eapply senv_preserved.
+ - eapply transf_initial_states.
+ - eapply transf_final_states.
+ - eapply step_simulation.
+Qed.
+
+End PRESERVATION.
diff --git a/common/AST.v b/common/AST.v
index a91138c9..7ffe355d 100644
--- a/common/AST.v
+++ b/common/AST.v
@@ -17,7 +17,7 @@
the abstract syntax trees of many of the intermediate languages. *)
Require Import String.
-Require Import Coqlib Maps Errors Integers Floats.
+Require Import Coqlib Maps Errors Integers Floats BinPos.
Require Archi.
Set Implicit Arguments.
@@ -630,11 +630,28 @@ Inductive builtin_arg (A: Type) : Type :=
| BA_splitlong (hi lo: builtin_arg A)
| BA_addptr (a1 a2: builtin_arg A).
+Definition builtin_arg_eq {A: Type}:
+ (forall x y : A, {x = y} + {x <> y}) ->
+ forall (ba1 ba2: (builtin_arg A)), {ba1=ba2} + {ba1<>ba2}.
+Proof.
+ intro. generalize Integers.int_eq int64_eq float_eq float32_eq chunk_eq ptrofs_eq ident_eq.
+ decide equality.
+Defined.
+Global Opaque builtin_arg_eq.
+
Inductive builtin_res (A: Type) : Type :=
| BR (x: A)
| BR_none
| BR_splitlong (hi lo: builtin_res A).
+Definition builtin_res_eq {A: Type}:
+ (forall x y : A, {x = y} + {x <> y}) ->
+ forall (a b: builtin_res A), {a=b} + {a<>b}.
+Proof.
+ intro. decide equality.
+Defined.
+Global Opaque builtin_res_eq.
+
Fixpoint globals_of_builtin_arg (A: Type) (a: builtin_arg A) : list ident :=
match a with
| BA_loadglobal chunk id ofs => id :: nil
diff --git a/configure b/configure
index 96e4a1fd..5dfac1a8 100755
--- a/configure
+++ b/configure
@@ -567,14 +567,14 @@ missingtools=false
echo "Testing Coq... " | tr -d '\n'
coq_ver=$(${COQBIN}coqc -v 2>/dev/null | sed -n -e 's/The Coq Proof Assistant, version \([^ ]*\).*$/\1/p')
case "$coq_ver" in
- 8.7.0|8.7.1|8.7.2|8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10)
+ 8.8.0|8.8.1|8.8.2|8.9.0|8.9.1|8.10)
echo "version $coq_ver -- good!";;
?*)
echo "version $coq_ver -- UNSUPPORTED"
if $ignore_coq_version; then
echo "Warning: this version of Coq is unsupported, proceed at your own risks."
else
- echo "Error: CompCert requires one of the following Coq versions: 8.10, 8.9.1, 8.9.0, 8.8.2, 8.8.1, 8.8.0, 8.7.2, 8.7.1, 8.7.0"
+ echo "Error: CompCert requires one of the following Coq versions: 8.10, 8.9.1, 8.9.0, 8.8.2, 8.8.1, 8.8.0"
missingtools=true
fi;;
"")
diff --git a/cparser/Machine.ml b/cparser/Machine.ml
index 0112a5ec..4999f0ac 100644
--- a/cparser/Machine.ml
+++ b/cparser/Machine.ml
@@ -18,6 +18,7 @@
type struct_passing_style =
| SP_ref_callee (* by reference, callee takes copy *)
| SP_ref_caller (* by reference, caller takes copy *)
+ | SP_value32_ref_callee (* by value if <= 32 bits, by ref_callee otherwise *)
| SP_split_args (* by value, as a sequence of ints *)
type struct_return_style =
@@ -268,12 +269,8 @@ let mppa_k1c =
bigendian = false;
bitfields_msb_first = false; (* TO CHECK *)
supports_unaligned_accesses = true;
- struct_passing_style = SP_split_args;
- struct_return_style = SR_int1248 }
-let aarch64 =
- { i32lpll64 with name = "aarch64";
- struct_passing_style = SP_ref_callee; (* Wrong *)
- struct_return_style = SR_ref } (* Wrong *)
+ struct_passing_style = SP_value32_ref_callee;
+ struct_return_style = SR_int1to4 }
(* Add GCC extensions re: sizeof and alignof *)
diff --git a/cparser/Machine.mli b/cparser/Machine.mli
index 31726d7f..ea25c4f6 100644
--- a/cparser/Machine.mli
+++ b/cparser/Machine.mli
@@ -17,6 +17,7 @@
type struct_passing_style =
| SP_ref_callee (* by reference, callee takes copy *)
| SP_ref_caller (* by reference, caller takes copy *)
+ | SP_value32_ref_callee (* by value if <= 32 bits, by ref_callee otherwise *)
| SP_split_args (* by value, as a sequence of ints *)
type struct_return_style =
diff --git a/cparser/StructPassing.ml b/cparser/StructPassing.ml
index 5c6454f0..3aff090e 100644
--- a/cparser/StructPassing.ml
+++ b/cparser/StructPassing.ml
@@ -68,7 +68,16 @@ let classify_param env ty =
match !struct_passing_style with
| SP_ref_callee -> Param_unchanged
| SP_ref_caller -> Param_ref_caller
- | _ ->
+ | SP_value32_ref_callee ->
+ (match sizeof env ty, alignof env ty with
+ | Some sz, Some al ->
+ if (sz <= 4) then
+ Param_flattened ((sz+3)/4, sz, al) (* FIXME - why (sz+3)/4 ? *)
+ else
+ Param_unchanged
+ | _, _ -> Param_unchanged (* when parsing prototype with incomplete structure definition *)
+ )
+ | SP_split_args ->
match sizeof env ty, alignof env ty with
| Some sz, Some al ->
Param_flattened ((sz + 3) / 4, sz, al)
diff --git a/driver/Compiler.v b/driver/Compiler.v
index c683c136..d7784f7a 100644
--- a/driver/Compiler.v
+++ b/driver/Compiler.v
@@ -38,6 +38,7 @@ Require RTLgen.
Require Tailcall.
Require Inlining.
Require Renumber.
+Require Duplicate.
Require Constprop.
Require CSE.
Require Deadcode.
@@ -59,6 +60,7 @@ Require RTLgenproof.
Require Tailcallproof.
Require Inliningproof.
Require Renumberproof.
+Require Duplicateproof.
Require Constpropproof.
Require CSEproof.
Require Deadcodeproof.
@@ -126,16 +128,18 @@ Definition transf_rtl_program (f: RTL.program) : res Asm.program :=
@@ print (print_RTL 2)
@@ time "Renumbering" Renumber.transf_program
@@ print (print_RTL 3)
- @@ total_if Compopts.optim_constprop (time "Constant propagation" Constprop.transf_program)
+ @@@ time "Duplicating" Duplicate.transf_program
@@ print (print_RTL 4)
- @@ total_if Compopts.optim_constprop (time "Renumbering" Renumber.transf_program)
+ @@ total_if Compopts.optim_constprop (time "Constant propagation" Constprop.transf_program)
@@ print (print_RTL 5)
- @@@ partial_if Compopts.optim_CSE (time "CSE" CSE.transf_program)
+ @@ total_if Compopts.optim_constprop (time "Renumbering" Renumber.transf_program)
@@ print (print_RTL 6)
- @@@ partial_if Compopts.optim_redundancy (time "Redundancy elimination" Deadcode.transf_program)
+ @@@ partial_if Compopts.optim_CSE (time "CSE" CSE.transf_program)
@@ print (print_RTL 7)
- @@@ time "Unused globals" Unusedglob.transform_program
+ @@@ partial_if Compopts.optim_redundancy (time "Redundancy elimination" Deadcode.transf_program)
@@ print (print_RTL 8)
+ @@@ time "Unused globals" Unusedglob.transform_program
+ @@ print (print_RTL 9)
@@@ time "Register allocation" Allocation.transf_program
@@ print print_LTL
@@ time "Branch tunneling" Tunneling.tunnel_program
@@ -238,6 +242,7 @@ Definition CompCert's_passes :=
::: mkpass (match_if Compopts.optim_tailcalls Tailcallproof.match_prog)
::: mkpass Inliningproof.match_prog
::: mkpass Renumberproof.match_prog
+ ::: mkpass Duplicateproof.match_prog
::: mkpass (match_if Compopts.optim_constprop Constpropproof.match_prog)
::: mkpass (match_if Compopts.optim_constprop Renumberproof.match_prog)
::: mkpass (match_if Compopts.optim_CSE CSEproof.match_prog)
@@ -281,17 +286,18 @@ Proof.
set (p7 := total_if optim_tailcalls Tailcall.transf_program p6) in *.
destruct (Inlining.transf_program p7) as [p8|e] eqn:P8; simpl in T; try discriminate.
set (p9 := Renumber.transf_program p8) in *.
- set (p10 := total_if optim_constprop Constprop.transf_program p9) in *.
- set (p11 := total_if optim_constprop Renumber.transf_program p10) in *.
- destruct (partial_if optim_CSE CSE.transf_program p11) as [p12|e] eqn:P12; simpl in T; try discriminate.
- destruct (partial_if optim_redundancy Deadcode.transf_program p12) as [p13|e] eqn:P13; simpl in T; try discriminate.
- destruct (Unusedglob.transform_program p13) as [p14|e] eqn:P14; simpl in T; try discriminate.
- destruct (Allocation.transf_program p14) as [p15|e] eqn:P15; simpl in T; try discriminate.
- set (p16 := Tunneling.tunnel_program p15) in *.
- destruct (Linearize.transf_program p16) as [p17|e] eqn:P17; simpl in T; try discriminate.
- set (p18 := CleanupLabels.transf_program p17) in *.
- destruct (partial_if debug Debugvar.transf_program p18) as [p19|e] eqn:P19; simpl in T; try discriminate.
- destruct (Stacking.transf_program p19) as [p20|e] eqn:P20; simpl in T; try discriminate.
+ destruct (Duplicate.transf_program p9) as [p10|e] eqn:P10; simpl in T; try discriminate.
+ set (p11 := total_if optim_constprop Constprop.transf_program p10) in *.
+ set (p12 := total_if optim_constprop Renumber.transf_program p11) in *.
+ destruct (partial_if optim_CSE CSE.transf_program p12) as [p13|e] eqn:P13; simpl in T; try discriminate.
+ destruct (partial_if optim_redundancy Deadcode.transf_program p13) as [p14|e] eqn:P14; simpl in T; try discriminate.
+ destruct (Unusedglob.transform_program p14) as [p15|e] eqn:P15; simpl in T; try discriminate.
+ destruct (Allocation.transf_program p15) as [p16|e] eqn:P16; simpl in T; try discriminate.
+ set (p17 := Tunneling.tunnel_program p16) in *.
+ destruct (Linearize.transf_program p17) as [p18|e] eqn:P18; simpl in T; try discriminate.
+ set (p19 := CleanupLabels.transf_program p18) in *.
+ destruct (partial_if debug Debugvar.transf_program p19) as [p20|e] eqn:P20; simpl in T; try discriminate.
+ destruct (Stacking.transf_program p20) as [p21|e] eqn:P21; simpl in T; try discriminate.
unfold match_prog; simpl.
exists p1; split. apply SimplExprproof.transf_program_match; auto.
exists p2; split. apply SimplLocalsproof.match_transf_program; auto.
@@ -302,17 +308,18 @@ Proof.
exists p7; split. apply total_if_match. apply Tailcallproof.transf_program_match.
exists p8; split. apply Inliningproof.transf_program_match; auto.
exists p9; split. apply Renumberproof.transf_program_match; auto.
- exists p10; split. apply total_if_match. apply Constpropproof.transf_program_match.
- exists p11; split. apply total_if_match. apply Renumberproof.transf_program_match.
- exists p12; split. eapply partial_if_match; eauto. apply CSEproof.transf_program_match.
- exists p13; split. eapply partial_if_match; eauto. apply Deadcodeproof.transf_program_match.
- exists p14; split. apply Unusedglobproof.transf_program_match; auto.
- exists p15; split. apply Allocproof.transf_program_match; auto.
- exists p16; split. apply Tunnelingproof.transf_program_match.
- exists p17; split. apply Linearizeproof.transf_program_match; auto.
- exists p18; split. apply CleanupLabelsproof.transf_program_match; auto.
- exists p19; split. eapply partial_if_match; eauto. apply Debugvarproof.transf_program_match.
- exists p20; split. apply Stackingproof.transf_program_match; auto.
+ exists p10; split. apply Duplicateproof.transf_program_match; auto.
+ exists p11; split. apply total_if_match. apply Constpropproof.transf_program_match.
+ exists p12; split. apply total_if_match. apply Renumberproof.transf_program_match.
+ exists p13; split. eapply partial_if_match; eauto. apply CSEproof.transf_program_match.
+ exists p14; split. eapply partial_if_match; eauto. apply Deadcodeproof.transf_program_match.
+ exists p15; split. apply Unusedglobproof.transf_program_match; auto.
+ exists p16; split. apply Allocproof.transf_program_match; auto.
+ exists p17; split. apply Tunnelingproof.transf_program_match.
+ exists p18; split. apply Linearizeproof.transf_program_match; auto.
+ exists p19; split. apply CleanupLabelsproof.transf_program_match; auto.
+ exists p20; split. eapply partial_if_match; eauto. apply Debugvarproof.transf_program_match.
+ exists p21; split. apply Stackingproof.transf_program_match; auto.
exists tp; split. apply Asmgenproof.transf_program_match; auto.
reflexivity.
Qed.
@@ -364,7 +371,7 @@ Ltac DestructM :=
destruct H as (p & M & MM); clear H
end.
repeat DestructM. subst tp.
- assert (F: forward_simulation (Cstrategy.semantics p) (Asm.semantics p21)).
+ assert (F: forward_simulation (Cstrategy.semantics p) (Asm.semantics p22)).
{
eapply compose_forward_simulations.
eapply SimplExprproof.transl_program_correct; eassumption.
@@ -383,6 +390,7 @@ Ltac DestructM :=
eapply compose_forward_simulations.
eapply Inliningproof.transf_program_correct; eassumption.
eapply compose_forward_simulations. eapply Renumberproof.transf_program_correct; eassumption.
+ eapply compose_forward_simulations. eapply Duplicateproof.transf_program_correct; eassumption.
eapply compose_forward_simulations.
eapply match_if_simulation. eassumption. exact Constpropproof.transf_program_correct.
eapply compose_forward_simulations.
diff --git a/lib/Floats.v b/lib/Floats.v
index 13350dd0..272efa52 100644
--- a/lib/Floats.v
+++ b/lib/Floats.v
@@ -16,7 +16,7 @@
(** Formalization of floating-point numbers, using the Flocq library. *)
-Require Import Coqlib Zbits Integers.
+Require Import Coqlib Zbits Integers Axioms.
(*From Flocq*)
Require Import Binary Bits Core.
Require Import IEEE754_extra.
@@ -27,8 +27,69 @@ Close Scope R_scope.
Open Scope Z_scope.
Definition float := binary64. (**r the type of IEE754 double-precision FP numbers *)
+
+Definition float_eq: forall (i1 i2: float), {i1=i2} + {i1<>i2}.
+Proof.
+ intros. destruct i1.
+(* B754_zero *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + apply eqb_prop in BEQ. subst. left. reflexivity.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_infinity *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + apply eqb_prop in BEQ. subst. left. reflexivity.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_nan *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + generalize (Pos.eq_dec pl pl0). intro. inv H.
+ ++ left. apply eqb_prop in BEQ. subst.
+ assert (e = e0) by (apply proof_irr). congruence.
+ ++ right. intro. inv H. contradiction.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_finite *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ; [apply eqb_prop in BEQ | apply eqb_false_iff in BEQ].
+ generalize (Pos.eq_dec m m0). intro. inv H.
+ generalize (Z.eq_dec e e1). intro. inv H.
+ 1: { left. assert (e0 = e2) by (apply proof_irr). congruence. }
+ all: right; intro; inv H; contradiction.
+Qed.
+
Definition float32 := binary32. (**r the type of IEE754 single-precision FP numbers *)
+Definition float32_eq: forall (i1 i2: float32), {i1=i2} + {i1<>i2}.
+Proof.
+ intros. destruct i1.
+(* B754_zero *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + apply eqb_prop in BEQ. subst. left. reflexivity.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_infinity *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + apply eqb_prop in BEQ. subst. left. reflexivity.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_nan *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ.
+ + generalize (Pos.eq_dec pl pl0). intro. inv H.
+ ++ left. apply eqb_prop in BEQ. subst.
+ assert (e = e0) by (apply proof_irr). congruence.
+ ++ right. intro. inv H. contradiction.
+ + apply eqb_false_iff in BEQ. right. intro. inv H. contradiction.
+(* B754_finite *)
+ - destruct i2; try (right; discriminate).
+ destruct (eqb s s0) eqn:BEQ; [apply eqb_prop in BEQ | apply eqb_false_iff in BEQ].
+ generalize (Pos.eq_dec m m0). intro. inv H.
+ generalize (Z.eq_dec e e1). intro. inv H.
+ 1: { left. assert (e0 = e2) by (apply proof_irr). congruence. }
+ all: right; intro; inv H; contradiction.
+Qed.
+
(** Boolean-valued comparisons *)
Definition cmp_of_comparison (c: comparison) (x: option Datatypes.comparison) : bool :=
diff --git a/lib/Integers.v b/lib/Integers.v
index 3b6c35eb..3e78ee60 100644
--- a/lib/Integers.v
+++ b/lib/Integers.v
@@ -16,7 +16,7 @@
(** Formalizations of machine integers modulo $2^N$ #2<sup>N</sup>#. *)
Require Import Eqdep_dec Zquot Zwf.
-Require Import Coqlib Zbits.
+Require Import Coqlib Zbits Axioms.
Require Archi.
(** * Comparisons *)
@@ -29,6 +29,11 @@ Inductive comparison : Type :=
| Cgt : comparison (**r greater than *)
| Cge : comparison. (**r greater than or equal *)
+Definition comparison_eq: forall (x y: comparison), {x = y} + {x <> y}.
+Proof.
+ decide equality.
+Defined.
+
Definition negate_comparison (c: comparison): comparison :=
match c with
| Ceq => Cne
@@ -4535,8 +4540,26 @@ End Int64.
Strategy 0 [Wordsize_64.wordsize].
+Definition int_eq: forall (i1 i2: int), {i1=i2} + {i1<>i2}.
+Proof.
+ generalize Z.eq_dec. intros.
+ destruct i1. destruct i2. generalize (H intval intval0). intro.
+ inversion H0.
+ - subst. left. assert (intrange = intrange0) by (apply proof_irr). congruence.
+ - right. intro. inversion H2. contradiction.
+Qed.
+
Notation int64 := Int64.int.
+Definition int64_eq: forall (i1 i2: int64), {i1=i2} + {i1<>i2}.
+Proof.
+ generalize Z.eq_dec. intros.
+ destruct i1. destruct i2. generalize (H intval intval0). intro.
+ inversion H0.
+ - subst. left. assert (intrange = intrange0) by (apply proof_irr). congruence.
+ - right. intro. inversion H2. contradiction.
+Qed.
+
Global Opaque Int.repr Int64.repr Byte.repr.
(** * Specialization to offsets in pointer values *)
@@ -4813,6 +4836,15 @@ Strategy 0 [Wordsize_Ptrofs.wordsize].
Notation ptrofs := Ptrofs.int.
+Definition ptrofs_eq: forall (i1 i2: ptrofs), {i1=i2} + {i1<>i2}.
+Proof.
+ generalize Z.eq_dec. intros.
+ destruct i1. destruct i2. generalize (H intval intval0). intro.
+ inversion H0.
+ - subst. left. assert (intrange = intrange0) by (apply proof_irr). congruence.
+ - right. intro. inversion H2. contradiction.
+Qed.
+
Global Opaque Ptrofs.repr.
Hint Resolve Int.modulus_pos Int.eqm_refl Int.eqm_refl2 Int.eqm_sym Int.eqm_trans
diff --git a/mppa_k1c/Archi.v b/mppa_k1c/Archi.v
index cdcf58c3..69b32c7c 100644
--- a/mppa_k1c/Archi.v
+++ b/mppa_k1c/Archi.v
@@ -17,7 +17,6 @@
(** Architecture-dependent parameters for MPPA K1c. Mostly copied from the Risc-V backend *)
Require Import ZArith List.
-(*From Flocq*)
Require Import Binary Bits.
Definition ptr64 := true.
diff --git a/mppa_k1c/Asm.v b/mppa_k1c/Asm.v
index f09aa99c..e27ff40c 100644
--- a/mppa_k1c/Asm.v
+++ b/mppa_k1c/Asm.v
@@ -1,753 +1,753 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* Prashanth Mundkur, SRI International *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* The contributions by Prashanth Mundkur are reused and adapted *)
-(* under the terms of a Contributor License Agreement between *)
-(* SRI International and INRIA. *)
-(* *)
-(* *********************************************************************)
-
-(** * Abstract syntax for K1c textual assembly language.
-
- Each emittable instruction is defined here. ';;' is also defined as an instruction.
- The goal of this representation is to stay compatible with the rest of the generic backend of CompCert
- We define [unfold : list bblock -> list instruction]
- An Asm function is then defined as : [fn_sig], [fn_blocks], [fn_code], and a proof of [unfold fn_blocks = fn_code]
- [fn_code] has no semantic. Instead, the semantic of Asm is given by using the AsmVLIW semantic on [fn_blocks] *)
-
-Require Import Coqlib.
-Require Import Maps.
-Require Import AST.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import ExtValues.
-Require Import Memory.
-Require Import Events.
-Require Import Globalenvs.
-Require Import Smallstep.
-Require Import Locations.
-Require Stacklayout.
-Require Import Conventions.
-Require Import Asmvliw.
-Require Import Linking.
-Require Import Errors.
-
-(** Definitions for OCaml code *)
-Definition label := positive.
-Definition preg := preg.
-
-Inductive addressing : Type :=
- | AOff (ofs: offset)
- | AReg (ro: ireg)
- | ARegXS (ro: ireg)
-.
-
-(** Syntax *)
-Inductive instruction : Type :=
- (** pseudo instructions *)
- | Pallocframe (sz: Z) (pos: ptrofs) (**r allocate new stack frame *)
- | Pfreeframe (sz: Z) (pos: ptrofs) (**r deallocate stack frame and restore previous frame *)
- | Plabel (lbl: label) (**r define a code label *)
- | Ploadsymbol (rd: ireg) (id: ident) (ofs: ptrofs) (**r load the address of a symbol *)
- | Pbuiltin: external_function -> list (builtin_arg preg)
- -> builtin_res preg -> instruction (**r built-in function (pseudo) *)
- | Psemi (**r semi colon separating bundles *)
- | Pnop (**r instruction that does nothing *)
-
- (** Control flow instructions *)
- | Pget (rd: ireg) (rs: preg) (**r get system register *)
- | Pset (rd: preg) (rs: ireg) (**r set system register *)
- | Pret (**r return *)
- | Pcall (l: label) (**r function call *)
- | Picall (rs: ireg) (**r function call on register *)
- (* Pgoto is for tailcalls, Pj_l is for jumping to a particular label *)
- | Pgoto (l: label) (**r goto *)
- | Pigoto (rs: ireg) (**r goto from register *)
- | Pj_l (l: label) (**r jump to label *)
- | Pcb (bt: btest) (r: ireg) (l: label) (**r branch based on btest *)
- | Pcbu (bt: btest) (r: ireg) (l: label) (**r branch based on btest with unsigned semantics *)
- | Pjumptable (r: ireg) (labels: list label)
-
- (* For builtins *)
- | Ploopdo (count: ireg) (loopend: label)
- | Pgetn (n: int) (dst: ireg)
- | Psetn (n: int) (src: ireg)
- | Pwfxl (n: int) (src: ireg)
- | Pwfxm (n: int) (src: ireg)
- | Pldu (dst: ireg) (addr: ireg)
- | Plbzu (dst: ireg) (addr: ireg)
- | Plhzu (dst: ireg) (addr: ireg)
- | Plwzu (dst: ireg) (addr: ireg)
- | Pawait
- | Psleep
- | Pstop
- | Pbarrier
- | Pfence
- | Pdinval
- | Pdinvall (addr: ireg)
- | Pdtouchl (addr: ireg)
- | Piinval
- | Piinvals (addr: ireg)
- | Pitouchl (addr: ireg)
- | Pdzerol (addr: ireg)
-(*| Pafaddd (addr: ireg) (incr_res: ireg)
- | Pafaddw (addr: ireg) (incr_res: ireg) *) (* see #157 *)
- | Palclrd (dst: ireg) (addr: ireg)
- | Palclrw (dst: ireg) (addr: ireg)
- | Pclzll (rd rs: ireg)
- | Pstsud (rd rs1 rs2: ireg)
-
- (** Loads **)
- | Plb (rd: ireg) (ra: ireg) (ofs: addressing) (**r load byte *)
- | Plbu (rd: ireg) (ra: ireg) (ofs: addressing) (**r load byte unsigned *)
- | Plh (rd: ireg) (ra: ireg) (ofs: addressing) (**r load half word *)
- | Plhu (rd: ireg) (ra: ireg) (ofs: addressing) (**r load half word unsigned *)
- | Plw (rd: ireg) (ra: ireg) (ofs: addressing) (**r load int32 *)
- | Plw_a (rd: ireg) (ra: ireg) (ofs: addressing) (**r load any32 *)
- | Pld (rd: ireg) (ra: ireg) (ofs: addressing) (**r load int64 *)
- | Pld_a (rd: ireg) (ra: ireg) (ofs: addressing) (**r load any64 *)
- | Pfls (rd: freg) (ra: ireg) (ofs: addressing) (**r load float *)
- | Pfld (rd: freg) (ra: ireg) (ofs: addressing) (**r load 64-bit float *)
- | Plq (rs: gpreg_q) (ra: ireg) (ofs: addressing) (**r load 2*64-bit *)
- | Plo (rs: gpreg_o) (ra: ireg) (ofs: addressing) (**r load 4*64-bit *)
-
- (** Stores **)
- | Psb (rs: ireg) (ra: ireg) (ofs: addressing) (**r store byte *)
- | Psh (rs: ireg) (ra: ireg) (ofs: addressing) (**r store half byte *)
- | Psw (rs: ireg) (ra: ireg) (ofs: addressing) (**r store int32 *)
- | Psw_a (rs: ireg) (ra: ireg) (ofs: addressing) (**r store any32 *)
- | Psd (rs: ireg) (ra: ireg) (ofs: addressing) (**r store int64 *)
- | Psd_a (rs: ireg) (ra: ireg) (ofs: addressing) (**r store any64 *)
- | Pfss (rs: freg) (ra: ireg) (ofs: addressing) (**r store float *)
- | Pfsd (rs: freg) (ra: ireg) (ofs: addressing) (**r store 64-bit float *)
-
- | Psq (rs: gpreg_q) (ra: ireg) (ofs: addressing) (**r store 2*64-bit *)
- | Pso (rs: gpreg_o) (ra: ireg) (ofs: addressing) (**r store 2*64-bit *)
-
- (** Arith RR *)
- | Pmv (rd rs: ireg) (**r register move *)
- | Pnegw (rd rs: ireg) (**r negate word *)
- | Pnegl (rd rs: ireg) (**r negate long *)
- | Pcvtl2w (rd rs: ireg) (**r Convert Long to Word *)
- | Psxwd (rd rs: ireg) (**r Sign Extend Word to Double Word *)
- | Pzxwd (rd rs: ireg) (**r Zero Extend Word to Double Word *)
-
- | Pextfz (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r extract bitfields unsigned *)
- | Pextfs (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r extract bitfields signed *)
-
- | Pextfzl (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r extract bitfields unsigned *)
- | Pextfsl (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r extract bitfields signed *)
-
- | Pinsf (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r insert bitfield *)
- | Pinsfl (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r insert bitfield *)
-
- | Pfabsd (rd rs: ireg) (**r float absolute double *)
- | Pfabsw (rd rs: ireg) (**r float absolute word *)
- | Pfnegd (rd rs: ireg) (**r float negate double *)
- | Pfnegw (rd rs: ireg) (**r float negate word *)
- | Pfnarrowdw (rd rs: ireg) (**r float narrow 64 -> 32 bits *)
- | Pfwidenlwd (rd rs: ireg) (**r float widen 32 -> 64 bits *)
- | Pfloatwrnsz (rd rs: ireg) (**r Floating Point Conversion from integer (32 -> 32) *)
- | Pfloatuwrnsz (rd rs: ireg) (**r Floating Point Conversion from integer (u32 -> 32) *)
- | Pfloatudrnsz (rd rs: ireg) (**r Floating Point Conversion from unsigned integer (64 bits) *)
- | Pfloatdrnsz (rd rs: ireg) (**r Floating Point Conversion from integer (64 bits) *)
- | Pfixedwrzz (rd rs: ireg) (**r Integer conversion from floating point *)
- | Pfixeduwrzz (rd rs: ireg) (**r Integer conversion from floating point (f32 -> 32 bits unsigned *)
- | Pfixeddrzz (rd rs: ireg) (**r Integer conversion from floating point (i64 -> 64 bits) *)
- | Pfixeddrzz_i32 (rd rs: ireg) (**r Integer conversion from floating point (i32 -> f64) *)
- | Pfixedudrzz (rd rs: ireg) (**r unsigned Integer conversion from floating point (u64 -> 64 bits) *)
- | Pfixedudrzz_i32 (rd rs: ireg) (**r unsigned Integer conversion from floating point (u32 -> 64 bits) *)
-
- (** Arith RI32 *)
- | Pmake (rd: ireg) (imm: int) (**r load immediate *)
-
- (** Arith RI64 *)
- | Pmakel (rd: ireg) (imm: int64) (**r load immediate long *)
-
- (** Arith RF32 *)
- | Pmakefs (rd: ireg) (imm: float32)
-
- (** Arith RF64 *)
- | Pmakef (rd: ireg) (imm: float)
-
- (** Arith RRR *)
- | Pcompw (it: itest) (rd rs1 rs2: ireg) (**r comparison word *)
- | Pcompl (it: itest) (rd rs1 rs2: ireg) (**r comparison long *)
- | Pfcompw (ft: ftest) (rd rs1 rs2: ireg) (**r comparison float *)
- | Pfcompl (ft: ftest) (rd rs1 rs2: ireg) (**r comparison float64 *)
-
- | Paddw (rd rs1 rs2: ireg) (**r add word *)
- | Paddxw (shift : shift1_4) (rd rs1 rs2: ireg) (**r add word *)
- | Psubw (rd rs1 rs2: ireg) (**r sub word *)
- | Prevsubxw (shift : shift1_4) (rd rs1 rs2: ireg) (**r add word *)
- | Pmulw (rd rs1 rs2: ireg) (**r mul word *)
- | Pandw (rd rs1 rs2: ireg) (**r and word *)
- | Pnandw (rd rs1 rs2: ireg) (**r nand word *)
- | Porw (rd rs1 rs2: ireg) (**r or word *)
- | Pnorw (rd rs1 rs2: ireg) (**r nor word *)
- | Pxorw (rd rs1 rs2: ireg) (**r xor word *)
- | Pnxorw (rd rs1 rs2: ireg) (**r xor word *)
- | Pandnw (rd rs1 rs2: ireg) (**r andn word *)
- | Pornw (rd rs1 rs2: ireg) (**r orn word *)
- | Psraw (rd rs1 rs2: ireg) (**r shift right arithmetic word *)
- | Psrxw (rd rs1 rs2: ireg) (**r shift right arithmetic word round to 0*)
- | Psrlw (rd rs1 rs2: ireg) (**r shift right logical word *)
- | Psllw (rd rs1 rs2: ireg) (**r shift left logical word *)
- | Pmaddw (rd rs1 rs2: ireg) (**r multiply-add words *)
- | Pmsubw (rd rs1 rs2: ireg) (**r multiply-add words *)
- | Pfmaddfw (rd rs1 rs2: ireg) (**r float fused multiply-add words *)
- | Pfmsubfw (rd rs1 rs2: ireg) (**r float fused multiply-subtract words *)
- | Pfmaddfl (rd rs1 rs2: ireg) (**r float fused multiply-add longs *)
- | Pfmsubfl (rd rs1 rs2: ireg) (**r float fused multiply-subtract longs *)
-
- | Paddl (rd rs1 rs2: ireg) (**r add long *)
- | Paddxl (shift : shift1_4) (rd rs1 rs2: ireg) (**r add long shift *)
- | Psubl (rd rs1 rs2: ireg) (**r sub long *)
- | Prevsubxl (shift : shift1_4) (rd rs1 rs2: ireg) (**r sub long shift *)
- | Pandl (rd rs1 rs2: ireg) (**r and long *)
- | Pnandl (rd rs1 rs2: ireg) (**r nand long *)
- | Porl (rd rs1 rs2: ireg) (**r or long *)
- | Pnorl (rd rs1 rs2: ireg) (**r nor long *)
- | Pxorl (rd rs1 rs2: ireg) (**r xor long *)
- | Pnxorl (rd rs1 rs2: ireg) (**r nxor long *)
- | Pandnl (rd rs1 rs2: ireg) (**r andn long *)
- | Pornl (rd rs1 rs2: ireg) (**r orn long *)
- | Pmull (rd rs1 rs2: ireg) (**r mul long (low part) *)
- | Pslll (rd rs1 rs2: ireg) (**r shift left logical long *)
- | Psrll (rd rs1 rs2: ireg) (**r shift right logical long *)
- | Psral (rd rs1 rs2: ireg) (**r shift right arithmetic long *)
- | Psrxl (rd rs1 rs2: ireg) (**r shift right arithmetic long round to 0*)
- | Pmaddl (rd rs1 rs2: ireg) (**r multiply-add long *)
- | Pmsubl (rd rs1 rs2: ireg) (**r multiply-add long *)
-
- | Pfaddd (rd rs1 rs2: ireg) (**r Float addition double *)
- | Pfaddw (rd rs1 rs2: ireg) (**r Float addition word *)
- | Pfsbfd (rd rs1 rs2: ireg) (**r Float sub double *)
- | Pfsbfw (rd rs1 rs2: ireg) (**r Float sub word *)
- | Pfmuld (rd rs1 rs2: ireg) (**r Float mul double *)
- | Pfmulw (rd rs1 rs2: ireg) (**r Float mul word *)
- | Pfmind (rd rs1 rs2: ireg) (**r Float min double *)
- | Pfminw (rd rs1 rs2: ireg) (**r Float min word *)
- | Pfmaxd (rd rs1 rs2: ireg) (**r Float max double *)
- | Pfmaxw (rd rs1 rs2: ireg) (**r Float max word *)
- | Pfinvw (rd rs1: ireg) (**r Float invert word *)
-
- (** Arith RRI32 *)
- | Pcompiw (it: itest) (rd rs: ireg) (imm: int) (**r comparison imm word *)
-
- | Paddiw (rd rs: ireg) (imm: int) (**r add imm word *)
- | Paddxiw (shift : shift1_4) (rd rs: ireg) (imm: int) (**r add imm word *)
- | Prevsubiw (rd rs: ireg) (imm: int) (**r subtract imm word *)
- | Prevsubxiw (shift : shift1_4) (rd rs: ireg) (imm: int) (**r subtract imm word *)
- | Pmuliw (rd rs: ireg) (imm: int) (**r mul imm word *)
- | Pandiw (rd rs: ireg) (imm: int) (**r and imm word *)
- | Pnandiw (rd rs: ireg) (imm: int) (**r nand imm word *)
- | Poriw (rd rs: ireg) (imm: int) (**r or imm word *)
- | Pnoriw (rd rs: ireg) (imm: int) (**r nor imm word *)
- | Pxoriw (rd rs: ireg) (imm: int) (**r xor imm word *)
- | Pnxoriw (rd rs: ireg) (imm: int) (**r nxor imm word *)
- | Pandniw (rd rs: ireg) (imm: int) (**r andn imm word *)
- | Porniw (rd rs: ireg) (imm: int) (**r orn imm word *)
- | Psraiw (rd rs: ireg) (imm: int) (**r shift right arithmetic imm word *)
- | Psrxiw (rd rs: ireg) (imm: int) (**r shift right arithmetic imm word round to 0*)
- | Psrliw (rd rs: ireg) (imm: int) (**r shift right logical imm word *)
- | Pslliw (rd rs: ireg) (imm: int) (**r shift left logical imm word *)
- | Proriw (rd rs: ireg) (imm: int) (**r rotate right imm word *)
- | Pmaddiw (rd rs: ireg) (imm: int) (**r multiply add imm word *)
- | Psllil (rd rs: ireg) (imm: int) (**r shift left logical immediate long *)
- | Psrxil (rd rs: ireg) (imm: int) (**r shift right arithmetic imm word round to 0*)
- | Psrlil (rd rs: ireg) (imm: int) (**r shift right logical immediate long *)
- | Psrail (rd rs: ireg) (imm: int) (**r shift right arithmetic immediate long *)
-
- (** Arith RRI64 *)
- | Pcompil (it: itest) (rd rs: ireg) (imm: int64) (**r comparison imm long *)
- | Paddil (rd rs: ireg) (imm: int64) (**r add immediate long *)
- | Paddxil (shift : shift1_4) (rd rs: ireg) (imm: int64) (**r add immediate long *)
- | Prevsubil (rd rs: ireg) (imm: int64) (**r subtract imm long *)
- | Prevsubxil (shift : shift1_4) (rd rs: ireg) (imm: int64) (**r subtract imm long *)
- | Pmulil (rd rs: ireg) (imm: int64) (**r add immediate long *)
- | Pandil (rd rs: ireg) (imm: int64) (**r and immediate long *)
- | Pnandil (rd rs: ireg) (imm: int64) (**r and immediate long *)
- | Poril (rd rs: ireg) (imm: int64) (**r or immediate long *)
- | Pnoril (rd rs: ireg) (imm: int64) (**r and immediate long *)
- | Pxoril (rd rs: ireg) (imm: int64) (**r xor immediate long *)
- | Pnxoril (rd rs: ireg) (imm: int64) (**r xor immediate long *)
- | Pandnil (rd rs: ireg) (imm: int64) (**r andn long *)
- | Pornil (rd rs: ireg) (imm: int64) (**r orn long *)
- | Pmaddil (rd rs: ireg) (imm: int64) (**r multiply add imm long *)
- | Pcmove (bt: btest) (rcond rd rs : ireg) (** conditional move *)
- | Pcmoveu (bt: btest) (rcond rd rs : ireg) (** conditional move, unsigned semantics *)
- | Pcmoveiw (bt: btest) (rcond rd : ireg) (imm: int) (** conditional move *)
- | Pcmoveuiw (bt: btest) (rcond rd : ireg) (imm: int) (** conditional move, unsigned semantics *)
- | Pcmoveil (bt: btest) (rcond rd : ireg) (imm: int64) (** conditional move *)
- | Pcmoveuil (bt: btest) (rcond rd : ireg) (imm: int64) (** conditional move, unsigned semantics *)
-.
-
-(** Correspondance between Asmblock and Asm *)
-
-Definition control_to_instruction (c: control) :=
- match c with
- | PExpand (Asmvliw.Pbuiltin ef args res) => Pbuiltin ef args res
- | PCtlFlow Asmvliw.Pret => Pret
- | PCtlFlow (Asmvliw.Pcall l) => Pcall l
- | PCtlFlow (Asmvliw.Picall r) => Picall r
- | PCtlFlow (Asmvliw.Pgoto l) => Pgoto l
- | PCtlFlow (Asmvliw.Pigoto l) => Pigoto l
- | PCtlFlow (Asmvliw.Pj_l l) => Pj_l l
- | PCtlFlow (Asmvliw.Pcb bt r l) => Pcb bt r l
- | PCtlFlow (Asmvliw.Pcbu bt r l) => Pcbu bt r l
- | PCtlFlow (Asmvliw.Pjumptable r label) => Pjumptable r label
- end.
-
-Definition basic_to_instruction (b: basic) :=
- match b with
- (** Special basics *)
- | Asmvliw.Pget rd rs => Pget rd rs
- | Asmvliw.Pset rd rs => Pset rd rs
- | Asmvliw.Pnop => Pnop
- | Asmvliw.Pallocframe sz pos => Pallocframe sz pos
- | Asmvliw.Pfreeframe sz pos => Pfreeframe sz pos
-
- (** PArith basics *)
- (* R *)
- | PArithR (Asmvliw.Ploadsymbol id ofs) r => Ploadsymbol r id ofs
-
- (* RR *)
- | PArithRR Asmvliw.Pmv rd rs => Pmv rd rs
- | PArithRR Asmvliw.Pnegw rd rs => Pnegw rd rs
- | PArithRR Asmvliw.Pnegl rd rs => Pnegl rd rs
- | PArithRR Asmvliw.Pcvtl2w rd rs => Pcvtl2w rd rs
- | PArithRR Asmvliw.Psxwd rd rs => Psxwd rd rs
- | PArithRR Asmvliw.Pzxwd rd rs => Pzxwd rd rs
- | PArithRR (Asmvliw.Pextfz stop start) rd rs => Pextfz rd rs stop start
- | PArithRR (Asmvliw.Pextfs stop start) rd rs => Pextfs rd rs stop start
- | PArithRR (Asmvliw.Pextfzl stop start) rd rs => Pextfzl rd rs stop start
- | PArithRR (Asmvliw.Pextfsl stop start) rd rs => Pextfsl rd rs stop start
- | PArithRR Asmvliw.Pfabsd rd rs => Pfabsd rd rs
- | PArithRR Asmvliw.Pfabsw rd rs => Pfabsw rd rs
- | PArithRR Asmvliw.Pfnegd rd rs => Pfnegd rd rs
- | PArithRR Asmvliw.Pfnegw rd rs => Pfnegw rd rs
- | PArithRR Asmvliw.Pfinvw rd rs => Pfinvw rd rs
- | PArithRR Asmvliw.Pfnarrowdw rd rs => Pfnarrowdw rd rs
- | PArithRR Asmvliw.Pfwidenlwd rd rs => Pfwidenlwd rd rs
- | PArithRR Asmvliw.Pfloatuwrnsz rd rs => Pfloatuwrnsz rd rs
- | PArithRR Asmvliw.Pfloatwrnsz rd rs => Pfloatwrnsz rd rs
- | PArithRR Asmvliw.Pfloatudrnsz rd rs => Pfloatudrnsz rd rs
- | PArithRR Asmvliw.Pfloatdrnsz rd rs => Pfloatdrnsz rd rs
- | PArithRR Asmvliw.Pfixedwrzz rd rs => Pfixedwrzz rd rs
- | PArithRR Asmvliw.Pfixeduwrzz rd rs => Pfixeduwrzz rd rs
- | PArithRR Asmvliw.Pfixeddrzz rd rs => Pfixeddrzz rd rs
- | PArithRR Asmvliw.Pfixedudrzz rd rs => Pfixedudrzz rd rs
- | PArithRR Asmvliw.Pfixeddrzz_i32 rd rs => Pfixeddrzz_i32 rd rs
- | PArithRR Asmvliw.Pfixedudrzz_i32 rd rs => Pfixedudrzz_i32 rd rs
-
- (* RI32 *)
- | PArithRI32 Asmvliw.Pmake rd imm => Pmake rd imm
-
- (* RI64 *)
- | PArithRI64 Asmvliw.Pmakel rd imm => Pmakel rd imm
-
- (* RF32 *)
- | PArithRF32 Asmvliw.Pmakefs rd imm => Pmakefs rd imm
-
- (* RF64 *)
- | PArithRF64 Asmvliw.Pmakef rd imm => Pmakef rd imm
-
- (* RRR *)
- | PArithRRR (Asmvliw.Pcompw it) rd rs1 rs2 => Pcompw it rd rs1 rs2
- | PArithRRR (Asmvliw.Pcompl it) rd rs1 rs2 => Pcompl it rd rs1 rs2
- | PArithRRR (Asmvliw.Pfcompw ft) rd rs1 rs2 => Pfcompw ft rd rs1 rs2
- | PArithRRR (Asmvliw.Pfcompl ft) rd rs1 rs2 => Pfcompl ft rd rs1 rs2
- | PArithRRR Asmvliw.Paddw rd rs1 rs2 => Paddw rd rs1 rs2
- | PArithRRR (Asmvliw.Paddxw shift) rd rs1 rs2 => Paddxw shift rd rs1 rs2
- | PArithRRR Asmvliw.Psubw rd rs1 rs2 => Psubw rd rs1 rs2
- | PArithRRR (Asmvliw.Prevsubxw shift) rd rs1 rs2 => Prevsubxw shift rd rs1 rs2
- | PArithRRR Asmvliw.Pmulw rd rs1 rs2 => Pmulw rd rs1 rs2
- | PArithRRR Asmvliw.Pandw rd rs1 rs2 => Pandw rd rs1 rs2
- | PArithRRR Asmvliw.Pnandw rd rs1 rs2 => Pnandw rd rs1 rs2
- | PArithRRR Asmvliw.Porw rd rs1 rs2 => Porw rd rs1 rs2
- | PArithRRR Asmvliw.Pnorw rd rs1 rs2 => Pnorw rd rs1 rs2
- | PArithRRR Asmvliw.Pxorw rd rs1 rs2 => Pxorw rd rs1 rs2
- | PArithRRR Asmvliw.Pnxorw rd rs1 rs2 => Pnxorw rd rs1 rs2
- | PArithRRR Asmvliw.Pandnw rd rs1 rs2 => Pandnw rd rs1 rs2
- | PArithRRR Asmvliw.Pornw rd rs1 rs2 => Pornw rd rs1 rs2
- | PArithRRR Asmvliw.Psraw rd rs1 rs2 => Psraw rd rs1 rs2
- | PArithRRR Asmvliw.Psrxw rd rs1 rs2 => Psrxw rd rs1 rs2
- | PArithRRR Asmvliw.Psrlw rd rs1 rs2 => Psrlw rd rs1 rs2
- | PArithRRR Asmvliw.Psllw rd rs1 rs2 => Psllw rd rs1 rs2
-
- | PArithRRR Asmvliw.Paddl rd rs1 rs2 => Paddl rd rs1 rs2
- | PArithRRR (Asmvliw.Paddxl shift) rd rs1 rs2 => Paddxl shift rd rs1 rs2
- | PArithRRR Asmvliw.Psubl rd rs1 rs2 => Psubl rd rs1 rs2
- | PArithRRR (Asmvliw.Prevsubxl shift) rd rs1 rs2 => Prevsubxl shift rd rs1 rs2
- | PArithRRR Asmvliw.Pandl rd rs1 rs2 => Pandl rd rs1 rs2
- | PArithRRR Asmvliw.Pnandl rd rs1 rs2 => Pnandl rd rs1 rs2
- | PArithRRR Asmvliw.Porl rd rs1 rs2 => Porl rd rs1 rs2
- | PArithRRR Asmvliw.Pnorl rd rs1 rs2 => Pnorl rd rs1 rs2
- | PArithRRR Asmvliw.Pxorl rd rs1 rs2 => Pxorl rd rs1 rs2
- | PArithRRR Asmvliw.Pnxorl rd rs1 rs2 => Pnxorl rd rs1 rs2
- | PArithRRR Asmvliw.Pandnl rd rs1 rs2 => Pandnl rd rs1 rs2
- | PArithRRR Asmvliw.Pornl rd rs1 rs2 => Pornl rd rs1 rs2
- | PArithRRR Asmvliw.Pmull rd rs1 rs2 => Pmull rd rs1 rs2
- | PArithRRR Asmvliw.Pslll rd rs1 rs2 => Pslll rd rs1 rs2
- | PArithRRR Asmvliw.Psrll rd rs1 rs2 => Psrll rd rs1 rs2
- | PArithRRR Asmvliw.Psral rd rs1 rs2 => Psral rd rs1 rs2
- | PArithRRR Asmvliw.Psrxl rd rs1 rs2 => Psrxl rd rs1 rs2
-
- | PArithRRR Asmvliw.Pfaddd rd rs1 rs2 => Pfaddd rd rs1 rs2
- | PArithRRR Asmvliw.Pfaddw rd rs1 rs2 => Pfaddw rd rs1 rs2
- | PArithRRR Asmvliw.Pfsbfd rd rs1 rs2 => Pfsbfd rd rs1 rs2
- | PArithRRR Asmvliw.Pfsbfw rd rs1 rs2 => Pfsbfw rd rs1 rs2
- | PArithRRR Asmvliw.Pfmuld rd rs1 rs2 => Pfmuld rd rs1 rs2
- | PArithRRR Asmvliw.Pfmulw rd rs1 rs2 => Pfmulw rd rs1 rs2
- | PArithRRR Asmvliw.Pfmind rd rs1 rs2 => Pfmind rd rs1 rs2
- | PArithRRR Asmvliw.Pfminw rd rs1 rs2 => Pfminw rd rs1 rs2
- | PArithRRR Asmvliw.Pfmaxd rd rs1 rs2 => Pfmaxd rd rs1 rs2
- | PArithRRR Asmvliw.Pfmaxw rd rs1 rs2 => Pfmaxw rd rs1 rs2
-
- (* RRI32 *)
- | PArithRRI32 (Asmvliw.Pcompiw it) rd rs imm => Pcompiw it rd rs imm
- | PArithRRI32 Asmvliw.Paddiw rd rs imm => Paddiw rd rs imm
- | PArithRRI32 (Asmvliw.Paddxiw shift) rd rs imm => Paddxiw shift rd rs imm
- | PArithRRI32 Asmvliw.Prevsubiw rd rs imm => Prevsubiw rd rs imm
- | PArithRRI32 (Asmvliw.Prevsubxiw shift) rd rs imm => Prevsubxiw shift rd rs imm
- | PArithRRI32 Asmvliw.Pmuliw rd rs imm => Pmuliw rd rs imm
- | PArithRRI32 Asmvliw.Pandiw rd rs imm => Pandiw rd rs imm
- | PArithRRI32 Asmvliw.Pnandiw rd rs imm => Pnandiw rd rs imm
- | PArithRRI32 Asmvliw.Poriw rd rs imm => Poriw rd rs imm
- | PArithRRI32 Asmvliw.Pnoriw rd rs imm => Pnoriw rd rs imm
- | PArithRRI32 Asmvliw.Pxoriw rd rs imm => Pxoriw rd rs imm
- | PArithRRI32 Asmvliw.Pnxoriw rd rs imm => Pnxoriw rd rs imm
- | PArithRRI32 Asmvliw.Pandniw rd rs imm => Pandniw rd rs imm
- | PArithRRI32 Asmvliw.Porniw rd rs imm => Porniw rd rs imm
- | PArithRRI32 Asmvliw.Psraiw rd rs imm => Psraiw rd rs imm
- | PArithRRI32 Asmvliw.Psrxiw rd rs imm => Psrxiw rd rs imm
- | PArithRRI32 Asmvliw.Psrliw rd rs imm => Psrliw rd rs imm
- | PArithRRI32 Asmvliw.Pslliw rd rs imm => Pslliw rd rs imm
- | PArithRRI32 Asmvliw.Proriw rd rs imm => Proriw rd rs imm
- | PArithRRI32 Asmvliw.Psllil rd rs imm => Psllil rd rs imm
- | PArithRRI32 Asmvliw.Psrlil rd rs imm => Psrlil rd rs imm
- | PArithRRI32 Asmvliw.Psrxil rd rs imm => Psrxil rd rs imm
- | PArithRRI32 Asmvliw.Psrail rd rs imm => Psrail rd rs imm
-
- (* RRI64 *)
- | PArithRRI64 (Asmvliw.Pcompil it) rd rs imm => Pcompil it rd rs imm
- | PArithRRI64 Asmvliw.Paddil rd rs imm => Paddil rd rs imm
- | PArithRRI64 (Asmvliw.Paddxil shift) rd rs imm => Paddxil shift rd rs imm
- | PArithRRI64 Asmvliw.Prevsubil rd rs imm => Prevsubil rd rs imm
- | PArithRRI64 (Asmvliw.Prevsubxil shift) rd rs imm => Prevsubxil shift rd rs imm
- | PArithRRI64 Asmvliw.Pmulil rd rs imm => Pmulil rd rs imm
- | PArithRRI64 Asmvliw.Pandil rd rs imm => Pandil rd rs imm
- | PArithRRI64 Asmvliw.Pnandil rd rs imm => Pnandil rd rs imm
- | PArithRRI64 Asmvliw.Poril rd rs imm => Poril rd rs imm
- | PArithRRI64 Asmvliw.Pnoril rd rs imm => Pnoril rd rs imm
- | PArithRRI64 Asmvliw.Pxoril rd rs imm => Pxoril rd rs imm
- | PArithRRI64 Asmvliw.Pnxoril rd rs imm => Pnxoril rd rs imm
- | PArithRRI64 Asmvliw.Pandnil rd rs imm => Pandnil rd rs imm
- | PArithRRI64 Asmvliw.Pornil rd rs imm => Pornil rd rs imm
-
- (** ARRR *)
- | PArithARRR Asmvliw.Pmaddw rd rs1 rs2 => Pmaddw rd rs1 rs2
- | PArithARRR Asmvliw.Pmaddl rd rs1 rs2 => Pmaddl rd rs1 rs2
- | PArithARRR Asmvliw.Pmsubw rd rs1 rs2 => Pmsubw rd rs1 rs2
- | PArithARRR Asmvliw.Pmsubl rd rs1 rs2 => Pmsubl rd rs1 rs2
- | PArithARRR Asmvliw.Pfmaddfw rd rs1 rs2 => Pfmaddfw rd rs1 rs2
- | PArithARRR Asmvliw.Pfmaddfl rd rs1 rs2 => Pfmaddfl rd rs1 rs2
- | PArithARRR Asmvliw.Pfmsubfw rd rs1 rs2 => Pfmsubfw rd rs1 rs2
- | PArithARRR Asmvliw.Pfmsubfl rd rs1 rs2 => Pfmsubfl rd rs1 rs2
- | PArithARRR (Asmvliw.Pcmove cond) rd rs1 rs2=> Pcmove cond rd rs1 rs2
- | PArithARRR (Asmvliw.Pcmoveu cond) rd rs1 rs2=> Pcmoveu cond rd rs1 rs2
-
- (** ARR *)
- | PArithARR (Asmvliw.Pinsf stop start) rd rs => Pinsf rd rs stop start
- | PArithARR (Asmvliw.Pinsfl stop start) rd rs => Pinsfl rd rs stop start
-
- (** ARRI32 *)
- | PArithARRI32 Asmvliw.Pmaddiw rd rs1 imm => Pmaddiw rd rs1 imm
- | PArithARRI32 (Asmvliw.Pcmoveiw cond) rd rs1 imm => Pcmoveiw cond rd rs1 imm
- | PArithARRI32 (Asmvliw.Pcmoveuiw cond) rd rs1 imm => Pcmoveuiw cond rd rs1 imm
-
- (** ARRI64 *)
- | PArithARRI64 Asmvliw.Pmaddil rd rs1 imm => Pmaddil rd rs1 imm
- | PArithARRI64 (Asmvliw.Pcmoveil cond) rd rs1 imm => Pcmoveil cond rd rs1 imm
- | PArithARRI64 (Asmvliw.Pcmoveuil cond) rd rs1 imm => Pcmoveuil cond rd rs1 imm
- (** Load *)
- | PLoadRRO Asmvliw.Plb rd ra ofs => Plb rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Plbu rd ra ofs => Plbu rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Plh rd ra ofs => Plh rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Plhu rd ra ofs => Plhu rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Plw rd ra ofs => Plw rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Plw_a rd ra ofs => Plw_a rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Pld rd ra ofs => Pld rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Pld_a rd ra ofs => Pld_a rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Pfls rd ra ofs => Pfls rd ra (AOff ofs)
- | PLoadRRO Asmvliw.Pfld rd ra ofs => Pfld rd ra (AOff ofs)
-
- | PLoadQRRO qrs ra ofs => Plq qrs ra (AOff ofs)
- | PLoadORRO qrs ra ofs => Plo qrs ra (AOff ofs)
-
- | PLoadRRR Asmvliw.Plb rd ra ro => Plb rd ra (AReg ro)
- | PLoadRRR Asmvliw.Plbu rd ra ro => Plbu rd ra (AReg ro)
- | PLoadRRR Asmvliw.Plh rd ra ro => Plh rd ra (AReg ro)
- | PLoadRRR Asmvliw.Plhu rd ra ro => Plhu rd ra (AReg ro)
- | PLoadRRR Asmvliw.Plw rd ra ro => Plw rd ra (AReg ro)
- | PLoadRRR Asmvliw.Plw_a rd ra ro => Plw_a rd ra (AReg ro)
- | PLoadRRR Asmvliw.Pld rd ra ro => Pld rd ra (AReg ro)
- | PLoadRRR Asmvliw.Pld_a rd ra ro => Pld_a rd ra (AReg ro)
- | PLoadRRR Asmvliw.Pfls rd ra ro => Pfls rd ra (AReg ro)
- | PLoadRRR Asmvliw.Pfld rd ra ro => Pfld rd ra (AReg ro)
-
- | PLoadRRRXS Asmvliw.Plb rd ra ro => Plb rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Plbu rd ra ro => Plbu rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Plh rd ra ro => Plh rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Plhu rd ra ro => Plhu rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Plw rd ra ro => Plw rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Plw_a rd ra ro => Plw_a rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Pld rd ra ro => Pld rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Pld_a rd ra ro => Pld_a rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Pfls rd ra ro => Pfls rd ra (ARegXS ro)
- | PLoadRRRXS Asmvliw.Pfld rd ra ro => Pfld rd ra (ARegXS ro)
-
- (** Store *)
- | PStoreRRO Asmvliw.Psb rd ra ofs => Psb rd ra (AOff ofs)
- | PStoreRRO Asmvliw.Psh rd ra ofs => Psh rd ra (AOff ofs)
- | PStoreRRO Asmvliw.Psw rd ra ofs => Psw rd ra (AOff ofs)
- | PStoreRRO Asmvliw.Psw_a rd ra ofs => Psw_a rd ra (AOff ofs)
- | PStoreRRO Asmvliw.Psd rd ra ofs => Psd rd ra (AOff ofs)
- | PStoreRRO Asmvliw.Psd_a rd ra ofs => Psd_a rd ra (AOff ofs)
- | PStoreRRO Asmvliw.Pfss rd ra ofs => Pfss rd ra (AOff ofs)
- | PStoreRRO Asmvliw.Pfsd rd ra ofs => Pfsd rd ra (AOff ofs)
-
- | PStoreRRR Asmvliw.Psb rd ra ro => Psb rd ra (AReg ro)
- | PStoreRRR Asmvliw.Psh rd ra ro => Psh rd ra (AReg ro)
- | PStoreRRR Asmvliw.Psw rd ra ro => Psw rd ra (AReg ro)
- | PStoreRRR Asmvliw.Psw_a rd ra ro => Psw_a rd ra (AReg ro)
- | PStoreRRR Asmvliw.Psd rd ra ro => Psd rd ra (AReg ro)
- | PStoreRRR Asmvliw.Psd_a rd ra ro => Psd_a rd ra (AReg ro)
- | PStoreRRR Asmvliw.Pfss rd ra ro => Pfss rd ra (AReg ro)
- | PStoreRRR Asmvliw.Pfsd rd ra ro => Pfsd rd ra (AReg ro)
-
- | PStoreRRRXS Asmvliw.Psb rd ra ro => Psb rd ra (ARegXS ro)
- | PStoreRRRXS Asmvliw.Psh rd ra ro => Psh rd ra (ARegXS ro)
- | PStoreRRRXS Asmvliw.Psw rd ra ro => Psw rd ra (ARegXS ro)
- | PStoreRRRXS Asmvliw.Psw_a rd ra ro => Psw_a rd ra (ARegXS ro)
- | PStoreRRRXS Asmvliw.Psd rd ra ro => Psd rd ra (ARegXS ro)
- | PStoreRRRXS Asmvliw.Psd_a rd ra ro => Psd_a rd ra (ARegXS ro)
- | PStoreRRRXS Asmvliw.Pfss rd ra ro => Pfss rd ra (ARegXS ro)
- | PStoreRRRXS Asmvliw.Pfsd rd ra ro => Pfsd rd ra (ARegXS ro)
-
- | PStoreQRRO qrs ra ofs => Psq qrs ra (AOff ofs)
- | PStoreORRO qrs ra ofs => Pso qrs ra (AOff ofs)
- end.
-
-Section RELSEM.
-
-Definition code := list instruction.
-
-Fixpoint unfold_label (ll: list label) :=
- match ll with
- | nil => nil
- | l :: ll => Plabel l :: unfold_label ll
- end.
-
-Fixpoint unfold_body (lb: list basic) :=
- match lb with
- | nil => nil
- | b :: lb => basic_to_instruction b :: unfold_body lb
- end.
-
-Definition unfold_exit (oc: option control) :=
- match oc with
- | None => nil
- | Some c => control_to_instruction c :: nil
- end.
-
-Definition unfold_bblock (b: bblock) := unfold_label (header b) ++
- (match (body b), (exit b) with
- | (((Asmvliw.Pfreeframe _ _ | Asmvliw.Pallocframe _ _)::nil) as bo), None =>
- unfold_body bo
- | bo, ex => unfold_body bo ++ unfold_exit ex ++ Psemi :: nil
- end).
-
-Fixpoint unfold (lb: bblocks) :=
- match lb with
- | nil => nil
- | b :: lb => (unfold_bblock b) ++ unfold lb
- end.
-
-Record function : Type := mkfunction { fn_sig: signature; fn_blocks: bblocks; fn_code: code;
- correct: unfold fn_blocks = fn_code }.
-
-Definition fundef := AST.fundef function.
-Definition program := AST.program fundef unit.
-Definition genv := Genv.t fundef unit.
-
-Definition function_proj (f: function) := Asmvliw.mkfunction (fn_sig f) (fn_blocks f).
-
-Definition fundef_proj (fu: fundef) : Asmvliw.fundef :=
- match fu with
- | Internal f => Internal (function_proj f)
- | External ef => External ef
- end.
-
-Definition globdef_proj (gd: globdef fundef unit) : globdef Asmvliw.fundef unit :=
- match gd with
- | Gfun f => Gfun (fundef_proj f)
- | Gvar gu => Gvar gu
- end.
-
-Program Definition genv_trans (ge: genv) : Asmvliw.genv :=
- {| Genv.genv_public := Genv.genv_public ge;
- Genv.genv_symb := Genv.genv_symb ge;
- 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.
-Qed. Next Obligation.
- destruct ge; simpl in *.
- rewrite PTree.gmap1 in H.
- destruct (genv_defs ! b) eqn:GEN.
- - eauto.
- - discriminate.
-Qed. Next Obligation.
- destruct ge; simpl in *.
- eauto.
-Qed.
-
-Fixpoint prog_defs_proj (l: list (ident * globdef fundef unit))
- : list (ident * globdef Asmvliw.fundef unit) :=
- match l with
- | nil => nil
- | (i, gd) :: l => (i, globdef_proj gd) :: prog_defs_proj l
- end.
-
-Definition program_proj (p: program) : Asmvliw.program :=
- {| prog_defs := prog_defs_proj (prog_defs p);
- prog_public := prog_public p;
- prog_main := prog_main p
- |}.
-
-End RELSEM.
-
-Definition semantics (p: program) := Asmvliw.semantics (program_proj p).
-
-(** Determinacy of the [Asm] semantics. *)
-
-Lemma semantics_determinate: forall p, determinate (semantics p).
-Proof.
- intros. apply semantics_determinate.
-Qed.
-
-(** transf_program *)
-
-Program Definition transf_function (f: Asmvliw.function) : function :=
- {| fn_sig := Asmvliw.fn_sig f; fn_blocks := Asmvliw.fn_blocks f;
- fn_code := unfold (Asmvliw.fn_blocks f) |}.
-
-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.
-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.
- rewrite transf_function_proj. auto.
-Qed.
-
-Definition transf_program : Asmvliw.program -> program := transform_program transf_fundef.
-
-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 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.
- 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.
-
-Definition match_prog (p: Asmvliw.program) (tp: program) :=
- match_program (fun _ f tf => tf = transf_fundef f) eq p tp.
-
-Lemma transf_program_match:
- forall p tp, transf_program p = tp -> match_prog p tp.
-Proof.
- intros. rewrite <- H. eapply match_transform_program; eauto.
-Qed.
-
-Lemma cons_extract {A: Type} : forall (l: list A) a b, a = b -> a::l = b::l.
-Proof.
- intros. congruence.
-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: Asmvliw.program.
-Variable tprog: program.
-Hypothesis TRANSF: match_prog prog tprog.
-Let ge := Genv.globalenv prog.
-Let tge := Genv.globalenv tprog.
-
-Definition match_states (s1 s2: state) := s1 = s2.
-
-Lemma symbols_preserved:
- forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof (Genv.find_symbol_match TRANSF).
-
-Lemma senv_preserved:
- Senv.equiv ge tge.
-Proof (Genv.senv_match TRANSF).
-
-
-Theorem transf_program_correct:
- forward_simulation (Asmvliw.semantics prog) (semantics tprog).
-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.
- - intros. exists s1. split; auto. congruence.
- - intros. inv H. auto.
- - intros. exists s1'. inv H0. split; auto. congruence.
-Qed.
-
-End PRESERVATION.
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris-Rocquencourt *)
+(* Prashanth Mundkur, SRI International *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* The contributions by Prashanth Mundkur are reused and adapted *)
+(* under the terms of a Contributor License Agreement between *)
+(* SRI International and INRIA. *)
+(* *)
+(* *********************************************************************)
+
+(** * Abstract syntax for K1c textual assembly language.
+
+ Each emittable instruction is defined here. ';;' is also defined as an instruction.
+ The goal of this representation is to stay compatible with the rest of the generic backend of CompCert
+ We define [unfold : list bblock -> list instruction]
+ An Asm function is then defined as : [fn_sig], [fn_blocks], [fn_code], and a proof of [unfold fn_blocks = fn_code]
+ [fn_code] has no semantic. Instead, the semantic of Asm is given by using the AsmVLIW semantic on [fn_blocks] *)
+
+Require Import Coqlib.
+Require Import Maps.
+Require Import AST.
+Require Import Integers.
+Require Import Floats.
+Require Import Values.
+Require Import ExtValues.
+Require Import Memory.
+Require Import Events.
+Require Import Globalenvs.
+Require Import Smallstep.
+Require Import Locations.
+Require Stacklayout.
+Require Import Conventions.
+Require Import Asmvliw.
+Require Import Linking.
+Require Import Errors.
+
+(** Definitions for OCaml code *)
+Definition label := positive.
+Definition preg := preg.
+
+Inductive addressing : Type :=
+ | AOff (ofs: offset)
+ | AReg (ro: ireg)
+ | ARegXS (ro: ireg)
+.
+
+(** Syntax *)
+Inductive instruction : Type :=
+ (** pseudo instructions *)
+ | Pallocframe (sz: Z) (pos: ptrofs) (**r allocate new stack frame *)
+ | Pfreeframe (sz: Z) (pos: ptrofs) (**r deallocate stack frame and restore previous frame *)
+ | Plabel (lbl: label) (**r define a code label *)
+ | Ploadsymbol (rd: ireg) (id: ident) (ofs: ptrofs) (**r load the address of a symbol *)
+ | Pbuiltin: external_function -> list (builtin_arg preg)
+ -> builtin_res preg -> instruction (**r built-in function (pseudo) *)
+ | Psemi (**r semi colon separating bundles *)
+ | Pnop (**r instruction that does nothing *)
+
+ (** Control flow instructions *)
+ | Pget (rd: ireg) (rs: preg) (**r get system register *)
+ | Pset (rd: preg) (rs: ireg) (**r set system register *)
+ | Pret (**r return *)
+ | Pcall (l: label) (**r function call *)
+ | Picall (rs: ireg) (**r function call on register *)
+ (* Pgoto is for tailcalls, Pj_l is for jumping to a particular label *)
+ | Pgoto (l: label) (**r goto *)
+ | Pigoto (rs: ireg) (**r goto from register *)
+ | Pj_l (l: label) (**r jump to label *)
+ | Pcb (bt: btest) (r: ireg) (l: label) (**r branch based on btest *)
+ | Pcbu (bt: btest) (r: ireg) (l: label) (**r branch based on btest with unsigned semantics *)
+ | Pjumptable (r: ireg) (labels: list label)
+
+ (* For builtins *)
+ | Ploopdo (count: ireg) (loopend: label)
+ | Pgetn (n: int) (dst: ireg)
+ | Psetn (n: int) (src: ireg)
+ | Pwfxl (n: int) (src: ireg)
+ | Pwfxm (n: int) (src: ireg)
+ | Pldu (dst: ireg) (addr: ireg)
+ | Plbzu (dst: ireg) (addr: ireg)
+ | Plhzu (dst: ireg) (addr: ireg)
+ | Plwzu (dst: ireg) (addr: ireg)
+ | Pawait
+ | Psleep
+ | Pstop
+ | Pbarrier
+ | Pfence
+ | Pdinval
+ | Pdinvall (addr: ireg)
+ | Pdtouchl (addr: ireg)
+ | Piinval
+ | Piinvals (addr: ireg)
+ | Pitouchl (addr: ireg)
+ | Pdzerol (addr: ireg)
+(*| Pafaddd (addr: ireg) (incr_res: ireg)
+ | Pafaddw (addr: ireg) (incr_res: ireg) *) (* see #157 *)
+ | Palclrd (dst: ireg) (addr: ireg)
+ | Palclrw (dst: ireg) (addr: ireg)
+ | Pclzll (rd rs: ireg)
+ | Pstsud (rd rs1 rs2: ireg)
+
+ (** Loads **)
+ | Plb (rd: ireg) (ra: ireg) (ofs: addressing) (**r load byte *)
+ | Plbu (rd: ireg) (ra: ireg) (ofs: addressing) (**r load byte unsigned *)
+ | Plh (rd: ireg) (ra: ireg) (ofs: addressing) (**r load half word *)
+ | Plhu (rd: ireg) (ra: ireg) (ofs: addressing) (**r load half word unsigned *)
+ | Plw (rd: ireg) (ra: ireg) (ofs: addressing) (**r load int32 *)
+ | Plw_a (rd: ireg) (ra: ireg) (ofs: addressing) (**r load any32 *)
+ | Pld (rd: ireg) (ra: ireg) (ofs: addressing) (**r load int64 *)
+ | Pld_a (rd: ireg) (ra: ireg) (ofs: addressing) (**r load any64 *)
+ | Pfls (rd: freg) (ra: ireg) (ofs: addressing) (**r load float *)
+ | Pfld (rd: freg) (ra: ireg) (ofs: addressing) (**r load 64-bit float *)
+ | Plq (rs: gpreg_q) (ra: ireg) (ofs: addressing) (**r load 2*64-bit *)
+ | Plo (rs: gpreg_o) (ra: ireg) (ofs: addressing) (**r load 4*64-bit *)
+
+ (** Stores **)
+ | Psb (rs: ireg) (ra: ireg) (ofs: addressing) (**r store byte *)
+ | Psh (rs: ireg) (ra: ireg) (ofs: addressing) (**r store half byte *)
+ | Psw (rs: ireg) (ra: ireg) (ofs: addressing) (**r store int32 *)
+ | Psw_a (rs: ireg) (ra: ireg) (ofs: addressing) (**r store any32 *)
+ | Psd (rs: ireg) (ra: ireg) (ofs: addressing) (**r store int64 *)
+ | Psd_a (rs: ireg) (ra: ireg) (ofs: addressing) (**r store any64 *)
+ | Pfss (rs: freg) (ra: ireg) (ofs: addressing) (**r store float *)
+ | Pfsd (rs: freg) (ra: ireg) (ofs: addressing) (**r store 64-bit float *)
+
+ | Psq (rs: gpreg_q) (ra: ireg) (ofs: addressing) (**r store 2*64-bit *)
+ | Pso (rs: gpreg_o) (ra: ireg) (ofs: addressing) (**r store 2*64-bit *)
+
+ (** Arith RR *)
+ | Pmv (rd rs: ireg) (**r register move *)
+ | Pnegw (rd rs: ireg) (**r negate word *)
+ | Pnegl (rd rs: ireg) (**r negate long *)
+ | Pcvtl2w (rd rs: ireg) (**r Convert Long to Word *)
+ | Psxwd (rd rs: ireg) (**r Sign Extend Word to Double Word *)
+ | Pzxwd (rd rs: ireg) (**r Zero Extend Word to Double Word *)
+
+ | Pextfz (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r extract bitfields unsigned *)
+ | Pextfs (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r extract bitfields signed *)
+
+ | Pextfzl (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r extract bitfields unsigned *)
+ | Pextfsl (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r extract bitfields signed *)
+
+ | Pinsf (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r insert bitfield *)
+ | Pinsfl (rd : ireg) (rs : ireg) (stop : Z) (start : Z) (**r insert bitfield *)
+
+ | Pfabsd (rd rs: ireg) (**r float absolute double *)
+ | Pfabsw (rd rs: ireg) (**r float absolute word *)
+ | Pfnegd (rd rs: ireg) (**r float negate double *)
+ | Pfnegw (rd rs: ireg) (**r float negate word *)
+ | Pfnarrowdw (rd rs: ireg) (**r float narrow 64 -> 32 bits *)
+ | Pfwidenlwd (rd rs: ireg) (**r float widen 32 -> 64 bits *)
+ | Pfloatwrnsz (rd rs: ireg) (**r Floating Point Conversion from integer (32 -> 32) *)
+ | Pfloatuwrnsz (rd rs: ireg) (**r Floating Point Conversion from integer (u32 -> 32) *)
+ | Pfloatudrnsz (rd rs: ireg) (**r Floating Point Conversion from unsigned integer (64 bits) *)
+ | Pfloatdrnsz (rd rs: ireg) (**r Floating Point Conversion from integer (64 bits) *)
+ | Pfixedwrzz (rd rs: ireg) (**r Integer conversion from floating point *)
+ | Pfixeduwrzz (rd rs: ireg) (**r Integer conversion from floating point (f32 -> 32 bits unsigned *)
+ | Pfixeddrzz (rd rs: ireg) (**r Integer conversion from floating point (i64 -> 64 bits) *)
+ | Pfixeddrzz_i32 (rd rs: ireg) (**r Integer conversion from floating point (i32 -> f64) *)
+ | Pfixedudrzz (rd rs: ireg) (**r unsigned Integer conversion from floating point (u64 -> 64 bits) *)
+ | Pfixedudrzz_i32 (rd rs: ireg) (**r unsigned Integer conversion from floating point (u32 -> 64 bits) *)
+
+ (** Arith RI32 *)
+ | Pmake (rd: ireg) (imm: int) (**r load immediate *)
+
+ (** Arith RI64 *)
+ | Pmakel (rd: ireg) (imm: int64) (**r load immediate long *)
+
+ (** Arith RF32 *)
+ | Pmakefs (rd: ireg) (imm: float32)
+
+ (** Arith RF64 *)
+ | Pmakef (rd: ireg) (imm: float)
+
+ (** Arith RRR *)
+ | Pcompw (it: itest) (rd rs1 rs2: ireg) (**r comparison word *)
+ | Pcompl (it: itest) (rd rs1 rs2: ireg) (**r comparison long *)
+ | Pfcompw (ft: ftest) (rd rs1 rs2: ireg) (**r comparison float *)
+ | Pfcompl (ft: ftest) (rd rs1 rs2: ireg) (**r comparison float64 *)
+
+ | Paddw (rd rs1 rs2: ireg) (**r add word *)
+ | Paddxw (shift : shift1_4) (rd rs1 rs2: ireg) (**r add word *)
+ | Psubw (rd rs1 rs2: ireg) (**r sub word *)
+ | Prevsubxw (shift : shift1_4) (rd rs1 rs2: ireg) (**r add word *)
+ | Pmulw (rd rs1 rs2: ireg) (**r mul word *)
+ | Pandw (rd rs1 rs2: ireg) (**r and word *)
+ | Pnandw (rd rs1 rs2: ireg) (**r nand word *)
+ | Porw (rd rs1 rs2: ireg) (**r or word *)
+ | Pnorw (rd rs1 rs2: ireg) (**r nor word *)
+ | Pxorw (rd rs1 rs2: ireg) (**r xor word *)
+ | Pnxorw (rd rs1 rs2: ireg) (**r xor word *)
+ | Pandnw (rd rs1 rs2: ireg) (**r andn word *)
+ | Pornw (rd rs1 rs2: ireg) (**r orn word *)
+ | Psraw (rd rs1 rs2: ireg) (**r shift right arithmetic word *)
+ | Psrxw (rd rs1 rs2: ireg) (**r shift right arithmetic word round to 0*)
+ | Psrlw (rd rs1 rs2: ireg) (**r shift right logical word *)
+ | Psllw (rd rs1 rs2: ireg) (**r shift left logical word *)
+ | Pmaddw (rd rs1 rs2: ireg) (**r multiply-add words *)
+ | Pmsubw (rd rs1 rs2: ireg) (**r multiply-add words *)
+ | Pfmaddfw (rd rs1 rs2: ireg) (**r float fused multiply-add words *)
+ | Pfmsubfw (rd rs1 rs2: ireg) (**r float fused multiply-subtract words *)
+ | Pfmaddfl (rd rs1 rs2: ireg) (**r float fused multiply-add longs *)
+ | Pfmsubfl (rd rs1 rs2: ireg) (**r float fused multiply-subtract longs *)
+
+ | Paddl (rd rs1 rs2: ireg) (**r add long *)
+ | Paddxl (shift : shift1_4) (rd rs1 rs2: ireg) (**r add long shift *)
+ | Psubl (rd rs1 rs2: ireg) (**r sub long *)
+ | Prevsubxl (shift : shift1_4) (rd rs1 rs2: ireg) (**r sub long shift *)
+ | Pandl (rd rs1 rs2: ireg) (**r and long *)
+ | Pnandl (rd rs1 rs2: ireg) (**r nand long *)
+ | Porl (rd rs1 rs2: ireg) (**r or long *)
+ | Pnorl (rd rs1 rs2: ireg) (**r nor long *)
+ | Pxorl (rd rs1 rs2: ireg) (**r xor long *)
+ | Pnxorl (rd rs1 rs2: ireg) (**r nxor long *)
+ | Pandnl (rd rs1 rs2: ireg) (**r andn long *)
+ | Pornl (rd rs1 rs2: ireg) (**r orn long *)
+ | Pmull (rd rs1 rs2: ireg) (**r mul long (low part) *)
+ | Pslll (rd rs1 rs2: ireg) (**r shift left logical long *)
+ | Psrll (rd rs1 rs2: ireg) (**r shift right logical long *)
+ | Psral (rd rs1 rs2: ireg) (**r shift right arithmetic long *)
+ | Psrxl (rd rs1 rs2: ireg) (**r shift right arithmetic long round to 0*)
+ | Pmaddl (rd rs1 rs2: ireg) (**r multiply-add long *)
+ | Pmsubl (rd rs1 rs2: ireg) (**r multiply-add long *)
+
+ | Pfaddd (rd rs1 rs2: ireg) (**r Float addition double *)
+ | Pfaddw (rd rs1 rs2: ireg) (**r Float addition word *)
+ | Pfsbfd (rd rs1 rs2: ireg) (**r Float sub double *)
+ | Pfsbfw (rd rs1 rs2: ireg) (**r Float sub word *)
+ | Pfmuld (rd rs1 rs2: ireg) (**r Float mul double *)
+ | Pfmulw (rd rs1 rs2: ireg) (**r Float mul word *)
+ | Pfmind (rd rs1 rs2: ireg) (**r Float min double *)
+ | Pfminw (rd rs1 rs2: ireg) (**r Float min word *)
+ | Pfmaxd (rd rs1 rs2: ireg) (**r Float max double *)
+ | Pfmaxw (rd rs1 rs2: ireg) (**r Float max word *)
+ | Pfinvw (rd rs1: ireg) (**r Float invert word *)
+
+ (** Arith RRI32 *)
+ | Pcompiw (it: itest) (rd rs: ireg) (imm: int) (**r comparison imm word *)
+
+ | Paddiw (rd rs: ireg) (imm: int) (**r add imm word *)
+ | Paddxiw (shift : shift1_4) (rd rs: ireg) (imm: int) (**r add imm word *)
+ | Prevsubiw (rd rs: ireg) (imm: int) (**r subtract imm word *)
+ | Prevsubxiw (shift : shift1_4) (rd rs: ireg) (imm: int) (**r subtract imm word *)
+ | Pmuliw (rd rs: ireg) (imm: int) (**r mul imm word *)
+ | Pandiw (rd rs: ireg) (imm: int) (**r and imm word *)
+ | Pnandiw (rd rs: ireg) (imm: int) (**r nand imm word *)
+ | Poriw (rd rs: ireg) (imm: int) (**r or imm word *)
+ | Pnoriw (rd rs: ireg) (imm: int) (**r nor imm word *)
+ | Pxoriw (rd rs: ireg) (imm: int) (**r xor imm word *)
+ | Pnxoriw (rd rs: ireg) (imm: int) (**r nxor imm word *)
+ | Pandniw (rd rs: ireg) (imm: int) (**r andn imm word *)
+ | Porniw (rd rs: ireg) (imm: int) (**r orn imm word *)
+ | Psraiw (rd rs: ireg) (imm: int) (**r shift right arithmetic imm word *)
+ | Psrxiw (rd rs: ireg) (imm: int) (**r shift right arithmetic imm word round to 0*)
+ | Psrliw (rd rs: ireg) (imm: int) (**r shift right logical imm word *)
+ | Pslliw (rd rs: ireg) (imm: int) (**r shift left logical imm word *)
+ | Proriw (rd rs: ireg) (imm: int) (**r rotate right imm word *)
+ | Pmaddiw (rd rs: ireg) (imm: int) (**r multiply add imm word *)
+ | Psllil (rd rs: ireg) (imm: int) (**r shift left logical immediate long *)
+ | Psrxil (rd rs: ireg) (imm: int) (**r shift right arithmetic imm word round to 0*)
+ | Psrlil (rd rs: ireg) (imm: int) (**r shift right logical immediate long *)
+ | Psrail (rd rs: ireg) (imm: int) (**r shift right arithmetic immediate long *)
+
+ (** Arith RRI64 *)
+ | Pcompil (it: itest) (rd rs: ireg) (imm: int64) (**r comparison imm long *)
+ | Paddil (rd rs: ireg) (imm: int64) (**r add immediate long *)
+ | Paddxil (shift : shift1_4) (rd rs: ireg) (imm: int64) (**r add immediate long *)
+ | Prevsubil (rd rs: ireg) (imm: int64) (**r subtract imm long *)
+ | Prevsubxil (shift : shift1_4) (rd rs: ireg) (imm: int64) (**r subtract imm long *)
+ | Pmulil (rd rs: ireg) (imm: int64) (**r add immediate long *)
+ | Pandil (rd rs: ireg) (imm: int64) (**r and immediate long *)
+ | Pnandil (rd rs: ireg) (imm: int64) (**r and immediate long *)
+ | Poril (rd rs: ireg) (imm: int64) (**r or immediate long *)
+ | Pnoril (rd rs: ireg) (imm: int64) (**r and immediate long *)
+ | Pxoril (rd rs: ireg) (imm: int64) (**r xor immediate long *)
+ | Pnxoril (rd rs: ireg) (imm: int64) (**r xor immediate long *)
+ | Pandnil (rd rs: ireg) (imm: int64) (**r andn long *)
+ | Pornil (rd rs: ireg) (imm: int64) (**r orn long *)
+ | Pmaddil (rd rs: ireg) (imm: int64) (**r multiply add imm long *)
+ | Pcmove (bt: btest) (rcond rd rs : ireg) (** conditional move *)
+ | Pcmoveu (bt: btest) (rcond rd rs : ireg) (** conditional move, unsigned semantics *)
+ | Pcmoveiw (bt: btest) (rcond rd : ireg) (imm: int) (** conditional move *)
+ | Pcmoveuiw (bt: btest) (rcond rd : ireg) (imm: int) (** conditional move, unsigned semantics *)
+ | Pcmoveil (bt: btest) (rcond rd : ireg) (imm: int64) (** conditional move *)
+ | Pcmoveuil (bt: btest) (rcond rd : ireg) (imm: int64) (** conditional move, unsigned semantics *)
+.
+
+(** Correspondance between Asmblock and Asm *)
+
+Definition control_to_instruction (c: control) :=
+ match c with
+ | PExpand (Asmvliw.Pbuiltin ef args res) => Pbuiltin ef args res
+ | PCtlFlow Asmvliw.Pret => Pret
+ | PCtlFlow (Asmvliw.Pcall l) => Pcall l
+ | PCtlFlow (Asmvliw.Picall r) => Picall r
+ | PCtlFlow (Asmvliw.Pgoto l) => Pgoto l
+ | PCtlFlow (Asmvliw.Pigoto l) => Pigoto l
+ | PCtlFlow (Asmvliw.Pj_l l) => Pj_l l
+ | PCtlFlow (Asmvliw.Pcb bt r l) => Pcb bt r l
+ | PCtlFlow (Asmvliw.Pcbu bt r l) => Pcbu bt r l
+ | PCtlFlow (Asmvliw.Pjumptable r label) => Pjumptable r label
+ end.
+
+Definition basic_to_instruction (b: basic) :=
+ match b with
+ (** Special basics *)
+ | Asmvliw.Pget rd rs => Pget rd rs
+ | Asmvliw.Pset rd rs => Pset rd rs
+ | Asmvliw.Pnop => Pnop
+ | Asmvliw.Pallocframe sz pos => Pallocframe sz pos
+ | Asmvliw.Pfreeframe sz pos => Pfreeframe sz pos
+
+ (** PArith basics *)
+ (* R *)
+ | PArithR (Asmvliw.Ploadsymbol id ofs) r => Ploadsymbol r id ofs
+
+ (* RR *)
+ | PArithRR Asmvliw.Pmv rd rs => Pmv rd rs
+ | PArithRR Asmvliw.Pnegw rd rs => Pnegw rd rs
+ | PArithRR Asmvliw.Pnegl rd rs => Pnegl rd rs
+ | PArithRR Asmvliw.Pcvtl2w rd rs => Pcvtl2w rd rs
+ | PArithRR Asmvliw.Psxwd rd rs => Psxwd rd rs
+ | PArithRR Asmvliw.Pzxwd rd rs => Pzxwd rd rs
+ | PArithRR (Asmvliw.Pextfz stop start) rd rs => Pextfz rd rs stop start
+ | PArithRR (Asmvliw.Pextfs stop start) rd rs => Pextfs rd rs stop start
+ | PArithRR (Asmvliw.Pextfzl stop start) rd rs => Pextfzl rd rs stop start
+ | PArithRR (Asmvliw.Pextfsl stop start) rd rs => Pextfsl rd rs stop start
+ | PArithRR Asmvliw.Pfabsd rd rs => Pfabsd rd rs
+ | PArithRR Asmvliw.Pfabsw rd rs => Pfabsw rd rs
+ | PArithRR Asmvliw.Pfnegd rd rs => Pfnegd rd rs
+ | PArithRR Asmvliw.Pfnegw rd rs => Pfnegw rd rs
+ | PArithRR Asmvliw.Pfinvw rd rs => Pfinvw rd rs
+ | PArithRR Asmvliw.Pfnarrowdw rd rs => Pfnarrowdw rd rs
+ | PArithRR Asmvliw.Pfwidenlwd rd rs => Pfwidenlwd rd rs
+ | PArithRR Asmvliw.Pfloatuwrnsz rd rs => Pfloatuwrnsz rd rs
+ | PArithRR Asmvliw.Pfloatwrnsz rd rs => Pfloatwrnsz rd rs
+ | PArithRR Asmvliw.Pfloatudrnsz rd rs => Pfloatudrnsz rd rs
+ | PArithRR Asmvliw.Pfloatdrnsz rd rs => Pfloatdrnsz rd rs
+ | PArithRR Asmvliw.Pfixedwrzz rd rs => Pfixedwrzz rd rs
+ | PArithRR Asmvliw.Pfixeduwrzz rd rs => Pfixeduwrzz rd rs
+ | PArithRR Asmvliw.Pfixeddrzz rd rs => Pfixeddrzz rd rs
+ | PArithRR Asmvliw.Pfixedudrzz rd rs => Pfixedudrzz rd rs
+ | PArithRR Asmvliw.Pfixeddrzz_i32 rd rs => Pfixeddrzz_i32 rd rs
+ | PArithRR Asmvliw.Pfixedudrzz_i32 rd rs => Pfixedudrzz_i32 rd rs
+
+ (* RI32 *)
+ | PArithRI32 Asmvliw.Pmake rd imm => Pmake rd imm
+
+ (* RI64 *)
+ | PArithRI64 Asmvliw.Pmakel rd imm => Pmakel rd imm
+
+ (* RF32 *)
+ | PArithRF32 Asmvliw.Pmakefs rd imm => Pmakefs rd imm
+
+ (* RF64 *)
+ | PArithRF64 Asmvliw.Pmakef rd imm => Pmakef rd imm
+
+ (* RRR *)
+ | PArithRRR (Asmvliw.Pcompw it) rd rs1 rs2 => Pcompw it rd rs1 rs2
+ | PArithRRR (Asmvliw.Pcompl it) rd rs1 rs2 => Pcompl it rd rs1 rs2
+ | PArithRRR (Asmvliw.Pfcompw ft) rd rs1 rs2 => Pfcompw ft rd rs1 rs2
+ | PArithRRR (Asmvliw.Pfcompl ft) rd rs1 rs2 => Pfcompl ft rd rs1 rs2
+ | PArithRRR Asmvliw.Paddw rd rs1 rs2 => Paddw rd rs1 rs2
+ | PArithRRR (Asmvliw.Paddxw shift) rd rs1 rs2 => Paddxw shift rd rs1 rs2
+ | PArithRRR Asmvliw.Psubw rd rs1 rs2 => Psubw rd rs1 rs2
+ | PArithRRR (Asmvliw.Prevsubxw shift) rd rs1 rs2 => Prevsubxw shift rd rs1 rs2
+ | PArithRRR Asmvliw.Pmulw rd rs1 rs2 => Pmulw rd rs1 rs2
+ | PArithRRR Asmvliw.Pandw rd rs1 rs2 => Pandw rd rs1 rs2
+ | PArithRRR Asmvliw.Pnandw rd rs1 rs2 => Pnandw rd rs1 rs2
+ | PArithRRR Asmvliw.Porw rd rs1 rs2 => Porw rd rs1 rs2
+ | PArithRRR Asmvliw.Pnorw rd rs1 rs2 => Pnorw rd rs1 rs2
+ | PArithRRR Asmvliw.Pxorw rd rs1 rs2 => Pxorw rd rs1 rs2
+ | PArithRRR Asmvliw.Pnxorw rd rs1 rs2 => Pnxorw rd rs1 rs2
+ | PArithRRR Asmvliw.Pandnw rd rs1 rs2 => Pandnw rd rs1 rs2
+ | PArithRRR Asmvliw.Pornw rd rs1 rs2 => Pornw rd rs1 rs2
+ | PArithRRR Asmvliw.Psraw rd rs1 rs2 => Psraw rd rs1 rs2
+ | PArithRRR Asmvliw.Psrxw rd rs1 rs2 => Psrxw rd rs1 rs2
+ | PArithRRR Asmvliw.Psrlw rd rs1 rs2 => Psrlw rd rs1 rs2
+ | PArithRRR Asmvliw.Psllw rd rs1 rs2 => Psllw rd rs1 rs2
+
+ | PArithRRR Asmvliw.Paddl rd rs1 rs2 => Paddl rd rs1 rs2
+ | PArithRRR (Asmvliw.Paddxl shift) rd rs1 rs2 => Paddxl shift rd rs1 rs2
+ | PArithRRR Asmvliw.Psubl rd rs1 rs2 => Psubl rd rs1 rs2
+ | PArithRRR (Asmvliw.Prevsubxl shift) rd rs1 rs2 => Prevsubxl shift rd rs1 rs2
+ | PArithRRR Asmvliw.Pandl rd rs1 rs2 => Pandl rd rs1 rs2
+ | PArithRRR Asmvliw.Pnandl rd rs1 rs2 => Pnandl rd rs1 rs2
+ | PArithRRR Asmvliw.Porl rd rs1 rs2 => Porl rd rs1 rs2
+ | PArithRRR Asmvliw.Pnorl rd rs1 rs2 => Pnorl rd rs1 rs2
+ | PArithRRR Asmvliw.Pxorl rd rs1 rs2 => Pxorl rd rs1 rs2
+ | PArithRRR Asmvliw.Pnxorl rd rs1 rs2 => Pnxorl rd rs1 rs2
+ | PArithRRR Asmvliw.Pandnl rd rs1 rs2 => Pandnl rd rs1 rs2
+ | PArithRRR Asmvliw.Pornl rd rs1 rs2 => Pornl rd rs1 rs2
+ | PArithRRR Asmvliw.Pmull rd rs1 rs2 => Pmull rd rs1 rs2
+ | PArithRRR Asmvliw.Pslll rd rs1 rs2 => Pslll rd rs1 rs2
+ | PArithRRR Asmvliw.Psrll rd rs1 rs2 => Psrll rd rs1 rs2
+ | PArithRRR Asmvliw.Psral rd rs1 rs2 => Psral rd rs1 rs2
+ | PArithRRR Asmvliw.Psrxl rd rs1 rs2 => Psrxl rd rs1 rs2
+
+ | PArithRRR Asmvliw.Pfaddd rd rs1 rs2 => Pfaddd rd rs1 rs2
+ | PArithRRR Asmvliw.Pfaddw rd rs1 rs2 => Pfaddw rd rs1 rs2
+ | PArithRRR Asmvliw.Pfsbfd rd rs1 rs2 => Pfsbfd rd rs1 rs2
+ | PArithRRR Asmvliw.Pfsbfw rd rs1 rs2 => Pfsbfw rd rs1 rs2
+ | PArithRRR Asmvliw.Pfmuld rd rs1 rs2 => Pfmuld rd rs1 rs2
+ | PArithRRR Asmvliw.Pfmulw rd rs1 rs2 => Pfmulw rd rs1 rs2
+ | PArithRRR Asmvliw.Pfmind rd rs1 rs2 => Pfmind rd rs1 rs2
+ | PArithRRR Asmvliw.Pfminw rd rs1 rs2 => Pfminw rd rs1 rs2
+ | PArithRRR Asmvliw.Pfmaxd rd rs1 rs2 => Pfmaxd rd rs1 rs2
+ | PArithRRR Asmvliw.Pfmaxw rd rs1 rs2 => Pfmaxw rd rs1 rs2
+
+ (* RRI32 *)
+ | PArithRRI32 (Asmvliw.Pcompiw it) rd rs imm => Pcompiw it rd rs imm
+ | PArithRRI32 Asmvliw.Paddiw rd rs imm => Paddiw rd rs imm
+ | PArithRRI32 (Asmvliw.Paddxiw shift) rd rs imm => Paddxiw shift rd rs imm
+ | PArithRRI32 Asmvliw.Prevsubiw rd rs imm => Prevsubiw rd rs imm
+ | PArithRRI32 (Asmvliw.Prevsubxiw shift) rd rs imm => Prevsubxiw shift rd rs imm
+ | PArithRRI32 Asmvliw.Pmuliw rd rs imm => Pmuliw rd rs imm
+ | PArithRRI32 Asmvliw.Pandiw rd rs imm => Pandiw rd rs imm
+ | PArithRRI32 Asmvliw.Pnandiw rd rs imm => Pnandiw rd rs imm
+ | PArithRRI32 Asmvliw.Poriw rd rs imm => Poriw rd rs imm
+ | PArithRRI32 Asmvliw.Pnoriw rd rs imm => Pnoriw rd rs imm
+ | PArithRRI32 Asmvliw.Pxoriw rd rs imm => Pxoriw rd rs imm
+ | PArithRRI32 Asmvliw.Pnxoriw rd rs imm => Pnxoriw rd rs imm
+ | PArithRRI32 Asmvliw.Pandniw rd rs imm => Pandniw rd rs imm
+ | PArithRRI32 Asmvliw.Porniw rd rs imm => Porniw rd rs imm
+ | PArithRRI32 Asmvliw.Psraiw rd rs imm => Psraiw rd rs imm
+ | PArithRRI32 Asmvliw.Psrxiw rd rs imm => Psrxiw rd rs imm
+ | PArithRRI32 Asmvliw.Psrliw rd rs imm => Psrliw rd rs imm
+ | PArithRRI32 Asmvliw.Pslliw rd rs imm => Pslliw rd rs imm
+ | PArithRRI32 Asmvliw.Proriw rd rs imm => Proriw rd rs imm
+ | PArithRRI32 Asmvliw.Psllil rd rs imm => Psllil rd rs imm
+ | PArithRRI32 Asmvliw.Psrlil rd rs imm => Psrlil rd rs imm
+ | PArithRRI32 Asmvliw.Psrxil rd rs imm => Psrxil rd rs imm
+ | PArithRRI32 Asmvliw.Psrail rd rs imm => Psrail rd rs imm
+
+ (* RRI64 *)
+ | PArithRRI64 (Asmvliw.Pcompil it) rd rs imm => Pcompil it rd rs imm
+ | PArithRRI64 Asmvliw.Paddil rd rs imm => Paddil rd rs imm
+ | PArithRRI64 (Asmvliw.Paddxil shift) rd rs imm => Paddxil shift rd rs imm
+ | PArithRRI64 Asmvliw.Prevsubil rd rs imm => Prevsubil rd rs imm
+ | PArithRRI64 (Asmvliw.Prevsubxil shift) rd rs imm => Prevsubxil shift rd rs imm
+ | PArithRRI64 Asmvliw.Pmulil rd rs imm => Pmulil rd rs imm
+ | PArithRRI64 Asmvliw.Pandil rd rs imm => Pandil rd rs imm
+ | PArithRRI64 Asmvliw.Pnandil rd rs imm => Pnandil rd rs imm
+ | PArithRRI64 Asmvliw.Poril rd rs imm => Poril rd rs imm
+ | PArithRRI64 Asmvliw.Pnoril rd rs imm => Pnoril rd rs imm
+ | PArithRRI64 Asmvliw.Pxoril rd rs imm => Pxoril rd rs imm
+ | PArithRRI64 Asmvliw.Pnxoril rd rs imm => Pnxoril rd rs imm
+ | PArithRRI64 Asmvliw.Pandnil rd rs imm => Pandnil rd rs imm
+ | PArithRRI64 Asmvliw.Pornil rd rs imm => Pornil rd rs imm
+
+ (** ARRR *)
+ | PArithARRR Asmvliw.Pmaddw rd rs1 rs2 => Pmaddw rd rs1 rs2
+ | PArithARRR Asmvliw.Pmaddl rd rs1 rs2 => Pmaddl rd rs1 rs2
+ | PArithARRR Asmvliw.Pmsubw rd rs1 rs2 => Pmsubw rd rs1 rs2
+ | PArithARRR Asmvliw.Pmsubl rd rs1 rs2 => Pmsubl rd rs1 rs2
+ | PArithARRR Asmvliw.Pfmaddfw rd rs1 rs2 => Pfmaddfw rd rs1 rs2
+ | PArithARRR Asmvliw.Pfmaddfl rd rs1 rs2 => Pfmaddfl rd rs1 rs2
+ | PArithARRR Asmvliw.Pfmsubfw rd rs1 rs2 => Pfmsubfw rd rs1 rs2
+ | PArithARRR Asmvliw.Pfmsubfl rd rs1 rs2 => Pfmsubfl rd rs1 rs2
+ | PArithARRR (Asmvliw.Pcmove cond) rd rs1 rs2=> Pcmove cond rd rs1 rs2
+ | PArithARRR (Asmvliw.Pcmoveu cond) rd rs1 rs2=> Pcmoveu cond rd rs1 rs2
+
+ (** ARR *)
+ | PArithARR (Asmvliw.Pinsf stop start) rd rs => Pinsf rd rs stop start
+ | PArithARR (Asmvliw.Pinsfl stop start) rd rs => Pinsfl rd rs stop start
+
+ (** ARRI32 *)
+ | PArithARRI32 Asmvliw.Pmaddiw rd rs1 imm => Pmaddiw rd rs1 imm
+ | PArithARRI32 (Asmvliw.Pcmoveiw cond) rd rs1 imm => Pcmoveiw cond rd rs1 imm
+ | PArithARRI32 (Asmvliw.Pcmoveuiw cond) rd rs1 imm => Pcmoveuiw cond rd rs1 imm
+
+ (** ARRI64 *)
+ | PArithARRI64 Asmvliw.Pmaddil rd rs1 imm => Pmaddil rd rs1 imm
+ | PArithARRI64 (Asmvliw.Pcmoveil cond) rd rs1 imm => Pcmoveil cond rd rs1 imm
+ | PArithARRI64 (Asmvliw.Pcmoveuil cond) rd rs1 imm => Pcmoveuil cond rd rs1 imm
+ (** Load *)
+ | PLoadRRO Asmvliw.Plb rd ra ofs => Plb rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Plbu rd ra ofs => Plbu rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Plh rd ra ofs => Plh rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Plhu rd ra ofs => Plhu rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Plw rd ra ofs => Plw rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Plw_a rd ra ofs => Plw_a rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Pld rd ra ofs => Pld rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Pld_a rd ra ofs => Pld_a rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Pfls rd ra ofs => Pfls rd ra (AOff ofs)
+ | PLoadRRO Asmvliw.Pfld rd ra ofs => Pfld rd ra (AOff ofs)
+
+ | PLoadQRRO qrs ra ofs => Plq qrs ra (AOff ofs)
+ | PLoadORRO qrs ra ofs => Plo qrs ra (AOff ofs)
+
+ | PLoadRRR Asmvliw.Plb rd ra ro => Plb rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Plbu rd ra ro => Plbu rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Plh rd ra ro => Plh rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Plhu rd ra ro => Plhu rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Plw rd ra ro => Plw rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Plw_a rd ra ro => Plw_a rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Pld rd ra ro => Pld rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Pld_a rd ra ro => Pld_a rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Pfls rd ra ro => Pfls rd ra (AReg ro)
+ | PLoadRRR Asmvliw.Pfld rd ra ro => Pfld rd ra (AReg ro)
+
+ | PLoadRRRXS Asmvliw.Plb rd ra ro => Plb rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Plbu rd ra ro => Plbu rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Plh rd ra ro => Plh rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Plhu rd ra ro => Plhu rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Plw rd ra ro => Plw rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Plw_a rd ra ro => Plw_a rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Pld rd ra ro => Pld rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Pld_a rd ra ro => Pld_a rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Pfls rd ra ro => Pfls rd ra (ARegXS ro)
+ | PLoadRRRXS Asmvliw.Pfld rd ra ro => Pfld rd ra (ARegXS ro)
+
+ (** Store *)
+ | PStoreRRO Asmvliw.Psb rd ra ofs => Psb rd ra (AOff ofs)
+ | PStoreRRO Asmvliw.Psh rd ra ofs => Psh rd ra (AOff ofs)
+ | PStoreRRO Asmvliw.Psw rd ra ofs => Psw rd ra (AOff ofs)
+ | PStoreRRO Asmvliw.Psw_a rd ra ofs => Psw_a rd ra (AOff ofs)
+ | PStoreRRO Asmvliw.Psd rd ra ofs => Psd rd ra (AOff ofs)
+ | PStoreRRO Asmvliw.Psd_a rd ra ofs => Psd_a rd ra (AOff ofs)
+ | PStoreRRO Asmvliw.Pfss rd ra ofs => Pfss rd ra (AOff ofs)
+ | PStoreRRO Asmvliw.Pfsd rd ra ofs => Pfsd rd ra (AOff ofs)
+
+ | PStoreRRR Asmvliw.Psb rd ra ro => Psb rd ra (AReg ro)
+ | PStoreRRR Asmvliw.Psh rd ra ro => Psh rd ra (AReg ro)
+ | PStoreRRR Asmvliw.Psw rd ra ro => Psw rd ra (AReg ro)
+ | PStoreRRR Asmvliw.Psw_a rd ra ro => Psw_a rd ra (AReg ro)
+ | PStoreRRR Asmvliw.Psd rd ra ro => Psd rd ra (AReg ro)
+ | PStoreRRR Asmvliw.Psd_a rd ra ro => Psd_a rd ra (AReg ro)
+ | PStoreRRR Asmvliw.Pfss rd ra ro => Pfss rd ra (AReg ro)
+ | PStoreRRR Asmvliw.Pfsd rd ra ro => Pfsd rd ra (AReg ro)
+
+ | PStoreRRRXS Asmvliw.Psb rd ra ro => Psb rd ra (ARegXS ro)
+ | PStoreRRRXS Asmvliw.Psh rd ra ro => Psh rd ra (ARegXS ro)
+ | PStoreRRRXS Asmvliw.Psw rd ra ro => Psw rd ra (ARegXS ro)
+ | PStoreRRRXS Asmvliw.Psw_a rd ra ro => Psw_a rd ra (ARegXS ro)
+ | PStoreRRRXS Asmvliw.Psd rd ra ro => Psd rd ra (ARegXS ro)
+ | PStoreRRRXS Asmvliw.Psd_a rd ra ro => Psd_a rd ra (ARegXS ro)
+ | PStoreRRRXS Asmvliw.Pfss rd ra ro => Pfss rd ra (ARegXS ro)
+ | PStoreRRRXS Asmvliw.Pfsd rd ra ro => Pfsd rd ra (ARegXS ro)
+
+ | PStoreQRRO qrs ra ofs => Psq qrs ra (AOff ofs)
+ | PStoreORRO qrs ra ofs => Pso qrs ra (AOff ofs)
+ end.
+
+Section RELSEM.
+
+Definition code := list instruction.
+
+Fixpoint unfold_label (ll: list label) :=
+ match ll with
+ | nil => nil
+ | l :: ll => Plabel l :: unfold_label ll
+ end.
+
+Fixpoint unfold_body (lb: list basic) :=
+ match lb with
+ | nil => nil
+ | b :: lb => basic_to_instruction b :: unfold_body lb
+ end.
+
+Definition unfold_exit (oc: option control) :=
+ match oc with
+ | None => nil
+ | Some c => control_to_instruction c :: nil
+ end.
+
+Definition unfold_bblock (b: bblock) := unfold_label (header b) ++
+ (match (body b), (exit b) with
+ | (((Asmvliw.Pfreeframe _ _ | Asmvliw.Pallocframe _ _)::nil) as bo), None =>
+ unfold_body bo
+ | bo, ex => unfold_body bo ++ unfold_exit ex ++ Psemi :: nil
+ end).
+
+Fixpoint unfold (lb: bblocks) :=
+ match lb with
+ | nil => nil
+ | b :: lb => (unfold_bblock b) ++ unfold lb
+ end.
+
+Record function : Type := mkfunction { fn_sig: signature; fn_blocks: bblocks; fn_code: code;
+ correct: unfold fn_blocks = fn_code }.
+
+Definition fundef := AST.fundef function.
+Definition program := AST.program fundef unit.
+Definition genv := Genv.t fundef unit.
+
+Definition function_proj (f: function) := Asmvliw.mkfunction (fn_sig f) (fn_blocks f).
+
+Definition fundef_proj (fu: fundef) : Asmvliw.fundef :=
+ match fu with
+ | Internal f => Internal (function_proj f)
+ | External ef => External ef
+ end.
+
+Definition globdef_proj (gd: globdef fundef unit) : globdef Asmvliw.fundef unit :=
+ match gd with
+ | Gfun f => Gfun (fundef_proj f)
+ | Gvar gu => Gvar gu
+ end.
+
+Program Definition genv_trans (ge: genv) : Asmvliw.genv :=
+ {| Genv.genv_public := Genv.genv_public ge;
+ Genv.genv_symb := Genv.genv_symb ge;
+ 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.
+Qed. Next Obligation.
+ destruct ge; simpl in *.
+ rewrite PTree.gmap1 in H.
+ destruct (genv_defs ! b) eqn:GEN.
+ - eauto.
+ - discriminate.
+Qed. Next Obligation.
+ destruct ge; simpl in *.
+ eauto.
+Qed.
+
+Fixpoint prog_defs_proj (l: list (ident * globdef fundef unit))
+ : list (ident * globdef Asmvliw.fundef unit) :=
+ match l with
+ | nil => nil
+ | (i, gd) :: l => (i, globdef_proj gd) :: prog_defs_proj l
+ end.
+
+Definition program_proj (p: program) : Asmvliw.program :=
+ {| prog_defs := prog_defs_proj (prog_defs p);
+ prog_public := prog_public p;
+ prog_main := prog_main p
+ |}.
+
+End RELSEM.
+
+Definition semantics (p: program) := Asmvliw.semantics (program_proj p).
+
+(** Determinacy of the [Asm] semantics. *)
+
+Lemma semantics_determinate: forall p, determinate (semantics p).
+Proof.
+ intros. apply semantics_determinate.
+Qed.
+
+(** transf_program *)
+
+Program Definition transf_function (f: Asmvliw.function) : function :=
+ {| fn_sig := Asmvliw.fn_sig f; fn_blocks := Asmvliw.fn_blocks f;
+ fn_code := unfold (Asmvliw.fn_blocks f) |}.
+
+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.
+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.
+ rewrite transf_function_proj. auto.
+Qed.
+
+Definition transf_program : Asmvliw.program -> program := transform_program transf_fundef.
+
+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 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.
+ 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.
+
+Definition match_prog (p: Asmvliw.program) (tp: program) :=
+ match_program (fun _ f tf => tf = transf_fundef f) eq p tp.
+
+Lemma transf_program_match:
+ forall p tp, transf_program p = tp -> match_prog p tp.
+Proof.
+ intros. rewrite <- H. eapply match_transform_program; eauto.
+Qed.
+
+Lemma cons_extract {A: Type} : forall (l: list A) a b, a = b -> a::l = b::l.
+Proof.
+ intros. congruence.
+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: Asmvliw.program.
+Variable tprog: program.
+Hypothesis TRANSF: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Definition match_states (s1 s2: state) := s1 = s2.
+
+Lemma symbols_preserved:
+ forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof (Genv.find_symbol_match TRANSF).
+
+Lemma senv_preserved:
+ Senv.equiv ge tge.
+Proof (Genv.senv_match TRANSF).
+
+
+Theorem transf_program_correct:
+ forward_simulation (Asmvliw.semantics prog) (semantics tprog).
+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.
+ - intros. exists s1. split; auto. congruence.
+ - intros. inv H. auto.
+ - intros. exists s1'. inv H0. split; auto. congruence.
+Qed.
+
+End PRESERVATION.
diff --git a/mppa_k1c/Asmaux.v b/mppa_k1c/Asmaux.v
index 94b39f4e..891d1068 100644
--- a/mppa_k1c/Asmaux.v
+++ b/mppa_k1c/Asmaux.v
@@ -2,4 +2,4 @@ Require Import Asm.
Require Import AST.
(** Constant only needed by Asmexpandaux.ml *)
-Program Definition dummy_function := {| fn_code := nil; fn_sig := signature_main; fn_blocks := nil |}.
+Program Definition dummy_function := {| fn_code := nil; fn_sig := signature_main; fn_blocks := nil |}.
diff --git a/mppa_k1c/Asmblockgen.v b/mppa_k1c/Asmblockgen.v
index ade84e7b..bbe24fec 100644
--- a/mppa_k1c/Asmblockgen.v
+++ b/mppa_k1c/Asmblockgen.v
@@ -1111,10 +1111,12 @@ Definition transl_instr_control (f: Machblock.function) (oi: option Machblock.co
Definition fp_is_parent (before: bool) (i: Machblock.basic_inst) : bool :=
match i with
+ | MBgetstack ofs ty dst => before && negb (mreg_eq dst MFP)
| MBsetstack src ofs ty => before
| MBgetparam ofs ty dst => negb (mreg_eq dst MFP)
| MBop op args res => before && negb (mreg_eq res MFP)
- | _ => false
+ | MBload chunk addr args dst => before && negb (mreg_eq dst MFP)
+ | MBstore chunk addr args res => before
end.
(** This is the naive definition, which is not tail-recursive unlike the other backends *)
diff --git a/mppa_k1c/Asmblockgenproof.v b/mppa_k1c/Asmblockgenproof.v
index c44ef3ff..bd2dc985 100644
--- a/mppa_k1c/Asmblockgenproof.v
+++ b/mppa_k1c/Asmblockgenproof.v
@@ -1,1797 +1,1752 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(** Correctness proof for RISC-V generation: main proof. *)
-
-Require Import Coqlib Errors.
-Require Import Integers Floats AST Linking.
-Require Import Values Memory Events Globalenvs Smallstep.
-Require Import Op Locations Machblock Conventions Asmblock.
-Require Import Asmblockgen Asmblockgenproof0 Asmblockgenproof1.
-Require Import Axioms.
-
-Module MB := Machblock.
-Module AB := Asmvliw.
-
-Definition match_prog (p: Machblock.program) (tp: Asmvliw.program) :=
- match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
-
-Lemma transf_program_match:
- forall p tp, transf_program p = OK tp -> match_prog p tp.
-Proof.
- intros. eapply match_transform_partial_program; eauto.
-Qed.
-
-Section PRESERVATION.
-
-Variable prog: Machblock.program.
-Variable tprog: Asmvliw.program.
-Hypothesis TRANSF: match_prog prog tprog.
-Let ge := Genv.globalenv prog.
-Let tge := Genv.globalenv tprog.
-
-Lemma symbols_preserved:
- forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
-Proof (Genv.find_symbol_match TRANSF).
-
-Lemma senv_preserved:
- Senv.equiv ge tge.
-Proof (Genv.senv_match TRANSF).
-
-
-Lemma functions_translated:
- forall b f,
- Genv.find_funct_ptr ge b = Some f ->
- exists tf,
- Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = OK tf.
-Proof (Genv.find_funct_ptr_transf_partial TRANSF).
-
-Lemma functions_transl:
- forall fb f tf,
- Genv.find_funct_ptr ge fb = Some (Internal f) ->
- transf_function f = OK tf ->
- Genv.find_funct_ptr tge fb = Some (Internal tf).
-Proof.
- intros. exploit functions_translated; eauto. intros [tf' [A B]].
- monadInv B. rewrite H0 in EQ; inv EQ; auto.
-Qed.
-
-(** * Properties of control flow *)
-
-Lemma transf_function_no_overflow:
- forall f tf,
- transf_function f = OK tf -> size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned.
-Proof.
- intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (size_blocks x.(fn_blocks))); inv EQ0.
- omega.
-Qed.
-
-(** The following lemmas show that the translation from Mach to Asm
- preserves labels, in the sense that the following diagram commutes:
-<<
- translation
- Mach code ------------------------ Asm instr sequence
- | |
- | Mach.find_label lbl find_label lbl |
- | |
- v v
- Mach code tail ------------------- Asm instr seq tail
- translation
->>
- The proof demands many boring lemmas showing that Asm constructor
- functions do not introduce new labels.
-*)
-
-Section TRANSL_LABEL.
-
-Lemma gen_bblocks_label:
- forall hd bdy ex tbb tc,
- gen_bblocks hd bdy ex = tbb::tc ->
- header tbb = hd.
-Proof.
- intros until tc. intros GENB. unfold gen_bblocks in GENB.
- destruct (extract_ctl ex); try destruct c; try destruct i; try destruct bdy.
- all: inv GENB; simpl; auto.
-Qed.
-
-Lemma gen_bblocks_label2:
- forall hd bdy ex tbb1 tbb2,
- gen_bblocks hd bdy ex = tbb1::tbb2::nil ->
- header tbb2 = nil.
-Proof.
- intros until tbb2. intros GENB. unfold gen_bblocks in GENB.
- destruct (extract_ctl ex); try destruct c; try destruct i; try destruct bdy.
- all: inv GENB; simpl; auto.
-Qed.
-
-Lemma in_dec_transl:
- forall lbl hd,
- (if in_dec lbl hd then true else false) = (if MB.in_dec lbl hd then true else false).
-Proof.
- intros. destruct (in_dec lbl hd), (MB.in_dec lbl hd). all: tauto.
-Qed.
-
-Lemma transl_is_label:
- forall lbl bb tbb f ep tc,
- transl_block f bb ep = OK (tbb::tc) ->
- is_label lbl tbb = MB.is_label lbl bb.
-Proof.
- intros until tc. intros TLB.
- destruct tbb as [thd tbdy tex]; simpl in *.
- monadInv TLB.
- unfold is_label. simpl.
- apply gen_bblocks_label in H0. simpl in H0. subst.
- rewrite in_dec_transl. auto.
-Qed.
-
-Lemma transl_is_label_false2:
- forall lbl bb f ep tbb1 tbb2,
- transl_block f bb ep = OK (tbb1::tbb2::nil) ->
- is_label lbl tbb2 = false.
-Proof.
- intros until tbb2. intros TLB.
- destruct tbb2 as [thd tbdy tex]; simpl in *.
- monadInv TLB. apply gen_bblocks_label2 in H0. simpl in H0. subst.
- apply is_label_correct_false. simpl. auto.
-Qed.
-
-Lemma transl_is_label2:
- forall f bb ep tbb1 tbb2 lbl,
- transl_block f bb ep = OK (tbb1::tbb2::nil) ->
- is_label lbl tbb1 = MB.is_label lbl bb
- /\ is_label lbl tbb2 = false.
-Proof.
- intros. split. eapply transl_is_label; eauto. eapply transl_is_label_false2; eauto.
-Qed.
-
-Lemma transl_block_nonil:
- forall f c ep tc,
- transl_block f c ep = OK tc ->
- tc <> nil.
-Proof.
- intros. monadInv H. unfold gen_bblocks.
- destruct (extract_ctl x0); try destruct c0; try destruct x; try destruct i.
- all: discriminate.
-Qed.
-
-Lemma transl_block_limit: forall f bb ep tbb1 tbb2 tbb3 tc,
- ~transl_block f bb ep = OK (tbb1 :: tbb2 :: tbb3 :: tc).
-Proof.
- intros. intro. monadInv H.
- unfold gen_bblocks in H0.
- destruct (extract_ctl x0); try destruct x; try destruct c; try destruct i.
- all: discriminate.
-Qed.
-
-Lemma find_label_transl_false:
- forall x f lbl bb ep x',
- transl_block f bb ep = OK x ->
- MB.is_label lbl bb = false ->
- find_label lbl (x++x') = find_label lbl x'.
-Proof.
- intros until x'. intros TLB MBis; simpl; auto.
- destruct x as [|x0 x1]; simpl; auto.
- destruct x1 as [|x1 x2]; simpl; auto.
- - erewrite <- transl_is_label in MBis; eauto. rewrite MBis. auto.
- - destruct x2 as [|x2 x3]; simpl; auto.
- + erewrite <- transl_is_label in MBis; eauto. rewrite MBis.
- erewrite transl_is_label_false2; eauto.
- + apply transl_block_limit in TLB. destruct TLB.
-Qed.
-
-Lemma transl_blocks_label:
- forall lbl f c tc ep,
- transl_blocks f c ep = OK tc ->
- match MB.find_label lbl c with
- | None => find_label lbl tc = None
- | Some c' => exists tc', find_label lbl tc = Some tc' /\ transl_blocks f c' false = OK tc'
- end.
-Proof.
- induction c; simpl; intros.
- inv H. auto.
- monadInv H.
- destruct (MB.is_label lbl a) eqn:MBis.
- - destruct x as [|tbb tc]. { apply transl_block_nonil in EQ. contradiction. }
- simpl find_label. exploit transl_is_label; eauto. intros ABis. rewrite MBis in ABis.
- rewrite ABis.
- eexists. eexists. split; eauto. simpl transl_blocks.
- assert (MB.header a <> nil).
- { apply MB.is_label_correct_true in MBis.
- destruct (MB.header a). contradiction. discriminate. }
- destruct (MB.header a); try contradiction.
- rewrite EQ. simpl. rewrite EQ1. simpl. auto.
- - apply IHc in EQ1. destruct (MB.find_label lbl c).
- + destruct EQ1 as (tc' & FIND & TLBS). exists tc'; eexists; auto.
- erewrite find_label_transl_false; eauto.
- + erewrite find_label_transl_false; eauto.
-Qed.
-
-Lemma find_label_nil:
- forall bb lbl c,
- header bb = nil ->
- find_label lbl (bb::c) = find_label lbl c.
-Proof.
- intros. destruct bb as [hd bdy ex]; simpl in *. subst.
- assert (is_label lbl {| AB.header := nil; AB.body := bdy; AB.exit := ex; AB.correct := correct |} = false).
- { erewrite <- is_label_correct_false. simpl. auto. }
- rewrite H. auto.
-Qed.
-
-Lemma transl_find_label:
- forall lbl f tf,
- transf_function f = OK tf ->
- match MB.find_label lbl f.(MB.fn_code) with
- | None => find_label lbl tf.(fn_blocks) = None
- | Some c => exists tc, find_label lbl tf.(fn_blocks) = Some tc /\ transl_blocks f c false = OK tc
- end.
-Proof.
- intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (size_blocks (fn_blocks x))); inv EQ0. clear g.
- monadInv EQ. unfold make_prologue. simpl fn_blocks. repeat (rewrite find_label_nil); simpl; auto.
- eapply transl_blocks_label; eauto.
-Qed.
-
-End TRANSL_LABEL.
-
-(** A valid branch in a piece of Mach code translates to a valid ``go to''
- transition in the generated Asm code. *)
-
-Lemma find_label_goto_label:
- forall f tf lbl rs m c' b ofs,
- Genv.find_funct_ptr ge b = Some (Internal f) ->
- transf_function f = OK tf ->
- rs PC = Vptr b ofs ->
- MB.find_label lbl f.(MB.fn_code) = Some c' ->
- exists tc', exists rs',
- goto_label tf lbl rs m = Next rs' m
- /\ transl_code_at_pc ge (rs' PC) b f c' false tf tc'
- /\ forall r, r <> PC -> rs'#r = rs#r.
-Proof.
- intros. exploit (transl_find_label lbl f tf); eauto. rewrite H2.
- intros (tc & A & B).
- exploit label_pos_code_tail; eauto. instantiate (1 := 0).
- intros [pos' [P [Q R]]].
- exists tc; exists (rs#PC <- (Vptr b (Ptrofs.repr pos'))).
- split. unfold goto_label. unfold par_goto_label. rewrite P. rewrite H1. auto.
- split. rewrite Pregmap.gss. constructor; auto.
- rewrite Ptrofs.unsigned_repr. replace (pos' - 0) with pos' in Q.
- auto. omega.
- generalize (transf_function_no_overflow _ _ H0). omega.
- intros. apply Pregmap.gso; auto.
-Qed.
-
-(** Existence of return addresses *)
-
-(* NB: the hypothesis in comment on [b] is not needed in the proof !
-*)
-Lemma return_address_exists:
- forall b f (* sg ros *) c, (* b.(MB.exit) = Some (MBcall sg ros) -> *) is_tail (b :: c) f.(MB.fn_code) ->
- exists ra, return_address_offset f c ra.
-Proof.
- intros. eapply Asmblockgenproof0.return_address_exists; eauto.
-
-- intros. monadInv H0.
- destruct (zlt Ptrofs.max_unsigned (size_blocks x.(fn_blocks))); inv EQ0. monadInv EQ. simpl.
-(* rewrite transl_code'_transl_code in EQ0. *)
- exists x; exists true; split; auto. (* unfold fn_code. *)
- repeat constructor.
- - exact transf_function_no_overflow.
-Qed.
-
-(** * Proof of semantic preservation *)
-
-(** Semantic preservation is proved using simulation diagrams
- of the following form.
-<<
- st1 --------------- st2
- | |
- t| *|t
- | |
- v v
- st1'--------------- st2'
->>
- The invariant is the [match_states] predicate below, which includes:
-- The Asm code pointed by the PC register is the translation of
- the current Mach code sequence.
-- Mach register values and Asm register values agree.
-*)
-
-(** We need to show that, in the simulation diagram, we cannot
- take infinitely many Mach transitions that correspond to zero
- transitions on the Asm side. Actually, all Mach transitions
- correspond to at least one Asm transition, except the
- transition from [Machsem.Returnstate] to [Machsem.State].
- So, the following integer measure will suffice to rule out
- the unwanted behaviour. *)
-
-
-Remark preg_of_not_FP: forall r, negb (mreg_eq r MFP) = true -> IR FP <> preg_of r.
-Proof.
- intros. change (IR FP) with (preg_of MFP). red; intros.
- exploit preg_of_injective; eauto. intros; subst r; discriminate.
-Qed.
-
-Inductive match_states: Machblock.state -> Asmvliw.state -> Prop :=
- | match_states_intro:
- forall s fb sp c ep ms m m' rs f tf tc
- (STACKS: match_stack ge s)
- (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
- (MEXT: Mem.extends m m')
- (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc)
- (AG: agree ms sp rs)
- (DXP: ep = true -> rs#FP = parent_sp s),
- match_states (Machblock.State s fb sp c ms m)
- (Asmvliw.State rs m')
- | match_states_call:
- forall s fb ms m m' rs
- (STACKS: match_stack ge s)
- (MEXT: Mem.extends m m')
- (AG: agree ms (parent_sp s) rs)
- (ATPC: rs PC = Vptr fb Ptrofs.zero)
- (ATLR: rs RA = parent_ra s),
- match_states (Machblock.Callstate s fb ms m)
- (Asmvliw.State rs m')
- | match_states_return:
- forall s ms m m' rs
- (STACKS: match_stack ge s)
- (MEXT: Mem.extends m m')
- (AG: agree ms (parent_sp s) rs)
- (ATPC: rs PC = parent_ra s),
- match_states (Machblock.Returnstate s ms m)
- (Asmvliw.State rs m').
-
-Record codestate :=
- Codestate { pstate: state;
- pheader: list label;
- pbody1: list basic;
- pbody2: list basic;
- pctl: option control;
- fpok: bool;
- rem: list AB.bblock;
- cur: option bblock }.
-
-(* | Codestate: state -> list AB.bblock -> option bblock -> codestate. *)
-
-Inductive match_codestate fb: Machblock.state -> codestate -> Prop :=
- | match_codestate_intro:
- forall s sp ms m rs0 m0 f tc ep c bb tbb tbc tbi
- (STACKS: match_stack ge s)
- (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
- (MEXT: Mem.extends m m0)
- (TBC: transl_basic_code f (MB.body bb) (if MB.header bb then ep else false) = OK tbc)
- (TIC: transl_instr_control f (MB.exit bb) = OK tbi)
- (TBLS: transl_blocks f c false = OK tc)
-(* (TRANS: transl_blocks f (bb::c) ep = OK (tbb::tc)) *)
- (AG: agree ms sp rs0)
- (DXP: (if MB.header bb then ep else false) = true -> rs0#FP = parent_sp s)
- ,
- match_codestate fb (Machblock.State s fb sp (bb::c) ms m)
- {| pstate := (Asmvliw.State rs0 m0);
- pheader := (MB.header bb);
- pbody1 := tbc;
- pbody2 := (extract_basic tbi);
- pctl := extract_ctl tbi;
- fpok := ep;
- rem := tc;
- cur := Some tbb
- |}
-.
-
-Inductive match_asmstate fb: codestate -> Asmvliw.state -> Prop :=
- | match_asmstate_some:
- forall rs f tf tc m tbb ofs ep tbdy tex lhd
- (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
- (TRANSF: transf_function f = OK tf)
- (PCeq: rs PC = Vptr fb ofs)
- (TAIL: code_tail (Ptrofs.unsigned ofs) (fn_blocks tf) (tbb::tc))
-(* (HDROK: header tbb = lhd) *)
- ,
- match_asmstate fb
- {| pstate := (Asmvliw.State rs m);
- pheader := lhd;
- pbody1 := tbdy;
- pbody2 := extract_basic tex;
- pctl := extract_ctl tex;
- fpok := ep;
- rem := tc;
- cur := Some tbb |}
- (Asmvliw.State rs m)
-.
-
-Ltac exploreInst :=
- repeat match goal with
- | [ H : match ?var with | _ => _ end = _ |- _ ] => destruct var
- | [ H : OK _ = OK _ |- _ ] => monadInv H
- | [ |- context[if ?b then _ else _] ] => destruct b
- | [ |- context[match ?m with | _ => _ end] ] => destruct m
- | [ |- context[match ?m as _ return _ with | _ => _ end]] => destruct m
- | [ H : bind _ _ = OK _ |- _ ] => monadInv H
- | [ H : Error _ = OK _ |- _ ] => inversion H
- end.
-
-Lemma transl_blocks_nonil:
- forall f bb c tc ep,
- transl_blocks f (bb::c) ep = OK tc ->
- exists tbb tc', tc = tbb :: tc'.
-Proof.
- intros until ep. intros TLBS. monadInv TLBS. monadInv EQ. unfold gen_bblocks.
- destruct (extract_ctl x2).
- - destruct c0; destruct i; simpl; eauto. destruct x1; simpl; eauto.
- - destruct x1; simpl; eauto.
-Qed.
-
-Lemma no_builtin_preserved:
- forall f ex x2,
- (forall ef args res, ex <> Some (MBbuiltin ef args res)) ->
- transl_instr_control f ex = OK x2 ->
- (exists i, extract_ctl x2 = Some (PCtlFlow i))
- \/ extract_ctl x2 = None.
-Proof.
- intros until x2. intros Hbuiltin TIC.
- destruct ex.
- - destruct c.
- (* MBcall *)
- + simpl in TIC. exploreInst; simpl; eauto.
- (* MBtailcall *)
- + simpl in TIC. exploreInst; simpl; eauto.
- (* MBbuiltin *)
- + assert (H: Some (MBbuiltin e l b) <> Some (MBbuiltin e l b)).
- apply Hbuiltin. contradict H; auto.
- (* MBgoto *)
- + simpl in TIC. exploreInst; simpl; eauto.
- (* MBcond *)
- + simpl in TIC. unfold transl_cbranch in TIC. exploreInst; simpl; eauto.
- * unfold transl_opt_compuimm. exploreInst; simpl; eauto.
- * unfold transl_opt_compluimm. exploreInst; simpl; eauto.
- * unfold transl_comp_float64. exploreInst; simpl; eauto.
- * unfold transl_comp_notfloat64. exploreInst; simpl; eauto.
- * unfold transl_comp_float32. exploreInst; simpl; eauto.
- * unfold transl_comp_notfloat32. exploreInst; simpl; eauto.
- (* MBjumptable *)
- + simpl in TIC. exploreInst; simpl; eauto.
- (* MBreturn *)
- + simpl in TIC. monadInv TIC. simpl. eauto.
- - monadInv TIC. simpl; auto.
-Qed.
-
-Lemma transl_blocks_distrib:
- forall c f bb tbb tc ep,
- transl_blocks f (bb::c) ep = OK (tbb::tc)
- -> (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res))
- -> transl_block f bb (if MB.header bb then ep else false) = OK (tbb :: nil)
- /\ transl_blocks f c false = OK tc.
-Proof.
- intros until ep. intros TLBS Hbuiltin.
- destruct bb as [hd bdy ex].
- monadInv TLBS. monadInv EQ.
- exploit no_builtin_preserved; eauto. intros Hectl. destruct Hectl.
- - destruct H as [i Hectl].
- unfold gen_bblocks in H0. rewrite Hectl in H0. inv H0.
- simpl in *. unfold transl_block; simpl. rewrite EQ0. rewrite EQ. simpl.
- unfold gen_bblocks. rewrite Hectl. auto.
- - unfold gen_bblocks in H0. rewrite H in H0.
- destruct x1 as [|bi x1].
- + simpl in H0. inv H0. simpl in *. unfold transl_block; simpl. rewrite EQ0. rewrite EQ. simpl.
- unfold gen_bblocks. rewrite H. auto.
- + simpl in H0. inv H0. simpl in *. unfold transl_block; simpl. rewrite EQ0. rewrite EQ. simpl.
- unfold gen_bblocks. rewrite H. auto.
-Qed.
-
-Lemma gen_bblocks_nobuiltin:
- forall thd tbdy tex tbb,
- (tbdy <> nil \/ extract_ctl tex <> None) ->
- (forall ef args res, extract_ctl tex <> Some (PExpand (Pbuiltin ef args res))) ->
- gen_bblocks thd tbdy tex = tbb :: nil ->
- header tbb = thd
- /\ body tbb = tbdy ++ extract_basic tex
- /\ exit tbb = extract_ctl tex.
-Proof.
- intros until tbb. intros Hnonil Hnobuiltin GENB. unfold gen_bblocks in GENB.
- destruct (extract_ctl tex) eqn:ECTL.
- - destruct c.
- + destruct i; try (inv GENB; simpl; auto; fail).
- assert False. eapply Hnobuiltin. eauto. destruct H.
- + inv GENB. simpl. auto.
- - inversion Hnonil.
- + destruct tbdy as [|bi tbdy]; try (contradict H; simpl; auto; fail). inv GENB. auto.
- + contradict H; simpl; auto.
-Qed.
-
-Lemma transl_instr_basic_nonil:
- forall k f bi ep x,
- transl_instr_basic f bi ep k = OK x ->
- x <> nil.
-Proof.
- intros until x. intros TIB.
- destruct bi.
- - simpl in TIB. unfold loadind in TIB. exploreInst; try discriminate.
- - simpl in TIB. unfold storeind in TIB. exploreInst; try discriminate.
- - simpl in TIB. monadInv TIB. unfold loadind in EQ. exploreInst; try discriminate.
- - simpl in TIB. unfold transl_op in TIB. exploreInst; try discriminate.
- unfold transl_cond_op in EQ0. exploreInst; try discriminate.
- unfold transl_cond_float64. exploreInst; try discriminate.
- unfold transl_cond_notfloat64. exploreInst; try discriminate.
- unfold transl_cond_float32. exploreInst; try discriminate.
- unfold transl_cond_notfloat32. exploreInst; try discriminate.
- - simpl in TIB. unfold transl_load in TIB. exploreInst; try discriminate.
- all: monadInv TIB; unfold transl_memory_access in EQ0; unfold transl_memory_access2 in EQ0; unfold transl_memory_access2XS in EQ0; exploreInst; try discriminate.
- - simpl in TIB. unfold transl_store in TIB. exploreInst; try discriminate.
- all: monadInv TIB; unfold transl_memory_access in EQ0; unfold transl_memory_access2 in EQ0; unfold transl_memory_access2XS in EQ0; exploreInst; try discriminate.
-Qed.
-
-Lemma transl_basic_code_nonil:
- forall bdy f x ep,
- bdy <> nil ->
- transl_basic_code f bdy ep = OK x ->
- x <> nil.
-Proof.
- induction bdy as [|bi bdy].
- intros. contradict H0; auto.
- destruct bdy as [|bi2 bdy].
- - clear IHbdy. intros f x b _ TBC. simpl in TBC. eapply transl_instr_basic_nonil; eauto.
- - intros f x b Hnonil TBC. remember (bi2 :: bdy) as bdy'.
- monadInv TBC.
- assert (x0 <> nil).
- eapply IHbdy; eauto. subst bdy'. discriminate.
- eapply transl_instr_basic_nonil; eauto.
-Qed.
-
-Lemma transl_instr_control_nonil:
- forall ex f x,
- ex <> None ->
- transl_instr_control f ex = OK x ->
- extract_ctl x <> None.
-Proof.
- intros ex f x Hnonil TIC.
- destruct ex as [ex|].
- - clear Hnonil. destruct ex.
- all: try (simpl in TIC; exploreInst; discriminate).
- + simpl in TIC. unfold transl_cbranch in TIC. exploreInst; try discriminate.
- * unfold transl_opt_compuimm. exploreInst; try discriminate.
- * unfold transl_opt_compluimm. exploreInst; try discriminate.
- * unfold transl_comp_float64. exploreInst; try discriminate.
- * unfold transl_comp_notfloat64. exploreInst; try discriminate.
- * unfold transl_comp_float32. exploreInst; try discriminate.
- * unfold transl_comp_notfloat32. exploreInst; try discriminate.
- - contradict Hnonil; auto.
-Qed.
-
-Lemma transl_instr_control_nobuiltin:
- forall f ex x,
- (forall ef args res, ex <> Some (MBbuiltin ef args res)) ->
- transl_instr_control f ex = OK x ->
- (forall ef args res, extract_ctl x <> Some (PExpand (Pbuiltin ef args res))).
-Proof.
- intros until x. intros Hnobuiltin TIC. intros until res.
- unfold transl_instr_control in TIC. exploreInst.
- all: try discriminate.
- - assert False. eapply Hnobuiltin; eauto. destruct H.
- - unfold transl_cbranch in TIC. exploreInst.
- all: try discriminate.
- * unfold transl_opt_compuimm. exploreInst. all: try discriminate.
- * unfold transl_opt_compluimm. exploreInst. all: try discriminate.
- * unfold transl_comp_float64. exploreInst; try discriminate.
- * unfold transl_comp_notfloat64. exploreInst; try discriminate.
- * unfold transl_comp_float32. exploreInst; try discriminate.
- * unfold transl_comp_notfloat32. exploreInst; try discriminate.
-Qed.
-
-Theorem match_state_codestate:
- forall mbs abs s fb sp bb c ms m,
- (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
- (MB.body bb <> nil \/ MB.exit bb <> None) ->
- mbs = (Machblock.State s fb sp (bb::c) ms m) ->
- match_states mbs abs ->
- exists cs fb f tbb tc ep,
- match_codestate fb mbs cs /\ match_asmstate fb cs abs
- /\ Genv.find_funct_ptr ge fb = Some (Internal f)
- /\ transl_blocks f (bb::c) ep = OK (tbb::tc)
- /\ body tbb = pbody1 cs ++ pbody2 cs
- /\ exit tbb = pctl cs
- /\ cur cs = Some tbb /\ rem cs = tc
- /\ pstate cs = abs.
-Proof.
- intros until m. intros Hnobuiltin Hnotempty Hmbs MS. subst. inv MS.
- inv AT. clear H0. exploit transl_blocks_nonil; eauto. intros (tbb & tc' & Htc). subst.
- exploit transl_blocks_distrib; eauto. intros (TLB & TLBS). clear H2.
- monadInv TLB. exploit gen_bblocks_nobuiltin; eauto.
- { inversion Hnotempty.
- - destruct (MB.body bb) as [|bi bdy]; try (contradict H0; simpl; auto; fail).
- left. eapply transl_basic_code_nonil; eauto.
- - destruct (MB.exit bb) as [ei|]; try (contradict H0; simpl; auto; fail).
- right. eapply transl_instr_control_nonil; eauto. }
- eapply transl_instr_control_nobuiltin; eauto.
- intros (Hth & Htbdy & Htexit).
- exists {| pstate := (State rs m'); pheader := (Machblock.header bb); pbody1 := x; pbody2 := extract_basic x0;
- pctl := extract_ctl x0; fpok := ep; rem := tc'; cur := Some tbb |}, fb, f, tbb, tc', ep.
- repeat split. 1-2: econstructor; eauto.
- { destruct (MB.header bb). eauto. discriminate. } eauto.
- unfold transl_blocks. fold transl_blocks. unfold transl_block. rewrite EQ. simpl. rewrite EQ1; simpl.
- rewrite TLBS. simpl. rewrite H2.
- all: simpl; auto.
-Qed.
-
-Definition mb_remove_body (bb: MB.bblock) :=
- {| MB.header := MB.header bb; MB.body := nil; MB.exit := MB.exit bb |}.
-
-Lemma exec_straight_pnil:
- forall c rs1 m1 rs2 m2,
- exec_straight tge c rs1 m1 (Pnop::gnil) rs2 m2 ->
- exec_straight tge c rs1 m1 nil rs2 m2.
-Proof.
- intros. eapply exec_straight_trans. eapply H. econstructor; eauto.
-Qed.
-
-Lemma transl_block_nobuiltin:
- forall f bb ep tbb,
- (MB.body bb <> nil \/ MB.exit bb <> None) ->
- (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
- transl_block f bb ep = OK (tbb :: nil) ->
- exists c c',
- transl_basic_code f (MB.body bb) ep = OK c
- /\ transl_instr_control f (MB.exit bb) = OK c'
- /\ body tbb = c ++ extract_basic c'
- /\ exit tbb = extract_ctl c'.
-Proof.
- intros until tbb. intros Hnonil Hnobuiltin TLB. monadInv TLB. destruct Hnonil.
- - eexists. eexists. split; eauto. split; eauto. eapply gen_bblocks_nobuiltin; eauto.
- left. eapply transl_basic_code_nonil; eauto. eapply transl_instr_control_nobuiltin; eauto.
- - eexists. eexists. split; eauto. split; eauto. eapply gen_bblocks_nobuiltin; eauto.
- right. eapply transl_instr_control_nonil; eauto. eapply transl_instr_control_nobuiltin; eauto.
-Qed.
-
-Lemma nextblock_preserves:
- forall rs rs' bb r,
- rs' = nextblock bb rs ->
- data_preg r = true ->
- rs r = rs' r.
-Proof.
- intros. destruct r; try discriminate.
- subst. Simpl.
-(* - subst. Simpl. *)
-Qed.
-
-Lemma cons3_app {A: Type}:
- forall a b c (l: list A),
- a :: b :: c :: l = (a :: b :: c :: nil) ++ l.
-Proof.
- intros. simpl. auto.
-Qed.
-
-Lemma exec_straight_opt_body2:
- forall c rs1 m1 c' rs2 m2,
- exec_straight_opt tge c rs1 m1 c' rs2 m2 ->
- exists body,
- exec_body tge body rs1 m1 = Next rs2 m2
- /\ (basics_to_code body) ++g c' = c.
-Proof.
- intros until m2. intros EXES.
- inv EXES.
- - exists nil. split; auto.
- - eapply exec_straight_body2. auto.
-Qed.
-
-Lemma extract_basics_to_code:
- forall lb c,
- extract_basic (basics_to_code lb ++ c) = lb ++ extract_basic c.
-Proof.
- induction lb; intros; simpl; congruence.
-Qed.
-
-Lemma extract_ctl_basics_to_code:
- forall lb c,
- extract_ctl (basics_to_code lb ++ c) = extract_ctl c.
-Proof.
- induction lb; intros; simpl; congruence.
-Qed.
-
-(* Lemma goto_label_inv:
- forall fn tbb l rs m b ofs,
- rs PC = Vptr b ofs ->
- goto_label fn l rs m = goto_label fn l (nextblock tbb rs) m.
-Proof.
- intros.
- unfold goto_label. rewrite nextblock_pc. unfold Val.offset_ptr. rewrite H.
- exploreInst; auto.
- unfold nextblock. rewrite Pregmap.gss.
-
-Qed.
-
-
-Lemma exec_control_goto_label_inv:
- exec_control tge fn (Some ctl) rs m = goto_label fn l rs m ->
- exec_control tge fn (Some ctl) (nextblock tbb rs) m = goto_label fn l (nextblock tbb rs) m.
-Proof.
-Qed. *)
-
-Theorem step_simu_control:
- forall bb' fb fn s sp c ms' m' rs2 m2 E0 S'' rs1 m1 tbb tbdy2 tex cs2,
- MB.body bb' = nil ->
- (forall ef args res, MB.exit bb' <> Some (MBbuiltin ef args res)) ->
- Genv.find_funct_ptr tge fb = Some (Internal fn) ->
- pstate cs2 = (Asmvliw.State rs2 m2) ->
- pbody1 cs2 = nil -> pbody2 cs2 = tbdy2 -> pctl cs2 = tex ->
- cur cs2 = Some tbb ->
- match_codestate fb (MB.State s fb sp (bb'::c) ms' m') cs2 ->
- match_asmstate fb cs2 (Asmvliw.State rs1 m1) ->
- exit_step return_address_offset ge (MB.exit bb') (MB.State s fb sp (bb'::c) ms' m') E0 S'' ->
- (exists rs3 m3 rs4 m4,
- exec_body tge tbdy2 rs2 m2 = Next rs3 m3
- /\ exec_control_rel tge fn tex tbb rs3 m3 rs4 m4
- /\ match_states S'' (State rs4 m4)).
-Proof.
- intros until cs2. intros Hbody Hbuiltin FIND Hpstate Hpbody1 Hpbody2 Hpctl Hcur MCS MAS ESTEP.
- inv ESTEP.
- - inv MCS. inv MAS. simpl in *.
- inv Hcur. inv Hpstate.
- destruct ctl.
- + (* MBcall *)
- destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
- inv TBC. inv TIC. inv H0.
-
- assert (f0 = f) by congruence. subst f0.
- assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
- eapply transf_function_no_overflow; eauto.
- destruct s1 as [rf|fid]; simpl in H7.
- * (* Indirect call *)
- monadInv H1.
- assert (ms' rf = Vptr f' Ptrofs.zero).
- { unfold find_function_ptr in H14. destruct (ms' rf); try discriminate.
- revert H14; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. }
- assert (rs2 x = Vptr f' Ptrofs.zero).
- { exploit ireg_val; eauto. rewrite H; intros LD; inv LD; auto. }
- generalize (code_tail_next_int _ _ _ _ NOOV TAIL). intro CT1.
- remember (Ptrofs.add _ _) as ofs'.
- assert (TCA: transl_code_at_pc ge (Vptr fb ofs') fb f c false tf tc).
- { econstructor; eauto. }
- assert (f1 = f) by congruence. subst f1.
- exploit return_address_offset_correct; eauto. intros; subst ra.
-
- repeat eexists.
- rewrite H6. econstructor; eauto.
- rewrite H7. econstructor; eauto.
- econstructor; eauto.
- econstructor; eauto. eapply agree_sp_def; eauto. simpl. eapply agree_exten; eauto. intros. Simpl.
- simpl. Simpl. rewrite PCeq. rewrite Heqofs'. simpl. auto.
-
- * (* Direct call *)
- monadInv H1.
- generalize (code_tail_next_int _ _ _ _ NOOV TAIL). intro CT1.
- remember (Ptrofs.add _ _) as ofs'.
- assert (TCA: transl_code_at_pc ge (Vptr fb ofs') fb f c false tf tc).
- econstructor; eauto.
- assert (f1 = f) by congruence. subst f1.
- exploit return_address_offset_correct; eauto. intros; subst ra.
- repeat eexists.
- rewrite H6. econstructor; eauto.
- rewrite H7. econstructor; eauto.
- econstructor; eauto.
- econstructor; eauto. eapply agree_sp_def; eauto. simpl. eapply agree_exten; eauto. intros. Simpl.
- Simpl. unfold Genv.symbol_address. rewrite symbols_preserved. simpl in H14. rewrite H14. auto.
- Simpl. simpl. subst. Simpl. simpl. unfold Val.offset_ptr. rewrite PCeq. auto.
- + (* MBtailcall *)
- destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
- inv TBC. inv TIC. inv H0.
-
- assert (f0 = f) by congruence. subst f0.
- assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
- eapply transf_function_no_overflow; eauto.
- exploit Mem.loadv_extends. eauto. eexact H15. auto. simpl. intros [parent' [A B]].
- destruct s1 as [rf|fid]; simpl in H13.
- * monadInv H1.
- assert (ms' rf = Vptr f' Ptrofs.zero).
- { destruct (ms' rf); try discriminate. revert H13. predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. }
- assert (rs2 x = Vptr f' Ptrofs.zero).
- { exploit ireg_val; eauto. rewrite H; intros LD; inv LD; auto. }
-
- assert (f = f1) by congruence. subst f1. clear FIND1. clear H14.
- exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
- exploit exec_straight_body; eauto.
- { simpl. eauto. }
- intros EXEB.
- repeat eexists.
- rewrite H6. simpl extract_basic. eauto.
- rewrite H7. simpl extract_ctl. simpl. reflexivity.
- econstructor; eauto.
- { apply agree_set_other.
- - econstructor; auto with asmgen.
- + apply V.
- + intro r. destruct r; apply V; auto.
- - eauto with asmgen. }
- assert (IR x <> IR GPR12 /\ IR x <> IR GPR32 /\ IR x <> IR GPR16).
- { clear - EQ. destruct x; repeat split; try discriminate.
- all: unfold ireg_of in EQ; destruct rf; try discriminate. }
- Simpl. inv H1. inv H3. rewrite Z; auto; try discriminate.
- * monadInv H1. assert (f = f1) by congruence. subst f1. clear FIND1. clear H14.
- exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
- exploit exec_straight_body; eauto.
- simpl. eauto.
- intros EXEB.
- repeat eexists.
- rewrite H6. simpl extract_basic. eauto.
- rewrite H7. simpl extract_ctl. simpl. reflexivity.
- econstructor; eauto.
- { apply agree_set_other.
- - econstructor; auto with asmgen.
- + apply V.
- + intro r. destruct r; apply V; auto.
- - eauto with asmgen. }
- { Simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H13. auto. }
- + (* MBbuiltin (contradiction) *)
- assert (MB.exit bb' <> Some (MBbuiltin e l b)) by (apply Hbuiltin).
- rewrite <- H in H1. contradict H1; auto.
- + (* MBgoto *)
- destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
- inv TBC. inv TIC. inv H0.
-
- assert (f0 = f) by congruence. subst f0. assert (f1 = f) by congruence. subst f1. clear H11.
- remember (nextblock tbb rs2) as rs2'.
- (* inv AT. monadInv H4. *)
- exploit functions_transl. eapply FIND0. eapply TRANSF0. intros FIND'.
- assert (tf = fn) by congruence. subst tf.
- exploit find_label_goto_label.
- eauto. eauto.
- instantiate (2 := rs2').
- { subst. unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite PCeq. eauto. }
- eauto.
- intros (tc' & rs' & GOTO & AT2 & INV).
-
- eexists. eexists. repeat eexists. repeat split.
- rewrite H6. simpl extract_basic. simpl. eauto.
- rewrite H7. simpl extract_ctl. simpl. rewrite <- Heqrs2'. eauto.
- econstructor; eauto.
- rewrite Heqrs2' in INV. unfold nextblock, incrPC in INV.
- eapply agree_exten; eauto with asmgen.
- assert (forall r : preg, r <> PC -> rs' r = rs2 r).
- { intros. destruct r.
- - destruct g. all: rewrite INV; Simpl; auto.
-(* - destruct g. all: rewrite INV; Simpl; auto. *)
- - rewrite INV; Simpl; auto.
- - contradiction. }
- eauto with asmgen.
- congruence.
- + (* MBcond *)
- destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
- inv TBC. inv TIC. inv H0.
-
- * (* MBcond true *)
- assert (f0 = f) by congruence. subst f0.
- exploit eval_condition_lessdef.
- eapply preg_vals; eauto.
- all: eauto.
- intros EC.
- exploit transl_cbranch_correct_true; eauto. intros (rs' & jmp & A & B & C).
- exploit exec_straight_opt_body2. eauto. intros (bdy & EXEB & BTC).
- assert (PCeq': rs2 PC = rs' PC). { inv A; auto. erewrite <- exec_straight_pc. 2: eapply H. eauto. }
- rewrite PCeq' in PCeq.
- assert (f1 = f) by congruence. subst f1.
- exploit find_label_goto_label.
- 4: eapply H16. 1-2: eauto. instantiate (2 := (nextblock tbb rs')). rewrite nextblock_pc.
- unfold Val.offset_ptr. rewrite PCeq. eauto.
- intros (tc' & rs3 & GOTOL & TLPC & Hrs3).
- exploit functions_transl. eapply FIND1. eapply TRANSF0. intros FIND'.
- assert (tf = fn) by congruence. subst tf.
-
- repeat eexists.
- rewrite H6. rewrite <- BTC. rewrite extract_basics_to_code. simpl. rewrite app_nil_r. eauto.
- rewrite H7. rewrite <- BTC. rewrite extract_ctl_basics_to_code. simpl extract_ctl. rewrite B. eauto.
-
- econstructor; eauto.
- eapply agree_exten with rs2; eauto with asmgen.
- { intros. destruct r; try destruct g; try discriminate.
- all: rewrite Hrs3; try discriminate; unfold nextblock, incrPC; Simpl. }
- intros. discriminate.
-
- * (* MBcond false *)
- assert (f0 = f) by congruence. subst f0.
- exploit eval_condition_lessdef.
- eapply preg_vals; eauto.
- all: eauto.
- intros EC.
-
- exploit transl_cbranch_correct_false; eauto. intros (rs' & jmp & A & B & C).
- exploit exec_straight_opt_body2. eauto. intros (bdy & EXEB & BTC).
- assert (PCeq': rs2 PC = rs' PC). { inv A; auto. erewrite <- exec_straight_pc. 2: eapply H. eauto. }
- rewrite PCeq' in PCeq.
- exploit functions_transl. eapply FIND1. eapply TRANSF0. intros FIND'.
- assert (tf = fn) by congruence. subst tf.
-
- assert (NOOV: size_blocks fn.(fn_blocks) <= Ptrofs.max_unsigned).
- eapply transf_function_no_overflow; eauto.
- generalize (code_tail_next_int _ _ _ _ NOOV TAIL). intro CT1.
-
- repeat eexists.
- rewrite H6. rewrite <- BTC. rewrite extract_basics_to_code. simpl. rewrite app_nil_r. eauto.
- rewrite H7. rewrite <- BTC. rewrite extract_ctl_basics_to_code. simpl extract_ctl. rewrite B. eauto.
-
- econstructor; eauto.
- unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite PCeq. econstructor; eauto.
- eapply agree_exten with rs2; eauto with asmgen.
- { intros. destruct r; try destruct g; try discriminate.
- all: rewrite <- C; try discriminate; unfold nextblock, incrPC; Simpl. }
- intros. discriminate.
- + (* MBjumptable *)
- destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
- inv TBC. inv TIC. inv H0.
-
- assert (f0 = f) by congruence. subst f0.
- monadInv H1.
- generalize (transf_function_no_overflow _ _ TRANSF0); intro NOOV.
- assert (f1 = f) by congruence. subst f1.
- exploit find_label_goto_label. 4: eapply H16. 1-2: eauto. instantiate (2 := (nextblock tbb rs2) # GPR62 <- Vundef # GPR63 <- Vundef).
- unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite PCeq. reflexivity.
- exploit functions_transl. eapply FIND0. eapply TRANSF0. intros FIND3. assert (fn = tf) by congruence. subst fn.
-
- intros [tc' [rs' [A [B C]]]].
- exploit ireg_val; eauto. rewrite H13. intros LD; inv LD.
-
- repeat eexists.
- rewrite H6. simpl extract_basic. simpl. eauto.
- rewrite H7. simpl extract_ctl. simpl. Simpl. rewrite <- H1. unfold Mach.label in H14. unfold label. rewrite H14. eapply A.
- econstructor; eauto.
- eapply agree_undef_regs; eauto. intros. rewrite C; auto with asmgen.
- { assert (destroyed_by_jumptable = R62 :: R63 :: nil) by auto. rewrite H2 in H0. simpl in H0. inv H0.
- destruct (preg_eq r' GPR63). subst. contradiction.
- destruct (preg_eq r' GPR62). subst. contradiction.
- destruct r'; Simpl. }
- discriminate.
- + (* MBreturn *)
- destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
- inv TBC. inv TIC. inv H0.
-
- assert (f0 = f) by congruence. subst f0.
- assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
- eapply transf_function_no_overflow; eauto.
- exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
- exploit exec_straight_body; eauto.
- simpl. eauto.
- intros EXEB.
- assert (f1 = f) by congruence. subst f1.
-
- repeat eexists.
- rewrite H6. simpl extract_basic. eauto.
- rewrite H7. simpl extract_ctl. simpl. reflexivity.
- econstructor; eauto.
- unfold nextblock, incrPC. repeat apply agree_set_other; auto with asmgen.
-
- - inv MCS. inv MAS. simpl in *. subst. inv Hpstate. inv Hcur.
-(* exploit transl_blocks_distrib; eauto. (* rewrite <- H2. discriminate. *)
- intros (TLB & TLBS).
- *) destruct bb' as [hd' bdy' ex']; simpl in *. subst.
-(* unfold transl_block in TLB. simpl in TLB. unfold gen_bblocks in TLB; simpl in TLB. inv TLB. *)
- monadInv TBC. monadInv TIC. simpl in *. rewrite H5. rewrite H6.
- simpl. repeat eexists.
- econstructor. 4: instantiate (3 := false). all:eauto.
- unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite PCeq.
- assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
- eapply transf_function_no_overflow; eauto.
- assert (f = f0) by congruence. subst f0. econstructor; eauto.
- generalize (code_tail_next_int _ _ _ _ NOOV TAIL). intro CT1. eauto.
- eapply agree_exten; eauto. intros. Simpl.
- discriminate.
-Qed.
-
-Definition mb_remove_first (bb: MB.bblock) :=
- {| MB.header := MB.header bb; MB.body := tail (MB.body bb); MB.exit := MB.exit bb |}.
-
-Lemma exec_straight_body:
- forall c c' lc rs1 m1 rs2 m2,
- exec_straight tge c rs1 m1 c' rs2 m2 ->
- code_to_basics c = Some lc ->
- exists l ll,
- c = l ++ c'
- /\ code_to_basics l = Some ll
- /\ exec_body tge ll rs1 m1 = Next rs2 m2.
-Proof.
- induction c; try (intros; inv H; fail).
- intros until m2. intros EXES CTB. inv EXES.
- - exists (i1 ::g nil),(i1::nil). repeat (split; simpl; auto). rewrite H6. auto.
- - inv CTB. destruct (code_to_basics c); try discriminate. inv H0.
- eapply IHc in H7; eauto. destruct H7 as (l' & ll & Hc & CTB & EXECB). subst.
- exists (i ::g l'),(i::ll). repeat (split; simpl; auto).
- rewrite CTB. auto.
- rewrite H1. auto.
-Qed.
-
-Lemma basics_to_code_app:
- forall c l x ll,
- basics_to_code c = l ++ basics_to_code x ->
- code_to_basics l = Some ll ->
- c = ll ++ x.
-Proof.
- intros. apply (f_equal code_to_basics) in H.
- erewrite code_to_basics_dist in H; eauto. 2: eapply code_to_basics_id.
- rewrite code_to_basics_id in H. inv H. auto.
-Qed.
-
-Lemma basics_to_code_app2:
- forall i c l x ll,
- (PBasic i) :: basics_to_code c = l ++ basics_to_code x ->
- code_to_basics l = Some ll ->
- i :: c = ll ++ x.
-Proof.
- intros until ll. intros.
- exploit basics_to_code_app. instantiate (3 := (i::c)). simpl.
- all: eauto.
-Qed.
-
-Lemma step_simu_basic:
- forall bb bb' s fb sp c ms m rs1 m1 ms' m' bi cs1 tbdy bdy,
- MB.header bb = nil -> MB.body bb = bi::(bdy) -> (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
- bb' = {| MB.header := nil; MB.body := bdy; MB.exit := MB.exit bb |} ->
- basic_step ge s fb sp ms m bi ms' m' ->
- pstate cs1 = (State rs1 m1) -> pbody1 cs1 = tbdy ->
- match_codestate fb (MB.State s fb sp (bb::c) ms m) cs1 ->
- (exists rs2 m2 l cs2 tbdy',
- cs2 = {| pstate := (State rs2 m2); pheader := nil; pbody1 := tbdy'; pbody2 := pbody2 cs1;
- pctl := pctl cs1; fpok := fp_is_parent (fpok cs1) bi; rem := rem cs1; cur := cur cs1 |}
- /\ tbdy = l ++ tbdy'
- /\ exec_body tge l rs1 m1 = Next rs2 m2
- /\ match_codestate fb (MB.State s fb sp (bb'::c) ms' m') cs2).
-Proof.
- intros until bdy. intros Hheader Hbody Hnobuiltin (* Hnotempty *) Hbb' BSTEP Hpstate Hpbody1 MCS. inv MCS.
- simpl in *. inv Hpstate.
- rewrite Hbody in TBC. monadInv TBC.
- inv BSTEP.
- - (* MBgetstack *)
- simpl in EQ0.
- unfold Mach.load_stack in H.
- exploit Mem.loadv_extends; eauto. intros [v' [A B]].
- rewrite (sp_val _ _ _ AG) in A.
- exploit loadind_correct; eauto with asmgen.
- intros (rs2 & EXECS & Hrs'1 & Hrs'2).
- eapply exec_straight_body in EXECS.
- 2: eapply code_to_basics_id; eauto.
- destruct EXECS as (l & Hlbi & BTC & CTB & EXECB).
- exists rs2, m1, Hlbi.
- eexists. eexists. split. instantiate (1 := x). eauto.
- repeat (split; auto).
- eapply basics_to_code_app; eauto.
- remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
-(* assert (Hheadereq: MB.header bb' = MB.header bb). { subst. simpl. auto. }
- rewrite <- Hheadereq. *) subst.
-
- eapply match_codestate_intro; eauto. simpl. simpl in EQ. (* { destruct (MB.header bb); auto. } *)
- eapply agree_set_mreg; eauto with asmgen.
- intro Hep. simpl in Hep. inv Hep.
- - (* MBsetstack *)
- simpl in EQ0.
- unfold Mach.store_stack in H.
- assert (Val.lessdef (ms src) (rs1 (preg_of src))). { eapply preg_val; eauto. }
- exploit Mem.storev_extends; eauto. intros [m2' [A B]].
- exploit storeind_correct; eauto with asmgen.
- rewrite (sp_val _ _ _ AG) in A. eauto. intros [rs' [P Q]].
-
- eapply exec_straight_body in P.
- 2: eapply code_to_basics_id; eauto.
- destruct P as (l & ll & TBC & CTB & EXECB).
- exists rs', m2', ll.
- eexists. eexists. split. instantiate (1 := x). eauto.
- repeat (split; auto).
- eapply basics_to_code_app; eauto.
- remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
-(* assert (Hheadereq: MB.header bb' = MB.header bb). { subst. auto. }
- rewrite <- Hheadereq. *) subst.
- eapply match_codestate_intro; eauto. simpl. simpl in EQ. rewrite Hheader in EQ. auto.
-
- eapply agree_undef_regs; eauto with asmgen.
- simpl; intros. rewrite Q; auto with asmgen. rewrite Hheader in DXP. auto.
- - (* MBgetparam *)
- simpl in EQ0.
-
- assert (f0 = f) by congruence; subst f0.
- unfold Mach.load_stack in *.
- exploit Mem.loadv_extends. eauto. eexact H0. auto.
- intros [parent' [A B]]. rewrite (sp_val _ _ _ AG) in A.
- exploit lessdef_parent_sp; eauto. clear B; intros B; subst parent'.
- exploit Mem.loadv_extends. eauto. eexact H1. auto.
- intros [v' [C D]].
-
- (* Opaque loadind. *)
-(* left; eapply exec_straight_steps; eauto; intros. monadInv TR. *)
- monadInv EQ0. rewrite Hheader. rewrite Hheader in DXP.
- destruct ep eqn:EPeq.
- (* RTMP contains parent *)
- + exploit loadind_correct. eexact EQ1.
- instantiate (2 := rs1). rewrite DXP; eauto.
- intros [rs2 [P [Q R]]].
-
- eapply exec_straight_body in P.
- 2: eapply code_to_basics_id; eauto.
- destruct P as (l & ll & BTC & CTB & EXECB).
- exists rs2, m1, ll. eexists.
- eexists. split. instantiate (1 := x). eauto.
- repeat (split; auto).
- { eapply basics_to_code_app; eauto. }
- remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
- assert (Hheadereq: MB.header bb' = MB.header bb). { subst. simpl. auto. }
- (* rewrite <- Hheadereq. *)subst.
- eapply match_codestate_intro; eauto.
-
- eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto with asmgen.
- simpl; intros. rewrite R; auto with asmgen.
- apply preg_of_not_FP; auto.
-
- (* GPR11 does not contain parent *)
- + rewrite chunk_of_Tptr in A.
- exploit loadind_ptr_correct. eexact A. intros [rs2 [P [Q R]]].
- exploit loadind_correct. eexact EQ1. instantiate (2 := rs2). rewrite Q. eauto.
- intros [rs3 [S [T U]]].
-
- exploit exec_straight_trans.
- eapply P.
- eapply S.
- intros EXES.
-
- eapply exec_straight_body in EXES.
- 2: simpl. 2: erewrite code_to_basics_id; eauto.
- destruct EXES as (l & ll & BTC & CTB & EXECB).
- exists rs3, m1, ll.
- eexists. eexists. split. instantiate (1 := x). eauto.
- repeat (split; auto).
- eapply basics_to_code_app2; eauto.
- remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
- assert (Hheadereq: MB.header bb' = MB.header bb). { subst. auto. }
- subst.
- eapply match_codestate_intro; eauto.
- eapply agree_set_mreg. eapply agree_set_mreg. eauto. eauto.
- instantiate (1 := rs2#FP <- (rs3#FP)). intros.
- rewrite Pregmap.gso; auto with asmgen.
- congruence.
- intros. unfold Pregmap.set. destruct (PregEq.eq r' FP). congruence. auto with asmgen.
- simpl; intros. rewrite U; auto with asmgen.
- apply preg_of_not_FP; auto.
- - (* MBop *)
- simpl in EQ0. rewrite Hheader in DXP.
-
- assert (eval_operation tge sp op (map ms args) m' = Some v).
- rewrite <- H. apply eval_operation_preserved. exact symbols_preserved.
- exploit eval_operation_lessdef.
- eapply preg_vals; eauto.
- 2: eexact H0.
- all: eauto.
- intros [v' [A B]]. rewrite (sp_val _ _ _ AG) in A.
- exploit transl_op_correct; eauto. intros [rs2 [P [Q R]]].
-
- eapply exec_straight_body in P.
- 2: eapply code_to_basics_id; eauto.
- destruct P as (l & ll & TBC & CTB & EXECB).
- exists rs2, m1, ll.
- eexists. eexists. split. instantiate (1 := x). eauto.
- repeat (split; auto).
- eapply basics_to_code_app; eauto.
- remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
-(* assert (Hheadereq: MB.header bb' = MB.header bb). { subst. auto. }
- rewrite <- Hheadereq. *) subst.
- eapply match_codestate_intro; eauto. simpl. simpl in EQ. rewrite Hheader in EQ. auto.
- apply agree_set_undef_mreg with rs1; auto.
- apply Val.lessdef_trans with v'; auto.
- simpl; intros. destruct (andb_prop _ _ H1); clear H1.
- rewrite R; auto. apply preg_of_not_FP; auto.
-Local Transparent destroyed_by_op.
- destruct op; simpl; auto; congruence.
- - (* MBload *)
- simpl in EQ0. rewrite Hheader in DXP.
-
- assert (eval_addressing tge sp addr (map ms args) = Some a).
- rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
- exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1.
- intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A.
- exploit Mem.loadv_extends; eauto. intros [v' [C D]].
- exploit transl_load_correct; eauto.
- intros [rs2 [P [Q R]]].
-
- eapply exec_straight_body in P.
- 2: eapply code_to_basics_id; eauto.
- destruct P as (l & ll & TBC & CTB & EXECB).
- exists rs2, m1, ll.
- eexists. eexists. split. instantiate (1 := x). eauto.
- repeat (split; auto).
- eapply basics_to_code_app; eauto.
- remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
-(* assert (Hheadereq: MB.header bb' = MB.header bb). { subst. auto. }
- rewrite <- Hheadereq. *) subst.
- eapply match_codestate_intro; eauto. simpl. simpl in EQ.
-
- eapply agree_set_undef_mreg; eauto. intros; auto with asmgen.
- simpl; congruence.
-
- - (* MBstore *)
- simpl in EQ0. rewrite Hheader in DXP.
-
- assert (eval_addressing tge sp addr (map ms args) = Some a).
- rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
- exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1.
- intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A.
- assert (Val.lessdef (ms src) (rs1 (preg_of src))). eapply preg_val; eauto.
- exploit Mem.storev_extends; eauto. intros [m2' [C D]].
- exploit transl_store_correct; eauto. intros [rs2 [P Q]].
-
- eapply exec_straight_body in P.
- 2: eapply code_to_basics_id; eauto.
- destruct P as (l & ll & TBC & CTB & EXECB).
- exists rs2, m2', ll.
- eexists. eexists. split. instantiate (1 := x). eauto.
- repeat (split; auto).
- eapply basics_to_code_app; eauto.
- remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
- subst.
- eapply match_codestate_intro; eauto. simpl. simpl in EQ.
-
- eapply agree_undef_regs; eauto with asmgen.
- simpl; congruence.
-Qed.
-
-Lemma exec_body_trans:
- forall l l' rs0 m0 rs1 m1 rs2 m2,
- exec_body tge l rs0 m0 = Next rs1 m1 ->
- exec_body tge l' rs1 m1 = Next rs2 m2 ->
- exec_body tge (l++l') rs0 m0 = Next rs2 m2.
-Proof.
- induction l.
- - simpl. congruence.
- - intros until m2. intros EXEB1 EXEB2.
- inv EXEB1. destruct (exec_basic_instr _) eqn:EBI; try discriminate.
- simpl. rewrite EBI. eapply IHl; eauto.
-Qed.
-
-Definition mb_remove_header bb := {| MB.header := nil; MB.body := MB.body bb; MB.exit := MB.exit bb |}.
-
-Program Definition remove_header tbb := {| header := nil; body := body tbb; exit := exit tbb |}.
-Next Obligation.
- destruct tbb. simpl. auto.
-Qed.
-
-Inductive exec_header: codestate -> codestate -> Prop :=
- | exec_header_cons: forall cs1,
- exec_header cs1 {| pstate := pstate cs1; pheader := nil; pbody1 := pbody1 cs1; pbody2 := pbody2 cs1;
- pctl := pctl cs1; fpok := (if pheader cs1 then fpok cs1 else false); rem := rem cs1;
- (* cur := match cur cs1 with None => None | Some bcur => Some (remove_header bcur) end *)
- cur := cur cs1 |}.
-
-Lemma step_simu_header:
- forall bb s fb sp c ms m rs1 m1 cs1,
-(* (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) -> *)
- pstate cs1 = (State rs1 m1) ->
- match_codestate fb (MB.State s fb sp (bb::c) ms m) cs1 ->
- (exists cs1',
- exec_header cs1 cs1'
- /\ match_codestate fb (MB.State s fb sp (mb_remove_header bb::c) ms m) cs1').
-Proof.
- intros until cs1. intros Hpstate MCS.
- eexists. split; eauto.
- econstructor; eauto.
- inv MCS. simpl in *. inv Hpstate.
- econstructor; eauto.
-Qed.
-
-Lemma step_matchasm_header:
- forall fb cs1 cs1' s1,
- match_asmstate fb cs1 s1 ->
- exec_header cs1 cs1' ->
- match_asmstate fb cs1' s1.
-Proof.
- intros until s1. intros MAS EXH.
- inv MAS. inv EXH.
- simpl. econstructor; eauto.
-Qed.
-
-Lemma step_simu_body:
- forall bb s fb sp c ms m rs1 m1 ms' cs1 m',
- MB.header bb = nil ->
- (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
- body_step ge s fb sp (MB.body bb) ms m ms' m' ->
- pstate cs1 = (State rs1 m1) ->
- match_codestate fb (MB.State s fb sp (bb::c) ms m) cs1 ->
- (exists rs2 m2 cs2 ep,
- cs2 = {| pstate := (State rs2 m2); pheader := nil; pbody1 := nil; pbody2 := pbody2 cs1;
- pctl := pctl cs1; fpok := ep; rem := rem cs1; cur := cur cs1 |}
- /\ exec_body tge (pbody1 cs1) rs1 m1 = Next rs2 m2
- /\ match_codestate fb (MB.State s fb sp ({| MB.header := nil; MB.body := nil; MB.exit := MB.exit bb |}::c) ms' m') cs2).
-Proof.
- intros bb. destruct bb as [hd bdy ex]; simpl; auto. induction bdy as [|bi bdy].
- - intros until m'. intros Hheader Hnobuiltin BSTEP Hpstate MCS.
- inv BSTEP.
- exists rs1, m1, cs1, (fpok cs1).
- inv MCS. inv Hpstate. simpl in *. monadInv TBC. repeat (split; simpl; auto).
- econstructor; eauto.
- - intros until m'. intros Hheader Hnobuiltin BSTEP Hpstate MCS. inv BSTEP.
- rename ms' into ms''. rename m' into m''. rename rs' into ms'. rename m'0 into m'.
- exploit (step_simu_basic); eauto. simpl. eauto. simpl; auto. simpl; auto.
- intros (rs2 & m2 & l & cs2 & tbdy' & Hcs2 & Happ & EXEB & MCS').
- simpl in *.
- exploit IHbdy. auto. 2: eapply H6. 3: eapply MCS'. all: eauto. subst; eauto. simpl; auto.
- intros (rs3 & m3 & cs3 & ep & Hcs3 & EXEB' & MCS'').
- exists rs3, m3, cs3, ep.
- repeat (split; simpl; auto). subst. simpl in *. auto.
- rewrite Happ. eapply exec_body_trans; eauto. rewrite Hcs2 in EXEB'; simpl in EXEB'. auto.
-Qed.
-
-(* Lemma exec_body_straight:
- forall l rs0 m0 rs1 m1,
- l <> nil ->
- exec_body tge l rs0 m0 = Next rs1 m1 ->
- exec_straight tge l rs0 m0 nil rs1 m1.
-Proof.
- induction l as [|i1 l].
- intros. contradict H; auto.
- destruct l as [|i2 l].
- - intros until m1. intros _ EXEB. simpl in EXEB.
- destruct (exec_basic_instr _ _ _) eqn:EBI; try discriminate.
- inv EXEB. econstructor; eauto.
- - intros until m1. intros _ EXEB. simpl in EXEB. simpl in IHl.
- destruct (exec_basic_instr tge i1 rs0 m0) eqn:EBI; try discriminate.
- econstructor; eauto. eapply IHl; eauto. discriminate.
-Qed. *)
-
-Lemma exec_body_pc:
- forall l rs1 m1 rs2 m2,
- exec_body tge l rs1 m1 = Next rs2 m2 ->
- rs2 PC = rs1 PC.
-Proof.
- induction l.
- - intros. inv H. auto.
- - intros until m2. intro EXEB.
- inv EXEB. destruct (exec_basic_instr _ _ _) eqn:EBI; try discriminate.
- eapply IHl in H0. rewrite H0.
- erewrite exec_basic_instr_pc; eauto.
-Qed.
-
-Lemma exec_body_control:
- forall b rs1 m1 rs2 m2 rs3 m3 fn,
- exec_body tge (body b) rs1 m1 = Next rs2 m2 ->
- exec_control_rel tge fn (exit b) b rs2 m2 rs3 m3 ->
- exec_bblock_rel tge fn b rs1 m1 rs3 m3.
-Proof.
- intros until fn. intros EXEB EXECTL.
- econstructor; eauto. inv EXECTL.
- unfold exec_bblock. rewrite EXEB. auto.
-Qed.
-
-Definition mbsize (bb: MB.bblock) := (length (MB.body bb) + length_opt (MB.exit bb))%nat.
-
-Lemma mbsize_eqz:
- forall bb, mbsize bb = 0%nat -> MB.body bb = nil /\ MB.exit bb = None.
-Proof.
- intros. destruct bb as [hd bdy ex]; simpl in *. unfold mbsize in H.
- remember (length _) as a. remember (length_opt _) as b.
- assert (a = 0%nat) by omega. assert (b = 0%nat) by omega. subst. clear H.
- inv H0. inv H1. destruct bdy; destruct ex; auto.
- all: try discriminate.
-Qed.
-
-Lemma mbsize_neqz:
- forall bb, mbsize bb <> 0%nat -> (MB.body bb <> nil \/ MB.exit bb <> None).
-Proof.
- intros. destruct bb as [hd bdy ex]; simpl in *.
- destruct bdy; destruct ex; try (right; discriminate); try (left; discriminate).
- contradict H. unfold mbsize. simpl. auto.
-Qed.
-
-(* Alternative form of step_simulation_bblock, easier to prove *)
-Lemma step_simulation_bblock':
- forall sf f sp bb bb' bb'' rs m rs' m' s'' c S1,
- bb' = mb_remove_header bb ->
- body_step ge sf f sp (Machblock.body bb') rs m rs' m' ->
- bb'' = mb_remove_body bb' ->
- (forall ef args res, MB.exit bb'' <> Some (MBbuiltin ef args res)) ->
- exit_step return_address_offset ge (Machblock.exit bb'') (Machblock.State sf f sp (bb'' :: c) rs' m') E0 s'' ->
- match_states (Machblock.State sf f sp (bb :: c) rs m) S1 ->
- exists S2 : state, plus step tge S1 E0 S2 /\ match_states s'' S2.
-Proof.
- intros until S1. intros Hbb' BSTEP Hbb'' Hbuiltin ESTEP MS.
- destruct (mbsize bb) eqn:SIZE.
- - apply mbsize_eqz in SIZE. destruct SIZE as (Hbody & Hexit).
- destruct bb as [hd bdy ex]; simpl in *; subst.
- inv MS. inv AT. exploit transl_blocks_nonil; eauto. intros (tbb & tc' & Htc). subst. rename tc' into tc.
- monadInv H2. simpl in *. inv ESTEP. inv BSTEP.
- eexists. split. eapply plus_one.
- exploit functions_translated; eauto. intros (tf0 & FIND' & TRANSF'). monadInv TRANSF'.
- assert (x = tf) by congruence. subst x.
- eapply exec_step_internal; eauto. eapply find_bblock_tail; eauto.
- unfold exec_bblock. simpl. eauto.
- econstructor. eauto. eauto. eauto.
- unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite <- H.
- assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
- eapply transf_function_no_overflow; eauto.
- econstructor; eauto.
- generalize (code_tail_next_int _ _ _ _ NOOV H3). intro CT1. eauto.
- eapply agree_exten; eauto. intros. Simpl.
- intros. discriminate.
- - subst. exploit mbsize_neqz. { instantiate (1 := bb). rewrite SIZE. discriminate. }
- intros Hnotempty.
-
- (* initial setting *)
- exploit match_state_codestate.
- 2: eapply Hnotempty.
- all: eauto.
- intros (cs1 & fb & f0 & tbb & tc & ep & MCS & MAS & FIND & TLBS & Hbody & Hexit & Hcur & Hrem & Hpstate).
-
- (* step_simu_header part *)
- assert (exists rs1 m1, pstate cs1 = State rs1 m1). { inv MAS. simpl. eauto. }
- destruct H as (rs1 & m1 & Hpstate2). subst.
- assert (f = fb). { inv MCS. auto. } subst fb.
- exploit step_simu_header.
- 2: eapply MCS.
- all: eauto.
- intros (cs1' & EXEH & MCS2).
-
- (* step_simu_body part *)
-(* assert (MB.body bb = MB.body (mb_remove_header bb)). { destruct bb; simpl; auto. }
- rewrite H in BSTEP. clear H. *)
- assert (Hpstate': pstate cs1' = pstate cs1). { inv EXEH; auto. }
- exploit step_simu_body.
- 3: eapply BSTEP.
- 4: eapply MCS2.
- all: eauto. rewrite Hpstate'. eauto.
- intros (rs2 & m2 & cs2 & ep' & Hcs2 & EXEB & MCS').
-
- (* step_simu_control part *)
- assert (exists tf, Genv.find_funct_ptr tge f = Some (Internal tf)).
- { exploit functions_translated; eauto. intros (tf & FIND' & TRANSF'). monadInv TRANSF'. eauto. }
- destruct H as (tf & FIND').
- assert (exists tex, pbody2 cs1 = extract_basic tex /\ pctl cs1 = extract_ctl tex).
- { inv MAS. simpl in *. eauto. }
- destruct H as (tex & Hpbody2 & Hpctl).
- inv EXEH. simpl in *.
- subst. exploit step_simu_control.
- 9: eapply MCS'. all: simpl.
- 10: eapply ESTEP.
- all: simpl; eauto.
- rewrite Hpbody2. rewrite Hpctl. rewrite Hcur.
- { inv MAS; simpl in *. inv Hcur. inv Hpstate2. eapply match_asmstate_some; eauto.
- erewrite exec_body_pc; eauto. }
- intros (rs3 & m3 & rs4 & m4 & EXEB' & EXECTL' & MS').
-
- (* bringing the pieces together *)
- exploit exec_body_trans.
- eapply EXEB.
- eauto.
- intros EXEB2.
- exploit exec_body_control; eauto.
- rewrite <- Hpbody2 in EXEB2. rewrite <- Hbody in EXEB2. eauto.
- rewrite Hexit. rewrite Hpctl. eauto.
- intros EXECB. inv EXECB.
- exists (State rs4 m4).
- split; auto. eapply plus_one. rewrite Hpstate2.
- assert (exists ofs, rs1 PC = Vptr f ofs).
- { rewrite Hpstate2 in MAS. inv MAS. simpl in *. eauto. }
- destruct H0 as (ofs & Hrs1pc).
- eapply exec_step_internal; eauto.
-
- (* proving the initial find_bblock *)
- rewrite Hpstate2 in MAS. inv MAS. simpl in *.
- assert (f1 = f0) by congruence. subst f0.
- rewrite PCeq in Hrs1pc. inv Hrs1pc.
- exploit functions_translated; eauto. intros (tf1 & FIND'' & TRANS''). rewrite FIND' in FIND''.
- inv FIND''. monadInv TRANS''. rewrite TRANSF0 in EQ. inv EQ. inv Hcur.
- eapply find_bblock_tail; eauto.
-Qed.
-
-Lemma step_simulation_bblock:
- forall sf f sp bb ms m ms' m' S2 c,
- body_step ge sf f sp (Machblock.body bb) ms m ms' m' ->
- (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
- exit_step return_address_offset ge (Machblock.exit bb) (Machblock.State sf f sp (bb :: c) ms' m') E0 S2 ->
- forall S1', match_states (Machblock.State sf f sp (bb :: c) ms m) S1' ->
- exists S2' : state, plus step tge S1' E0 S2' /\ match_states S2 S2'.
-Proof.
- intros until c. intros BSTEP Hbuiltin ESTEP S1' MS.
- eapply step_simulation_bblock'; eauto.
- all: destruct bb as [hd bdy ex]; simpl in *; eauto.
- inv ESTEP.
- - econstructor. inv H; try (econstructor; eauto; fail).
- - econstructor.
-Qed.
-
-Definition measure (s: MB.state) : nat :=
- match s with
- | MB.State _ _ _ _ _ _ => 0%nat
- | MB.Callstate _ _ _ _ => 0%nat
- | MB.Returnstate _ _ _ => 1%nat
- end.
-
-Definition split (c: MB.code) :=
- match c with
- | nil => nil
- | bb::c => {| MB.header := MB.header bb; MB.body := MB.body bb; MB.exit := None |}
- :: {| MB.header := nil; MB.body := nil; MB.exit := MB.exit bb |} :: c
- end.
-
-Lemma cons_ok_eq3 {A: Type} :
- forall (x:A) y z x' y' z',
- x = x' -> y = y' -> z = z' ->
- OK (x::y::z) = OK (x'::y'::z').
-Proof.
- intros. subst. auto.
-Qed.
-
-Lemma transl_blocks_split_builtin:
- forall bb c ep f ef args res,
- MB.exit bb = Some (MBbuiltin ef args res) -> MB.body bb <> nil ->
- transl_blocks f (split (bb::c)) ep = transl_blocks f (bb::c) ep.
-Proof.
- intros until res. intros Hexit Hbody. simpl split.
- unfold transl_blocks. fold transl_blocks. unfold transl_block.
- simpl. remember (transl_basic_code _ _ _) as tbc. remember (transl_instr_control _ _) as tbi.
- remember (transl_blocks _ _ _) as tlbs.
- destruct tbc; destruct tbi; destruct tlbs.
- all: try simpl; auto.
- - simpl. rewrite Hexit in Heqtbi. simpl in Heqtbi. monadInv Heqtbi. simpl.
- unfold gen_bblocks. simpl. destruct l.
- + exploit transl_basic_code_nonil; eauto. intro. destruct H.
- + simpl. rewrite app_nil_r. apply cons_ok_eq3. all: try eapply bblock_equality. all: simpl; auto.
-Qed.
-
-Lemma transl_code_at_pc_split_builtin:
- forall rs f f0 bb c ep tf tc ef args res,
- MB.body bb <> nil -> MB.exit bb = Some (MBbuiltin ef args res) ->
- transl_code_at_pc ge (rs PC) f f0 (bb :: c) ep tf tc ->
- transl_code_at_pc ge (rs PC) f f0 (split (bb :: c)) ep tf tc.
-Proof.
- intros until res. intros Hbody Hexit AT. inv AT.
- econstructor; eauto. erewrite transl_blocks_split_builtin; eauto.
-Qed.
-
-Theorem match_states_split_builtin:
- forall sf f sp bb c rs m ef args res S1,
- MB.body bb <> nil -> MB.exit bb = Some (MBbuiltin ef args res) ->
- match_states (Machblock.State sf f sp (bb :: c) rs m) S1 ->
- match_states (Machblock.State sf f sp (split (bb::c)) rs m) S1.
-Proof.
- intros until S1. intros Hbody Hexit MS.
- inv MS.
- econstructor; eauto.
- eapply transl_code_at_pc_split_builtin; eauto.
-Qed.
-
-Lemma step_simulation_builtin:
- forall ef args res bb sf f sp c ms m t S2,
- MB.body bb = nil -> MB.exit bb = Some (MBbuiltin ef args res) ->
- exit_step return_address_offset ge (MB.exit bb) (Machblock.State sf f sp (bb :: c) ms m) t S2 ->
- forall S1', match_states (Machblock.State sf f sp (bb :: c) ms m) S1' ->
- exists S2' : state, plus step tge S1' t S2' /\ match_states S2 S2'.
-Proof.
- intros until S2. intros Hbody Hexit ESTEP S1' MS.
- inv MS. inv AT. monadInv H2. monadInv EQ.
- rewrite Hbody in EQ0. monadInv EQ0.
- rewrite Hexit in EQ. monadInv EQ.
- rewrite Hexit in ESTEP. inv ESTEP. inv H4.
-
- exploit functions_transl; eauto. intro FN.
- generalize (transf_function_no_overflow _ _ H1); intro NOOV.
- exploit builtin_args_match; eauto. intros [vargs' [P Q]].
- exploit external_call_mem_extends; eauto.
- intros [vres' [m2' [A [B [C D]]]]].
- econstructor; split. apply plus_one.
- simpl in H3.
- eapply exec_step_builtin. eauto. eauto.
- eapply find_bblock_tail; eauto.
- simpl. eauto.
- erewrite <- sp_val by eauto.
- eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
- eapply external_call_symbols_preserved; eauto. apply senv_preserved.
- eauto.
- econstructor; eauto.
- instantiate (2 := tf); instantiate (1 := x0).
- unfold nextblock, incrPC. rewrite Pregmap.gss.
- rewrite set_res_other. rewrite undef_regs_other_2. rewrite Pregmap.gso by congruence.
- rewrite <- H. simpl. econstructor; eauto.
- eapply code_tail_next_int; eauto.
- rewrite preg_notin_charact. intros. auto with asmgen.
- auto with asmgen.
- apply agree_nextblock. eapply agree_set_res; auto.
- eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto.
- apply Pregmap.gso; auto with asmgen.
- congruence.
-Qed.
-
-Lemma next_sep:
- forall rs m rs' m', rs = rs' -> m = m' -> Next rs m = Next rs' m'.
-Proof.
- congruence.
-Qed.
-
-Theorem step_simulation:
- forall S1 t S2, MB.step return_address_offset ge S1 t S2 ->
- forall S1' (MS: match_states S1 S1'),
- (exists S2', plus step tge S1' t S2' /\ match_states S2 S2')
- \/ (measure S2 < measure S1 /\ t = E0 /\ match_states S2 S1')%nat.
-Proof.
- induction 1; intros.
-
-- (* bblock *)
- left. destruct (Machblock.exit bb) eqn:MBE; try destruct c0.
- all: try(inversion H0; subst; inv H2; eapply step_simulation_bblock;
- try (rewrite MBE; try discriminate); eauto).
- + (* MBbuiltin *)
- destruct (MB.body bb) eqn:MBB.
- * inv H. eapply step_simulation_builtin; eauto. rewrite MBE. eauto.
- * eapply match_states_split_builtin in MS; eauto.
- 2: rewrite MBB; discriminate.
- simpl split in MS.
- rewrite <- MBB in H.
- remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb1.
- assert (MB.body bb = MB.body bb1). { subst. simpl. auto. }
- rewrite H1 in H. subst.
- exploit step_simulation_bblock. eapply H.
- discriminate.
- simpl. constructor.
- eauto.
- intros (S2' & PLUS1 & MS').
- rewrite MBE in MS'.
- assert (exit_step return_address_offset ge (Some (MBbuiltin e l b))
- (MB.State sf f sp ({| MB.header := nil; MB.body := nil; MB.exit := Some (MBbuiltin e l b) |}::c)
- rs' m') t s').
- { inv H0. inv H3. econstructor. econstructor; eauto. }
- exploit step_simulation_builtin.
- 4: eapply MS'.
- all: simpl; eauto.
- intros (S3' & PLUS'' & MS'').
- exists S3'. split; eauto.
- eapply plus_trans. eapply PLUS1. eapply PLUS''. eauto.
- + inversion H0. subst. eapply step_simulation_bblock; try (rewrite MBE; try discriminate); eauto.
-
-- (* internal function *)
- inv MS.
- exploit functions_translated; eauto. intros [tf [A B]]. monadInv B.
- generalize EQ; intros EQ'. monadInv EQ'.
- destruct (zlt Ptrofs.max_unsigned (size_blocks x0.(fn_blocks))); inversion EQ1. clear EQ1. subst x0.
- unfold Mach.store_stack in *.
- exploit Mem.alloc_extends. eauto. eauto. apply Z.le_refl. apply Z.le_refl.
- intros [m1' [C D]].
- exploit Mem.storev_extends. eexact D. eexact H1. eauto. eauto.
- intros [m2' [F G]].
- simpl chunk_of_type in F.
- exploit Mem.storev_extends. eexact G. eexact H2. eauto. eauto.
- intros [m3' [P Q]].
- (* Execution of function prologue *)
- monadInv EQ0. (* rewrite transl_code'_transl_code in EQ1. *)
- set (tfbody := make_prologue f x0) in *.
- set (tf := {| fn_sig := MB.fn_sig f; fn_blocks := tfbody |}) in *.
- set (rs2 := rs0#FP <- (parent_sp s) #SP <- sp #RTMP <- Vundef).
- exploit (Pget_correct tge GPRA RA nil rs2 m2'); auto.
- intros (rs' & U' & V').
-(* exploit (exec_straight_through_singleinst); eauto.
- intro W'. remember (nextblock _ rs') as rs''. *)
- exploit (storeind_ptr_correct tge SP (fn_retaddr_ofs f) GPRA nil rs' m2').
- rewrite chunk_of_Tptr in P.
- assert (rs' GPRA = rs0 RA). { apply V'. }
- assert (rs' SP = rs2 SP). { apply V'; discriminate. }
- rewrite H4. rewrite H3.
- (* change (rs' GPRA) with (rs0 RA). *)
- rewrite ATLR.
- change (rs2 SP) with sp. eexact P.
- intros (rs3 & U & V).
-(* exploit (exec_straight_through_singleinst); eauto.
- intro W. *)
- assert (EXEC_PROLOGUE: exists rs3',
- exec_straight_blocks tge tf
- tf.(fn_blocks) rs0 m'
- x0 rs3' m3'
- /\ forall r, r <> PC -> rs3' r = rs3 r).
- { eexists. split.
- - change (fn_blocks tf) with tfbody; unfold tfbody.
- econstructor; eauto. unfold exec_bblock. simpl exec_body.
- rewrite C. fold sp. rewrite <- (sp_val _ _ _ AG). rewrite chunk_of_Tptr in F. simpl in F. rewrite F.
- Simpl. unfold parexec_store_offset. rewrite Ptrofs.of_int64_to_int64. unfold eval_offset.
- rewrite chunk_of_Tptr in P. Simpl. rewrite ATLR. unfold Mptr in P. assert (Archi.ptr64 = true) by auto. 2: auto. rewrite H3 in P. rewrite P.
- simpl. apply next_sep; eauto. reflexivity.
- - intros. destruct V' as (V'' & V'). destruct r.
- + Simpl.
- destruct (gpreg_eq g0 GPR16). { subst. Simpl. rewrite V; try discriminate. rewrite V''. subst rs2. Simpl. }
- destruct (gpreg_eq g0 GPR32). { subst. Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl. }
- destruct (gpreg_eq g0 GPR12). { subst. Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl. }
- destruct (gpreg_eq g0 GPR17). { subst. Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl. }
- Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl. { destruct g0; try discriminate. contradiction. }
- + Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl.
- + contradiction.
- } destruct EXEC_PROLOGUE as (rs3' & EXEC_PROLOGUE & Heqrs3').
- exploit exec_straight_steps_2; eauto using functions_transl.
- simpl fn_blocks. simpl fn_blocks in g. omega. constructor.
- intros (ofs' & X & Y).
- left; exists (State rs3' m3'); split.
- eapply exec_straight_steps_1; eauto.
- simpl fn_blocks. simpl fn_blocks in g. omega.
- constructor.
- econstructor; eauto.
- rewrite X; econstructor; eauto.
- apply agree_exten with rs2; eauto with asmgen.
- unfold rs2.
- apply agree_set_other; auto with asmgen.
- apply agree_change_sp with (parent_sp s).
- apply agree_undef_regs with rs0. auto.
-Local Transparent destroyed_at_function_entry.
- simpl; intros; Simpl.
- unfold sp; congruence.
-
- intros.
- assert (r <> RTMP). { contradict H3; rewrite H3; unfold data_preg; auto. }
- rewrite Heqrs3'. Simpl. rewrite V. inversion V'. rewrite H6. auto.
- assert (r <> GPRA). { contradict H3; rewrite H3; unfold data_preg; auto. }
- assert (forall r : preg, r <> PC -> r <> GPRA -> rs' r = rs2 r). { apply V'. }
- (* rewrite H8; auto. *)
- contradict H3; rewrite H3; unfold data_preg; auto.
- contradict H3; rewrite H3; unfold data_preg; auto.
- contradict H3; rewrite H3; unfold data_preg; auto.
- contradict H3; rewrite H3; unfold data_preg; auto.
- intros. rewrite Heqrs3'. rewrite V by auto with asmgen.
- assert (forall r : preg, r <> PC -> r <> GPRA -> rs' r = rs2 r). { apply V'. }
- rewrite H4 by auto with asmgen. reflexivity. discriminate.
-- (* external function *)
- inv MS.
- exploit functions_translated; eauto.
- intros [tf [A B]]. simpl in B. inv B.
- exploit extcall_arguments_match; eauto.
- intros [args' [C D]].
- exploit external_call_mem_extends; eauto.
- intros [res' [m2' [P [Q [R S]]]]].
- left; econstructor; split.
- apply plus_one. eapply exec_step_external; eauto.
- eapply external_call_symbols_preserved; eauto. apply senv_preserved.
- econstructor; eauto.
- unfold loc_external_result.
- apply agree_set_other; auto.
- apply agree_set_pair; auto.
- apply agree_undef_caller_save_regs; auto.
-
-- (* return *)
- inv MS.
- inv STACKS. simpl in *.
- right. split. omega. split. auto.
- rewrite <- ATPC in H5.
- econstructor; eauto. congruence.
-Qed.
-
-Lemma transf_initial_states:
- forall st1, MB.initial_state prog st1 ->
- exists st2, AB.initial_state tprog st2 /\ match_states st1 st2.
-Proof.
- intros. inversion H. unfold ge0 in *.
- econstructor; split.
- econstructor.
- eapply (Genv.init_mem_transf_partial TRANSF); eauto.
- replace (Genv.symbol_address (Genv.globalenv tprog) (prog_main tprog) Ptrofs.zero)
- with (Vptr fb Ptrofs.zero).
- econstructor; eauto.
- constructor.
- apply Mem.extends_refl.
- split. auto. simpl. unfold Vnullptr; destruct Archi.ptr64; congruence.
- intros. rewrite Mach.Regmap.gi. auto.
- unfold Genv.symbol_address.
- rewrite (match_program_main TRANSF).
- rewrite symbols_preserved.
- unfold ge; rewrite H1. auto.
-Qed.
-
-Lemma transf_final_states:
- forall st1 st2 r,
- match_states st1 st2 -> MB.final_state st1 r -> AB.final_state st2 r.
-Proof.
- intros. inv H0. inv H. constructor. assumption.
- compute in H1. inv H1.
- generalize (preg_val _ _ _ R0 AG). rewrite H2. intros LD; inv LD. auto.
-Qed.
-
-Definition return_address_offset : Machblock.function -> Machblock.code -> ptrofs -> Prop :=
- Asmblockgenproof0.return_address_offset.
-
-Theorem transf_program_correct:
- forward_simulation (MB.semantics return_address_offset prog) (Asmblock.semantics tprog).
-Proof.
- eapply forward_simulation_star with (measure := measure).
- - apply senv_preserved.
- - eexact transf_initial_states.
- - eexact transf_final_states.
- - exact step_simulation.
-Qed.
-
-End PRESERVATION.
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris-Rocquencourt *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Correctness proof for RISC-V generation: main proof. *)
+
+Require Import Coqlib Errors.
+Require Import Integers Floats AST Linking.
+Require Import Values Memory Events Globalenvs Smallstep.
+Require Import Op Locations Machblock Conventions Asmblock.
+Require Import Asmblockgen Asmblockgenproof0 Asmblockgenproof1.
+Require Import Axioms.
+
+Module MB := Machblock.
+Module AB := Asmvliw.
+
+Definition match_prog (p: Machblock.program) (tp: Asmvliw.program) :=
+ match_program (fun _ f tf => transf_fundef f = OK tf) eq p tp.
+
+Lemma transf_program_match:
+ forall p tp, transf_program p = OK tp -> match_prog p tp.
+Proof.
+ intros. eapply match_transform_partial_program; eauto.
+Qed.
+
+Section PRESERVATION.
+
+Variable prog: Machblock.program.
+Variable tprog: Asmvliw.program.
+Hypothesis TRANSF: match_prog prog tprog.
+Let ge := Genv.globalenv prog.
+Let tge := Genv.globalenv tprog.
+
+Lemma symbols_preserved:
+ forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
+Proof (Genv.find_symbol_match TRANSF).
+
+Lemma senv_preserved:
+ Senv.equiv ge tge.
+Proof (Genv.senv_match TRANSF).
+
+Lemma functions_translated:
+ forall b f,
+ Genv.find_funct_ptr ge b = Some f ->
+ exists tf,
+ Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = OK tf.
+Proof (Genv.find_funct_ptr_transf_partial TRANSF).
+
+Lemma functions_transl:
+ forall fb f tf,
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ transf_function f = OK tf ->
+ Genv.find_funct_ptr tge fb = Some (Internal tf).
+Proof.
+ intros. exploit functions_translated; eauto. intros [tf' [A B]].
+ monadInv B. rewrite H0 in EQ; inv EQ; auto.
+Qed.
+
+Lemma transf_function_no_overflow:
+ forall f tf,
+ transf_function f = OK tf -> size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned.
+Proof.
+ intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (size_blocks x.(fn_blocks))); inv EQ0.
+ omega.
+Qed.
+
+Section TRANSL_LABEL. (* Lemmas on translation of MB.is_label into AB.is_label *)
+
+Lemma gen_bblocks_label:
+ forall hd bdy ex tbb tc,
+ gen_bblocks hd bdy ex = tbb::tc ->
+ header tbb = hd.
+Proof.
+ intros until tc. intros GENB. unfold gen_bblocks in GENB.
+ destruct (extract_ctl ex); try destruct c; try destruct i; try destruct bdy.
+ all: inv GENB; simpl; auto.
+Qed.
+
+Lemma gen_bblocks_label2:
+ forall hd bdy ex tbb1 tbb2,
+ gen_bblocks hd bdy ex = tbb1::tbb2::nil ->
+ header tbb2 = nil.
+Proof.
+ intros until tbb2. intros GENB. unfold gen_bblocks in GENB.
+ destruct (extract_ctl ex); try destruct c; try destruct i; try destruct bdy.
+ all: inv GENB; simpl; auto.
+Qed.
+
+Remark in_dec_transl:
+ forall lbl hd,
+ (if in_dec lbl hd then true else false) = (if MB.in_dec lbl hd then true else false).
+Proof.
+ intros. destruct (in_dec lbl hd), (MB.in_dec lbl hd). all: tauto.
+Qed.
+
+Lemma transl_is_label:
+ forall lbl bb tbb f ep tc,
+ transl_block f bb ep = OK (tbb::tc) ->
+ is_label lbl tbb = MB.is_label lbl bb.
+Proof.
+ intros until tc. intros TLB.
+ destruct tbb as [thd tbdy tex]; simpl in *.
+ monadInv TLB.
+ unfold is_label. simpl.
+ apply gen_bblocks_label in H0. simpl in H0. subst.
+ rewrite in_dec_transl. auto.
+Qed.
+
+Lemma transl_is_label_false2:
+ forall lbl bb f ep tbb1 tbb2,
+ transl_block f bb ep = OK (tbb1::tbb2::nil) ->
+ is_label lbl tbb2 = false.
+Proof.
+ intros until tbb2. intros TLB.
+ destruct tbb2 as [thd tbdy tex]; simpl in *.
+ monadInv TLB. apply gen_bblocks_label2 in H0. simpl in H0. subst.
+ apply is_label_correct_false. simpl. auto.
+Qed.
+
+Lemma transl_is_label2:
+ forall f bb ep tbb1 tbb2 lbl,
+ transl_block f bb ep = OK (tbb1::tbb2::nil) ->
+ is_label lbl tbb1 = MB.is_label lbl bb
+ /\ is_label lbl tbb2 = false.
+Proof.
+ intros. split. eapply transl_is_label; eauto. eapply transl_is_label_false2; eauto.
+Qed.
+
+Lemma transl_block_nonil:
+ forall f c ep tc,
+ transl_block f c ep = OK tc ->
+ tc <> nil.
+Proof.
+ intros. monadInv H. unfold gen_bblocks.
+ destruct (extract_ctl x0); try destruct c0; try destruct x; try destruct i.
+ all: discriminate.
+Qed.
+
+Lemma transl_block_limit: forall f bb ep tbb1 tbb2 tbb3 tc,
+ ~transl_block f bb ep = OK (tbb1 :: tbb2 :: tbb3 :: tc).
+Proof.
+ intros. intro. monadInv H.
+ unfold gen_bblocks in H0.
+ destruct (extract_ctl x0); try destruct x; try destruct c; try destruct i.
+ all: discriminate.
+Qed.
+
+Lemma find_label_transl_false:
+ forall x f lbl bb ep x',
+ transl_block f bb ep = OK x ->
+ MB.is_label lbl bb = false ->
+ find_label lbl (x++x') = find_label lbl x'.
+Proof.
+ intros until x'. intros TLB MBis; simpl; auto.
+ destruct x as [|x0 x1]; simpl; auto.
+ destruct x1 as [|x1 x2]; simpl; auto.
+ - erewrite <- transl_is_label in MBis; eauto. rewrite MBis. auto.
+ - destruct x2 as [|x2 x3]; simpl; auto.
+ + erewrite <- transl_is_label in MBis; eauto. rewrite MBis.
+ erewrite transl_is_label_false2; eauto.
+ + apply transl_block_limit in TLB. destruct TLB.
+Qed.
+
+Lemma transl_blocks_label:
+ forall lbl f c tc ep,
+ transl_blocks f c ep = OK tc ->
+ match MB.find_label lbl c with
+ | None => find_label lbl tc = None
+ | Some c' => exists tc', find_label lbl tc = Some tc' /\ transl_blocks f c' false = OK tc'
+ end.
+Proof.
+ induction c; simpl; intros.
+ inv H. auto.
+ monadInv H.
+ destruct (MB.is_label lbl a) eqn:MBis.
+ - destruct x as [|tbb tc]. { apply transl_block_nonil in EQ. contradiction. }
+ simpl find_label. exploit transl_is_label; eauto. intros ABis. rewrite MBis in ABis.
+ rewrite ABis.
+ eexists. eexists. split; eauto. simpl transl_blocks.
+ assert (MB.header a <> nil).
+ { apply MB.is_label_correct_true in MBis.
+ destruct (MB.header a). contradiction. discriminate. }
+ destruct (MB.header a); try contradiction.
+ rewrite EQ. simpl. rewrite EQ1. simpl. auto.
+ - apply IHc in EQ1. destruct (MB.find_label lbl c).
+ + destruct EQ1 as (tc' & FIND & TLBS). exists tc'; eexists; auto.
+ erewrite find_label_transl_false; eauto.
+ + erewrite find_label_transl_false; eauto.
+Qed.
+
+Lemma find_label_nil:
+ forall bb lbl c,
+ header bb = nil ->
+ find_label lbl (bb::c) = find_label lbl c.
+Proof.
+ intros. destruct bb as [hd bdy ex]; simpl in *. subst.
+ assert (is_label lbl {| AB.header := nil; AB.body := bdy; AB.exit := ex; AB.correct := correct |} = false).
+ { erewrite <- is_label_correct_false. simpl. auto. }
+ rewrite H. auto.
+Qed.
+
+Theorem transl_find_label:
+ forall lbl f tf,
+ transf_function f = OK tf ->
+ match MB.find_label lbl f.(MB.fn_code) with
+ | None => find_label lbl tf.(fn_blocks) = None
+ | Some c => exists tc, find_label lbl tf.(fn_blocks) = Some tc /\ transl_blocks f c false = OK tc
+ end.
+Proof.
+ intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (size_blocks (fn_blocks x))); inv EQ0. clear g.
+ monadInv EQ. unfold make_prologue. simpl fn_blocks. repeat (rewrite find_label_nil); simpl; auto.
+ eapply transl_blocks_label; eauto.
+Qed.
+
+End TRANSL_LABEL.
+
+(** A valid branch in a piece of Machblock code translates to a valid ``go to''
+ transition in the generated Asmblock code. *)
+
+Lemma find_label_goto_label:
+ forall f tf lbl rs m c' b ofs,
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ transf_function f = OK tf ->
+ rs PC = Vptr b ofs ->
+ MB.find_label lbl f.(MB.fn_code) = Some c' ->
+ exists tc', exists rs',
+ goto_label tf lbl rs m = Next rs' m
+ /\ transl_code_at_pc ge (rs' PC) b f c' false tf tc'
+ /\ forall r, r <> PC -> rs'#r = rs#r.
+Proof.
+ intros. exploit (transl_find_label lbl f tf); eauto. rewrite H2.
+ intros (tc & A & B).
+ exploit label_pos_code_tail; eauto. instantiate (1 := 0).
+ intros [pos' [P [Q R]]].
+ exists tc; exists (rs#PC <- (Vptr b (Ptrofs.repr pos'))).
+ split. unfold goto_label. unfold par_goto_label. rewrite P. rewrite H1. auto.
+ split. rewrite Pregmap.gss. constructor; auto.
+ rewrite Ptrofs.unsigned_repr. replace (pos' - 0) with pos' in Q.
+ auto. omega.
+ generalize (transf_function_no_overflow _ _ H0). omega.
+ intros. apply Pregmap.gso; auto.
+Qed.
+
+(** Existence of return addresses *)
+
+Lemma return_address_exists:
+ forall b f c, is_tail (b :: c) f.(MB.fn_code) ->
+ exists ra, return_address_offset f c ra.
+Proof.
+ intros. eapply Asmblockgenproof0.return_address_exists; eauto.
+
+- intros. monadInv H0.
+ destruct (zlt Ptrofs.max_unsigned (size_blocks x.(fn_blocks))); inv EQ0. monadInv EQ. simpl.
+ exists x; exists true; split; auto.
+ repeat constructor.
+- exact transf_function_no_overflow.
+Qed.
+
+(** * Proof of semantic preservation *)
+
+(** Semantic preservation is proved using a complex simulation diagram
+ of the following form.
+<<
+ MB.step
+ ---------------------------------------->
+ header body exit
+ st1 -----> st2 -----> st3 ------------------> st4
+ | | | |
+ | (A) | (B) | (C) |
+ match_codestate | | | |
+ | header | body1 | body2 | match_states
+ cs1 -----> cs2 -----> cs3 ------> cs4 |
+ | / \ exit |
+ match_asmstate | --------------- --->--- |
+ | / match_asmstate \ |
+ st'1 ---------------------------------------> st'2
+ AB.step *
+>>
+ The invariant between each MB.step/AB.step is the [match_states] predicate below.
+ However, we also need to introduce an intermediary state [Codestate] which allows
+ us to reason on a finer grain, executing header, body and exit separately.
+
+ This [Codestate] consists in a state like [Asmblock.State], except that the
+ code is directly stored in the state, much like [Machblock.State]. It also features
+ additional useful elements to keep track of while executing a bblock.
+*)
+
+Remark preg_of_not_FP: forall r, negb (mreg_eq r MFP) = true -> IR FP <> preg_of r.
+Proof.
+ intros. change (IR FP) with (preg_of MFP). red; intros.
+ exploit preg_of_injective; eauto. intros; subst r; discriminate.
+Qed.
+
+Inductive match_states: Machblock.state -> Asmvliw.state -> Prop :=
+ | match_states_intro:
+ forall s fb sp c ep ms m m' rs f tf tc
+ (STACKS: match_stack ge s)
+ (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
+ (MEXT: Mem.extends m m')
+ (AT: transl_code_at_pc ge (rs PC) fb f c ep tf tc)
+ (AG: agree ms sp rs)
+ (DXP: ep = true -> rs#FP = parent_sp s),
+ match_states (Machblock.State s fb sp c ms m)
+ (Asmvliw.State rs m')
+ | match_states_call:
+ forall s fb ms m m' rs
+ (STACKS: match_stack ge s)
+ (MEXT: Mem.extends m m')
+ (AG: agree ms (parent_sp s) rs)
+ (ATPC: rs PC = Vptr fb Ptrofs.zero)
+ (ATLR: rs RA = parent_ra s),
+ match_states (Machblock.Callstate s fb ms m)
+ (Asmvliw.State rs m')
+ | match_states_return:
+ forall s ms m m' rs
+ (STACKS: match_stack ge s)
+ (MEXT: Mem.extends m m')
+ (AG: agree ms (parent_sp s) rs)
+ (ATPC: rs PC = parent_ra s),
+ match_states (Machblock.Returnstate s ms m)
+ (Asmvliw.State rs m').
+
+Record codestate :=
+ Codestate { pstate: state; (**r projection to Asmblock.state *)
+ pheader: list label;
+ pbody1: list basic; (**r list of basic instructions coming from the translation of the Machblock body *)
+ pbody2: list basic; (**r list of basic instructions coming from the translation of the Machblock exit *)
+ pctl: option control; (**r exit instruction, coming from the translation of the Machblock exit *)
+ ep: bool; (**r reflects the [ep] variable used in the translation *)
+ rem: list AB.bblock; (**r remaining bblocks to execute *)
+ cur: bblock (**r current bblock to execute - to keep track of its size when incrementing PC *)
+ }.
+
+(* The part that deals with Machblock <-> Codestate agreement
+ * Note about DXP: the property of [ep] only matters if the current block doesn't have a header, hence the condition *)
+Inductive match_codestate fb: Machblock.state -> codestate -> Prop :=
+ | match_codestate_intro:
+ forall s sp ms m rs0 m0 f tc ep c bb tbb tbc tbi
+ (STACKS: match_stack ge s)
+ (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
+ (MEXT: Mem.extends m m0)
+ (TBC: transl_basic_code f (MB.body bb) (if MB.header bb then ep else false) = OK tbc)
+ (TIC: transl_instr_control f (MB.exit bb) = OK tbi)
+ (TBLS: transl_blocks f c false = OK tc)
+ (AG: agree ms sp rs0)
+ (DXP: (if MB.header bb then ep else false) = true -> rs0#FP = parent_sp s)
+ ,
+ match_codestate fb (Machblock.State s fb sp (bb::c) ms m)
+ {| pstate := (Asmvliw.State rs0 m0);
+ pheader := (MB.header bb);
+ pbody1 := tbc;
+ pbody2 := extract_basic tbi;
+ pctl := extract_ctl tbi;
+ ep := ep;
+ rem := tc;
+ cur := tbb
+ |}
+.
+
+(* The part ensuring that the code in Codestate actually resides at [rs PC] *)
+Inductive match_asmstate fb: codestate -> Asmvliw.state -> Prop :=
+ | match_asmstate_some:
+ forall rs f tf tc m tbb ofs ep tbdy tex lhd
+ (FIND: Genv.find_funct_ptr ge fb = Some (Internal f))
+ (TRANSF: transf_function f = OK tf)
+ (PCeq: rs PC = Vptr fb ofs)
+ (TAIL: code_tail (Ptrofs.unsigned ofs) (fn_blocks tf) (tbb::tc))
+ ,
+ match_asmstate fb
+ {| pstate := (Asmvliw.State rs m);
+ pheader := lhd;
+ pbody1 := tbdy;
+ pbody2 := extract_basic tex;
+ pctl := extract_ctl tex;
+ ep := ep;
+ rem := tc;
+ cur := tbb |}
+ (Asmvliw.State rs m)
+.
+
+(* Useful for dealing with the many cases in some proofs *)
+Ltac exploreInst :=
+ repeat match goal with
+ | [ H : match ?var with | _ => _ end = _ |- _ ] => destruct var
+ | [ H : OK _ = OK _ |- _ ] => monadInv H
+ | [ |- context[if ?b then _ else _] ] => destruct b
+ | [ |- context[match ?m with | _ => _ end] ] => destruct m
+ | [ |- context[match ?m as _ return _ with | _ => _ end]] => destruct m
+ | [ H : bind _ _ = OK _ |- _ ] => monadInv H
+ | [ H : Error _ = OK _ |- _ ] => inversion H
+ end.
+
+(** Some translation properties *)
+
+Lemma transl_blocks_nonil:
+ forall f bb c tc ep,
+ transl_blocks f (bb::c) ep = OK tc ->
+ exists tbb tc', tc = tbb :: tc'.
+Proof.
+ intros until ep0. intros TLBS. monadInv TLBS. monadInv EQ. unfold gen_bblocks.
+ destruct (extract_ctl x2).
+ - destruct c0; destruct i; simpl; eauto. destruct x1; simpl; eauto.
+ - destruct x1; simpl; eauto.
+Qed.
+
+Lemma no_builtin_preserved:
+ forall f ex x2,
+ (forall ef args res, ex <> Some (MBbuiltin ef args res)) ->
+ transl_instr_control f ex = OK x2 ->
+ (exists i, extract_ctl x2 = Some (PCtlFlow i))
+ \/ extract_ctl x2 = None.
+Proof.
+ intros until x2. intros Hbuiltin TIC.
+ destruct ex.
+ - destruct c.
+ (* MBcall *)
+ + simpl in TIC. exploreInst; simpl; eauto.
+ (* MBtailcall *)
+ + simpl in TIC. exploreInst; simpl; eauto.
+ (* MBbuiltin *)
+ + assert (H: Some (MBbuiltin e l b) <> Some (MBbuiltin e l b)).
+ apply Hbuiltin. contradict H; auto.
+ (* MBgoto *)
+ + simpl in TIC. exploreInst; simpl; eauto.
+ (* MBcond *)
+ + simpl in TIC. unfold transl_cbranch in TIC. exploreInst; simpl; eauto.
+ * unfold transl_opt_compuimm. exploreInst; simpl; eauto.
+ * unfold transl_opt_compluimm. exploreInst; simpl; eauto.
+ * unfold transl_comp_float64. exploreInst; simpl; eauto.
+ * unfold transl_comp_notfloat64. exploreInst; simpl; eauto.
+ * unfold transl_comp_float32. exploreInst; simpl; eauto.
+ * unfold transl_comp_notfloat32. exploreInst; simpl; eauto.
+ (* MBjumptable *)
+ + simpl in TIC. exploreInst; simpl; eauto.
+ (* MBreturn *)
+ + simpl in TIC. monadInv TIC. simpl. eauto.
+ - monadInv TIC. simpl; auto.
+Qed.
+
+Lemma transl_blocks_distrib:
+ forall c f bb tbb tc ep,
+ transl_blocks f (bb::c) ep = OK (tbb::tc)
+ -> (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res))
+ -> transl_block f bb (if MB.header bb then ep else false) = OK (tbb :: nil)
+ /\ transl_blocks f c false = OK tc.
+Proof.
+ intros until ep0. intros TLBS Hbuiltin.
+ destruct bb as [hd bdy ex].
+ monadInv TLBS. monadInv EQ.
+ exploit no_builtin_preserved; eauto. intros Hectl. destruct Hectl.
+ - destruct H as [i Hectl].
+ unfold gen_bblocks in H0. rewrite Hectl in H0. inv H0.
+ simpl in *. unfold transl_block; simpl. rewrite EQ0. rewrite EQ. simpl.
+ unfold gen_bblocks. rewrite Hectl. auto.
+ - unfold gen_bblocks in H0. rewrite H in H0.
+ destruct x1 as [|bi x1].
+ + simpl in H0. inv H0. simpl in *. unfold transl_block; simpl. rewrite EQ0. rewrite EQ. simpl.
+ unfold gen_bblocks. rewrite H. auto.
+ + simpl in H0. inv H0. simpl in *. unfold transl_block; simpl. rewrite EQ0. rewrite EQ. simpl.
+ unfold gen_bblocks. rewrite H. auto.
+Qed.
+
+Lemma gen_bblocks_nobuiltin:
+ forall thd tbdy tex tbb,
+ (tbdy <> nil \/ extract_ctl tex <> None) ->
+ (forall ef args res, extract_ctl tex <> Some (PExpand (Pbuiltin ef args res))) ->
+ gen_bblocks thd tbdy tex = tbb :: nil ->
+ header tbb = thd
+ /\ body tbb = tbdy ++ extract_basic tex
+ /\ exit tbb = extract_ctl tex.
+Proof.
+ intros until tbb. intros Hnonil Hnobuiltin GENB. unfold gen_bblocks in GENB.
+ destruct (extract_ctl tex) eqn:ECTL.
+ - destruct c.
+ + destruct i; try (inv GENB; simpl; auto; fail).
+ assert False. eapply Hnobuiltin. eauto. destruct H.
+ + inv GENB. simpl. auto.
+ - inversion Hnonil.
+ + destruct tbdy as [|bi tbdy]; try (contradict H; simpl; auto; fail). inv GENB. auto.
+ + contradict H; simpl; auto.
+Qed.
+
+Lemma transl_instr_basic_nonil:
+ forall k f bi ep x,
+ transl_instr_basic f bi ep k = OK x ->
+ x <> nil.
+Proof.
+ intros until x. intros TIB.
+ destruct bi.
+ - simpl in TIB. unfold loadind in TIB. exploreInst; try discriminate.
+ - simpl in TIB. unfold storeind in TIB. exploreInst; try discriminate.
+ - simpl in TIB. monadInv TIB. unfold loadind in EQ. exploreInst; try discriminate.
+ - simpl in TIB. unfold transl_op in TIB. exploreInst; try discriminate.
+ unfold transl_cond_op in EQ0. exploreInst; try discriminate.
+ unfold transl_cond_float64. exploreInst; try discriminate.
+ unfold transl_cond_notfloat64. exploreInst; try discriminate.
+ unfold transl_cond_float32. exploreInst; try discriminate.
+ unfold transl_cond_notfloat32. exploreInst; try discriminate.
+ - simpl in TIB. unfold transl_load in TIB. exploreInst; try discriminate.
+ all: monadInv TIB; unfold transl_memory_access in EQ0; unfold transl_memory_access2 in EQ0; unfold transl_memory_access2XS in EQ0; exploreInst; try discriminate.
+ - simpl in TIB. unfold transl_store in TIB. exploreInst; try discriminate.
+ all: monadInv TIB; unfold transl_memory_access in EQ0; unfold transl_memory_access2 in EQ0; unfold transl_memory_access2XS in EQ0; exploreInst; try discriminate.
+Qed.
+
+Lemma transl_basic_code_nonil:
+ forall bdy f x ep,
+ bdy <> nil ->
+ transl_basic_code f bdy ep = OK x ->
+ x <> nil.
+Proof.
+ induction bdy as [|bi bdy].
+ intros. contradict H0; auto.
+ destruct bdy as [|bi2 bdy].
+ - clear IHbdy. intros f x b _ TBC. simpl in TBC. eapply transl_instr_basic_nonil; eauto.
+ - intros f x b Hnonil TBC. remember (bi2 :: bdy) as bdy'.
+ monadInv TBC.
+ assert (x0 <> nil).
+ eapply IHbdy; eauto. subst bdy'. discriminate.
+ eapply transl_instr_basic_nonil; eauto.
+Qed.
+
+Lemma transl_instr_control_nonil:
+ forall ex f x,
+ ex <> None ->
+ transl_instr_control f ex = OK x ->
+ extract_ctl x <> None.
+Proof.
+ intros ex f x Hnonil TIC.
+ destruct ex as [ex|].
+ - clear Hnonil. destruct ex.
+ all: try (simpl in TIC; exploreInst; discriminate).
+ + simpl in TIC. unfold transl_cbranch in TIC. exploreInst; try discriminate.
+ * unfold transl_opt_compuimm. exploreInst; try discriminate.
+ * unfold transl_opt_compluimm. exploreInst; try discriminate.
+ * unfold transl_comp_float64. exploreInst; try discriminate.
+ * unfold transl_comp_notfloat64. exploreInst; try discriminate.
+ * unfold transl_comp_float32. exploreInst; try discriminate.
+ * unfold transl_comp_notfloat32. exploreInst; try discriminate.
+ - contradict Hnonil; auto.
+Qed.
+
+Lemma transl_instr_control_nobuiltin:
+ forall f ex x,
+ (forall ef args res, ex <> Some (MBbuiltin ef args res)) ->
+ transl_instr_control f ex = OK x ->
+ (forall ef args res, extract_ctl x <> Some (PExpand (Pbuiltin ef args res))).
+Proof.
+ intros until x. intros Hnobuiltin TIC. intros until res.
+ unfold transl_instr_control in TIC. exploreInst.
+ all: try discriminate.
+ - assert False. eapply Hnobuiltin; eauto. destruct H.
+ - unfold transl_cbranch in TIC. exploreInst.
+ all: try discriminate.
+ * unfold transl_opt_compuimm. exploreInst. all: try discriminate.
+ * unfold transl_opt_compluimm. exploreInst. all: try discriminate.
+ * unfold transl_comp_float64. exploreInst; try discriminate.
+ * unfold transl_comp_notfloat64. exploreInst; try discriminate.
+ * unfold transl_comp_float32. exploreInst; try discriminate.
+ * unfold transl_comp_notfloat32. exploreInst; try discriminate.
+Qed.
+
+(* Proving that one can decompose a [match_state] relation into a [match_codestate]
+ and a [match_asmstate], along with some helpful properties tying both relations together *)
+
+Theorem match_state_codestate:
+ forall mbs abs s fb sp bb c ms m,
+ (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
+ (MB.body bb <> nil \/ MB.exit bb <> None) ->
+ mbs = (Machblock.State s fb sp (bb::c) ms m) ->
+ match_states mbs abs ->
+ exists cs fb f tbb tc ep,
+ match_codestate fb mbs cs /\ match_asmstate fb cs abs
+ /\ Genv.find_funct_ptr ge fb = Some (Internal f)
+ /\ transl_blocks f (bb::c) ep = OK (tbb::tc)
+ /\ body tbb = pbody1 cs ++ pbody2 cs
+ /\ exit tbb = pctl cs
+ /\ cur cs = tbb /\ rem cs = tc
+ /\ pstate cs = abs.
+Proof.
+ intros until m. intros Hnobuiltin Hnotempty Hmbs MS. subst. inv MS.
+ inv AT. clear H0. exploit transl_blocks_nonil; eauto. intros (tbb & tc' & Htc). subst.
+ exploit transl_blocks_distrib; eauto. intros (TLB & TLBS). clear H2.
+ monadInv TLB. exploit gen_bblocks_nobuiltin; eauto.
+ { inversion Hnotempty.
+ - destruct (MB.body bb) as [|bi bdy]; try (contradict H0; simpl; auto; fail).
+ left. eapply transl_basic_code_nonil; eauto.
+ - destruct (MB.exit bb) as [ei|]; try (contradict H0; simpl; auto; fail).
+ right. eapply transl_instr_control_nonil; eauto. }
+ eapply transl_instr_control_nobuiltin; eauto.
+ intros (Hth & Htbdy & Htexit).
+ exists {| pstate := (State rs m'); pheader := (Machblock.header bb); pbody1 := x; pbody2 := extract_basic x0;
+ pctl := extract_ctl x0; ep := ep0; rem := tc'; cur := tbb |}, fb, f, tbb, tc', ep0.
+ repeat split. 1-2: econstructor; eauto.
+ { destruct (MB.header bb). eauto. discriminate. } eauto.
+ unfold transl_blocks. fold transl_blocks. unfold transl_block. rewrite EQ. simpl. rewrite EQ1; simpl.
+ rewrite TLBS. simpl. rewrite H2.
+ all: simpl; auto.
+Qed.
+
+Definition mb_remove_body (bb: MB.bblock) :=
+ {| MB.header := MB.header bb; MB.body := nil; MB.exit := MB.exit bb |}.
+
+Lemma exec_straight_pnil:
+ forall c rs1 m1 rs2 m2,
+ exec_straight tge c rs1 m1 (Pnop ::g nil) rs2 m2 ->
+ exec_straight tge c rs1 m1 nil rs2 m2.
+Proof.
+ intros. eapply exec_straight_trans. eapply H. econstructor; eauto.
+Qed.
+
+Lemma transl_block_nobuiltin:
+ forall f bb ep tbb,
+ (MB.body bb <> nil \/ MB.exit bb <> None) ->
+ (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
+ transl_block f bb ep = OK (tbb :: nil) ->
+ exists c c',
+ transl_basic_code f (MB.body bb) ep = OK c
+ /\ transl_instr_control f (MB.exit bb) = OK c'
+ /\ body tbb = c ++ extract_basic c'
+ /\ exit tbb = extract_ctl c'.
+Proof.
+ intros until tbb. intros Hnonil Hnobuiltin TLB. monadInv TLB. destruct Hnonil.
+ - eexists. eexists. split; eauto. split; eauto. eapply gen_bblocks_nobuiltin; eauto.
+ left. eapply transl_basic_code_nonil; eauto. eapply transl_instr_control_nobuiltin; eauto.
+ - eexists. eexists. split; eauto. split; eauto. eapply gen_bblocks_nobuiltin; eauto.
+ right. eapply transl_instr_control_nonil; eauto. eapply transl_instr_control_nobuiltin; eauto.
+Qed.
+
+Lemma nextblock_preserves:
+ forall rs rs' bb r,
+ rs' = nextblock bb rs ->
+ data_preg r = true ->
+ rs r = rs' r.
+Proof.
+ intros. destruct r; try discriminate.
+ subst. Simpl.
+Qed.
+
+Remark cons3_app {A: Type}:
+ forall a b c (l: list A),
+ a :: b :: c :: l = (a :: b :: c :: nil) ++ l.
+Proof.
+ intros. simpl. auto.
+Qed.
+
+Lemma exec_straight_opt_body2:
+ forall c rs1 m1 c' rs2 m2,
+ exec_straight_opt tge c rs1 m1 c' rs2 m2 ->
+ exists body,
+ exec_body tge body rs1 m1 = Next rs2 m2
+ /\ (basics_to_code body) ++g c' = c.
+Proof.
+ intros until m2. intros EXES.
+ inv EXES.
+ - exists nil. split; auto.
+ - eapply exec_straight_body2. auto.
+Qed.
+
+Lemma extract_basics_to_code:
+ forall lb c,
+ extract_basic (basics_to_code lb ++ c) = lb ++ extract_basic c.
+Proof.
+ induction lb; intros; simpl; congruence.
+Qed.
+
+Lemma extract_ctl_basics_to_code:
+ forall lb c,
+ extract_ctl (basics_to_code lb ++ c) = extract_ctl c.
+Proof.
+ induction lb; intros; simpl; congruence.
+Qed.
+
+(* See (C) in the diagram. The proofs are mostly adapted from the previous Mach->Asm proofs, but are
+ unfortunately quite cumbersome. To reproduce them, it's best to have a Coq IDE with you and see by
+ yourself the steps *)
+Theorem step_simu_control:
+ forall bb' fb fn s sp c ms' m' rs2 m2 t S'' rs1 m1 tbb tbdy2 tex cs2,
+ MB.body bb' = nil ->
+ (forall ef args res, MB.exit bb' <> Some (MBbuiltin ef args res)) ->
+ Genv.find_funct_ptr tge fb = Some (Internal fn) ->
+ pstate cs2 = (Asmvliw.State rs2 m2) ->
+ pbody1 cs2 = nil -> pbody2 cs2 = tbdy2 -> pctl cs2 = tex ->
+ cur cs2 = tbb ->
+ match_codestate fb (MB.State s fb sp (bb'::c) ms' m') cs2 ->
+ match_asmstate fb cs2 (Asmvliw.State rs1 m1) ->
+ exit_step return_address_offset ge (MB.exit bb') (MB.State s fb sp (bb'::c) ms' m') t S'' ->
+ (exists rs3 m3 rs4 m4,
+ exec_body tge tbdy2 rs2 m2 = Next rs3 m3
+ /\ exec_control_rel tge fn tex tbb rs3 m3 rs4 m4
+ /\ match_states S'' (State rs4 m4)).
+Proof.
+ intros until cs2. intros Hbody Hbuiltin FIND Hpstate Hpbody1 Hpbody2 Hpctl Hcur MCS MAS ESTEP.
+ inv ESTEP.
+ - inv MCS. inv MAS. simpl in *.
+ inv Hpstate.
+ destruct ctl.
+ + (* MBcall *)
+ destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
+ inv TBC. inv TIC. inv H0.
+
+ assert (f0 = f) by congruence. subst f0.
+ assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto.
+ destruct s1 as [rf|fid]; simpl in H7.
+ * (* Indirect call *)
+ monadInv H1.
+ assert (ms' rf = Vptr f' Ptrofs.zero).
+ { unfold find_function_ptr in H14. destruct (ms' rf); try discriminate.
+ revert H14; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. }
+ assert (rs2 x = Vptr f' Ptrofs.zero).
+ { exploit ireg_val; eauto. rewrite H; intros LD; inv LD; auto. }
+ generalize (code_tail_next_int _ _ _ _ NOOV TAIL). intro CT1.
+ remember (Ptrofs.add _ _) as ofs'.
+ assert (TCA: transl_code_at_pc ge (Vptr fb ofs') fb f c false tf tc).
+ { econstructor; eauto. }
+ assert (f1 = f) by congruence. subst f1.
+ exploit return_address_offset_correct; eauto. intros; subst ra.
+
+ repeat eexists.
+ rewrite H6. econstructor; eauto.
+ rewrite H7. econstructor; eauto.
+ econstructor; eauto.
+ econstructor; eauto. eapply agree_sp_def; eauto. simpl. eapply agree_exten; eauto. intros. Simpl.
+ simpl. Simpl. rewrite PCeq. rewrite Heqofs'. simpl. auto.
+
+ * (* Direct call *)
+ monadInv H1.
+ generalize (code_tail_next_int _ _ _ _ NOOV TAIL). intro CT1.
+ remember (Ptrofs.add _ _) as ofs'.
+ assert (TCA: transl_code_at_pc ge (Vptr fb ofs') fb f c false tf tc).
+ econstructor; eauto.
+ assert (f1 = f) by congruence. subst f1.
+ exploit return_address_offset_correct; eauto. intros; subst ra.
+ repeat eexists.
+ rewrite H6. econstructor; eauto.
+ rewrite H7. econstructor; eauto.
+ econstructor; eauto.
+ econstructor; eauto. eapply agree_sp_def; eauto. simpl. eapply agree_exten; eauto. intros. Simpl.
+ Simpl. unfold Genv.symbol_address. rewrite symbols_preserved. simpl in H14. rewrite H14. auto.
+ Simpl. simpl. subst. Simpl. simpl. unfold Val.offset_ptr. rewrite PCeq. auto.
+ + (* MBtailcall *)
+ destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
+ inv TBC. inv TIC. inv H0.
+
+ assert (f0 = f) by congruence. subst f0.
+ assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto.
+ exploit Mem.loadv_extends. eauto. eexact H15. auto. simpl. intros [parent' [A B]].
+ destruct s1 as [rf|fid]; simpl in H13.
+ * monadInv H1.
+ assert (ms' rf = Vptr f' Ptrofs.zero).
+ { destruct (ms' rf); try discriminate. revert H13. predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence. }
+ assert (rs2 x = Vptr f' Ptrofs.zero).
+ { exploit ireg_val; eauto. rewrite H; intros LD; inv LD; auto. }
+
+ assert (f = f1) by congruence. subst f1. clear FIND1. clear H14.
+ exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
+ exploit exec_straight_body; eauto.
+ { simpl. eauto. }
+ intros EXEB.
+ repeat eexists.
+ rewrite H6. simpl extract_basic. eauto.
+ rewrite H7. simpl extract_ctl. simpl. reflexivity.
+ econstructor; eauto.
+ { apply agree_set_other.
+ - econstructor; auto with asmgen.
+ + apply V.
+ + intro r. destruct r; apply V; auto.
+ - eauto with asmgen. }
+ assert (IR x <> IR GPR12 /\ IR x <> IR GPR32 /\ IR x <> IR GPR16).
+ { clear - EQ. destruct x; repeat split; try discriminate.
+ all: unfold ireg_of in EQ; destruct rf; try discriminate. }
+ Simpl. inv H1. inv H3. rewrite Z; auto; try discriminate.
+ * monadInv H1. assert (f = f1) by congruence. subst f1. clear FIND1. clear H14.
+ exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
+ exploit exec_straight_body; eauto.
+ simpl. eauto.
+ intros EXEB.
+ repeat eexists.
+ rewrite H6. simpl extract_basic. eauto.
+ rewrite H7. simpl extract_ctl. simpl. reflexivity.
+ econstructor; eauto.
+ { apply agree_set_other.
+ - econstructor; auto with asmgen.
+ + apply V.
+ + intro r. destruct r; apply V; auto.
+ - eauto with asmgen. }
+ { Simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H13. auto. }
+ + (* MBbuiltin (contradiction) *)
+ assert (MB.exit bb' <> Some (MBbuiltin e l b)) by (apply Hbuiltin).
+ rewrite <- H in H1. contradict H1; auto.
+ + (* MBgoto *)
+ destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
+ inv TBC. inv TIC. inv H0.
+
+ assert (f0 = f) by congruence. subst f0. assert (f1 = f) by congruence. subst f1. clear H11.
+ remember (nextblock tbb rs2) as rs2'.
+ exploit functions_transl. eapply FIND0. eapply TRANSF0. intros FIND'.
+ assert (tf = fn) by congruence. subst tf.
+ exploit find_label_goto_label.
+ eauto. eauto.
+ instantiate (2 := rs2').
+ { subst. unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite PCeq. eauto. }
+ eauto.
+ intros (tc' & rs' & GOTO & AT2 & INV).
+
+ eexists. eexists. repeat eexists. repeat split.
+ rewrite H6. simpl extract_basic. simpl. eauto.
+ rewrite H7. simpl extract_ctl. simpl. rewrite <- Heqrs2'. eauto.
+ econstructor; eauto.
+ rewrite Heqrs2' in INV. unfold nextblock, incrPC in INV.
+ eapply agree_exten; eauto with asmgen.
+ assert (forall r : preg, r <> PC -> rs' r = rs2 r).
+ { intros. destruct r.
+ - destruct g. all: rewrite INV; Simpl; auto.
+ - rewrite INV; Simpl; auto.
+ - contradiction. }
+ eauto with asmgen.
+ congruence.
+ + (* MBcond *)
+ destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
+ inv TBC. inv TIC. inv H0.
+
+ * (* MBcond true *)
+ assert (f0 = f) by congruence. subst f0.
+ exploit eval_condition_lessdef.
+ eapply preg_vals; eauto.
+ all: eauto.
+ intros EC.
+ exploit transl_cbranch_correct_true; eauto. intros (rs' & jmp & A & B & C).
+ exploit exec_straight_opt_body2. eauto. intros (bdy & EXEB & BTC).
+ assert (PCeq': rs2 PC = rs' PC). { inv A; auto. erewrite <- exec_straight_pc. 2: eapply H. eauto. }
+ rewrite PCeq' in PCeq.
+ assert (f1 = f) by congruence. subst f1.
+ exploit find_label_goto_label.
+ 4: eapply H16. 1-2: eauto. instantiate (2 := (nextblock tbb rs')). rewrite nextblock_pc.
+ unfold Val.offset_ptr. rewrite PCeq. eauto.
+ intros (tc' & rs3 & GOTOL & TLPC & Hrs3).
+ exploit functions_transl. eapply FIND1. eapply TRANSF0. intros FIND'.
+ assert (tf = fn) by congruence. subst tf.
+
+ repeat eexists.
+ rewrite H6. rewrite <- BTC. rewrite extract_basics_to_code. simpl. rewrite app_nil_r. eauto.
+ rewrite H7. rewrite <- BTC. rewrite extract_ctl_basics_to_code. simpl extract_ctl. rewrite B. eauto.
+
+ econstructor; eauto.
+ eapply agree_exten with rs2; eauto with asmgen.
+ { intros. destruct r; try destruct g; try discriminate.
+ all: rewrite Hrs3; try discriminate; unfold nextblock, incrPC; Simpl. }
+ intros. discriminate.
+
+ * (* MBcond false *)
+ assert (f0 = f) by congruence. subst f0.
+ exploit eval_condition_lessdef.
+ eapply preg_vals; eauto.
+ all: eauto.
+ intros EC.
+
+ exploit transl_cbranch_correct_false; eauto. intros (rs' & jmp & A & B & C).
+ exploit exec_straight_opt_body2. eauto. intros (bdy & EXEB & BTC).
+ assert (PCeq': rs2 PC = rs' PC). { inv A; auto. erewrite <- exec_straight_pc. 2: eapply H. eauto. }
+ rewrite PCeq' in PCeq.
+ exploit functions_transl. eapply FIND1. eapply TRANSF0. intros FIND'.
+ assert (tf = fn) by congruence. subst tf.
+
+ assert (NOOV: size_blocks fn.(fn_blocks) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto.
+ generalize (code_tail_next_int _ _ _ _ NOOV TAIL). intro CT1.
+
+ repeat eexists.
+ rewrite H6. rewrite <- BTC. rewrite extract_basics_to_code. simpl. rewrite app_nil_r. eauto.
+ rewrite H7. rewrite <- BTC. rewrite extract_ctl_basics_to_code. simpl extract_ctl. rewrite B. eauto.
+
+ econstructor; eauto.
+ unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite PCeq. econstructor; eauto.
+ eapply agree_exten with rs2; eauto with asmgen.
+ { intros. destruct r; try destruct g; try discriminate.
+ all: rewrite <- C; try discriminate; unfold nextblock, incrPC; Simpl. }
+ intros. discriminate.
+ + (* MBjumptable *)
+ destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
+ inv TBC. inv TIC. inv H0.
+
+ assert (f0 = f) by congruence. subst f0.
+ monadInv H1.
+ generalize (transf_function_no_overflow _ _ TRANSF0); intro NOOV.
+ assert (f1 = f) by congruence. subst f1.
+ exploit find_label_goto_label. 4: eapply H16. 1-2: eauto. instantiate (2 := (nextblock tbb rs2) # GPR62 <- Vundef # GPR63 <- Vundef).
+ unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite PCeq. reflexivity.
+ exploit functions_transl. eapply FIND0. eapply TRANSF0. intros FIND3. assert (fn = tf) by congruence. subst fn.
+
+ intros [tc' [rs' [A [B C]]]].
+ exploit ireg_val; eauto. rewrite H13. intros LD; inv LD.
+
+ repeat eexists.
+ rewrite H6. simpl extract_basic. simpl. eauto.
+ rewrite H7. simpl extract_ctl. simpl. Simpl. rewrite <- H1. unfold Mach.label in H14. unfold label. rewrite H14. eapply A.
+ econstructor; eauto.
+ eapply agree_undef_regs; eauto. intros. rewrite C; auto with asmgen.
+ { assert (destroyed_by_jumptable = R62 :: R63 :: nil) by auto. rewrite H2 in H0. simpl in H0. inv H0.
+ destruct (preg_eq r' GPR63). subst. contradiction.
+ destruct (preg_eq r' GPR62). subst. contradiction.
+ destruct r'; Simpl. }
+ discriminate.
+ + (* MBreturn *)
+ destruct bb' as [mhd' mbdy' mex']; simpl in *. subst.
+ inv TBC. inv TIC. inv H0.
+
+ assert (f0 = f) by congruence. subst f0.
+ assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto.
+ exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
+ exploit exec_straight_body; eauto.
+ simpl. eauto.
+ intros EXEB.
+ assert (f1 = f) by congruence. subst f1.
+
+ repeat eexists.
+ rewrite H6. simpl extract_basic. eauto.
+ rewrite H7. simpl extract_ctl. simpl. reflexivity.
+ econstructor; eauto.
+ unfold nextblock, incrPC. repeat apply agree_set_other; auto with asmgen.
+
+ - inv MCS. inv MAS. simpl in *. subst. inv Hpstate.
+ destruct bb' as [hd' bdy' ex']; simpl in *. subst.
+ monadInv TBC. monadInv TIC. simpl in *. rewrite H5. rewrite H6.
+ simpl. repeat eexists.
+ econstructor. 4: instantiate (3 := false). all:eauto.
+ unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite PCeq.
+ assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto.
+ assert (f = f0) by congruence. subst f0. econstructor; eauto.
+ generalize (code_tail_next_int _ _ _ _ NOOV TAIL). intro CT1. eauto.
+ eapply agree_exten; eauto. intros. Simpl.
+ discriminate.
+Qed.
+
+Definition mb_remove_first (bb: MB.bblock) :=
+ {| MB.header := MB.header bb; MB.body := tail (MB.body bb); MB.exit := MB.exit bb |}.
+
+Lemma exec_straight_body:
+ forall c c' lc rs1 m1 rs2 m2,
+ exec_straight tge c rs1 m1 c' rs2 m2 ->
+ code_to_basics c = Some lc ->
+ exists l ll,
+ c = l ++ c'
+ /\ code_to_basics l = Some ll
+ /\ exec_body tge ll rs1 m1 = Next rs2 m2.
+Proof.
+ induction c; try (intros; inv H; fail).
+ intros until m2. intros EXES CTB. inv EXES.
+ - exists (i1 ::g nil),(i1::nil). repeat (split; simpl; auto). rewrite H6. auto.
+ - inv CTB. destruct (code_to_basics c); try discriminate. inv H0.
+ eapply IHc in H7; eauto. destruct H7 as (l' & ll & Hc & CTB & EXECB). subst.
+ exists (i ::g l'),(i::ll). repeat (split; simpl; auto).
+ rewrite CTB. auto.
+ rewrite H1. auto.
+Qed.
+
+Lemma basics_to_code_app:
+ forall c l x ll,
+ basics_to_code c = l ++ basics_to_code x ->
+ code_to_basics l = Some ll ->
+ c = ll ++ x.
+Proof.
+ intros. apply (f_equal code_to_basics) in H.
+ erewrite code_to_basics_dist in H; eauto. 2: eapply code_to_basics_id.
+ rewrite code_to_basics_id in H. inv H. auto.
+Qed.
+
+Lemma basics_to_code_app2:
+ forall i c l x ll,
+ (PBasic i) :: basics_to_code c = l ++ basics_to_code x ->
+ code_to_basics l = Some ll ->
+ i :: c = ll ++ x.
+Proof.
+ intros until ll. intros.
+ exploit basics_to_code_app. instantiate (3 := (i::c)). simpl.
+ all: eauto.
+Qed.
+
+(* Handling the individual instructions of theorem (B) in the above diagram. A bit less cumbersome, but still tough *)
+Theorem step_simu_basic:
+ forall bb bb' s fb sp c ms m rs1 m1 ms' m' bi cs1 tbdy bdy,
+ MB.header bb = nil -> MB.body bb = bi::(bdy) -> (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
+ bb' = {| MB.header := nil; MB.body := bdy; MB.exit := MB.exit bb |} ->
+ basic_step ge s fb sp ms m bi ms' m' ->
+ pstate cs1 = (State rs1 m1) -> pbody1 cs1 = tbdy ->
+ match_codestate fb (MB.State s fb sp (bb::c) ms m) cs1 ->
+ (exists rs2 m2 l cs2 tbdy',
+ cs2 = {| pstate := (State rs2 m2); pheader := nil; pbody1 := tbdy'; pbody2 := pbody2 cs1;
+ pctl := pctl cs1; ep := fp_is_parent (ep cs1) bi; rem := rem cs1; cur := cur cs1 |}
+ /\ tbdy = l ++ tbdy'
+ /\ exec_body tge l rs1 m1 = Next rs2 m2
+ /\ match_codestate fb (MB.State s fb sp (bb'::c) ms' m') cs2).
+Proof.
+ intros until bdy. intros Hheader Hbody Hnobuiltin (* Hnotempty *) Hbb' BSTEP Hpstate Hpbody1 MCS. inv MCS.
+ simpl in *. inv Hpstate.
+ rewrite Hbody in TBC. monadInv TBC.
+ inv BSTEP.
+
+ - (* MBgetstack *)
+ simpl in EQ0.
+ unfold Mach.load_stack in H.
+ exploit Mem.loadv_extends; eauto. intros [v' [A B]].
+ rewrite (sp_val _ _ _ AG) in A.
+ exploit loadind_correct; eauto with asmgen.
+ intros (rs2 & EXECS & Hrs'1 & Hrs'2).
+ eapply exec_straight_body in EXECS.
+ 2: eapply code_to_basics_id; eauto.
+ destruct EXECS as (l & Hlbi & BTC & CTB & EXECB).
+ exists rs2, m1, Hlbi.
+ eexists. eexists. split. instantiate (1 := x). eauto.
+ repeat (split; auto).
+ eapply basics_to_code_app; eauto.
+ remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
+ assert (Hheadereq: MB.header bb' = MB.header bb). { subst. simpl. auto. }
+ subst. simpl in Hheadereq.
+
+ eapply match_codestate_intro; eauto.
+ { simpl. simpl in EQ. rewrite <- Hheadereq in EQ. assumption. }
+ eapply agree_set_mreg; eauto with asmgen.
+ intro Hep. simpl in Hep.
+ destruct (andb_prop _ _ Hep). clear Hep.
+ rewrite <- Hheadereq in DXP. subst. rewrite <- DXP. rewrite Hrs'2. reflexivity.
+ discriminate. apply preg_of_not_FP; assumption. reflexivity.
+
+ - (* MBsetstack *)
+ simpl in EQ0.
+ unfold Mach.store_stack in H.
+ assert (Val.lessdef (ms src) (rs1 (preg_of src))). { eapply preg_val; eauto. }
+ exploit Mem.storev_extends; eauto. intros [m2' [A B]].
+ exploit storeind_correct; eauto with asmgen.
+ rewrite (sp_val _ _ _ AG) in A. eauto. intros [rs' [P Q]].
+
+ eapply exec_straight_body in P.
+ 2: eapply code_to_basics_id; eauto.
+ destruct P as (l & ll & TBC & CTB & EXECB).
+ exists rs', m2', ll.
+ eexists. eexists. split. instantiate (1 := x). eauto.
+ repeat (split; auto).
+ eapply basics_to_code_app; eauto.
+ remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
+ subst.
+ eapply match_codestate_intro; eauto. simpl. simpl in EQ. rewrite Hheader in EQ. auto.
+
+ eapply agree_undef_regs; eauto with asmgen.
+ simpl; intros. rewrite Q; auto with asmgen. rewrite Hheader in DXP. auto.
+ - (* MBgetparam *)
+ simpl in EQ0.
+
+ assert (f0 = f) by congruence; subst f0.
+ unfold Mach.load_stack in *.
+ exploit Mem.loadv_extends. eauto. eexact H0. auto.
+ intros [parent' [A B]]. rewrite (sp_val _ _ _ AG) in A.
+ exploit lessdef_parent_sp; eauto. clear B; intros B; subst parent'.
+ exploit Mem.loadv_extends. eauto. eexact H1. auto.
+ intros [v' [C D]].
+
+ monadInv EQ0. rewrite Hheader. rewrite Hheader in DXP.
+ destruct ep0 eqn:EPeq.
+
+ (* RTMP contains parent *)
+ + exploit loadind_correct. eexact EQ1.
+ instantiate (2 := rs1). rewrite DXP; eauto.
+ intros [rs2 [P [Q R]]].
+
+ eapply exec_straight_body in P.
+ 2: eapply code_to_basics_id; eauto.
+ destruct P as (l & ll & BTC & CTB & EXECB).
+ exists rs2, m1, ll. eexists.
+ eexists. split. instantiate (1 := x). eauto.
+ repeat (split; auto).
+ { eapply basics_to_code_app; eauto. }
+ remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
+ assert (Hheadereq: MB.header bb' = MB.header bb). { subst. simpl. auto. }
+ subst.
+ eapply match_codestate_intro; eauto.
+
+ eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto with asmgen.
+ simpl; intros. rewrite R; auto with asmgen.
+ apply preg_of_not_FP; auto.
+
+ (* RTMP does not contain parent *)
+ + rewrite chunk_of_Tptr in A.
+ exploit loadind_ptr_correct. eexact A. intros [rs2 [P [Q R]]].
+ exploit loadind_correct. eexact EQ1. instantiate (2 := rs2). rewrite Q. eauto.
+ intros [rs3 [S [T U]]].
+
+ exploit exec_straight_trans.
+ eapply P.
+ eapply S.
+ intros EXES.
+
+ eapply exec_straight_body in EXES.
+ 2: simpl. 2: erewrite code_to_basics_id; eauto.
+ destruct EXES as (l & ll & BTC & CTB & EXECB).
+ exists rs3, m1, ll.
+ eexists. eexists. split. instantiate (1 := x). eauto.
+ repeat (split; auto).
+ eapply basics_to_code_app2; eauto.
+ remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
+ assert (Hheadereq: MB.header bb' = MB.header bb). { subst. auto. }
+ subst.
+ eapply match_codestate_intro; eauto.
+ eapply agree_set_mreg. eapply agree_set_mreg. eauto. eauto.
+ instantiate (1 := rs2#FP <- (rs3#FP)). intros.
+ rewrite Pregmap.gso; auto with asmgen.
+ congruence.
+ intros. unfold Pregmap.set. destruct (PregEq.eq r' FP). congruence. auto with asmgen.
+ simpl; intros. rewrite U; auto with asmgen.
+ apply preg_of_not_FP; auto.
+ - (* MBop *)
+ simpl in EQ0. rewrite Hheader in DXP.
+
+ assert (eval_operation tge sp op (map ms args) m' = Some v).
+ rewrite <- H. apply eval_operation_preserved. exact symbols_preserved.
+ exploit eval_operation_lessdef.
+ eapply preg_vals; eauto.
+ 2: eexact H0.
+ all: eauto.
+ intros [v' [A B]]. rewrite (sp_val _ _ _ AG) in A.
+ exploit transl_op_correct; eauto. intros [rs2 [P [Q R]]].
+
+ eapply exec_straight_body in P.
+ 2: eapply code_to_basics_id; eauto.
+ destruct P as (l & ll & TBC & CTB & EXECB).
+ exists rs2, m1, ll.
+ eexists. eexists. split. instantiate (1 := x). eauto.
+ repeat (split; auto).
+ eapply basics_to_code_app; eauto.
+ remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
+ subst.
+ eapply match_codestate_intro; eauto. simpl. simpl in EQ. rewrite Hheader in EQ. auto.
+ apply agree_set_undef_mreg with rs1; auto.
+ apply Val.lessdef_trans with v'; auto.
+ simpl; intros. destruct (andb_prop _ _ H1); clear H1.
+ rewrite R; auto. apply preg_of_not_FP; auto.
+Local Transparent destroyed_by_op.
+ destruct op; simpl; auto; congruence.
+ - (* MBload *)
+ simpl in EQ0. rewrite Hheader in DXP.
+
+ assert (eval_addressing tge sp addr (map ms args) = Some a).
+ rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
+ exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1.
+ intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A.
+ exploit Mem.loadv_extends; eauto. intros [v' [C D]].
+ exploit transl_load_correct; eauto.
+ intros [rs2 [P [Q R]]].
+
+ eapply exec_straight_body in P.
+ 2: eapply code_to_basics_id; eauto.
+ destruct P as (l & ll & TBC & CTB & EXECB).
+ exists rs2, m1, ll.
+ eexists. eexists. split. instantiate (1 := x). eauto.
+ repeat (split; auto).
+ eapply basics_to_code_app; eauto.
+ remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
+ assert (Hheadereq: MB.header bb' = MB.header bb). { subst. auto. }
+ subst.
+ eapply match_codestate_intro; eauto. simpl. simpl in EQ.
+ rewrite <- Hheadereq in EQ. assumption.
+ eapply agree_set_mreg; eauto with asmgen.
+ intro Hep. simpl in Hep.
+ destruct (andb_prop _ _ Hep). clear Hep.
+ subst. rewrite <- DXP. rewrite R; try discriminate. reflexivity.
+ apply preg_of_not_FP; assumption. reflexivity.
+
+ - (* MBstore *)
+ simpl in EQ0. rewrite Hheader in DXP.
+
+ assert (eval_addressing tge sp addr (map ms args) = Some a).
+ rewrite <- H. apply eval_addressing_preserved. exact symbols_preserved.
+ exploit eval_addressing_lessdef. eapply preg_vals; eauto. eexact H1.
+ intros [a' [A B]]. rewrite (sp_val _ _ _ AG) in A.
+ assert (Val.lessdef (ms src) (rs1 (preg_of src))). eapply preg_val; eauto.
+ exploit Mem.storev_extends; eauto. intros [m2' [C D]].
+ exploit transl_store_correct; eauto. intros [rs2 [P Q]].
+
+ eapply exec_straight_body in P.
+ 2: eapply code_to_basics_id; eauto.
+ destruct P as (l & ll & TBC & CTB & EXECB).
+ exists rs2, m2', ll.
+ eexists. eexists. split. instantiate (1 := x). eauto.
+ repeat (split; auto).
+ eapply basics_to_code_app; eauto.
+ remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb'.
+ assert (Hheadereq: MB.header bb' = MB.header bb). { subst. auto. }
+ subst.
+ eapply match_codestate_intro; eauto. simpl. simpl in EQ.
+ rewrite <- Hheadereq in EQ. assumption.
+ eapply agree_undef_regs; eauto with asmgen.
+ intro Hep. simpl in Hep.
+ subst. rewrite <- DXP. rewrite Q; try discriminate. reflexivity. reflexivity.
+Qed.
+
+Lemma exec_body_trans:
+ forall l l' rs0 m0 rs1 m1 rs2 m2,
+ exec_body tge l rs0 m0 = Next rs1 m1 ->
+ exec_body tge l' rs1 m1 = Next rs2 m2 ->
+ exec_body tge (l++l') rs0 m0 = Next rs2 m2.
+Proof.
+ induction l.
+ - simpl. congruence.
+ - intros until m2. intros EXEB1 EXEB2.
+ inv EXEB1. destruct (exec_basic_instr _) eqn:EBI; try discriminate.
+ simpl. rewrite EBI. eapply IHl; eauto.
+Qed.
+
+Definition mb_remove_header bb := {| MB.header := nil; MB.body := MB.body bb; MB.exit := MB.exit bb |}.
+
+Program Definition remove_header tbb := {| header := nil; body := body tbb; exit := exit tbb |}.
+Next Obligation.
+ destruct tbb. simpl. auto.
+Qed.
+
+Inductive exec_header: codestate -> codestate -> Prop :=
+ | exec_header_cons: forall cs1,
+ exec_header cs1 {| pstate := pstate cs1; pheader := nil; pbody1 := pbody1 cs1; pbody2 := pbody2 cs1;
+ pctl := pctl cs1; ep := (if pheader cs1 then ep cs1 else false); rem := rem cs1;
+ cur := cur cs1 |}.
+
+(* Theorem (A) in the diagram, the easiest of all *)
+Theorem step_simu_header:
+ forall bb s fb sp c ms m rs1 m1 cs1,
+ pstate cs1 = (State rs1 m1) ->
+ match_codestate fb (MB.State s fb sp (bb::c) ms m) cs1 ->
+ (exists cs1',
+ exec_header cs1 cs1'
+ /\ match_codestate fb (MB.State s fb sp (mb_remove_header bb::c) ms m) cs1').
+Proof.
+ intros until cs1. intros Hpstate MCS.
+ eexists. split; eauto.
+ econstructor; eauto.
+ inv MCS. simpl in *. inv Hpstate.
+ econstructor; eauto.
+Qed.
+
+Lemma step_matchasm_header:
+ forall fb cs1 cs1' s1,
+ match_asmstate fb cs1 s1 ->
+ exec_header cs1 cs1' ->
+ match_asmstate fb cs1' s1.
+Proof.
+ intros until s1. intros MAS EXH.
+ inv MAS. inv EXH.
+ simpl. econstructor; eauto.
+Qed.
+
+(* Theorem (B) in the diagram, using step_simu_basic + induction on the Machblock body *)
+Theorem step_simu_body:
+ forall bb s fb sp c ms m rs1 m1 ms' cs1 m',
+ MB.header bb = nil ->
+ (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
+ body_step ge s fb sp (MB.body bb) ms m ms' m' ->
+ pstate cs1 = (State rs1 m1) ->
+ match_codestate fb (MB.State s fb sp (bb::c) ms m) cs1 ->
+ (exists rs2 m2 cs2 ep,
+ cs2 = {| pstate := (State rs2 m2); pheader := nil; pbody1 := nil; pbody2 := pbody2 cs1;
+ pctl := pctl cs1; ep := ep; rem := rem cs1; cur := cur cs1 |}
+ /\ exec_body tge (pbody1 cs1) rs1 m1 = Next rs2 m2
+ /\ match_codestate fb (MB.State s fb sp ({| MB.header := nil; MB.body := nil; MB.exit := MB.exit bb |}::c) ms' m') cs2).
+Proof.
+ intros bb. destruct bb as [hd bdy ex]; simpl; auto. induction bdy as [|bi bdy].
+ - intros until m'. intros Hheader Hnobuiltin BSTEP Hpstate MCS.
+ inv BSTEP.
+ exists rs1, m1, cs1, (ep cs1).
+ inv MCS. inv Hpstate. simpl in *. monadInv TBC. repeat (split; simpl; auto).
+ econstructor; eauto.
+ - intros until m'. intros Hheader Hnobuiltin BSTEP Hpstate MCS. inv BSTEP.
+ rename ms' into ms''. rename m' into m''. rename rs' into ms'. rename m'0 into m'.
+ exploit (step_simu_basic); eauto. simpl. eauto. simpl; auto. simpl; auto.
+ intros (rs2 & m2 & l & cs2 & tbdy' & Hcs2 & Happ & EXEB & MCS').
+ simpl in *.
+ exploit IHbdy. auto. 2: eapply H6. 3: eapply MCS'. all: eauto. subst; eauto. simpl; auto.
+ intros (rs3 & m3 & cs3 & ep & Hcs3 & EXEB' & MCS'').
+ exists rs3, m3, cs3, ep.
+ repeat (split; simpl; auto). subst. simpl in *. auto.
+ rewrite Happ. eapply exec_body_trans; eauto. rewrite Hcs2 in EXEB'; simpl in EXEB'. auto.
+Qed.
+
+Lemma exec_body_pc:
+ forall l rs1 m1 rs2 m2,
+ exec_body tge l rs1 m1 = Next rs2 m2 ->
+ rs2 PC = rs1 PC.
+Proof.
+ induction l.
+ - intros. inv H. auto.
+ - intros until m2. intro EXEB.
+ inv EXEB. destruct (exec_basic_instr _ _ _) eqn:EBI; try discriminate.
+ eapply IHl in H0. rewrite H0.
+ erewrite exec_basic_instr_pc; eauto.
+Qed.
+
+Lemma exec_body_control:
+ forall b rs1 m1 rs2 m2 rs3 m3 fn,
+ exec_body tge (body b) rs1 m1 = Next rs2 m2 ->
+ exec_control_rel tge fn (exit b) b rs2 m2 rs3 m3 ->
+ exec_bblock_rel tge fn b rs1 m1 rs3 m3.
+Proof.
+ intros until fn. intros EXEB EXECTL.
+ econstructor; eauto. inv EXECTL.
+ unfold exec_bblock. rewrite EXEB. auto.
+Qed.
+
+Definition mbsize (bb: MB.bblock) := (length (MB.body bb) + length_opt (MB.exit bb))%nat.
+
+Lemma mbsize_eqz:
+ forall bb, mbsize bb = 0%nat -> MB.body bb = nil /\ MB.exit bb = None.
+Proof.
+ intros. destruct bb as [hd bdy ex]; simpl in *. unfold mbsize in H.
+ remember (length _) as a. remember (length_opt _) as b.
+ assert (a = 0%nat) by omega. assert (b = 0%nat) by omega. subst. clear H.
+ inv H0. inv H1. destruct bdy; destruct ex; auto.
+ all: try discriminate.
+Qed.
+
+Lemma mbsize_neqz:
+ forall bb, mbsize bb <> 0%nat -> (MB.body bb <> nil \/ MB.exit bb <> None).
+Proof.
+ intros. destruct bb as [hd bdy ex]; simpl in *.
+ destruct bdy; destruct ex; try (right; discriminate); try (left; discriminate).
+ contradict H. unfold mbsize. simpl. auto.
+Qed.
+
+(* Bringing theorems (A), (B) and (C) together, for the case of the absence of builtin instruction *)
+(* This more general form is easier to prove, but the actual theorem is step_simulation_bblock further below *)
+Lemma step_simulation_bblock':
+ forall sf f sp bb bb' bb'' rs m rs' m' s'' c S1,
+ bb' = mb_remove_header bb ->
+ body_step ge sf f sp (Machblock.body bb') rs m rs' m' ->
+ bb'' = mb_remove_body bb' ->
+ (forall ef args res, MB.exit bb'' <> Some (MBbuiltin ef args res)) ->
+ exit_step return_address_offset ge (Machblock.exit bb'') (Machblock.State sf f sp (bb'' :: c) rs' m') E0 s'' ->
+ match_states (Machblock.State sf f sp (bb :: c) rs m) S1 ->
+ exists S2 : state, plus step tge S1 E0 S2 /\ match_states s'' S2.
+Proof.
+ intros until S1. intros Hbb' BSTEP Hbb'' Hbuiltin ESTEP MS.
+ destruct (mbsize bb) eqn:SIZE.
+ - apply mbsize_eqz in SIZE. destruct SIZE as (Hbody & Hexit).
+ destruct bb as [hd bdy ex]; simpl in *; subst.
+ inv MS. inv AT. exploit transl_blocks_nonil; eauto. intros (tbb & tc' & Htc). subst. rename tc' into tc.
+ monadInv H2. simpl in *. inv ESTEP. inv BSTEP.
+ eexists. split. eapply plus_one.
+ exploit functions_translated; eauto. intros (tf0 & FIND' & TRANSF'). monadInv TRANSF'.
+ assert (x = tf) by congruence. subst x.
+ eapply exec_step_internal; eauto. eapply find_bblock_tail; eauto.
+ unfold exec_bblock. simpl. eauto.
+ econstructor. eauto. eauto. eauto.
+ unfold nextblock, incrPC. Simpl. unfold Val.offset_ptr. rewrite <- H.
+ assert (NOOV: size_blocks tf.(fn_blocks) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto.
+ econstructor; eauto.
+ generalize (code_tail_next_int _ _ _ _ NOOV H3). intro CT1. eauto.
+ eapply agree_exten; eauto. intros. Simpl.
+ intros. discriminate.
+ - subst. exploit mbsize_neqz. { instantiate (1 := bb). rewrite SIZE. discriminate. }
+ intros Hnotempty.
+
+ (* initial setting *)
+ exploit match_state_codestate.
+ 2: eapply Hnotempty.
+ all: eauto.
+ intros (cs1 & fb & f0 & tbb & tc & ep & MCS & MAS & FIND & TLBS & Hbody & Hexit & Hcur & Hrem & Hpstate).
+
+ (* step_simu_header part *)
+ assert (exists rs1 m1, pstate cs1 = State rs1 m1). { inv MAS. simpl. eauto. }
+ destruct H as (rs1 & m1 & Hpstate2). subst.
+ assert (f = fb). { inv MCS. auto. } subst fb.
+ exploit step_simu_header.
+ 2: eapply MCS.
+ all: eauto.
+ intros (cs1' & EXEH & MCS2).
+
+ (* step_simu_body part *)
+ assert (Hpstate': pstate cs1' = pstate cs1). { inv EXEH; auto. }
+ exploit step_simu_body.
+ 3: eapply BSTEP.
+ 4: eapply MCS2.
+ all: eauto. rewrite Hpstate'. eauto.
+ intros (rs2 & m2 & cs2 & ep' & Hcs2 & EXEB & MCS').
+
+ (* step_simu_control part *)
+ assert (exists tf, Genv.find_funct_ptr tge f = Some (Internal tf)).
+ { exploit functions_translated; eauto. intros (tf & FIND' & TRANSF'). monadInv TRANSF'. eauto. }
+ destruct H as (tf & FIND').
+ assert (exists tex, pbody2 cs1 = extract_basic tex /\ pctl cs1 = extract_ctl tex).
+ { inv MAS. simpl in *. eauto. }
+ destruct H as (tex & Hpbody2 & Hpctl).
+ inv EXEH. simpl in *.
+ subst. exploit step_simu_control.
+ 9: eapply MCS'. all: simpl.
+ 10: eapply ESTEP.
+ all: simpl; eauto.
+ rewrite Hpbody2. rewrite Hpctl.
+ { inv MAS; simpl in *. inv Hpstate2. eapply match_asmstate_some; eauto.
+ erewrite exec_body_pc; eauto. }
+ intros (rs3 & m3 & rs4 & m4 & EXEB' & EXECTL' & MS').
+
+ (* bringing the pieces together *)
+ exploit exec_body_trans.
+ eapply EXEB.
+ eauto.
+ intros EXEB2.
+ exploit exec_body_control; eauto.
+ rewrite <- Hpbody2 in EXEB2. rewrite <- Hbody in EXEB2. eauto.
+ rewrite Hexit. rewrite Hpctl. eauto.
+ intros EXECB. inv EXECB.
+ exists (State rs4 m4).
+ split; auto. eapply plus_one. rewrite Hpstate2.
+ assert (exists ofs, rs1 PC = Vptr f ofs).
+ { rewrite Hpstate2 in MAS. inv MAS. simpl in *. eauto. }
+ destruct H0 as (ofs & Hrs1pc).
+ eapply exec_step_internal; eauto.
+
+ (* proving the initial find_bblock *)
+ rewrite Hpstate2 in MAS. inv MAS. simpl in *.
+ assert (f1 = f0) by congruence. subst f0.
+ rewrite PCeq in Hrs1pc. inv Hrs1pc.
+ exploit functions_translated; eauto. intros (tf1 & FIND'' & TRANS''). rewrite FIND' in FIND''.
+ inv FIND''. monadInv TRANS''. rewrite TRANSF0 in EQ. inv EQ.
+ eapply find_bblock_tail; eauto.
+Qed.
+
+Theorem step_simulation_bblock:
+ forall sf f sp bb ms m ms' m' S2 c,
+ body_step ge sf f sp (Machblock.body bb) ms m ms' m' ->
+ (forall ef args res, MB.exit bb <> Some (MBbuiltin ef args res)) ->
+ exit_step return_address_offset ge (Machblock.exit bb) (Machblock.State sf f sp (bb :: c) ms' m') E0 S2 ->
+ forall S1', match_states (Machblock.State sf f sp (bb :: c) ms m) S1' ->
+ exists S2' : state, plus step tge S1' E0 S2' /\ match_states S2 S2'.
+Proof.
+ intros until c. intros BSTEP Hbuiltin ESTEP S1' MS.
+ eapply step_simulation_bblock'; eauto.
+ all: destruct bb as [hd bdy ex]; simpl in *; eauto.
+ inv ESTEP.
+ - econstructor. inv H; try (econstructor; eauto; fail).
+ - econstructor.
+Qed.
+
+(** Dealing now with the builtin case *)
+
+Definition split (c: MB.code) :=
+ match c with
+ | nil => nil
+ | bb::c => {| MB.header := MB.header bb; MB.body := MB.body bb; MB.exit := None |}
+ :: {| MB.header := nil; MB.body := nil; MB.exit := MB.exit bb |} :: c
+ end.
+
+Lemma cons_ok_eq3 {A: Type} :
+ forall (x:A) y z x' y' z',
+ x = x' -> y = y' -> z = z' ->
+ OK (x::y::z) = OK (x'::y'::z').
+Proof.
+ intros. subst. auto.
+Qed.
+
+Lemma transl_blocks_split_builtin:
+ forall bb c ep f ef args res,
+ MB.exit bb = Some (MBbuiltin ef args res) -> MB.body bb <> nil ->
+ transl_blocks f (split (bb::c)) ep = transl_blocks f (bb::c) ep.
+Proof.
+ intros until res. intros Hexit Hbody. simpl split.
+ unfold transl_blocks. fold transl_blocks. unfold transl_block.
+ simpl. remember (transl_basic_code _ _ _) as tbc. remember (transl_instr_control _ _) as tbi.
+ remember (transl_blocks _ _ _) as tlbs.
+ destruct tbc; destruct tbi; destruct tlbs.
+ all: try simpl; auto.
+ - simpl. rewrite Hexit in Heqtbi. simpl in Heqtbi. monadInv Heqtbi. simpl.
+ unfold gen_bblocks. simpl. destruct l.
+ + exploit transl_basic_code_nonil; eauto. intro. destruct H.
+ + simpl. rewrite app_nil_r. apply cons_ok_eq3. all: try eapply bblock_equality. all: simpl; auto.
+Qed.
+
+Lemma transl_code_at_pc_split_builtin:
+ forall rs f f0 bb c ep tf tc ef args res,
+ MB.body bb <> nil -> MB.exit bb = Some (MBbuiltin ef args res) ->
+ transl_code_at_pc ge (rs PC) f f0 (bb :: c) ep tf tc ->
+ transl_code_at_pc ge (rs PC) f f0 (split (bb :: c)) ep tf tc.
+Proof.
+ intros until res. intros Hbody Hexit AT. inv AT.
+ econstructor; eauto. erewrite transl_blocks_split_builtin; eauto.
+Qed.
+
+Theorem match_states_split_builtin:
+ forall sf f sp bb c rs m ef args res S1,
+ MB.body bb <> nil -> MB.exit bb = Some (MBbuiltin ef args res) ->
+ match_states (Machblock.State sf f sp (bb :: c) rs m) S1 ->
+ match_states (Machblock.State sf f sp (split (bb::c)) rs m) S1.
+Proof.
+ intros until S1. intros Hbody Hexit MS.
+ inv MS.
+ econstructor; eauto.
+ eapply transl_code_at_pc_split_builtin; eauto.
+Qed.
+
+Theorem step_simulation_builtin:
+ forall ef args res bb sf f sp c ms m t S2,
+ MB.body bb = nil -> MB.exit bb = Some (MBbuiltin ef args res) ->
+ exit_step return_address_offset ge (MB.exit bb) (Machblock.State sf f sp (bb :: c) ms m) t S2 ->
+ forall S1', match_states (Machblock.State sf f sp (bb :: c) ms m) S1' ->
+ exists S2' : state, plus step tge S1' t S2' /\ match_states S2 S2'.
+Proof.
+ intros until S2. intros Hbody Hexit ESTEP S1' MS.
+ inv MS. inv AT. monadInv H2. monadInv EQ.
+ rewrite Hbody in EQ0. monadInv EQ0.
+ rewrite Hexit in EQ. monadInv EQ.
+ rewrite Hexit in ESTEP. inv ESTEP. inv H4.
+
+ exploit functions_transl; eauto. intro FN.
+ generalize (transf_function_no_overflow _ _ H1); intro NOOV.
+ exploit builtin_args_match; eauto. intros [vargs' [P Q]].
+ exploit external_call_mem_extends; eauto.
+ intros [vres' [m2' [A [B [C D]]]]].
+ econstructor; split. apply plus_one.
+ simpl in H3.
+ eapply exec_step_builtin. eauto. eauto.
+ eapply find_bblock_tail; eauto.
+ simpl. eauto.
+ erewrite <- sp_val by eauto.
+ eapply eval_builtin_args_preserved with (ge1 := ge); eauto. exact symbols_preserved.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ eauto.
+ econstructor; eauto.
+ instantiate (2 := tf); instantiate (1 := x0).
+ unfold nextblock, incrPC. rewrite Pregmap.gss.
+ rewrite set_res_other. rewrite undef_regs_other_2. rewrite Pregmap.gso by congruence.
+ rewrite <- H. simpl. econstructor; eauto.
+ eapply code_tail_next_int; eauto.
+ rewrite preg_notin_charact. intros. auto with asmgen.
+ auto with asmgen.
+ apply agree_nextblock. eapply agree_set_res; auto.
+ eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto.
+ apply Pregmap.gso; auto with asmgen.
+ congruence.
+Qed.
+
+Lemma next_sep:
+ forall rs m rs' m', rs = rs' -> m = m' -> Next rs m = Next rs' m'.
+Proof.
+ congruence.
+Qed.
+
+(* Measure to prove finite stuttering, see the other backends *)
+Definition measure (s: MB.state) : nat :=
+ match s with
+ | MB.State _ _ _ _ _ _ => 0%nat
+ | MB.Callstate _ _ _ _ => 0%nat
+ | MB.Returnstate _ _ _ => 1%nat
+ end.
+
+(* The actual MB.step/AB.step simulation, using the above theorems, plus extra proofs
+ for the internal and external function cases *)
+Theorem step_simulation:
+ forall S1 t S2, MB.step return_address_offset ge S1 t S2 ->
+ forall S1' (MS: match_states S1 S1'),
+ (exists S2', plus step tge S1' t S2' /\ match_states S2 S2')
+ \/ (measure S2 < measure S1 /\ t = E0 /\ match_states S2 S1')%nat.
+Proof.
+ induction 1; intros.
+
+- (* bblock *)
+ left. destruct (Machblock.exit bb) eqn:MBE; try destruct c0.
+ all: try(inversion H0; subst; inv H2; eapply step_simulation_bblock;
+ try (rewrite MBE; try discriminate); eauto).
+ + (* MBbuiltin *)
+ destruct (MB.body bb) eqn:MBB.
+ * inv H. eapply step_simulation_builtin; eauto. rewrite MBE. eauto.
+ * eapply match_states_split_builtin in MS; eauto.
+ 2: rewrite MBB; discriminate.
+ simpl split in MS.
+ rewrite <- MBB in H.
+ remember {| MB.header := _; MB.body := _; MB.exit := _ |} as bb1.
+ assert (MB.body bb = MB.body bb1). { subst. simpl. auto. }
+ rewrite H1 in H. subst.
+ exploit step_simulation_bblock. eapply H.
+ discriminate.
+ simpl. constructor.
+ eauto.
+ intros (S2' & PLUS1 & MS').
+ rewrite MBE in MS'.
+ assert (exit_step return_address_offset ge (Some (MBbuiltin e l b))
+ (MB.State sf f sp ({| MB.header := nil; MB.body := nil; MB.exit := Some (MBbuiltin e l b) |}::c)
+ rs' m') t s').
+ { inv H0. inv H3. econstructor. econstructor; eauto. }
+ exploit step_simulation_builtin.
+ 4: eapply MS'.
+ all: simpl; eauto.
+ intros (S3' & PLUS'' & MS'').
+ exists S3'. split; eauto.
+ eapply plus_trans. eapply PLUS1. eapply PLUS''. eauto.
+ + inversion H0. subst. eapply step_simulation_bblock; try (rewrite MBE; try discriminate); eauto.
+
+- (* internal function *)
+ inv MS.
+ exploit functions_translated; eauto. intros [tf [A B]]. monadInv B.
+ generalize EQ; intros EQ'. monadInv EQ'.
+ destruct (zlt Ptrofs.max_unsigned (size_blocks x0.(fn_blocks))); inversion EQ1. clear EQ1. subst x0.
+ unfold Mach.store_stack in *.
+ exploit Mem.alloc_extends. eauto. eauto. apply Z.le_refl. apply Z.le_refl.
+ intros [m1' [C D]].
+ exploit Mem.storev_extends. eexact D. eexact H1. eauto. eauto.
+ intros [m2' [F G]].
+ simpl chunk_of_type in F.
+ exploit Mem.storev_extends. eexact G. eexact H2. eauto. eauto.
+ intros [m3' [P Q]].
+ (* Execution of function prologue *)
+ monadInv EQ0.
+ set (tfbody := make_prologue f x0) in *.
+ set (tf := {| fn_sig := MB.fn_sig f; fn_blocks := tfbody |}) in *.
+ set (rs2 := rs0#FP <- (parent_sp s) #SP <- sp #RTMP <- Vundef).
+ exploit (Pget_correct tge GPRA RA nil rs2 m2'); auto.
+ intros (rs' & U' & V').
+ exploit (storeind_ptr_correct tge SP (fn_retaddr_ofs f) GPRA nil rs' m2').
+ { rewrite chunk_of_Tptr in P.
+ assert (rs' GPRA = rs0 RA). { apply V'. }
+ assert (rs' SP = rs2 SP). { apply V'; discriminate. }
+ rewrite H4. rewrite H3.
+ rewrite ATLR.
+ change (rs2 SP) with sp. eexact P. }
+ intros (rs3 & U & V).
+ assert (EXEC_PROLOGUE: exists rs3',
+ exec_straight_blocks tge tf
+ tf.(fn_blocks) rs0 m'
+ x0 rs3' m3'
+ /\ forall r, r <> PC -> rs3' r = rs3 r).
+ { eexists. split.
+ - change (fn_blocks tf) with tfbody; unfold tfbody.
+ econstructor; eauto. unfold exec_bblock. simpl exec_body.
+ rewrite C. fold sp. rewrite <- (sp_val _ _ _ AG). rewrite chunk_of_Tptr in F. simpl in F. rewrite F.
+ Simpl. unfold parexec_store_offset. rewrite Ptrofs.of_int64_to_int64. unfold eval_offset.
+ rewrite chunk_of_Tptr in P. Simpl. rewrite ATLR. unfold Mptr in P. assert (Archi.ptr64 = true) by auto. 2: auto. rewrite H3 in P. rewrite P.
+ simpl. apply next_sep; eauto. reflexivity.
+ - intros. destruct V' as (V'' & V'). destruct r.
+ + Simpl.
+ destruct (gpreg_eq g0 GPR16). { subst. Simpl. rewrite V; try discriminate. rewrite V''. subst rs2. Simpl. }
+ destruct (gpreg_eq g0 GPR32). { subst. Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl. }
+ destruct (gpreg_eq g0 GPR12). { subst. Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl. }
+ destruct (gpreg_eq g0 GPR17). { subst. Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl. }
+ Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl. { destruct g0; try discriminate. contradiction. }
+ + Simpl. rewrite V; try discriminate. rewrite V'; try discriminate. subst rs2. Simpl.
+ + contradiction.
+ } destruct EXEC_PROLOGUE as (rs3' & EXEC_PROLOGUE & Heqrs3').
+ exploit exec_straight_steps_2; eauto using functions_transl.
+ simpl fn_blocks. simpl fn_blocks in g. omega. constructor.
+ intros (ofs' & X & Y).
+ left; exists (State rs3' m3'); split.
+ eapply exec_straight_steps_1; eauto.
+ simpl fn_blocks. simpl fn_blocks in g. omega.
+ constructor.
+ econstructor; eauto.
+ rewrite X; econstructor; eauto.
+ apply agree_exten with rs2; eauto with asmgen.
+ unfold rs2.
+ apply agree_set_other; auto with asmgen.
+ apply agree_change_sp with (parent_sp s).
+ apply agree_undef_regs with rs0. auto.
+Local Transparent destroyed_at_function_entry.
+ simpl; intros; Simpl.
+ unfold sp; congruence.
+
+ intros.
+ assert (r <> RTMP). { contradict H3; rewrite H3; unfold data_preg; auto. }
+ rewrite Heqrs3'. Simpl. rewrite V. inversion V'. rewrite H6. auto.
+ assert (r <> GPRA). { contradict H3; rewrite H3; unfold data_preg; auto. }
+ assert (forall r : preg, r <> PC -> r <> GPRA -> rs' r = rs2 r). { apply V'. }
+ contradict H3; rewrite H3; unfold data_preg; auto.
+ contradict H3; rewrite H3; unfold data_preg; auto.
+ contradict H3; rewrite H3; unfold data_preg; auto.
+ contradict H3; rewrite H3; unfold data_preg; auto.
+ intros. rewrite Heqrs3'. rewrite V by auto with asmgen.
+ assert (forall r : preg, r <> PC -> r <> GPRA -> rs' r = rs2 r). { apply V'. }
+ rewrite H4 by auto with asmgen. reflexivity. discriminate.
+
+- (* external function *)
+ inv MS.
+ exploit functions_translated; eauto.
+ intros [tf [A B]]. simpl in B. inv B.
+ exploit extcall_arguments_match; eauto.
+ intros [args' [C D]].
+ exploit external_call_mem_extends; eauto.
+ intros [res' [m2' [P [Q [R S]]]]].
+ left; econstructor; split.
+ apply plus_one. eapply exec_step_external; eauto.
+ eapply external_call_symbols_preserved; eauto. apply senv_preserved.
+ econstructor; eauto.
+ unfold loc_external_result.
+ apply agree_set_other; auto.
+ apply agree_set_pair; auto.
+ apply agree_undef_caller_save_regs; auto.
+
+- (* return *)
+ inv MS.
+ inv STACKS. simpl in *.
+ right. split. omega. split. auto.
+ rewrite <- ATPC in H5.
+ econstructor; eauto. congruence.
+Qed.
+
+Lemma transf_initial_states:
+ forall st1, MB.initial_state prog st1 ->
+ exists st2, AB.initial_state tprog st2 /\ match_states st1 st2.
+Proof.
+ intros. inversion H. unfold ge0 in *.
+ econstructor; split.
+ econstructor.
+ eapply (Genv.init_mem_transf_partial TRANSF); eauto.
+ replace (Genv.symbol_address (Genv.globalenv tprog) (prog_main tprog) Ptrofs.zero)
+ with (Vptr fb Ptrofs.zero).
+ econstructor; eauto.
+ constructor.
+ apply Mem.extends_refl.
+ split. auto. simpl. unfold Vnullptr; destruct Archi.ptr64; congruence.
+ intros. rewrite Mach.Regmap.gi. auto.
+ unfold Genv.symbol_address.
+ rewrite (match_program_main TRANSF).
+ rewrite symbols_preserved.
+ unfold ge; rewrite H1. auto.
+Qed.
+
+Lemma transf_final_states:
+ forall st1 st2 r,
+ match_states st1 st2 -> MB.final_state st1 r -> AB.final_state st2 r.
+Proof.
+ intros. inv H0. inv H. constructor. assumption.
+ compute in H1. inv H1.
+ generalize (preg_val _ _ _ R0 AG). rewrite H2. intros LD; inv LD. auto.
+Qed.
+
+Definition return_address_offset : Machblock.function -> Machblock.code -> ptrofs -> Prop :=
+ Asmblockgenproof0.return_address_offset.
+
+Theorem transf_program_correct:
+ forward_simulation (MB.semantics return_address_offset prog) (Asmblock.semantics tprog).
+Proof.
+ eapply forward_simulation_star with (measure := measure).
+ - apply senv_preserved.
+ - eexact transf_initial_states.
+ - eexact transf_final_states.
+ - exact step_simulation.
+Qed.
+
+End PRESERVATION.
diff --git a/mppa_k1c/Asmexpand.ml b/mppa_k1c/Asmexpand.ml
index 9c256bd0..67ef6f52 100644
--- a/mppa_k1c/Asmexpand.ml
+++ b/mppa_k1c/Asmexpand.ml
@@ -507,8 +507,8 @@ let expand_instruction instr =
expand_storeind_ptr Asmvliw.GPR17 stack_pointer ofs;
emit Psemi;
let va_ofs =
- sz in
- (*Z.add full_sz (Z.of_sint ((n - _nbregargs_) * wordsize)) in *)
+ let extra_ofs = if n <= _nbregargs_ then 0 else ((n - _nbregargs_) * wordsize) in
+ Z.add sz (Z.of_sint extra_ofs) in
vararg_start_ofs := Some va_ofs;
save_arguments n va_ofs
end else begin
diff --git a/mppa_k1c/Op.v b/mppa_k1c/Op.v
index f9a774e8..ce9a5dcd 100644
--- a/mppa_k1c/Op.v
+++ b/mppa_k1c/Op.v
@@ -51,6 +51,12 @@ Inductive condition : Type :=
| Ccompfs (c: comparison) (**r 32-bit floating-point comparison *)
| Cnotcompfs (c: comparison). (**r negation of a floating-point comparison *)
+Definition condition_eq: forall (x y: condition), {x = y} + {x <> y}.
+Proof.
+ generalize comparison_eq int_eq int64_eq.
+ decide equality.
+Defined.
+
Inductive condition0 : Type :=
| Ccomp0 (c: comparison) (**r signed integer comparison with 0 *)
| Ccompu0 (c: comparison) (**r unsigned integer comparison with 0 *)
diff --git a/test/c/Makefile b/test/c/Makefile
index 72054179..72814390 100644
--- a/test/c/Makefile
+++ b/test/c/Makefile
@@ -23,6 +23,10 @@ PROGS?=fib integr qsort fft fftsp fftw sha1 sha3 aes almabench \
# * also removed bisect, who is exhibiting different float values on the Kalray
# architecture than using x86 GCC (for both CompCert and GCC ports) (tested with n=10)
##
+ifeq ($(ARCH),mppa_k1c)
+ PROGS:=$(filter-out knucleotide,$(PROGS))
+ PROGS:=$(filter-out bisect,$(PROGS))
+endif
all: $(PROGS:%=%.compcert)
@@ -41,20 +45,12 @@ all_gcc: $(PROGS:%=%.gcc)
test: all
@for i in $(PROGS); do \
- $(EXECUTE) ./$$i.compcert > $$i.compcert.out;\
- if cmp -s $$i.compcert.out Results/$$i; \
- then echo "$$i: passed"; \
- else echo "$$i: FAILED"; exit 2; \
- fi; \
+ SIMU='$(EXECUTE)' ./Runtest $$i ./$$i.compcert;\
done
test_gcc: all_gcc
@for i in $(PROGS); do \
- $(EXECUTE) ./$$i.gcc > $$i.gcc.out;\
- if cmp -s $$i.gcc.out Results/$$i; \
- then echo "$$i: passed"; \
- else echo "$$i: FAILED"; exit 2;\
- fi; \
+ SIMU='$(EXECUTE)' ./Runtest $$i ./$$i.gcc;\
done
bench_gcc: all_gcc
diff --git a/test/c/Results/binarytrees b/test/c/Results/binarytrees
index 72654db9..9dfe1355 100644
--- a/test/c/Results/binarytrees
+++ b/test/c/Results/binarytrees
@@ -1,4 +1,7 @@
-stretch tree of depth 7 check: -1
-128 trees of depth 4 check: -128
-32 trees of depth 6 check: -32
-long lived tree of depth 6 check: -1
+stretch tree of depth 13 check: -1
+8192 trees of depth 4 check: -8192
+2048 trees of depth 6 check: -2048
+512 trees of depth 8 check: -512
+128 trees of depth 10 check: -128
+32 trees of depth 12 check: -32
+long lived tree of depth 12 check: -1
diff --git a/test/c/Results/binarytrees-mppa_k1c b/test/c/Results/binarytrees-mppa_k1c
new file mode 100644
index 00000000..72654db9
--- /dev/null
+++ b/test/c/Results/binarytrees-mppa_k1c
@@ -0,0 +1,4 @@
+stretch tree of depth 7 check: -1
+128 trees of depth 4 check: -128
+32 trees of depth 6 check: -32
+long lived tree of depth 6 check: -1
diff --git a/test/c/Results/chomp b/test/c/Results/chomp
index 7898d32f..145b603a 100644
--- a/test/c/Results/chomp
+++ b/test/c/Results/chomp
@@ -1,4 +1,10 @@
player 0 plays at (1,1)
+player 1 plays at (6,0)
+player 0 plays at (0,6)
+player 1 plays at (5,0)
+player 0 plays at (0,5)
+player 1 plays at (4,0)
+player 0 plays at (0,4)
player 1 plays at (3,0)
player 0 plays at (0,3)
player 1 plays at (2,0)
diff --git a/test/c/Results/chomp-mppa_k1c b/test/c/Results/chomp-mppa_k1c
new file mode 100644
index 00000000..7898d32f
--- /dev/null
+++ b/test/c/Results/chomp-mppa_k1c
@@ -0,0 +1,9 @@
+player 0 plays at (1,1)
+player 1 plays at (3,0)
+player 0 plays at (0,3)
+player 1 plays at (2,0)
+player 0 plays at (0,2)
+player 1 plays at (1,0)
+player 0 plays at (0,1)
+player 1 plays at (0,0)
+player 1 loses
diff --git a/test/c/Results/fannkuch b/test/c/Results/fannkuch
index 09ecc715..be1815d4 100644
--- a/test/c/Results/fannkuch
+++ b/test/c/Results/fannkuch
@@ -1,31 +1,31 @@
-123456
-213456
-231456
-321456
-312456
-132456
-234156
-324156
-342156
-432156
-423156
-243156
-341256
-431256
-413256
-143256
-134256
-314256
-412356
-142356
-124356
-214356
-241356
-421356
-234516
-324516
-342516
-432516
-423516
-243516
-Pfannkuchen(6) = 10
+12345678910
+21345678910
+23145678910
+32145678910
+31245678910
+13245678910
+23415678910
+32415678910
+34215678910
+43215678910
+42315678910
+24315678910
+34125678910
+43125678910
+41325678910
+14325678910
+13425678910
+31425678910
+41235678910
+14235678910
+12435678910
+21435678910
+24135678910
+42135678910
+23451678910
+32451678910
+34251678910
+43251678910
+42351678910
+24351678910
+Pfannkuchen(10) = 38
diff --git a/test/c/Results/fannkuch-mppa_k1c b/test/c/Results/fannkuch-mppa_k1c
new file mode 100644
index 00000000..09ecc715
--- /dev/null
+++ b/test/c/Results/fannkuch-mppa_k1c
@@ -0,0 +1,31 @@
+123456
+213456
+231456
+321456
+312456
+132456
+234156
+324156
+342156
+432156
+423156
+243156
+341256
+431256
+413256
+143256
+134256
+314256
+412356
+142356
+124356
+214356
+241356
+421356
+234516
+324516
+342516
+432516
+423516
+243516
+Pfannkuchen(6) = 10
diff --git a/test/c/Results/fft b/test/c/Results/fft
index 0fc1c969..a48608b0 100644
--- a/test/c/Results/fft
+++ b/test/c/Results/fft
@@ -1 +1 @@
-1024 points, result OK
+262144 points, result OK
diff --git a/test/c/Results/fft-mppa_k1c b/test/c/Results/fft-mppa_k1c
new file mode 100644
index 00000000..0fc1c969
--- /dev/null
+++ b/test/c/Results/fft-mppa_k1c
@@ -0,0 +1 @@
+1024 points, result OK
diff --git a/test/c/Results/fftsp b/test/c/Results/fftsp
index 2b5711a6..cbeb0999 100644
--- a/test/c/Results/fftsp
+++ b/test/c/Results/fftsp
@@ -1 +1 @@
-8 points, result OK
+4096 points, result OK
diff --git a/test/c/Results/fftsp-mppa_k1c b/test/c/Results/fftsp-mppa_k1c
new file mode 100644
index 00000000..2b5711a6
--- /dev/null
+++ b/test/c/Results/fftsp-mppa_k1c
@@ -0,0 +1 @@
+8 points, result OK
diff --git a/test/c/Results/fftw-mppa_k1c b/test/c/Results/fftw-mppa_k1c
new file mode 100644
index 00000000..a1b6130c
--- /dev/null
+++ b/test/c/Results/fftw-mppa_k1c
@@ -0,0 +1,16 @@
+o[0] = 2.918193e+01
+o[1] = -3.230611e+01
+o[2] = 1.271687e+01
+o[3] = -1.099040e+01
+o[4] = 5.728673e+00
+o[5] = -4.918940e+00
+o[6] = 1.880764e+00
+o[7] = -1.292782e+00
+o[8] = 1.104073e+02
+o[9] = -5.867858e+01
+o[10] = 2.768382e+01
+o[11] = -2.073843e+01
+o[12] = 1.229410e+01
+o[13] = -9.195029e+00
+o[14] = 4.307537e+00
+o[15] = -2.080713e+00
diff --git a/test/c/Results/fib b/test/c/Results/fib
index 0e0fa4d1..84ce6474 100644
--- a/test/c/Results/fib
+++ b/test/c/Results/fib
@@ -1 +1 @@
-fib(15) = 987
+fib(35) = 14930352
diff --git a/test/c/Results/fib-mppa_k1c b/test/c/Results/fib-mppa_k1c
new file mode 100644
index 00000000..0e0fa4d1
--- /dev/null
+++ b/test/c/Results/fib-mppa_k1c
@@ -0,0 +1 @@
+fib(15) = 987
diff --git a/test/c/Results/integr b/test/c/Results/integr
index c61fdcc2..973806c9 100644
--- a/test/c/Results/integr
+++ b/test/c/Results/integr
@@ -1 +1 @@
-integr(square, 0.0, 1.0, 100000) = 0.333328
+integr(square, 0.0, 1.0, 10000000) = 0.333333
diff --git a/test/c/Results/integr-mppa_k1c b/test/c/Results/integr-mppa_k1c
new file mode 100644
index 00000000..c61fdcc2
--- /dev/null
+++ b/test/c/Results/integr-mppa_k1c
@@ -0,0 +1 @@
+integr(square, 0.0, 1.0, 100000) = 0.333328
diff --git a/test/c/Results/knucleotide b/test/c/Results/knucleotide
index e69de29b..d13ae7dc 100644
--- a/test/c/Results/knucleotide
+++ b/test/c/Results/knucleotide
@@ -0,0 +1,27 @@
+A 30.284
+T 29.796
+C 20.312
+G 19.608
+
+AA 9.212
+AT 8.950
+TT 8.948
+TA 8.936
+CA 6.166
+CT 6.100
+AC 6.086
+TC 6.042
+AG 6.036
+GA 5.968
+TG 5.868
+GT 5.798
+CC 4.140
+GC 4.044
+CG 3.906
+GG 3.798
+
+562 GGT
+152 GGTA
+15 GGTATT
+0 GGTATTTTAATT
+0 GGTATTTTAATTTATAGT
diff --git a/test/c/Results/knucleotide-mppa_k1c b/test/c/Results/knucleotide-mppa_k1c
new file mode 100644
index 00000000..e69de29b
--- /dev/null
+++ b/test/c/Results/knucleotide-mppa_k1c
diff --git a/test/c/Results/lists-mppa_k1c b/test/c/Results/lists-mppa_k1c
new file mode 100644
index 00000000..2c94e483
--- /dev/null
+++ b/test/c/Results/lists-mppa_k1c
@@ -0,0 +1,2 @@
+OK
+OK
diff --git a/test/c/Results/mandelbrot b/test/c/Results/mandelbrot
index 246f7ce1..b81e96bf 100644
--- a/test/c/Results/mandelbrot
+++ b/test/c/Results/mandelbrot
Binary files differ
diff --git a/test/c/Results/mandelbrot-mppa_k1c b/test/c/Results/mandelbrot-mppa_k1c
new file mode 100644
index 00000000..246f7ce1
--- /dev/null
+++ b/test/c/Results/mandelbrot-mppa_k1c
Binary files differ
diff --git a/test/c/Results/nbody b/test/c/Results/nbody
index 99ad4fd1..41b648fd 100644
--- a/test/c/Results/nbody
+++ b/test/c/Results/nbody
@@ -1,2 +1,2 @@
-0.169075164
--0.169050762
+-0.169086185
diff --git a/test/c/Results/nbody-mppa_k1c b/test/c/Results/nbody-mppa_k1c
new file mode 100644
index 00000000..99ad4fd1
--- /dev/null
+++ b/test/c/Results/nbody-mppa_k1c
@@ -0,0 +1,2 @@
+-0.169075164
+-0.169050762
diff --git a/test/c/Results/nsieve b/test/c/Results/nsieve
index 95fea812..bb9b87dc 100644
--- a/test/c/Results/nsieve
+++ b/test/c/Results/nsieve
@@ -1,3 +1,3 @@
-Primes up to 12800 1526
-Primes up to 6400 834
-Primes up to 3200 452
+Primes up to 5120000 356244
+Primes up to 2560000 187134
+Primes up to 1280000 98610
diff --git a/test/c/Results/nsieve-mppa_k1c b/test/c/Results/nsieve-mppa_k1c
new file mode 100644
index 00000000..95fea812
--- /dev/null
+++ b/test/c/Results/nsieve-mppa_k1c
@@ -0,0 +1,3 @@
+Primes up to 12800 1526
+Primes up to 6400 834
+Primes up to 3200 452
diff --git a/test/c/Results/nsievebits b/test/c/Results/nsievebits
index 2131804c..bb9b87dc 100644
--- a/test/c/Results/nsievebits
+++ b/test/c/Results/nsievebits
@@ -1,3 +1,3 @@
-Primes up to 40000 4203
-Primes up to 20000 2262
-Primes up to 10000 1229
+Primes up to 5120000 356244
+Primes up to 2560000 187134
+Primes up to 1280000 98610
diff --git a/test/c/Results/nsievebits-mppa_k1c b/test/c/Results/nsievebits-mppa_k1c
new file mode 100644
index 00000000..2131804c
--- /dev/null
+++ b/test/c/Results/nsievebits-mppa_k1c
@@ -0,0 +1,3 @@
+Primes up to 40000 4203
+Primes up to 20000 2262
+Primes up to 10000 1229
diff --git a/test/c/Results/perlin b/test/c/Results/perlin
index 8438b53c..4503fc1c 100644
--- a/test/c/Results/perlin
+++ b/test/c/Results/perlin
@@ -1 +1 @@
-6.0000e+00
+-4.0543e+03
diff --git a/test/c/Results/perlin-mppa_k1c b/test/c/Results/perlin-mppa_k1c
new file mode 100644
index 00000000..8438b53c
--- /dev/null
+++ b/test/c/Results/perlin-mppa_k1c
@@ -0,0 +1 @@
+6.0000e+00
diff --git a/test/c/Results/qsort-mppa_k1c b/test/c/Results/qsort-mppa_k1c
new file mode 100644
index 00000000..d86bac9d
--- /dev/null
+++ b/test/c/Results/qsort-mppa_k1c
@@ -0,0 +1 @@
+OK
diff --git a/test/c/Results/sha1-mppa_k1c b/test/c/Results/sha1-mppa_k1c
new file mode 100644
index 00000000..730d5406
--- /dev/null
+++ b/test/c/Results/sha1-mppa_k1c
@@ -0,0 +1,2 @@
+Test `abc': passed
+Test `abcdbcdecdefdefgefghfghighijhijkijkljklmklmnlmnomnopnopq': passed
diff --git a/test/c/Results/spectral b/test/c/Results/spectral
index b06cd560..1e35f7e0 100644
--- a/test/c/Results/spectral
+++ b/test/c/Results/spectral
@@ -1 +1 @@
-1.272359925
+1.274224148
diff --git a/test/c/Results/spectral-mppa_k1c b/test/c/Results/spectral-mppa_k1c
new file mode 100644
index 00000000..b06cd560
--- /dev/null
+++ b/test/c/Results/spectral-mppa_k1c
@@ -0,0 +1 @@
+1.272359925
diff --git a/test/c/Results/vmach b/test/c/Results/vmach
index a95237a6..9caa2e51 100644
--- a/test/c/Results/vmach
+++ b/test/c/Results/vmach
@@ -1,2 +1,2 @@
-fib(15) = 987
-tak(12, 9, 6) = 9
+fib(30) = 1346269
+tak(18, 12, 6) = 7
diff --git a/test/c/Results/vmach-mppa_k1c b/test/c/Results/vmach-mppa_k1c
new file mode 100644
index 00000000..a95237a6
--- /dev/null
+++ b/test/c/Results/vmach-mppa_k1c
@@ -0,0 +1,2 @@
+fib(15) = 987
+tak(12, 9, 6) = 9
diff --git a/test/c/Runtest b/test/c/Runtest
new file mode 100755
index 00000000..f693219a
--- /dev/null
+++ b/test/c/Runtest
@@ -0,0 +1,71 @@
+#!/bin/sh
+
+# The name of the test
+name="$1"
+shift
+
+# The temp file for output
+out="test$$.log"
+rm -f $out
+trap "rm -f $out" 0 INT QUIT
+
+# Is the test expected to fail?
+expect_fail=false
+
+# The architecture and the bitsize
+arch=`sed -n -e 's/^ARCH=//p' ../../Makefile.config`
+bits=`sed -n -e 's/^BITSIZE=//p' ../../Makefile.config`
+
+# The reference output
+if test -f "Results/$name-$arch-$bits"; then
+ ref="Results/$name-$arch-$bits"
+elif test -f "Results/$name-$arch"; then
+ ref="Results/$name-$arch"
+elif test -f "Results/$name-$bits"; then
+ ref="Results/$name-$bits"
+elif test -f "Results/$name"; then
+ ref="Results/$name"
+else
+ ref=""
+fi
+
+# Special conditions
+
+if test -f "$name.cond"; then
+ RUN=0 SKIP=1 EXPECT_FAIL=2 sh "$name.cond"
+ case "$?" in
+ 1) echo "$name: skipped"; exit 0;;
+ 2) expect_fail=true;;
+ esac
+fi
+
+# Administer the test
+if $SIMU $* > $out
+then
+ if $expect_fail; then
+ echo "$name: ERROR (should have failed but did not)"
+ exit 2
+ elif test -n "$ref"; then
+ if cmp -s "$out" "$ref"; then
+ echo "$name: passed"
+ exit 0
+ else
+ echo "$name: WRONG OUTPUT (diff follows)"
+ diff -u "$ref" "$out"
+ exit 2
+ fi
+ else
+ echo "$name: passed"
+ exit 0
+ fi
+else
+ retcode=$?
+ if $expect_fail; then
+ echo "$name: passed (failed as expected)"
+ exit 0
+ else
+ echo "$name: EXECUTION FAILED (status $retcode)"
+ exit 2
+ fi
+fi
+
diff --git a/test/c/binarytrees.c b/test/c/binarytrees.c
index adc0d7b3..becae164 100644
--- a/test/c/binarytrees.c
+++ b/test/c/binarytrees.c
@@ -75,7 +75,11 @@ int main(int argc, char* argv[])
unsigned N, depth, minDepth, maxDepth, stretchDepth;
treeNode *stretchTree, *longLivedTree, *tempTree;
+#ifdef __K1C__
N = argc < 2 ? 6 : atol(argv[1]);
+#else
+ N = argc < 2 ? 12 : atol(argv[1]);
+#endif
minDepth = 4;
diff --git a/test/c/chomp.c b/test/c/chomp.c
index 5b3c9cfb..7e2f62c1 100644
--- a/test/c/chomp.c
+++ b/test/c/chomp.c
@@ -338,8 +338,13 @@ int main(void)
struct _play *tree;
+#ifdef __K1C__
ncol = 4;
nrow = 4;
+#else
+ ncol = 7;
+ nrow = 7;
+#endif
tree = make_play(1); /* create entire tree structure, not just the */
player = 0; /* needed part for first move */
current = make_data(nrow,ncol); /* start play at full board */
diff --git a/test/c/fannkuch.c b/test/c/fannkuch.c
index 6fb7912f..befccd8d 100644
--- a/test/c/fannkuch.c
+++ b/test/c/fannkuch.c
@@ -102,7 +102,11 @@ fannkuch( int n )
int
main( int argc, char* argv[] )
{
+#ifdef __K1C__
int n = (argc>1) ? atoi(argv[1]) : 6;
+#else
+ int n = (argc>1) ? atoi(argv[1]) : 10;
+#endif
printf("Pfannkuchen(%d) = %ld\n", n, fannkuch(n));
return 0;
diff --git a/test/c/fft.c b/test/c/fft.c
index 429181b1..8ab59c9a 100644
--- a/test/c/fft.c
+++ b/test/c/fft.c
@@ -152,7 +152,11 @@ int main(int argc, char ** argv)
double enp, t, y, z, zr, zi, zm, a;
double * xr, * xi, * pxr, * pxi;
+#ifdef __K1C__
if (argc >= 2) n = atoi(argv[1]); else n = 10;
+#else
+ if (argc >= 2) n = atoi(argv[1]); else n = 18;
+#endif
np = 1 << n;
enp = np;
npm = np / 2 - 1;
diff --git a/test/c/fftsp.c b/test/c/fftsp.c
index 2dcd4ad2..d327a74c 100644
--- a/test/c/fftsp.c
+++ b/test/c/fftsp.c
@@ -153,7 +153,11 @@ int main(int argc, char ** argv)
float enp, t, y, z, zr, zi, zm, a;
float * xr, * xi, * pxr, * pxi;
+#ifdef __K1C__
if (argc >= 2) n = atoi(argv[1]); else n = 3;
+#else
+ if (argc >= 2) n = atoi(argv[1]); else n = 12;
+#endif
np = 1 << n;
enp = np;
npm = np / 2 - 1;
diff --git a/test/c/fftw.c b/test/c/fftw.c
index 755bac7f..04d896ad 100644
--- a/test/c/fftw.c
+++ b/test/c/fftw.c
@@ -74,7 +74,11 @@ const E KP1_847759065 = ((E) +1.847759065022573512256366378793576573644833252);
/* Test harness */
+#ifdef __K1C__
#define NRUNS (10 * 10)
+#else
+#define NRUNS (100 * 1000)
+#endif
int main()
{
diff --git a/test/c/fib.c b/test/c/fib.c
index 439b908c..168626bc 100644
--- a/test/c/fib.c
+++ b/test/c/fib.c
@@ -12,7 +12,11 @@ int fib(int n)
int main(int argc, char ** argv)
{
int n, r;
+#ifdef __K1C__
if (argc >= 2) n = atoi(argv[1]); else n = 15;
+#else
+ if (argc >= 2) n = atoi(argv[1]); else n = 35;
+#endif
r = fib(n);
printf("fib(%d) = %d\n", n, r);
return 0;
diff --git a/test/c/integr.c b/test/c/integr.c
index 19ea78ab..cd0521f5 100644
--- a/test/c/integr.c
+++ b/test/c/integr.c
@@ -25,7 +25,11 @@ double test(int n)
int main(int argc, char ** argv)
{
int n; double r;
+#ifdef __K1C__
if (argc >= 2) n = atoi(argv[1]); else n = 100000;
+#else
+ if (argc >= 2) n = atoi(argv[1]); else n = 10000000;
+#endif
r = test(n);
printf("integr(square, 0.0, 1.0, %d) = %g\n", n, r);
return 0;
diff --git a/test/c/lists.c b/test/c/lists.c
index a72ec398..8deb0f37 100644
--- a/test/c/lists.c
+++ b/test/c/lists.c
@@ -61,8 +61,13 @@ int main(int argc, char ** argv)
int n, niter, i;
struct list * l;
+#ifdef __K1C__
if (argc >= 2) n = atoi(argv[1]); else n = 500;
if (argc >= 3) niter = atoi(argv[1]); else niter = 100;
+#else
+ if (argc >= 2) n = atoi(argv[1]); else n = 1000;
+ if (argc >= 3) niter = atoi(argv[1]); else niter = 20000;
+#endif
l = buildlist(n);
if (checklist(n, reverselist(l))) {
printf("OK\n");
diff --git a/test/c/mandelbrot.c b/test/c/mandelbrot.c
index 84fae15c..133d55c5 100644
--- a/test/c/mandelbrot.c
+++ b/test/c/mandelbrot.c
@@ -17,12 +17,20 @@ int main (int argc, char **argv)
{
int w, h, bit_num = 0;
char byte_acc = 0;
+#ifdef __K1C__
int i, iter = 30;
+#else
+ int i, iter = 50;
+#endif
double x, y, limit = 2.0;
double Zr, Zi, Cr, Ci, Tr, Ti;
if (argc < 2) {
+#ifdef __K1C__
w = h = 50;
+#else
+ w = h = 1000;
+#endif
} else {
w = h = atoi(argv[1]);
}
diff --git a/test/c/nbody.c b/test/c/nbody.c
index 1ce8d788..ab0ebabe 100644
--- a/test/c/nbody.c
+++ b/test/c/nbody.c
@@ -140,7 +140,11 @@ void setup_bodies(void)
int main(int argc, char ** argv)
{
+#ifdef __K1C__
int n = argc < 2 ? 100 : atoi(argv[1]);
+#else
+ int n = argc < 2 ? 1000000 : atoi(argv[1]);
+#endif
int i;
setup_bodies();
diff --git a/test/c/nsieve.c b/test/c/nsieve.c
index fc79ba69..3954bcbe 100644
--- a/test/c/nsieve.c
+++ b/test/c/nsieve.c
@@ -29,10 +29,18 @@ static unsigned int nsieve(int m) {
#define NITER 2
int main(int argc, char * argv[]) {
+#ifdef __K1C__
int m = argc < 2 ? 6 : atoi(argv[1]);
+#else
+ int m = argc < 2 ? 9 : atoi(argv[1]);
+#endif
int i, j;
for (i = 0; i < 3; i++) {
+#ifdef __K1C__
int n = 200 << (m-i);
+#else
+ int n = 10000 << (m-i);
+#endif
unsigned count;
for (j = 0; j < NITER; j++) { count = nsieve(n); }
printf("Primes up to %8d %8u\n", n, count);
diff --git a/test/c/nsievebits.c b/test/c/nsievebits.c
index 15d31417..e3b7fd43 100644
--- a/test/c/nsievebits.c
+++ b/test/c/nsievebits.c
@@ -30,7 +30,11 @@ nsieve(unsigned int m)
return (count);
}
+#ifdef __K1C__
#define NITER 1
+#else
+#define NITER 2
+#endif
static void
test(unsigned int n)
@@ -48,7 +52,11 @@ main(int ac, char **av)
{
unsigned int n;
+#ifdef __K1C__
n = ac < 2 ? 2 : atoi(av[1]);
+#else
+ n = ac < 2 ? 9 : atoi(av[1]);
+#endif
test(n);
if (n >= 1)
test(n - 1);
diff --git a/test/c/perlin.c b/test/c/perlin.c
index 79d49d62..29ebf964 100644
--- a/test/c/perlin.c
+++ b/test/c/perlin.c
@@ -63,9 +63,15 @@ static void init(void) {
p[256+i] = p[i] = permutation[i];
}
+#ifdef __K1C__
#define INCREMENT 0.5
#define MIN -3.0
#define MAX 3.0
+#else
+#define INCREMENT 0.1
+#define MIN -5.0
+#define MAX 5.0
+#endif
int main(int argc, char ** argv) {
init();
diff --git a/test/c/qsort.c b/test/c/qsort.c
index 8a035eb5..1ebe1e11 100644
--- a/test/c/qsort.c
+++ b/test/c/qsort.c
@@ -34,7 +34,11 @@ int main(int argc, char ** argv)
int n, i, j;
int * a, * b;
+#ifdef __K1C__
if (argc >= 2) n = atoi(argv[1]); else n = 500;
+#else
+ if (argc >= 2) n = atoi(argv[1]); else n = 100000;
+#endif
a = malloc(n * sizeof(int));
b = malloc(n * sizeof(int));
for (j = 0; j < NITER; j++) {
diff --git a/test/c/spectral.c b/test/c/spectral.c
index b3a34070..dca78fe0 100644
--- a/test/c/spectral.c
+++ b/test/c/spectral.c
@@ -43,7 +43,11 @@ void eval_AtA_times_u(int N, const double u[], double AtAu[])
int main(int argc, char *argv[])
{
int i;
+#ifdef __K1C__
int N = ((argc == 2) ? atoi(argv[1]) : 11);
+#else
+ int N = ((argc == 2) ? atoi(argv[1]) : 1000);
+#endif
double * u, * v, vBv, vv;
u = malloc(N * sizeof(double));
v = malloc(N * sizeof(double));
diff --git a/test/c/vmach.c b/test/c/vmach.c
index 4e6848d2..5858d4d6 100644
--- a/test/c/vmach.c
+++ b/test/c/vmach.c
@@ -159,7 +159,11 @@ long wordcode_interp(unsigned int* code)
#define I(a,b,c,d) ((a) + ((b) << 8) + ((c) << 16) + ((d) << 24))
+#ifdef __K1C__
#define FIBSIZE 15
+#else
+#define FIBSIZE 30
+#endif
unsigned int wordcode_fib[] = {
/* 0 */ I(WCONST, FIBSIZE, 0, 0),
@@ -178,9 +182,15 @@ unsigned int wordcode_fib[] = {
/* 13 */ I(WRETURN, 0, 2, 0)
};
+#ifdef __K1C__
#define TAKSIZE1 6
#define TAKSIZE2 9
#define TAKSIZE3 12
+#else
+#define TAKSIZE1 6
+#define TAKSIZE2 12
+#define TAKSIZE3 18
+#endif
unsigned int wordcode_tak[] = {
/* 0 */ I(WCONST, TAKSIZE1, 0, 0),
diff --git a/test/monniaux/README.md b/test/monniaux/README.md
index dbb3f337..f2af67fb 100644
--- a/test/monniaux/README.md
+++ b/test/monniaux/README.md
@@ -37,6 +37,7 @@ float_mat c3, 1504675, 751514, 553235, 1929369, 1372441
- `K1C_CC`: GCC compiler (default k1-cos-gcc)
- `K1C_CCOMP`: compcert compiler (default ccomp)
- `EXECUTE_CYCLES`: running command (default `k1-cluster` with some options)
+- `EXECUTE_ARGS`: execution arguments
- `GCCiFLAGS` with i from 0 to 4: the wanted optimizations. If one of these flags is empty, nothing is done. Same for `CCOMPiFLAGS`. For now, the default values:
```
# You can define up to GCC4FLAGS and CCOMP4FLAGS
@@ -68,4 +69,10 @@ The `PREFIX` are the prefixes to add to the .s, .o, etc.. You should be careful
Assembly files will be generated in `asm/`, objects in `obj/`, binaries in `bin/` and outputs in `out/`.
-To compile and execute all the benches : `make`
+To compile and execute all the benches : `make` while in the `monniaux` directory (without any `-j` flag).
+
+To compile and/or execute a single bench, `cd` to the bench directory, then:
+- `make` for compiling the bench
+- `make run` for running it
+
+You can use `-j` flag when in a single bench directory
diff --git a/test/monniaux/benches.sh b/test/monniaux/benches.sh
index 2365063a..6014f628 100644
--- a/test/monniaux/benches.sh
+++ b/test/monniaux/benches.sh
@@ -1,3 +1,3 @@
-benches="binary_search bitsliced-aes bitsliced-tea complex float_mat glibc_qsort heapsort idea number_theoretic_transform quicksort sha-2 tacle-bench-lift tacle-bench-powerwindow too_slow heptagon_radio_transmitter lustrev4_lustrec_heater_control lustrev4_lv4_heater_control lustrev4_lv6-en-2cgc_heater_control lustrev6-convertible-en-2cgc xor_and_mat"
+benches="binary_search bitsliced-aes bitsliced-tea complex float_mat glibc_qsort heapsort idea number_theoretic_transform quicksort sha-2 tacle-bench-lift tacle-bench-powerwindow too_slow heptagon_radio_transmitter lustrev4_lustrec_heater_control lustrev4_lv4_heater_control lustrev4_lv6-en-2cgc_heater_control lustrev6-convertible-en-2cgc xor_and_mat glpk-4.65 picosat-965 genann"
# Removed for now : ternary
diff --git a/test/monniaux/genann/Makefile b/test/monniaux/genann/Makefile
new file mode 100644
index 00000000..2e76ec63
--- /dev/null
+++ b/test/monniaux/genann/Makefile
@@ -0,0 +1,4 @@
+ALL_CFILES= example4shorter.c genann.c
+TARGET=genann4
+
+include ../rules.mk
diff --git a/test/monniaux/genann/make.proto b/test/monniaux/genann/make.proto
deleted file mode 100644
index 7c4248bf..00000000
--- a/test/monniaux/genann/make.proto
+++ /dev/null
@@ -1,2 +0,0 @@
-sources: example4shorter.c genann.c
-target: genann4 \ No newline at end of file
diff --git a/test/monniaux/glpk-4.65/Makefile b/test/monniaux/glpk-4.65/Makefile
index a0ab40dc..eaa3f4b0 100644
--- a/test/monniaux/glpk-4.65/Makefile
+++ b/test/monniaux/glpk-4.65/Makefile
@@ -1,39 +1,6 @@
ALL_CFLAGS += -I src/amd -I src/colamd -I src/mpl -I src/simplex -I src/api -I src/intopt -I src/minisat -I src/npp -I src/zlib -I src/bflib -I src/env -I src/misc -I src/draft -I src
-
-include ../rules.mk
-
-LIBS = -lm
-
-src=examples/glpsol.c $(wildcard src/*/*.c)
-
-PRODUCTS?=glpsol.gcc.host glpsol.ccomp.host glpsol.gcc.k1c glpsol.gcc.o1.k1c glpsol.ccomp.k1c
-PRODUCTS_OUT=$(addsuffix .out,$(PRODUCTS))
-
-all: $(PRODUCTS)
-
-.PHONY:
-run: measures.csv
-
-
-glpsol.gcc.host: $(src:.c=.gcc.host.o) ../clock.gcc.host.o
- $(CC) $(CFLAGS) $+ $(LIBS) -o $@
-glpsol.ccomp.host: $(src:.c=.ccomp.host.o) ../clock.gcc.host.o
- $(CCOMP) $(CCOMPFLAGS) $+ $(LIBS) -o $@
-glpsol.gcc.k1c: $(src:.c=.gcc.k1c.o) ../clock.gcc.k1c.o
- $(K1C_CC) $(K1C_CFLAGS) $+ $(LIBS) -o $@
-glpsol.gcc.o1.k1c: $(src:.c=.gcc.o1.k1c.o) ../clock.gcc.k1c.o
- $(K1C_CC) $(K1C_CFLAGS_O1) $+ $(LIBS) -o $@
-glpsol.ccomp.k1c: $(src:.c=.ccomp.k1c.o) ../clock.gcc.k1c.o
- $(K1C_CCOMP) $(K1C_CCOMPFLAGS) $+ $(LIBS) -o $@
-
+ALL_CFILES=examples/glpsol.c $(wildcard src/*/*.c)
+TARGET=glpk
EXECUTE_ARGS=--math examples/prod.mod
-measures.csv: $(PRODUCTS_OUT)
- echo "benches, gcc host,ccomp host,gcc k1c,gcc o1 k1c,ccomp k1c" > $@
-
-.SECONDARY:
-
-.PHONY:
-clean:
- rm -f *.o *.s *.k1c *.csv
-
+include ../rules.mk
diff --git a/test/monniaux/picosat-965/Makefile b/test/monniaux/picosat-965/Makefile
index 991278ff..a887c0de 100644
--- a/test/monniaux/picosat-965/Makefile
+++ b/test/monniaux/picosat-965/Makefile
@@ -1,37 +1,11 @@
EXECUTE_ARGS=sudoku.sat
-
-include ../rules.mk
-
-#ALL_CFLAGS = -DNDEBUG
ALL_CFLAGS = -DNALARM -DNZIP -DNGETRUSAGE -DNDEBUG
-K1C_CFLAGS += $(EMBEDDED_CFLAGS)
-K1C_CCOMPFLAGS += $(EMBEDDED_CFLAGS)
-CCOMPFLAGS += -fbitfields
-K1C_CCOMPFLAGS += -fbitfields # -fno-if-conversion
-
-K1C_CFLAGS += $(ALL_CFLAGS)
-K1C_CCOMPFLAGS += $(ALL_CFLAGS)
-CCOMPFLAGS += $(ALL_CFLAGS)
-CFLAGS += $(ALL_CFLAGS)
-
-all: picosat.ccomp.k1c.s version.ccomp.k1c.s app.ccomp.k1c.s main.ccomp.k1c.s picosat.gcc.k1c.s version.gcc.k1c.s app.gcc.k1c.s main.gcc.k1c.s picosat.ccomp.k1c.out picosat.gcc.o1.k1c.out picosat.gcc.k1c.out picosat picosat.ccomp.host.out picosat.gcc.host.out
-
-picosat.ccomp.k1c : picosat.ccomp.k1c.s version.ccomp.k1c.s app.ccomp.k1c.s main.ccomp.k1c.s ../clock.gcc.k1c.o
- $(K1C_CCOMP) $(K1C_CCOMPFLAGS) $+ -o $@
+ALL_CCOMPFLAGS += -fbitfields # -fno-if-conversion
+TARGET=picosat
+ALL_CFILES=picosat.c version.c app.c main.c
-picosat.gcc.k1c : picosat.gcc.k1c.s version.gcc.k1c.s app.gcc.k1c.s main.gcc.k1c.s ../clock.gcc.k1c.o
- $(K1C_CC) $(K1C_CFLAGS) $+ -o $@
-
-picosat.gcc.o1.k1c : picosat.gcc.o1.k1c.s version.gcc.o1.k1c.s app.gcc.o1.k1c.s main.gcc.o1.k1c.s ../clock.gcc.k1c.o
- $(K1C_CC) $(K1C_CFLAGS_O1) $+ -o $@
-
-picosat.ccomp.host : picosat.ccomp.host.s version.ccomp.host.s app.ccomp.host.s main.ccomp.host.s ../clock.gcc.host.o
- $(CCOMP) $(CCOMPFLAGS) $+ -o $@
-
-picosat.gcc.host : picosat.gcc.host.s version.gcc.host.s app.gcc.host.s main.gcc.host.s ../clock.gcc.host.o
- $(CC) $(FLAGS) $+ -o $@
-
-clean:
- -rm -f *.s *.k1c *.out
+include ../rules.mk
-.PHONY: clean
+# FIXME - what were these for?
+#K1C_CFLAGS += $(EMBEDDED_CFLAGS)
+#K1C_CCOMPFLAGS += $(EMBEDDED_CFLAGS)
diff --git a/test/monniaux/rules.mk b/test/monniaux/rules.mk
index 9d05b4d6..079606e6 100644
--- a/test/monniaux/rules.mk
+++ b/test/monniaux/rules.mk
@@ -7,6 +7,9 @@ ALL_CFILES?=$(wildcard *.c)
# Name of the target
TARGET?=toto
+# Arguments of execution
+EXECUTE_ARGS?=
+
# Name of the clock object
CLOCK=../clock
@@ -92,7 +95,7 @@ obj/%.o: asm/%.s
out/%.out: bin/%.bin
@mkdir -p $(@D)
- $(EXECUTE_CYCLES) $< | tee $@
+ $(EXECUTE_CYCLES) $< $(EXECUTE_ARGS) | tee $@
##
# Generating the rules for all the compiler/flags..
diff --git a/test/regression/Makefile b/test/regression/Makefile
index aeabc7e1..585ffb1c 100644
--- a/test/regression/Makefile
+++ b/test/regression/Makefile
@@ -26,15 +26,18 @@ TESTS_COMP?=attribs1 bitfields1 bitfields2 bitfields3 bitfields4 \
varargs1 varargs2 varargs3 sections alias aligned\
packedstruct1 packedstruct2
-# varargs2 ---> bug in k1-cos-gcc
ifeq ($(ARCH),mppa_k1c)
- TESTS_COMP:=$(filter-out varargs2,$(TESTS_COMP))
+ TESTS_COMP:=$(filter-out packedstruct1,$(TESTS_COMP))
+ TESTS_COMP:=$(filter-out packedstruct2,$(TESTS_COMP))
endif
# Can run, both in compiled mode and in interpreter mode,
# but produce processor-dependent results, so no reference output in Results
TESTS_DIFF=NaNs
+ifeq ($(ARCH),mppa_k1c)
+ TESTS_DIFF:=$(filter-out NaNs,$(TESTS_DIFF))
+endif
# Other tests: should compile to .s without errors (but expect warnings)
diff --git a/test/regression/Results/varargs2-mppa_k1c b/test/regression/Results/varargs2-mppa_k1c
new file mode 100644
index 00000000..f954927e
--- /dev/null
+++ b/test/regression/Results/varargs2-mppa_k1c
@@ -0,0 +1,11 @@
+An int: 42
+A long long: 123456789012345
+A string: Hello world
+A double: 3.141592654
+A mixture: x & Hello, world! & 42 & 123456789012345 & 3.141592654 & 2.718281746
+Twice: -1 1.23
+Twice: -1 1.23
+With va_copy: -1 1.23
+With va_copy: -1 1.23
+With extra args: x & Hello, world! & 42 & 123456789012345 & 3.141592654 & 2.718281746
+va_list compatibility: x & Hello, world! & 42 & 123456789012345 & 3.141592654 & 2.718281746
diff --git a/test/regression/builtins-mppa_k1c.c b/test/regression/builtins-mppa_k1c.c
new file mode 100644
index 00000000..cbf51387
--- /dev/null
+++ b/test/regression/builtins-mppa_k1c.c
@@ -0,0 +1,72 @@
+/* Fun with builtins */
+
+#include <stdio.h>
+#include <math.h>
+
+char * check_relative_error(double exact, double actual, double precision)
+{
+ double relative_error = (actual - exact) / exact;
+ return fabs(relative_error) <= precision ? "OK" : "ERROR";
+}
+
+//unsigned int x = 0x12345678;
+//unsigned int y = 0xDEADBEEF;
+//unsigned long long xx = 0x1234567812345678ULL;
+//double a = 3.14159;
+//double b = 2.718;
+//double c = 1.414;
+//unsigned short s = 0x1234;
+
+int main(int argc, char ** argv)
+{
+ unsigned z;
+
+ //printf("mulhw(%x, %x) = %x\n", x, y, __builtin_mulhw(x, y));
+ //printf("mulhwu(%x, %x) = %x\n", x, y, __builtin_mulhwu(x, y));
+ //printf("clz(%x) = %d\n", x, __builtin_clz(x));
+ //printf("clzll(%llx) = %d\n", (unsigned long long) x, __builtin_clzll(x));
+ //printf("clzll(%llx) = %d\n", xx, __builtin_clzll(xx));
+ //z = __builtin_bswap(x);
+ //printf("clzll(%lx) = %d\n", z, __builtin_clzll(z));
+ //printf("bswap(%x) = %x\n", x, __builtin_bswap(x));
+ //printf("bswap16(%x) = %x\n", s, __builtin_bswap16(s));
+
+ //printf("fmadd(%f, %f, %f) = %f\n", a, b, c, __builtin_fmadd(a, b, c));
+ //printf("fmsub(%f, %f, %f) = %f\n", a, b, c, __builtin_fmsub(a, b, c));
+ //printf("fabs(%f) = %f\n", a, __builtin_fabs(a));
+ //printf("fabs(%f) = %f\n", -a, __builtin_fabs(-a));
+ //printf("fsqrt(%f) = %f\n", a, __builtin_fsqrt(a));
+ //printf("frsqrte(%f) = %s\n",
+ // a, check_relative_error(1.0 / sqrt(a), __builtin_frsqrte(a), 1./32.));
+ //printf("fres(%f) = %s\n",
+ // a, check_relative_error(1.0 / a, __builtin_fres(a), 1./256.));
+ //printf("fsel(%f, %f, %f) = %f\n", a, b, c, __builtin_fsel(a, b, c));
+ //printf("fsel(%f, %f, %f) = %f\n", -a, b, c, __builtin_fsel(-a, b, c));
+ //printf("fcti(%f) = %d\n", a, __builtin_fcti(a));
+ //printf("fcti(%f) = %d\n", b, __builtin_fcti(b));
+ //printf("fcti(%f) = %d\n", c, __builtin_fcti(c));
+ //__builtin_eieio();
+ //__builtin_sync();
+ //__builtin_isync();
+ //printf("isel(%d, %d, %d) = %d\n", 0, x, y, __builtin_isel(0, x, y));
+ //printf("isel(%d, %d, %d) = %d\n", 42, x, y, __builtin_isel(42, x, y));
+ //printf ("read_16_rev = %x\n", __builtin_read16_reversed(&s));
+ //printf ("read_32_rev = %x\n", __builtin_read32_reversed(&y));
+ //__builtin_write16_reversed(&s, 0x789A);
+ //printf ("after write_16_rev: %x\n", s);
+ //__builtin_write32_reversed(&y, 0x12345678);
+ //printf ("after write_32_rev: %x\n", y);
+ //y = 0;
+ //__builtin_write32_reversed(&y, 0x12345678);
+ //printf ("CSE write_32_rev: %s\n", y == 0x78563412 ? "ok" : "ERROR");
+ ///* Make sure that ignoring the result of a builtin
+ // doesn't cause an internal error */
+ //(void) __builtin_bswap(x);
+ //(void) __builtin_fsqrt(a);
+ return 0;
+}
+
+
+
+
+
diff --git a/test/regression/varargs2.c b/test/regression/varargs2.c
index b96d1940..84860ef3 100644
--- a/test/regression/varargs2.c
+++ b/test/regression/varargs2.c
@@ -115,15 +115,27 @@ void printf_compat(const char * fmt, ...)
}
/* The test harness */
-
int main()
{
miniprintf("An int: %d\n", 42);
miniprintf("A long long: %l\n", 123456789012345LL);
miniprintf("A string: %s\n", "Hello world");
miniprintf("A double: %e\n", 3.141592654);
+
+#ifndef __K1C__
miniprintf("A small struct: %y\n", (struct Y) { 'x', 12 });
miniprintf("A bigger struct: %z\n", (struct Z) { 123, 456, 789 });
+#endif
+
+#ifdef __K1C__
+ miniprintf("A mixture: %c & %s & %d & %l & %e & %f\n",
+ 'x',
+ "Hello, world!",
+ 42,
+ 123456789012345LL,
+ 3.141592654,
+ 2.71828182);
+#else
miniprintf("A mixture: %c & %s & %y & %d & %l & %e & %f\n",
'x',
"Hello, world!",
@@ -132,6 +144,8 @@ int main()
123456789012345LL,
3.141592654,
2.71828182);
+#endif
+
miniprintf2("Twice: %d %e\n", -1, 1.23);
miniprintf3("With va_copy: %d %e\n", -1, 1.23);
miniprintf_extra(0, 1, 2, 3, 4, 5, 6, 7,