aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2017-04-28 15:56:59 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2017-04-28 16:05:51 +0200
commitf642817f0dc761e51c3bd362f75b0068a8d4b0c8 (patch)
treeb5830bb772611d2271c4b7d26f162d5c200dd788
parent2fbdb0c45f0913b9fd8e95606c525fc5bfb3bc6d (diff)
downloadcompcert-kvx-f642817f0dc761e51c3bd362f75b0068a8d4b0c8.tar.gz
compcert-kvx-f642817f0dc761e51c3bd362f75b0068a8d4b0c8.zip
RISC-V port and assorted changes
This commits adds code generation for the RISC-V architecture, both in 32- and 64-bit modes. The generated code was lightly tested using the simulator and cross-binutils from https://riscv.org/software-tools/ This port required the following additional changes: - Integers: More properties about shrx - SelectOp: now provides smart constructors for mulhs and mulhu - SelectDiv, 32-bit integer division and modulus: implement constant propagation, use the new smart constructors mulhs and mulhu. - Runtime library: if no asm implementation is provided, run the reference C implementation through CompCert. Since CompCert rejects the definitions of names of special functions such as __i64_shl, the reference implementation now uses "i64_" names, e.g. "i64_shl", and a renaming "i64_ -> __i64_" is performed over the generated assembly file, before assembling and building the runtime library. - test/: add SIMU make variable to run tests through a simulator - test/regression/alignas.c: make sure _Alignas and _Alignof are not #define'd by C headers commit da14495c01cf4f66a928c2feff5c53f09bde837f Author: Xavier Leroy <xavier.leroy@inria.fr> Date: Thu Apr 13 17:36:10 2017 +0200 RISC-V port, continued Now working on Asmgen. commit 36f36eb3a5abfbb8805960443d087b6a83e86005 Author: Xavier Leroy <xavier.leroy@inria.fr> Date: Wed Apr 12 17:26:39 2017 +0200 RISC-V port, first steps This port is based on Prashanth Mundkur's experimental RV32 port and brings it up to date with CompCert, and adds 64-bit support (RV64). Work in progress.
-rw-r--r--.gitignore3
-rw-r--r--Changelog2
-rw-r--r--arm/SelectOp.vp3
-rw-r--r--arm/SelectOpproof.v10
-rw-r--r--backend/PrintAsmaux.ml7
-rw-r--r--backend/SelectDiv.vp58
-rw-r--r--backend/SelectDivproof.v78
-rw-r--r--backend/ValueDomain.v17
-rwxr-xr-xconfigure27
-rw-r--r--cparser/Machine.ml5
-rw-r--r--cparser/Machine.mli2
-rw-r--r--driver/Configuration.ml2
-rw-r--r--driver/Driver.ml3
-rw-r--r--lib/Camlcoq.ml2
-rw-r--r--lib/Integers.v26
-rw-r--r--powerpc/SelectOp.vp3
-rw-r--r--powerpc/SelectOpproof.v10
-rw-r--r--riscV/Archi.v65
-rw-r--r--riscV/Asm.v1172
-rw-r--r--riscV/AsmToJSON.ml18
-rw-r--r--riscV/Asmexpand.ml605
-rw-r--r--riscV/Asmgen.v936
-rw-r--r--riscV/Asmgenproof.v1028
-rw-r--r--riscV/Asmgenproof1.v1411
-rw-r--r--riscV/CBuiltins.ml69
-rw-r--r--riscV/CombineOp.v138
-rw-r--r--riscV/CombineOpproof.v172
-rw-r--r--riscV/ConstpropOp.vp291
-rw-r--r--riscV/ConstpropOpproof.v715
-rw-r--r--riscV/Conventions1.v436
-rw-r--r--riscV/Machregs.v253
-rw-r--r--riscV/Machregsaux.ml33
-rw-r--r--riscV/Machregsaux.mli18
-rw-r--r--riscV/NeedOp.v173
-rw-r--r--riscV/Op.v1340
-rw-r--r--riscV/PrintOp.ml166
-rw-r--r--riscV/SelectLong.vp364
-rw-r--r--riscV/SelectLongproof.v619
-rw-r--r--riscV/SelectOp.vp446
-rw-r--r--riscV/SelectOpproof.v915
-rw-r--r--riscV/Stacklayout.v147
-rw-r--r--riscV/TargetPrinter.ml695
-rw-r--r--riscV/ValueAOp.v218
-rw-r--r--riscV/extractionMachdep.v24
-rw-r--r--runtime/Makefile13
-rw-r--r--runtime/c/i64.h10
-rw-r--r--runtime/c/i64_dtos.c6
-rw-r--r--runtime/c/i64_dtou.c6
-rw-r--r--runtime/c/i64_sar.c2
-rw-r--r--runtime/c/i64_sdiv.c4
-rw-r--r--runtime/c/i64_shl.c2
-rw-r--r--runtime/c/i64_shr.c2
-rw-r--r--runtime/c/i64_smod.c4
-rw-r--r--runtime/c/i64_smulh.c4
-rw-r--r--runtime/c/i64_stod.c2
-rw-r--r--runtime/c/i64_stof.c2
-rw-r--r--runtime/c/i64_udiv.c4
-rw-r--r--runtime/c/i64_udivmod.c22
-rw-r--r--runtime/c/i64_umod.c4
-rw-r--r--runtime/c/i64_umulh.c2
-rw-r--r--runtime/c/i64_utod.c2
-rw-r--r--runtime/c/i64_utof.c2
-rw-r--r--runtime/riscV/sysdeps.h63
-rw-r--r--runtime/riscV/vararg.S90
-rw-r--r--test/Makefile2
-rw-r--r--test/c/Makefile2
-rw-r--r--test/c/aes.c3
-rw-r--r--test/compression/Makefile4
-rw-r--r--test/raytracer/Makefile2
-rw-r--r--test/regression/Makefile6
-rw-r--r--test/regression/Results/builtins-riscV12
-rwxr-xr-xtest/regression/Runtest2
-rw-r--r--test/regression/alignas.c7
-rw-r--r--test/regression/builtins-riscV.c30
-rw-r--r--test/regression/extasm.c2
-rw-r--r--test/regression/floats-basics.c3
-rw-r--r--test/regression/floats.c3
-rw-r--r--test/spass/Makefile2
-rw-r--r--x86/SelectOp.vp3
-rw-r--r--x86/SelectOpproof.v10
80 files changed, 12962 insertions, 102 deletions
diff --git a/.gitignore b/.gitignore
index 638906af..4d19db5d 100644
--- a/.gitignore
+++ b/.gitignore
@@ -37,6 +37,9 @@ powerpc/SelectLong.v
arm/ConstpropOp.v
arm/SelectOp.v
arm/SelectLong.v
+riscV/ConstpropOp.v
+riscV/SelectOp.v
+riscV/SelectLong.v
backend/SelectDiv.v
backend/SplitLong.v
cparser/Parser.v
diff --git a/Changelog b/Changelog
index c55449f5..d5558d61 100644
--- a/Changelog
+++ b/Changelog
@@ -1,3 +1,5 @@
+- New port targeting the RISC-V architecture, in 32- and 64-bit modes.
+
- Removed the compilation of '.cm' files written in Cminor concrete syntax.
Bug fixing:
diff --git a/arm/SelectOp.vp b/arm/SelectOp.vp
index 3d4e8661..4ea1e1a1 100644
--- a/arm/SelectOp.vp
+++ b/arm/SelectOp.vp
@@ -194,6 +194,9 @@ Nondetfunction mul (e1: expr) (e2: expr) :=
| _, _ => Eop Omul (e1:::e2:::Enil)
end.
+Definition mulhs (e1: expr) (e2: expr) := Eop Omulhs (e1 ::: e2 ::: Enil).
+Definition mulhu (e1: expr) (e2: expr) := Eop Omulhu (e1 ::: e2 ::: Enil).
+
(** ** Bitwise and, or, xor *)
Nondetfunction andimm (n1: int) (e2: expr) :=
diff --git a/arm/SelectOpproof.v b/arm/SelectOpproof.v
index dd194498..f025e345 100644
--- a/arm/SelectOpproof.v
+++ b/arm/SelectOpproof.v
@@ -337,6 +337,16 @@ Proof.
TrivialExists.
Qed.
+Theorem eval_mulhs: binary_constructor_sound mulhs Val.mulhs.
+Proof.
+ unfold mulhs; red; intros; TrivialExists.
+Qed.
+
+Theorem eval_mulhu: binary_constructor_sound mulhu Val.mulhu.
+Proof.
+ unfold mulhu; red; intros; TrivialExists.
+Qed.
+
Theorem eval_andimm:
forall n, unary_constructor_sound (andimm n) (fun x => Val.and x (Vint n)).
Proof.
diff --git a/backend/PrintAsmaux.ml b/backend/PrintAsmaux.ml
index ff276af1..09630e29 100644
--- a/backend/PrintAsmaux.ml
+++ b/backend/PrintAsmaux.ml
@@ -91,8 +91,8 @@ let elf_symbol oc symb =
let elf_symbol_offset oc (symb, ofs) =
elf_symbol oc symb;
- let ofs = camlint_of_coqint ofs in
- if ofs <> 0l then fprintf oc " + %ld" ofs
+ let ofs = camlint64_of_ptrofs ofs in
+ if ofs <> 0L then fprintf oc " + %Ld" ofs
(* Functions for fun and var info *)
let elf_print_fun_info oc name =
@@ -142,6 +142,9 @@ let coqint oc n =
let coqint64 oc n =
fprintf oc "%Ld" (camlint64_of_coqint n)
+let ptrofs oc n =
+ fprintf oc "%Ld" (camlint64_of_ptrofs n)
+
(** Programmer-supplied annotations (__builtin_annot). *)
let re_annot_param = Str.regexp "%%\\|%[1-9][0-9]*"
diff --git a/backend/SelectDiv.vp b/backend/SelectDiv.vp
index dc85fb25..96b07e28 100644
--- a/backend/SelectDiv.vp
+++ b/backend/SelectDiv.vp
@@ -19,6 +19,12 @@ Require Import Op CminorSel SelectOp SplitLong SelectLong.
Local Open Scope cminorsel_scope.
+Definition is_intconst (e: expr) : option int :=
+ match e with
+ | Eop (Ointconst n) _ => Some n
+ | _ => None
+ end.
+
(** We try to turn divisions by a constant into a multiplication by
a pseudo-inverse of the divisor. The approach is described in
- Torbjörn Granlund, Peter L. Montgomery: "Division by Invariant
@@ -101,7 +107,7 @@ Definition divlu_mul_params (d: Z) : option (Z * Z) :=
end.
Definition divu_mul (p: Z) (m: Z) :=
- shruimm (Eop Omulhu (Eletvar O ::: Eop (Ointconst (Int.repr m)) Enil ::: Enil))
+ shruimm (mulhu (Eletvar O) (Eop (Ointconst (Int.repr m)) Enil))
(Int.repr p).
Definition divuimm (e1: expr) (n2: int) :=
@@ -117,10 +123,14 @@ Definition divuimm (e1: expr) (n2: int) :=
end
end.
-Nondetfunction divu (e1: expr) (e2: expr) :=
- match e2 with
- | Eop (Ointconst n2) Enil => divuimm e1 n2
- | _ => divu_base e1 e2
+Definition divu (e1: expr) (e2: expr) :=
+ match is_intconst e2, is_intconst e1 with
+ | Some n2, Some n1 =>
+ if Int.eq n2 Int.zero
+ then divu_base e1 e2
+ else Eop (Ointconst (Int.divu n1 n2)) Enil
+ | Some n2, _ => divuimm e1 n2
+ | _, _ => divu_base e1 e2
end.
Definition mod_from_div (equo: expr) (n: int) :=
@@ -139,15 +149,19 @@ Definition moduimm (e1: expr) (n2: int) :=
end
end.
-Nondetfunction modu (e1: expr) (e2: expr) :=
- match e2 with
- | Eop (Ointconst n2) Enil => moduimm e1 n2
- | _ => modu_base e1 e2
+Definition modu (e1: expr) (e2: expr) :=
+ match is_intconst e2, is_intconst e1 with
+ | Some n2, Some n1 =>
+ if Int.eq n2 Int.zero
+ then modu_base e1 e2
+ else Eop (Ointconst (Int.modu n1 n2)) Enil
+ | Some n2, _ => moduimm e1 n2
+ | _, _ => modu_base e1 e2
end.
Definition divs_mul (p: Z) (m: Z) :=
let e2 :=
- Eop Omulhs (Eletvar O ::: Eop (Ointconst (Int.repr m)) Enil ::: Enil) in
+ mulhs (Eletvar O) (Eop (Ointconst (Int.repr m)) Enil) in
let e3 :=
if zlt m Int.half_modulus then e2 else add e2 (Eletvar O) in
add (shrimm e3 (Int.repr p))
@@ -169,10 +183,14 @@ Definition divsimm (e1: expr) (n2: int) :=
end
end.
-Nondetfunction divs (e1: expr) (e2: expr) :=
- match e2 with
- | Eop (Ointconst n2) Enil => divsimm e1 n2
- | _ => divs_base e1 e2
+Definition divs (e1: expr) (e2: expr) :=
+ match is_intconst e2, is_intconst e1 with
+ | Some n2, Some n1 =>
+ if Int.eq n2 Int.zero
+ then divs_base e1 e2
+ else Eop (Ointconst (Int.divs n1 n2)) Enil
+ | Some n2, _ => divsimm e1 n2
+ | _, _ => divs_base e1 e2
end.
Definition modsimm (e1: expr) (n2: int) :=
@@ -191,10 +209,14 @@ Definition modsimm (e1: expr) (n2: int) :=
end
end.
-Nondetfunction mods (e1: expr) (e2: expr) :=
- match e2 with
- | Eop (Ointconst n2) Enil => modsimm e1 n2
- | _ => mods_base e1 e2
+Definition mods (e1: expr) (e2: expr) :=
+ match is_intconst e2, is_intconst e1 with
+ | Some n2, Some n1 =>
+ if Int.eq n2 Int.zero
+ then mods_base e1 e2
+ else Eop (Ointconst (Int.mods n1 n2)) Enil
+ | Some n2, _ => modsimm e1 n2
+ | _, _ => mods_base e1 e2
end.
(** 64-bit integer divisions *)
diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v
index 2ca30e52..5704b32b 100644
--- a/backend/SelectDivproof.v
+++ b/backend/SelectDivproof.v
@@ -488,6 +488,14 @@ Variable sp: val.
Variable e: env.
Variable m: mem.
+Lemma is_intconst_sound:
+ forall v a n le,
+ is_intconst a = Some n -> eval_expr ge sp e m le a v -> v = Vint n.
+Proof with (try discriminate).
+ intros. unfold is_intconst in *.
+ destruct a... destruct o... inv H. inv H0. destruct vl; inv H5. auto.
+Qed.
+
Lemma eval_divu_mul:
forall le x y p M,
divu_mul_params (Int.unsigned y) = Some(p, M) ->
@@ -495,12 +503,10 @@ Lemma eval_divu_mul:
eval_expr ge sp e m le (divu_mul p M) (Vint (Int.divu x y)).
Proof.
intros. unfold divu_mul. exploit (divu_mul_shift x); eauto. intros [A B].
- assert (eval_expr ge sp e m le
- (Eop Omulhu (Eletvar 0 ::: Eop (Ointconst (Int.repr M)) Enil ::: Enil))
- (Vint (Int.mulhu x (Int.repr M)))).
- { EvalOp. econstructor. econstructor; eauto. econstructor. EvalOp. simpl; reflexivity. constructor.
- auto. }
- exploit eval_shruimm. eexact H1. instantiate (1 := Int.repr p).
+ assert (C: eval_expr ge sp e m le (Eletvar 0) (Vint x)) by (apply eval_Eletvar; eauto).
+ assert (D: eval_expr ge sp e m le (Eop (Ointconst (Int.repr M)) Enil) (Vint (Int.repr M))) by EvalOp.
+ exploit eval_mulhu. eexact C. eexact D. intros (v & E & F). simpl in F. inv F.
+ exploit eval_shruimm. eexact E. instantiate (1 := Int.repr p).
intros [v [P Q]]. simpl in Q.
replace (Int.ltu (Int.repr p) Int.iwordsize) with true in Q.
inv Q. rewrite B. auto.
@@ -537,8 +543,15 @@ Theorem eval_divu:
Val.divu x y = Some z ->
exists v, eval_expr ge sp e m le (divu a b) v /\ Val.lessdef z v.
Proof.
- unfold divu; intros until b. destruct (divu_match b); intros.
-- inv H0. inv H5. simpl in H7. inv H7. eapply eval_divuimm; eauto.
+ unfold divu; intros.
+ destruct (is_intconst b) as [n2|] eqn:B.
+- exploit is_intconst_sound; eauto. intros EB; clear B.
+ destruct (is_intconst a) as [n1|] eqn:A.
++ exploit is_intconst_sound; eauto. intros EA; clear A.
+ destruct (Int.eq n2 Int.zero) eqn:Z. eapply eval_divu_base; eauto.
+ subst. simpl in H1. rewrite Z in H1; inv H1.
+ TrivialExists.
++ subst. eapply eval_divuimm; eauto.
- eapply eval_divu_base; eauto.
Qed.
@@ -585,8 +598,15 @@ Theorem eval_modu:
Val.modu x y = Some z ->
exists v, eval_expr ge sp e m le (modu a b) v /\ Val.lessdef z v.
Proof.
- unfold modu; intros until b. destruct (modu_match b); intros.
-- inv H0. inv H5. simpl in H7. inv H7. eapply eval_moduimm; eauto.
+ unfold modu; intros.
+ destruct (is_intconst b) as [n2|] eqn:B.
+- exploit is_intconst_sound; eauto. intros EB; clear B.
+ destruct (is_intconst a) as [n1|] eqn:A.
++ exploit is_intconst_sound; eauto. intros EA; clear A.
+ destruct (Int.eq n2 Int.zero) eqn:Z. eapply eval_modu_base; eauto.
+ subst. simpl in H1. rewrite Z in H1; inv H1.
+ TrivialExists.
++ subst. eapply eval_moduimm; eauto.
- eapply eval_modu_base; eauto.
Qed.
@@ -597,14 +617,10 @@ Lemma eval_divs_mul:
eval_expr ge sp e m le (divs_mul p M) (Vint (Int.divs x y)).
Proof.
intros. unfold divs_mul.
- assert (V: eval_expr ge sp e m le (Eletvar O) (Vint x)).
- { constructor; auto. }
- assert (X: eval_expr ge sp e m le
- (Eop Omulhs (Eletvar 0 ::: Eop (Ointconst (Int.repr M)) Enil ::: Enil))
- (Vint (Int.mulhs x (Int.repr M)))).
- { EvalOp. econstructor. eauto. econstructor. EvalOp. simpl; reflexivity. constructor.
- auto. }
- exploit eval_shruimm. eexact V. instantiate (1 := Int.repr (Int.zwordsize - 1)).
+ assert (C: eval_expr ge sp e m le (Eletvar 0) (Vint x)) by (apply eval_Eletvar; eauto).
+ assert (D: eval_expr ge sp e m le (Eop (Ointconst (Int.repr M)) Enil) (Vint (Int.repr M))) by EvalOp.
+ exploit eval_mulhs. eexact C. eexact D. intros (v & X & F). simpl in F; inv F.
+ exploit eval_shruimm. eexact C. instantiate (1 := Int.repr (Int.zwordsize - 1)).
intros [v1 [Y LD]]. simpl in LD.
change (Int.ltu (Int.repr 31) Int.iwordsize) with true in LD.
simpl in LD. inv LD.
@@ -619,7 +635,7 @@ Proof.
simpl in LD. inv LD.
rewrite B. exact W.
- exploit (divs_mul_shift_2 x); eauto. intros [A B].
- exploit eval_add. eexact X. eexact V. intros [v1 [Z LD]].
+ exploit eval_add. eexact X. eexact C. intros [v1 [Z LD]].
simpl in LD. inv LD.
exploit eval_shrimm. eexact Z. instantiate (1 := Int.repr p). intros [v1 [U LD]].
simpl in LD. rewrite RANGE in LD by auto. inv LD.
@@ -657,8 +673,16 @@ Theorem eval_divs:
Val.divs x y = Some z ->
exists v, eval_expr ge sp e m le (divs a b) v /\ Val.lessdef z v.
Proof.
- unfold divs; intros until b. destruct (divs_match b); intros.
-- inv H0. inv H5. simpl in H7. inv H7. eapply eval_divsimm; eauto.
+ unfold divs; intros.
+ destruct (is_intconst b) as [n2|] eqn:B.
+- exploit is_intconst_sound; eauto. intros EB; clear B.
+ destruct (is_intconst a) as [n1|] eqn:A.
++ exploit is_intconst_sound; eauto. intros EA; clear A.
+ destruct (Int.eq n2 Int.zero) eqn:Z. eapply eval_divs_base; eauto.
+ subst. simpl in H1.
+ destruct (Int.eq n2 Int.zero || Int.eq n1 (Int.repr Int.min_signed) && Int.eq n2 Int.mone); inv H1.
+ TrivialExists.
++ subst. eapply eval_divsimm; eauto.
- eapply eval_divs_base; eauto.
Qed.
@@ -700,8 +724,16 @@ Theorem eval_mods:
Val.mods x y = Some z ->
exists v, eval_expr ge sp e m le (mods a b) v /\ Val.lessdef z v.
Proof.
- unfold mods; intros until b. destruct (mods_match b); intros.
-- inv H0. inv H5. simpl in H7. inv H7. eapply eval_modsimm; eauto.
+ unfold mods; intros.
+ destruct (is_intconst b) as [n2|] eqn:B.
+- exploit is_intconst_sound; eauto. intros EB; clear B.
+ destruct (is_intconst a) as [n1|] eqn:A.
++ exploit is_intconst_sound; eauto. intros EA; clear A.
+ destruct (Int.eq n2 Int.zero) eqn:Z. eapply eval_mods_base; eauto.
+ subst. simpl in H1.
+ destruct (Int.eq n2 Int.zero || Int.eq n1 (Int.repr Int.min_signed) && Int.eq n2 Int.mone); inv H1.
+ TrivialExists.
++ subst. eapply eval_modsimm; eauto.
- eapply eval_mods_base; eauto.
Qed.
diff --git a/backend/ValueDomain.v b/backend/ValueDomain.v
index be8bcccc..4b782286 100644
--- a/backend/ValueDomain.v
+++ b/backend/ValueDomain.v
@@ -1966,6 +1966,22 @@ Proof.
rewrite LTU; auto with va.
Qed.
+(** Pointer operations *)
+
+Definition offset_ptr (v: aval) (n: ptrofs) :=
+ match v with
+ | Ptr p => Ptr (padd p n)
+ | Ifptr p => Ifptr (padd p n)
+ | _ => ntop1 v
+ end.
+
+Lemma offset_ptr_sound:
+ forall v x n, vmatch v x -> vmatch (Val.offset_ptr v n) (offset_ptr x n).
+Proof.
+ intros. unfold Val.offset_ptr, offset_ptr.
+ inv H; constructor; apply padd_sound; assumption.
+Qed.
+
(** Floating-point arithmetic operations *)
Definition negf := unop_float Float.neg.
@@ -4574,6 +4590,7 @@ Hint Resolve cnot_sound symbol_address_sound
negl_sound addl_sound subl_sound
mull_sound mullhs_sound mullhu_sound
divls_sound divlu_sound modls_sound modlu_sound shrxl_sound
+ offset_ptr_sound
negf_sound absf_sound
addf_sound subf_sound mulf_sound divf_sound
negfs_sound absfs_sound
diff --git a/configure b/configure
index b08f95bd..43109e04 100755
--- a/configure
+++ b/configure
@@ -131,6 +131,10 @@ case "$target" in
arch="powerpc"; model="ppc64"; endianness="big"; bitsize=32;;
e5500-*)
arch="powerpc"; model="e5500"; endianness="big"; bitsize=32;;
+ rv32-*)
+ arch="riscV"; model="32"; endianness="little"; bitsize=32;;
+ rv64-*)
+ arch="riscV"; model="64"; endianness="little"; bitsize=64;;
manual)
;;
"")
@@ -367,6 +371,29 @@ if test "$arch" = "x86" -a "$bitsize" = "64"; then
fi
#
+# RISC-V Target Configuration
+#
+if test "$arch" = "riscV"; then
+ if test "$model" = "64"; then
+ model_options="-march=rv64imafd -mabi=lp64d"
+ else
+ model_options="-march=rv32imafd -mabi=ilp32d"
+ fi
+ abi="standard"
+ casm="${toolprefix}gcc"
+ casm_options="$model_options -c"
+ cc="${toolprefix}gcc $model_options"
+ clinker="${toolprefix}gcc"
+ clinker_options="$model_options"
+ cprepro="${toolprefix}gcc"
+ cprepro_options="$model_options -std=c99 -U__GNUC__ -E"
+ libmath="-lm"
+ struct_passing="ref-callee" # wrong!
+ struct_return="ref" # to check!
+ system="linux"
+fi
+
+#
# Finalize Target Configuration
#
if test -z "$casmruntime"; then casmruntime="$casm $casm_options"; fi
diff --git a/cparser/Machine.ml b/cparser/Machine.ml
index 364ebf28..4df80125 100644
--- a/cparser/Machine.ml
+++ b/cparser/Machine.ml
@@ -180,6 +180,11 @@ let arm_bigendian =
{ arm_littleendian with bigendian = true;
bitfields_msb_first = true }
+let rv32 =
+ { ilp32ll64 with name = "rv32" }
+let rv64 =
+ { i32lpll64 with name = "rv64" }
+
(* Add GCC extensions re: sizeof and alignof *)
let gcc_extensions c =
diff --git a/cparser/Machine.mli b/cparser/Machine.mli
index 8ca1e989..b971958d 100644
--- a/cparser/Machine.mli
+++ b/cparser/Machine.mli
@@ -65,6 +65,8 @@ val ppc_32_bigendian : t
val ppc_32_diab_bigendian : t
val arm_littleendian : t
val arm_bigendian : t
+val rv32 : t
+val rv64 : t
val gcc_extensions : t -> t
val compcert_interpreter : t -> t
diff --git a/driver/Configuration.ml b/driver/Configuration.ml
index 87e72f1c..58583330 100644
--- a/driver/Configuration.ml
+++ b/driver/Configuration.ml
@@ -123,7 +123,7 @@ let get_bool_config key =
let arch =
match get_config_string "arch" with
- | "powerpc"|"arm"|"x86" as a -> a
+ | "powerpc"|"arm"|"x86"|"riscV" as a -> a
| v -> bad_config "arch" [v]
let model = get_config_string "model"
let abi = get_config_string "abi"
diff --git a/driver/Driver.ml b/driver/Driver.ml
index a91bb67c..310e4209 100644
--- a/driver/Driver.ml
+++ b/driver/Driver.ml
@@ -461,6 +461,9 @@ let _ =
if Configuration.abi = "macosx"
then Machine.x86_32_macosx
else Machine.x86_32
+ | "riscV" -> if Configuration.model = "64"
+ then Machine.rv64
+ else Machine.rv32
| _ -> assert false
end;
Builtins.set C2C.builtins;
diff --git a/lib/Camlcoq.ml b/lib/Camlcoq.ml
index 63a31181..5c25796e 100644
--- a/lib/Camlcoq.ml
+++ b/lib/Camlcoq.ml
@@ -272,6 +272,8 @@ let coqint_of_camlint : int32 -> Integers.Int.int = Z.of_uint32
let camlint64_of_coqint : Integers.Int64.int -> int64 = Z.to_int64
let coqint_of_camlint64 : int64 -> Integers.Int64.int = Z.of_uint64
(* interpret the int64 as unsigned so that result Z is in range for int *)
+let camlint64_of_ptrofs : Integers.Ptrofs.int -> int64 =
+ fun x -> Z.to_int64 (Integers.Ptrofs.signed x)
(* Atoms (positive integers representing strings) *)
diff --git a/lib/Integers.v b/lib/Integers.v
index 8fd09dd1..b1fa982d 100644
--- a/lib/Integers.v
+++ b/lib/Integers.v
@@ -796,6 +796,12 @@ Proof.
unfold signed. rewrite unsigned_zero. apply zlt_true. generalize half_modulus_pos; omega.
Qed.
+Theorem signed_one: zwordsize > 1 -> signed one = 1.
+Proof.
+ intros. unfold signed. rewrite unsigned_one. apply zlt_true.
+ change 1 with (two_p 0). rewrite half_modulus_power. apply two_p_monotone_strict. omega.
+Qed.
+
Theorem signed_mone: signed mone = -1.
Proof.
unfold signed. rewrite unsigned_mone.
@@ -1844,6 +1850,15 @@ Proof.
destruct (testbit x i); destruct (testbit y i); reflexivity || discriminate.
Qed.
+Theorem xor_is_zero: forall x y, eq (xor x y) zero = eq x y.
+Proof.
+ intros. predSpec eq eq_spec (xor x y) zero.
+- apply xor_zero_equal in H. subst y. rewrite eq_true; auto.
+- predSpec eq eq_spec x y.
++ elim H; subst y; apply xor_idem.
++ auto.
+Qed.
+
Theorem and_xor_distrib:
forall x y z,
and x (xor y z) = xor (and x y) (and x z).
@@ -2933,6 +2948,13 @@ Proof.
- apply Zquot_Zdiv_pos; omega.
Qed.
+Theorem shrx_zero:
+ forall x, zwordsize > 1 -> shrx x zero = x.
+Proof.
+ intros. unfold shrx. rewrite shl_zero. unfold divs. rewrite signed_one by auto.
+ rewrite Z.quot_1_r. apply repr_signed.
+Qed.
+
Theorem shrx_shr:
forall x y,
ltu y (repr (zwordsize - 1)) = true ->
@@ -4080,9 +4102,7 @@ Qed.
Theorem shrx'_zero:
forall x, shrx' x Int.zero = x.
Proof.
- intros. unfold shrx'. rewrite shl'_one_two_p. unfold divs.
- change (signed (repr (two_p (Int.unsigned Int.zero)))) with 1.
- rewrite Z.quot_1_r. apply repr_signed.
+ intros. change (shrx' x Int.zero) with (shrx x zero). apply shrx_zero. compute; auto.
Qed.
Theorem shrx'_shr_2:
diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp
index b5e3ed7e..0a4b3ef6 100644
--- a/powerpc/SelectOp.vp
+++ b/powerpc/SelectOp.vp
@@ -234,6 +234,9 @@ Nondetfunction mul (e1: expr) (e2: expr) :=
| _, _ => Eop Omul (e1:::e2:::Enil)
end.
+Definition mulhs (e1: expr) (e2: expr) := Eop Omulhs (e1 ::: e2 ::: Enil).
+Definition mulhu (e1: expr) (e2: expr) := Eop Omulhu (e1 ::: e2 ::: Enil).
+
(** ** Bitwise and, or, xor *)
Nondetfunction andimm (n1: int) (e2: expr) :=
diff --git a/powerpc/SelectOpproof.v b/powerpc/SelectOpproof.v
index 548fbce2..7f3da409 100644
--- a/powerpc/SelectOpproof.v
+++ b/powerpc/SelectOpproof.v
@@ -367,6 +367,16 @@ Proof.
TrivialExists.
Qed.
+Theorem eval_mulhs: binary_constructor_sound mulhs Val.mulhs.
+Proof.
+ unfold mulhs; red; intros; TrivialExists.
+Qed.
+
+Theorem eval_mulhu: binary_constructor_sound mulhu Val.mulhu.
+Proof.
+ unfold mulhu; red; intros; TrivialExists.
+Qed.
+
Theorem eval_andimm:
forall n, unary_constructor_sound (andimm n) (fun x => Val.and x (Vint n)).
Proof.
diff --git a/riscV/Archi.v b/riscV/Archi.v
new file mode 100644
index 00000000..a1664262
--- /dev/null
+++ b/riscV/Archi.v
@@ -0,0 +1,65 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris *)
+(* Jacques-Henri Jourdan, INRIA Paris *)
+(* *)
+(* Copyright Institut National de Recherche en Informatique et en *)
+(* Automatique. All rights reserved. This file is distributed *)
+(* under the terms of the GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(** Architecture-dependent parameters for RISC-V *)
+
+Require Import ZArith.
+Require Import Fappli_IEEE.
+Require Import Fappli_IEEE_bits.
+
+Parameter ptr64 : bool.
+
+Definition big_endian := false.
+
+Definition align_int64 := 8%Z.
+Definition align_float64 := 8%Z.
+
+Definition splitlong := negb ptr64.
+
+Lemma splitlong_ptr32: splitlong = true -> ptr64 = false.
+Proof.
+ unfold splitlong. destruct ptr64; simpl; congruence.
+Qed.
+
+(** Section 7.3: "Except when otherwise stated, if the result of a
+ floating-point operation is NaN, it is the canonical NaN. The
+ canonical NaN has a positive sign and all significand bits clear
+ except the MSB, a.k.a. the quiet bit."
+ We need to extend the [choose_binop_pl] functions to account for
+ this case. *)
+
+Program Definition default_pl_64 : bool * nan_pl 53 :=
+ (false, iter_nat 51 _ xO xH).
+
+Definition choose_binop_pl_64 (s1: bool) (pl1: nan_pl 53) (s2: bool) (pl2: nan_pl 53) :=
+ false. (**r always choose first NaN *)
+
+Program Definition default_pl_32 : bool * nan_pl 24 :=
+ (false, iter_nat 22 _ xO xH).
+
+Definition choose_binop_pl_32 (s1: bool) (pl1: nan_pl 24) (s2: bool) (pl2: nan_pl 24) :=
+ false. (**r always choose first NaN *)
+
+Definition float_of_single_preserves_sNaN := false.
+
+Global Opaque ptr64 big_endian splitlong
+ default_pl_64 choose_binop_pl_64
+ default_pl_32 choose_binop_pl_32
+ float_of_single_preserves_sNaN.
+
+(** Whether to generate position-independent code or not *)
+
+Parameter pic_code: unit -> bool.
diff --git a/riscV/Asm.v b/riscV/Asm.v
new file mode 100644
index 00000000..4cd3b1fd
--- /dev/null
+++ b/riscV/Asm.v
@@ -0,0 +1,1172 @@
+(* *********************************************************************)
+(* *)
+(* 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 and semantics for RISC-V assembly language. *)
+
+Require Import Coqlib.
+Require Import Maps.
+Require Import AST.
+Require Import Integers.
+Require Import Floats.
+Require Import Values.
+Require Import Memory.
+Require Import Events.
+Require Import Globalenvs.
+Require Import Smallstep.
+Require Import Locations.
+Require Stacklayout.
+Require Import Conventions.
+
+(** * Abstract syntax *)
+
+(** Integer registers. X0 is treated specially because it always reads
+ as zero and is never used as a destination of an instruction. *)
+
+Inductive ireg: Type :=
+ | X1: ireg | X2: ireg | X3: ireg | X4: ireg | X5: ireg
+ | X6: ireg | X7: ireg | X8: ireg | X9: ireg | X10: ireg
+ | X11: ireg | X12: ireg | X13: ireg | X14: ireg | X15: ireg
+ | X16: ireg | X17: ireg | X18: ireg | X19: ireg | X20: ireg
+ | X21: ireg | X22: ireg | X23: ireg | X24: ireg | X25: ireg
+ | X26: ireg | X27: ireg | X28: ireg | X29: ireg | X30: ireg
+ | X31: ireg.
+
+Inductive ireg0: Type :=
+ | X0: ireg0 | X: ireg -> ireg0.
+
+Coercion X: ireg >-> ireg0.
+
+(** Floating-point registers *)
+
+Inductive freg: Type :=
+ | F0: freg | F1: freg | F2: freg | F3: freg
+ | F4: freg | F5: freg | F6: freg | F7: freg
+ | F8: freg | F9: freg | F10: freg | F11: freg
+ | F12: freg | F13: freg | F14: freg | F15: freg
+ | F16: freg | F17: freg | F18: freg | F19: freg
+ | F20: freg | F21: freg | F22: freg | F23: freg
+ | F24: freg | F25: freg | F26: freg | F27: freg
+ | F28: freg | F29: freg | F30: freg | F31: freg.
+
+Lemma ireg_eq: forall (x y: ireg), {x=y} + {x<>y}.
+Proof. decide equality. Defined.
+
+Lemma ireg0_eq: forall (x y: ireg0), {x=y} + {x<>y}.
+Proof. decide equality. apply ireg_eq. Defined.
+
+Lemma freg_eq: forall (x y: freg), {x=y} + {x<>y}.
+Proof. decide equality. Defined.
+
+(** We model the following registers of the RISC-V architecture. *)
+
+Inductive preg: Type :=
+ | IR: ireg -> preg (**r integer registers *)
+ | FR: freg -> preg (**r double-precision float registers *)
+ | PC: preg. (**r program counter *)
+
+Coercion IR: ireg >-> preg.
+Coercion FR: freg >-> preg.
+
+Lemma preg_eq: forall (x y: preg), {x=y} + {x<>y}.
+Proof. decide equality. apply ireg_eq. apply freg_eq. Defined.
+
+Module PregEq.
+ Definition t := preg.
+ Definition eq := preg_eq.
+End PregEq.
+
+Module Pregmap := EMap(PregEq).
+
+(** Conventional names for stack pointer ([SP]) and return address ([RA]). *)
+
+Notation "'SP'" := X2 (only parsing) : asm.
+Notation "'RA'" := X1 (only parsing) : asm.
+
+(** Offsets for load and store instructions. An offset is either an
+ immediate integer or the low part of a symbol. *)
+
+Inductive offset : Type :=
+ | Ofsimm (ofs: ptrofs)
+ | Ofslow (id: ident) (ofs: ptrofs).
+
+(** The RISC-V instruction set is composed of several subsets. We model
+ the "32I" (32-bit integers), "64I" (64-bit integers),
+ "M" (multiplication and division),
+ "F" (single-precision floating-point)
+ and "D" (double-precision floating-point) subsets.
+
+ For 32- and 64-bit integer arithmetic, the RISC-V instruction set comprises
+ generic integer operations such as ADD that operate over the full width
+ of an integer register (either 32 or 64 bit), plus specific instructions
+ such as ADDW that normalize their results to signed 32-bit integers.
+ Other instructions such as AND work equally well over 32- and 64-bit
+ integers, with the convention that 32-bit integers are represented
+ sign-extended in 64-bit registers.
+
+ This clever design is challenging to formalize in the CompCert value
+ model. As a first step, we follow a more traditional approach,
+ also used in the x86 port, whereas we have two sets of (pseudo-)
+ instructions, one for 32-bit integer arithmetic, with suffix W,
+ the other for 64-bit integer arithmetic, with suffix L. The mapping
+ to actual instructions is done when printing assembly code, as follows:
+ - In 32-bit mode:
+ ADDW becomes ADD, ADDL is an error, ANDW becomes AND, ANDL is an error.
+ - In 64-bit mode:
+ ADDW becomes ADDW, ADDL becomes ADD, ANDW and ANDL both become AND.
+*)
+
+Definition label := positive.
+
+(** A note on immediates: there are various constraints on immediate
+ operands to RISC-V instructions. We do not attempt to capture these
+ restrictions in the abstract syntax nor in the semantics. The
+ assembler will emit an error if immediate operands exceed the
+ representable range. Of course, our RISC-V generator (file
+ [Asmgen]) is careful to respect this range. *)
+
+Inductive instruction : Type :=
+ | Pmv (rd: ireg) (rs: ireg) (**r integer move *)
+
+(** 32-bit integer register-immediate instructions *)
+ | Paddiw (rd: ireg) (rs: ireg0) (imm: int) (**r add immediate *)
+ | Psltiw (rd: ireg) (rs: ireg0) (imm: int) (**r set-less-than immediate *)
+ | Psltiuw (rd: ireg) (rs: ireg0) (imm: int) (**r set-less-than unsigned immediate *)
+ | Pandiw (rd: ireg) (rs: ireg0) (imm: int) (**r and immediate *)
+ | Poriw (rd: ireg) (rs: ireg0) (imm: int) (**r or immediate *)
+ | Pxoriw (rd: ireg) (rs: ireg0) (imm: int) (**r xor immediate *)
+ | Pslliw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-left-logical immediate *)
+ | Psrliw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-logical immediate *)
+ | Psraiw (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-arith immediate *)
+ | Pluiw (rd: ireg) (imm: int) (**r load upper-immediate *)
+(** 32-bit integer register-register instructions *)
+ | Paddw (rd: ireg) (rs1 rs2: ireg0) (**r integer addition *)
+ | Psubw (rd: ireg) (rs1 rs2: ireg0) (**r integer subtraction *)
+
+ | Pmulw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply low *)
+ | Pmulhw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high signed *)
+ | Pmulhuw (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high unsigned *)
+ | Pdivw (rd: ireg) (rs1 rs2: ireg0) (**r integer division *)
+ | Pdivuw (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer division *)
+ | Premw (rd: ireg) (rs1 rs2: ireg0) (**r integer remainder *)
+ | Premuw (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer remainder *)
+ | Psltw (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than *)
+ | Psltuw (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than unsigned *)
+ | Pseqw (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 == rs2] (pseudo) *)
+ | Psnew (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 != rs2] (pseudo) *)
+ | Pandw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise and *)
+ | Porw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise or *)
+ | Pxorw (rd: ireg) (rs1 rs2: ireg0) (**r bitwise xor *)
+ | Psllw (rd: ireg) (rs1 rs2: ireg0) (**r shift-left-logical *)
+ | Psrlw (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-logical *)
+ | Psraw (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-arith *)
+
+(** 64-bit integer register-immediate instructions *)
+ | Paddil (rd: ireg) (rs: ireg0) (imm: int64) (**r add immediate *)
+ | Psltil (rd: ireg) (rs: ireg0) (imm: int64) (**r set-less-than immediate *)
+ | Psltiul (rd: ireg) (rs: ireg0) (imm: int64) (**r set-less-than unsigned immediate *)
+ | Pandil (rd: ireg) (rs: ireg0) (imm: int64) (**r and immediate *)
+ | Poril (rd: ireg) (rs: ireg0) (imm: int64) (**r or immediate *)
+ | Pxoril (rd: ireg) (rs: ireg0) (imm: int64) (**r xor immediate *)
+ | Psllil (rd: ireg) (rs: ireg0) (imm: int) (**r shift-left-logical immediate *)
+ | Psrlil (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-logical immediate *)
+ | Psrail (rd: ireg) (rs: ireg0) (imm: int) (**r shift-right-arith immediate *)
+ | Pluil (rd: ireg) (imm: int64) (**r load upper-immediate *)
+(** 64-bit integer register-register instructions *)
+ | Paddl (rd: ireg) (rs1 rs2: ireg0) (**r integer addition *)
+ | Psubl (rd: ireg) (rs1 rs2: ireg0) (**r integer subtraction *)
+
+ | Pmull (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply low *)
+ | Pmulhl (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high signed *)
+ | Pmulhul (rd: ireg) (rs1 rs2: ireg0) (**r integer multiply high unsigned *)
+ | Pdivl (rd: ireg) (rs1 rs2: ireg0) (**r integer division *)
+ | Pdivul (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer division *)
+ | Preml (rd: ireg) (rs1 rs2: ireg0) (**r integer remainder *)
+ | Premul (rd: ireg) (rs1 rs2: ireg0) (**r unsigned integer remainder *)
+ | Psltl (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than *)
+ | Psltul (rd: ireg) (rs1 rs2: ireg0) (**r set-less-than unsigned *)
+ | Pseql (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 == rs2] (pseudo) *)
+ | Psnel (rd: ireg) (rs1 rs2: ireg0) (**r [rd <- rs1 != rs2] (pseudo) *)
+ | Pandl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise and *)
+ | Porl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise or *)
+ | Pxorl (rd: ireg) (rs1 rs2: ireg0) (**r bitwise xor *)
+ | Pslll (rd: ireg) (rs1 rs2: ireg0) (**r shift-left-logical *)
+ | Psrll (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-logical *)
+ | Psral (rd: ireg) (rs1 rs2: ireg0) (**r shift-right-arith *)
+
+ | Pcvtl2w (rd: ireg) (rs: ireg0) (**r int64->int32 (pseudo) *)
+ | Pcvtw2l (r: ireg) (**r int32 signed -> int64 (pseudo) *)
+
+ (* Unconditional jumps. Links are always to X1/RA. *)
+ | Pj_l (l: label) (**r jump to label *)
+ | Pj_s (symb: ident) (sg: signature) (**r jump to symbol *)
+ | Pj_r (r: ireg) (sg: signature) (**r jump register *)
+ | Pjal_s (symb: ident) (sg: signature) (**r jump-and-link symbol *)
+ | Pjal_r (r: ireg) (sg: signature) (**r jump-and-link register *)
+
+ (* Conditional branches, 32-bit comparisons *)
+ | Pbeqw (rs1 rs2: ireg0) (l: label) (**r branch-if-equal *)
+ | Pbnew (rs1 rs2: ireg0) (l: label) (**r branch-if-not-equal signed *)
+ | Pbltw (rs1 rs2: ireg0) (l: label) (**r branch-if-less signed *)
+ | Pbltuw (rs1 rs2: ireg0) (l: label) (**r branch-if-less unsigned *)
+ | Pbgew (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal signed *)
+ | Pbgeuw (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal unsigned *)
+
+ (* Conditional branches, 64-bit comparisons *)
+ | Pbeql (rs1 rs2: ireg0) (l: label) (**r branch-if-equal *)
+ | Pbnel (rs1 rs2: ireg0) (l: label) (**r branch-if-not-equal signed *)
+ | Pbltl (rs1 rs2: ireg0) (l: label) (**r branch-if-less signed *)
+ | Pbltul (rs1 rs2: ireg0) (l: label) (**r branch-if-less unsigned *)
+ | Pbgel (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal signed *)
+ | Pbgeul (rs1 rs2: ireg0) (l: label) (**r branch-if-greater-or-equal unsigned *)
+
+ (* Loads and stores *)
+ | Plb (rd: ireg) (ra: ireg) (ofs: offset) (**r load signed int8 *)
+ | Plbu (rd: ireg) (ra: ireg) (ofs: offset) (**r load unsigned int8 *)
+ | Plh (rd: ireg) (ra: ireg) (ofs: offset) (**r load signed int16 *)
+ | Plhu (rd: ireg) (ra: ireg) (ofs: offset) (**r load unsigned int16 *)
+ | Plw (rd: ireg) (ra: ireg) (ofs: offset) (**r load int32 *)
+ | Plw_a (rd: ireg) (ra: ireg) (ofs: offset) (**r load any32 *)
+ | Pld (rd: ireg) (ra: ireg) (ofs: offset) (**r load int64 *)
+ | Pld_a (rd: ireg) (ra: ireg) (ofs: offset) (**r load any64 *)
+
+ | Psb (rs: ireg) (ra: ireg) (ofs: offset) (**r store int8 *)
+ | Psh (rs: ireg) (ra: ireg) (ofs: offset) (**r store int16 *)
+ | Psw (rs: ireg) (ra: ireg) (ofs: offset) (**r store int32 *)
+ | Psw_a (rs: ireg) (ra: ireg) (ofs: offset) (**r store any32 *)
+ | Psd (rs: ireg) (ra: ireg) (ofs: offset) (**r store int64 *)
+ | Psd_a (rs: ireg) (ra: ireg) (ofs: offset) (**r store any64 *)
+
+ (* Synchronization *)
+ | Pfence (**r fence *)
+
+ (* floating point register move *)
+ | Pfmv (rd: freg) (rs: freg) (**r move *)
+ | Pfmvxs (rd: ireg) (rs: freg) (**r move FP single to integer register *)
+ | Pfmvxd (rd: ireg) (rs: freg) (**r move FP double to integer register *)
+
+ (* 32-bit (single-precision) floating point *)
+ | Pfls (rd: freg) (ra: ireg) (ofs: offset) (**r load float *)
+ | Pfss (rs: freg) (ra: ireg) (ofs: offset) (**r store float *)
+
+ | Pfnegs (rd: freg) (rs: freg) (**r negation *)
+ | Pfabss (rd: freg) (rs: freg) (**r absolute value *)
+
+ | Pfadds (rd: freg) (rs1 rs2: freg) (**r addition *)
+ | Pfsubs (rd: freg) (rs1 rs2: freg) (**r subtraction *)
+ | Pfmuls (rd: freg) (rs1 rs2: freg) (**r multiplication *)
+ | Pfdivs (rd: freg) (rs1 rs2: freg) (**r division *)
+ | Pfmins (rd: freg) (rs1 rs2: freg) (**r minimum *)
+ | Pfmaxs (rd: freg) (rs1 rs2: freg) (**r maximum *)
+
+ | Pfeqs (rd: ireg) (rs1 rs2: freg) (**r compare equal *)
+ | Pflts (rd: ireg) (rs1 rs2: freg) (**r compare less-than *)
+ | Pfles (rd: ireg) (rs1 rs2: freg) (**r compare less-than/equal *)
+
+ | Pfsqrts (rd: freg) (rs: freg) (**r square-root *)
+
+ | Pfmadds (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-add *)
+ | Pfmsubs (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-sub *)
+ | Pfnmadds (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-add *)
+ | Pfnmsubs (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-sub *)
+
+ | Pfcvtws (rd: ireg) (rs: freg) (**r float32 -> int32 conversion *)
+ | Pfcvtwus (rd: ireg) (rs: freg) (**r float32 -> unsigned int32 conversion *)
+ | Pfcvtsw (rd: freg) (rs: ireg0) (**r int32 -> float32 conversion *)
+ | Pfcvtswu (rd: freg) (rs: ireg0) (**r unsigned int32 -> float32 conversion *)
+
+ | Pfcvtls (rd: ireg) (rs: freg) (**r float32 -> int64 conversion *)
+ | Pfcvtlus (rd: ireg) (rs: freg) (**r float32 -> unsigned int64 conversion *)
+ | Pfcvtsl (rd: freg) (rs: ireg0) (**r int64 -> float32 conversion *)
+ | Pfcvtslu (rd: freg) (rs: ireg0) (**r unsigned int 64-> float32 conversion *)
+
+ (* 64-bit (double-precision) floating point *)
+ | Pfld (rd: freg) (ra: ireg) (ofs: offset) (**r load 64-bit float *)
+ | Pfld_a (rd: freg) (ra: ireg) (ofs: offset) (**r load any64 *)
+ | Pfsd (rd: freg) (ra: ireg) (ofs: offset) (**r store 64-bit float *)
+ | Pfsd_a (rd: freg) (ra: ireg) (ofs: offset) (**r store any64 *)
+
+ | Pfnegd (rd: freg) (rs: freg) (**r negation *)
+ | Pfabsd (rd: freg) (rs: freg) (**r absolute value *)
+
+ | Pfaddd (rd: freg) (rs1 rs2: freg) (**r addition *)
+ | Pfsubd (rd: freg) (rs1 rs2: freg) (**r subtraction *)
+ | Pfmuld (rd: freg) (rs1 rs2: freg) (**r multiplication *)
+ | Pfdivd (rd: freg) (rs1 rs2: freg) (**r division *)
+ | Pfmind (rd: freg) (rs1 rs2: freg) (**r minimum *)
+ | Pfmaxd (rd: freg) (rs1 rs2: freg) (**r maximum *)
+
+ | Pfeqd (rd: ireg) (rs1 rs2: freg) (**r compare equal *)
+ | Pfltd (rd: ireg) (rs1 rs2: freg) (**r compare less-than *)
+ | Pfled (rd: ireg) (rs1 rs2: freg) (**r compare less-than/equal *)
+
+ | Pfsqrtd (rd: freg) (rs: freg) (**r square-root *)
+
+ | Pfmaddd (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-add *)
+ | Pfmsubd (rd: freg) (rs1 rs2 rs3: freg) (**r fused multiply-sub *)
+ | Pfnmaddd (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-add *)
+ | Pfnmsubd (rd: freg) (rs1 rs2 rs3: freg) (**r fused negated multiply-sub *)
+
+ | Pfcvtwd (rd: ireg) (rs: freg) (**r float -> int32 conversion *)
+ | Pfcvtwud (rd: ireg) (rs: freg) (**r float -> unsigned int32 conversion *)
+ | Pfcvtdw (rd: freg) (rs: ireg0) (**r int32 -> float conversion *)
+ | Pfcvtdwu (rd: freg) (rs: ireg0) (**r unsigned int32 -> float conversion *)
+
+ | Pfcvtld (rd: ireg) (rs: freg) (**r float -> int64 conversion *)
+ | Pfcvtlud (rd: ireg) (rs: freg) (**r float -> unsigned int64 conversion *)
+ | Pfcvtdl (rd: freg) (rs: ireg0) (**r int64 -> float conversion *)
+ | Pfcvtdlu (rd: freg) (rs: ireg0) (**r unsigned int64 -> float conversion *)
+
+ | Pfcvtds (rd: freg) (rs: freg) (**r float32 -> float *)
+ | Pfcvtsd (rd: freg) (rs: freg) (**r float -> float32 *)
+
+ (* 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 *)
+ | Ploadsymbol_high (rd: ireg) (id: ident) (ofs: ptrofs) (**r load the high part of the address of a symbol *)
+ | Ploadli (rd: ireg) (i: int64) (**r load an immediate int64 *)
+ | Ploadfi (rd: freg) (f: float) (**r load an immediate float *)
+ | Ploadsi (rd: freg) (f: float32) (**r load an immediate single *)
+ | Pbtbl (r: ireg) (tbl: list label) (**r N-way branch through a jump table *)
+ | Pbuiltin: external_function -> list (builtin_arg preg)
+ -> builtin_res preg -> instruction. (**r built-in function (pseudo) *)
+
+
+(** The pseudo-instructions are the following:
+
+- [Plabel]: define a code label at the current program point.
+
+- [Ploadsymbol]: load the address of a symbol in an integer register.
+ Expands to the [la] assembler pseudo-instruction, which does the right
+ thing even if we are in PIC mode.
+
+- [Ploadli rd ival]: load an immediate 64-bit integer into an integer
+ register rd. Expands to a load from an address in the constant data section,
+ using the integer register x31 as temporary:
+<<
+ lui x31, %hi(lbl)
+ ld rd, %lo(lbl)(x31)
+lbl:
+ .quad ival
+>>
+
+- [Ploadfi rd fval]: similar to [Ploadli] but loads a double FP constant fval
+ into a float register rd.
+
+- [Ploadsi rd fval]: similar to [Ploadli] but loads a singe FP constant fval
+ into a float register rd.
+
+- [Pallocframe sz pos]: in the formal semantics, this
+ pseudo-instruction allocates a memory block with bounds [0] and
+ [sz], stores the value of the stack pointer at offset [pos] in this
+ block, and sets the stack pointer to the address of the bottom of
+ this block.
+ In the printed ASM assembly code, this allocation is:
+<<
+ mv x30, sp
+ sub sp, sp, #sz
+ sw x30, #pos(sp)
+>>
+ This cannot be expressed in our memory model, which does not reflect
+ the fact that stack frames are adjacent and allocated/freed
+ following a stack discipline.
+
+- [Pfreeframe sz pos]: in the formal semantics, this pseudo-instruction
+ reads the word at [pos] of the block pointed by the stack pointer,
+ frees this block, and sets the stack pointer to the value read.
+ In the printed ASM assembly code, this freeing is just an increment of [sp]:
+<<
+ add sp, sp, #sz
+>>
+ Again, our memory model cannot comprehend that this operation
+ frees (logically) the current stack frame.
+
+- [Pbtbl reg table]: this is a N-way branch, implemented via a jump table
+ as follows:
+<<
+ la x31, table
+ add x31, x31, reg
+ jr x31
+table: .long table[0], table[1], ...
+>>
+ Note that [reg] contains 4 times the index of the desired table entry.
+
+- [Pseq rd rs1 rs2]: since unsigned comparisons have particular
+ semantics for pointers, we cannot encode equality directly using
+ xor/sub etc, which have only integer semantics.
+<<
+ xor rd, rs1, rs2
+ sltiu rd, rd, 1
+>>
+ The [xor] is omitted if one of [rs1] and [rs2] is [x0].
+
+- [Psne rd rs1 rs2]: similarly for unsigned inequality.
+<<
+ xor rd, rs1, rs2
+ sltu rd, x0, rd
+>>
+*)
+
+Definition code := list instruction.
+Record function : Type := mkfunction { fn_sig: signature; fn_code: code }.
+Definition fundef := AST.fundef function.
+Definition program := AST.program fundef unit.
+
+(** * Operational semantics *)
+
+(** The semantics operates over a single mapping from registers
+ (type [preg]) to values. We maintain
+ the convention that integer registers are mapped to values of
+ type [Tint] or [Tlong] (in 64 bit mode),
+ and float registers to values of type [Tsingle] or [Tfloat]. *)
+
+Definition regset := Pregmap.t val.
+Definition genv := Genv.t fundef unit.
+
+Definition get0w (rs: regset) (r: ireg0) : val :=
+ match r with
+ | X0 => Vint Int.zero
+ | X r => rs r
+ end.
+
+Definition get0l (rs: regset) (r: ireg0) : val :=
+ match r with
+ | X0 => Vlong Int64.zero
+ | X r => rs r
+ end.
+
+Notation "a # b" := (a b) (at level 1, only parsing) : asm.
+Notation "a ## b" := (get0w a b) (at level 1) : asm.
+Notation "a ### b" := (get0l a b) (at level 1) : asm.
+Notation "a # b <- c" := (Pregmap.set b c a) (at level 1, b at next level) : asm.
+
+Open Scope asm.
+
+(** Undefining some registers *)
+
+Fixpoint undef_regs (l: list preg) (rs: regset) : regset :=
+ match l with
+ | nil => rs
+ | r :: l' => undef_regs l' (rs#r <- Vundef)
+ end.
+
+(** Assigning a register pair *)
+
+Definition set_pair (p: rpair preg) (v: val) (rs: regset) : regset :=
+ match p with
+ | One r => rs#r <- v
+ | Twolong rhi rlo => rs#rhi <- (Val.hiword v) #rlo <- (Val.loword v)
+ end.
+
+(** Assigning multiple registers *)
+
+Fixpoint set_regs (rl: list preg) (vl: list val) (rs: regset) : regset :=
+ match rl, vl with
+ | r1 :: rl', v1 :: vl' => set_regs rl' vl' (rs#r1 <- v1)
+ | _, _ => rs
+ end.
+
+(** Assigning the result of a builtin *)
+
+Fixpoint set_res (res: builtin_res preg) (v: val) (rs: regset) : regset :=
+ match res with
+ | BR r => rs#r <- v
+ | BR_none => rs
+ | BR_splitlong hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs)
+ end.
+
+Section RELSEM.
+
+(** Looking up instructions in a code sequence by position. *)
+
+Fixpoint find_instr (pos: Z) (c: code) {struct c} : option instruction :=
+ match c with
+ | nil => None
+ | i :: il => if zeq pos 0 then Some i else find_instr (pos - 1) il
+ end.
+
+(** Position corresponding to a label *)
+
+Definition is_label (lbl: label) (instr: instruction) : bool :=
+ match instr with
+ | Plabel lbl' => if peq lbl lbl' then true else false
+ | _ => false
+ end.
+
+Lemma is_label_correct:
+ forall lbl instr,
+ if is_label lbl instr then instr = Plabel lbl else instr <> Plabel lbl.
+Proof.
+ intros. destruct instr; simpl; try discriminate.
+ case (peq lbl lbl0); intro; congruence.
+Qed.
+
+Fixpoint label_pos (lbl: label) (pos: Z) (c: code) {struct c} : option Z :=
+ match c with
+ | nil => None
+ | instr :: c' =>
+ if is_label lbl instr then Some (pos + 1) else label_pos lbl (pos + 1) c'
+ end.
+
+Variable ge: genv.
+
+(** The two functions below axiomatize how the linker processes
+ symbolic references [symbol + offset)] and splits their
+ actual values into a 20-bit high part [%hi(symbol + offset)] and
+ a 12-bit low part [%lo(symbol + offset)]. *)
+
+Parameter low_half: genv -> ident -> ptrofs -> ptrofs.
+Parameter high_half: genv -> ident -> ptrofs -> val.
+
+(** The fundamental property of these operations is that, when applied
+ to the address of a symbol, their results can be recombined by
+ addition, rebuilding the original address. *)
+
+Axiom low_high_half:
+ forall id ofs,
+ Val.offset_ptr (high_half ge id ofs) (low_half ge id ofs) = Genv.symbol_address ge id ofs.
+
+(** The semantics is purely small-step and defined as a function
+ from the current state (a register set + a memory state)
+ to either [Next rs' m'] where [rs'] and [m'] are the updated register
+ set and memory state after execution of the instruction at [rs#PC],
+ or [Stuck] if the processor is stuck. *)
+
+Inductive outcome: Type :=
+ | Next: regset -> mem -> outcome
+ | Stuck: outcome.
+
+(** Manipulations over the [PC] register: continuing with the next
+ instruction ([nextinstr]) or branching to a label ([goto_label]). *)
+
+Definition nextinstr (rs: regset) :=
+ rs#PC <- (Val.offset_ptr rs#PC Ptrofs.one).
+
+Definition goto_label (f: function) (lbl: label) (rs: regset) (m: mem) :=
+ match label_pos lbl 0 (fn_code f) with
+ | None => Stuck
+ | Some pos =>
+ match rs#PC with
+ | Vptr b ofs => Next (rs#PC <- (Vptr b (Ptrofs.repr pos))) m
+ | _ => Stuck
+ end
+ end.
+
+(** Auxiliaries for memory accesses *)
+
+Definition eval_offset (ofs: offset) : ptrofs :=
+ match ofs with
+ | Ofsimm n => n
+ | Ofslow id delta => low_half ge id delta
+ end.
+
+Definition exec_load (chunk: memory_chunk) (rs: regset) (m: mem)
+ (d: preg) (a: ireg) (ofs: offset) :=
+ match Mem.loadv chunk m (Val.offset_ptr (rs a) (eval_offset ofs)) with
+ | None => Stuck
+ | Some v => Next (nextinstr (rs#d <- v)) m
+ end.
+
+Definition exec_store (chunk: memory_chunk) (rs: regset) (m: mem)
+ (s: preg) (a: ireg) (ofs: offset) :=
+ match Mem.storev chunk m (Val.offset_ptr (rs a) (eval_offset ofs)) (rs s) with
+ | None => Stuck
+ | Some m' => Next (nextinstr rs) m'
+ end.
+
+(** Evaluating a branch *)
+
+Definition eval_branch (f: function) (l: label) (rs: regset) (m: mem) (res: option bool) : outcome :=
+ match res with
+ | Some true => goto_label f l rs m
+ | Some false => Next (nextinstr rs) m
+ | None => Stuck
+ end.
+
+(** Execution of a single instruction [i] in initial state [rs] and
+ [m]. Return updated state. For instructions that correspond to
+ actual RISC-V instructions, the cases are straightforward
+ transliterations of the informal descriptions given in the RISC-V
+ user-mode specification. For pseudo-instructions, refer to the
+ informal descriptions given above.
+
+ Note that we set to [Vundef] the registers used as temporaries by
+ the expansions of the pseudo-instructions, so that the RISC-V code
+ we generate cannot use those registers to hold values that must
+ survive the execution of the pseudo-instruction. *)
+
+Definition exec_instr (f: function) (i: instruction) (rs: regset) (m: mem) : outcome :=
+ match i with
+ | Pmv d s =>
+ Next (nextinstr (rs#d <- (rs#s))) m
+
+(** 32-bit integer register-immediate instructions *)
+ | Paddiw d s i =>
+ Next (nextinstr (rs#d <- (Val.add rs##s (Vint i)))) m
+ | Psltiw d s i =>
+ Next (nextinstr (rs#d <- (Val.cmp Clt rs##s (Vint i)))) m
+ | Psltiuw d s i =>
+ Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Clt rs##s (Vint i)))) m
+ | Pandiw d s i =>
+ Next (nextinstr (rs#d <- (Val.and rs##s (Vint i)))) m
+ | Poriw d s i =>
+ Next (nextinstr (rs#d <- (Val.or rs##s (Vint i)))) m
+ | Pxoriw d s i =>
+ Next (nextinstr (rs#d <- (Val.xor rs##s (Vint i)))) m
+ | Pslliw d s i =>
+ Next (nextinstr (rs#d <- (Val.shl rs##s (Vint i)))) m
+ | Psrliw d s i =>
+ Next (nextinstr (rs#d <- (Val.shru rs##s (Vint i)))) m
+ | Psraiw d s i =>
+ Next (nextinstr (rs#d <- (Val.shr rs##s (Vint i)))) m
+ | Pluiw d i =>
+ Next (nextinstr (rs#d <- (Vint (Int.shl i (Int.repr 12))))) m
+
+(** 32-bit integer register-register instructions *)
+ | Paddw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.add rs##s1 rs##s2))) m
+ | Psubw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.sub rs##s1 rs##s2))) m
+ | Pmulw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.mul rs##s1 rs##s2))) m
+ | Pmulhw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.mulhs rs##s1 rs##s2))) m
+ | Pmulhuw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.mulhu rs##s1 rs##s2))) m
+ | Pdivw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.divs rs##s1 rs##s2)))) m
+ | Pdivuw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.divu rs##s1 rs##s2)))) m
+ | Premw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.mods rs##s1 rs##s2)))) m
+ | Premuw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.modu rs##s1 rs##s2)))) m
+ | Psltw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmp Clt rs##s1 rs##s2))) m
+ | Psltuw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Clt rs##s1 rs##s2))) m
+ | Pseqw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Ceq rs##s1 rs##s2))) m
+ | Psnew d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpu (Mem.valid_pointer m) Cne rs##s1 rs##s2))) m
+ | Pandw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.and rs##s1 rs##s2))) m
+ | Porw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.or rs##s1 rs##s2))) m
+ | Pxorw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.xor rs##s1 rs##s2))) m
+ | Psllw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.shl rs##s1 rs##s2))) m
+ | Psrlw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.shru rs##s1 rs##s2))) m
+ | Psraw d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.shr rs##s1 rs##s2))) m
+
+(** 64-bit integer register-immediate instructions *)
+ | Paddil d s i =>
+ Next (nextinstr (rs#d <- (Val.addl rs###s (Vlong i)))) m
+ | Psltil d s i =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.cmpl Clt rs###s (Vlong i))))) m
+ | Psltiul d s i =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt rs###s (Vlong i))))) m
+ | Pandil d s i =>
+ Next (nextinstr (rs#d <- (Val.andl rs###s (Vlong i)))) m
+ | Poril d s i =>
+ Next (nextinstr (rs#d <- (Val.orl rs###s (Vlong i)))) m
+ | Pxoril d s i =>
+ Next (nextinstr (rs#d <- (Val.xorl rs###s (Vlong i)))) m
+ | Psllil d s i =>
+ Next (nextinstr (rs#d <- (Val.shll rs###s (Vint i)))) m
+ | Psrlil d s i =>
+ Next (nextinstr (rs#d <- (Val.shrlu rs###s (Vint i)))) m
+ | Psrail d s i =>
+ Next (nextinstr (rs#d <- (Val.shrl rs###s (Vint i)))) m
+ | Pluil d i =>
+ Next (nextinstr (rs#d <- (Vlong (Int64.sign_ext 32 (Int64.shl i (Int64.repr 12)))))) m
+
+(** 64-bit integer register-register instructions *)
+ | Paddl d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.addl rs###s1 rs###s2))) m
+ | Psubl d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.subl rs###s1 rs###s2))) m
+ | Pmull d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.mull rs###s1 rs###s2))) m
+ | Pmulhl d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.mullhs rs###s1 rs###s2))) m
+ | Pmulhul d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.mullhu rs###s1 rs###s2))) m
+ | Pdivl d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.divls rs###s1 rs###s2)))) m
+ | Pdivul d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.divlu rs###s1 rs###s2)))) m
+ | Preml d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.modls rs###s1 rs###s2)))) m
+ | Premul d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.modlu rs###s1 rs###s2)))) m
+ | Psltl d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.cmpl Clt rs###s1 rs###s2)))) m
+ | Psltul d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt rs###s1 rs###s2)))) m
+ | Pseql d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Ceq rs###s1 rs###s2)))) m
+ | Psnel d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.cmplu (Mem.valid_pointer m) Cne rs###s1 rs###s2)))) m
+ | Pandl d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.andl rs###s1 rs###s2))) m
+ | Porl d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.orl rs###s1 rs###s2))) m
+ | Pxorl d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.xorl rs###s1 rs###s2))) m
+ | Pslll d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.shll rs###s1 rs###s2))) m
+ | Psrll d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.shrlu rs###s1 rs###s2))) m
+ | Psral d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.shrl rs###s1 rs###s2))) m
+
+ | Pcvtl2w d s =>
+ Next (nextinstr (rs#d <- (Val.loword rs##s))) m
+ | Pcvtw2l r =>
+ Next (nextinstr (rs#r <- (Val.longofint rs#r))) m
+
+(** Unconditional jumps. Links are always to X1/RA. *)
+ | Pj_l l =>
+ goto_label f l rs m
+ | Pj_s s sg =>
+ Next (rs#PC <- (Genv.symbol_address ge s Ptrofs.zero)) m
+ | Pj_r r sg =>
+ Next (rs#PC <- (rs#r)) m
+ | Pjal_s s sg =>
+ Next (rs#PC <- (Genv.symbol_address ge s Ptrofs.zero)
+ #RA <- (Val.offset_ptr rs#PC Ptrofs.one)
+ ) m
+ | Pjal_r r sg =>
+ Next (rs#PC <- (rs#r)
+ #RA <- (Val.offset_ptr rs#PC Ptrofs.one)
+ ) m
+(** Conditional branches, 32-bit comparisons *)
+ | Pbeqw s1 s2 l =>
+ eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Ceq rs##s1 rs##s2)
+ | Pbnew s1 s2 l =>
+ eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Cne rs##s1 rs##s2)
+ | Pbltw s1 s2 l =>
+ eval_branch f l rs m (Val.cmp_bool Clt rs##s1 rs##s2)
+ | Pbltuw s1 s2 l =>
+ eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Clt rs##s1 rs##s2)
+ | Pbgew s1 s2 l =>
+ eval_branch f l rs m (Val.cmp_bool Cge rs##s1 rs##s2)
+ | Pbgeuw s1 s2 l =>
+ eval_branch f l rs m (Val.cmpu_bool (Mem.valid_pointer m) Cge rs##s1 rs##s2)
+
+(** Conditional branches, 64-bit comparisons *)
+ | Pbeql s1 s2 l =>
+ eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Ceq rs###s1 rs###s2)
+ | Pbnel s1 s2 l =>
+ eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Cne rs###s1 rs###s2)
+ | Pbltl s1 s2 l =>
+ eval_branch f l rs m (Val.cmpl_bool Clt rs###s1 rs###s2)
+ | Pbltul s1 s2 l =>
+ eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Clt rs###s1 rs###s2)
+ | Pbgel s1 s2 l =>
+ eval_branch f l rs m (Val.cmpl_bool Cge rs###s1 rs###s2)
+ | Pbgeul s1 s2 l =>
+ eval_branch f l rs m (Val.cmplu_bool (Mem.valid_pointer m) Cge rs###s1 rs###s2)
+
+(** Loads and stores *)
+ | Plb d a ofs =>
+ exec_load Mint8signed rs m d a ofs
+ | Plbu d a ofs =>
+ exec_load Mint8unsigned rs m d a ofs
+ | Plh d a ofs =>
+ exec_load Mint16signed rs m d a ofs
+ | Plhu d a ofs =>
+ exec_load Mint16unsigned rs m d a ofs
+ | Plw d a ofs =>
+ exec_load Mint32 rs m d a ofs
+ | Plw_a d a ofs =>
+ exec_load Many32 rs m d a ofs
+ | Pld d a ofs =>
+ exec_load Mint64 rs m d a ofs
+ | Pld_a d a ofs =>
+ exec_load Many64 rs m d a ofs
+ | Psb s a ofs =>
+ exec_store Mint8unsigned rs m s a ofs
+ | Psh s a ofs =>
+ exec_store Mint16unsigned rs m s a ofs
+ | Psw s a ofs =>
+ exec_store Mint32 rs m s a ofs
+ | Psw_a s a ofs =>
+ exec_store Many32 rs m s a ofs
+ | Psd s a ofs =>
+ exec_store Mint64 rs m s a ofs
+ | Psd_a s a ofs =>
+ exec_store Many64 rs m s a ofs
+
+(** Floating point register move *)
+ | Pfmv d s =>
+ Next (nextinstr (rs#d <- (rs#s))) m
+
+(** 32-bit (single-precision) floating point *)
+ | Pfls d a ofs =>
+ exec_load Mfloat32 rs m d a ofs
+ | Pfss s a ofs =>
+ exec_store Mfloat32 rs m s a ofs
+
+ | Pfnegs d s =>
+ Next (nextinstr (rs#d <- (Val.negfs rs#s))) m
+ | Pfabss d s =>
+ Next (nextinstr (rs#d <- (Val.absfs rs#s))) m
+
+ | Pfadds d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.addfs rs#s1 rs#s2))) m
+ | Pfsubs d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.subfs rs#s1 rs#s2))) m
+ | Pfmuls d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.mulfs rs#s1 rs#s2))) m
+ | Pfdivs d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.divfs rs#s1 rs#s2))) m
+ | Pfeqs d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpfs Ceq rs#s1 rs#s2))) m
+ | Pflts d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpfs Clt rs#s1 rs#s2))) m
+ | Pfles d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpfs Cle rs#s1 rs#s2))) m
+
+ | Pfcvtws d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.intofsingle rs#s)))) m
+ | Pfcvtwus d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.intuofsingle rs#s)))) m
+ | Pfcvtsw d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.singleofint rs##s)))) m
+ | Pfcvtswu d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.singleofintu rs##s)))) m
+
+ | Pfcvtls d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.longofsingle rs#s)))) m
+ | Pfcvtlus d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.longuofsingle rs#s)))) m
+ | Pfcvtsl d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.singleoflong rs###s)))) m
+ | Pfcvtslu d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.singleoflongu rs###s)))) m
+
+(** 64-bit (double-precision) floating point *)
+ | Pfld d a ofs =>
+ exec_load Mfloat64 rs m d a ofs
+ | Pfld_a d a ofs =>
+ exec_load Many64 rs m d a ofs
+ | Pfsd s a ofs =>
+ exec_store Mfloat64 rs m s a ofs
+ | Pfsd_a s a ofs =>
+ exec_store Many64 rs m s a ofs
+
+ | Pfnegd d s =>
+ Next (nextinstr (rs#d <- (Val.negf rs#s))) m
+ | Pfabsd d s =>
+ Next (nextinstr (rs#d <- (Val.absf rs#s))) m
+
+ | Pfaddd d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.addf rs#s1 rs#s2))) m
+ | Pfsubd d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.subf rs#s1 rs#s2))) m
+ | Pfmuld d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.mulf rs#s1 rs#s2))) m
+ | Pfdivd d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.divf rs#s1 rs#s2))) m
+ | Pfeqd d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpf Ceq rs#s1 rs#s2))) m
+ | Pfltd d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpf Clt rs#s1 rs#s2))) m
+ | Pfled d s1 s2 =>
+ Next (nextinstr (rs#d <- (Val.cmpf Cle rs#s1 rs#s2))) m
+
+ | Pfcvtwd d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.intoffloat rs#s)))) m
+ | Pfcvtwud d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.intuoffloat rs#s)))) m
+ | Pfcvtdw d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.floatofint rs##s)))) m
+ | Pfcvtdwu d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.floatofintu rs##s)))) m
+
+ | Pfcvtld d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.longoffloat rs#s)))) m
+ | Pfcvtlud d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.longuoffloat rs#s)))) m
+ | Pfcvtdl d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.floatoflong rs###s)))) m
+ | Pfcvtdlu d s =>
+ Next (nextinstr (rs#d <- (Val.maketotal (Val.floatoflongu rs###s)))) m
+
+ | Pfcvtds d s =>
+ Next (nextinstr (rs#d <- (Val.floatofsingle rs#s))) m
+ | Pfcvtsd d s =>
+ Next (nextinstr (rs#d <- (Val.singleoffloat rs#s))) m
+
+(** Pseudo-instructions *)
+ | Pallocframe sz pos =>
+ let (m1, stk) := Mem.alloc m 0 sz in
+ let sp := (Vptr stk Ptrofs.zero) in
+ match Mem.storev Mptr m1 (Val.offset_ptr sp pos) rs#SP with
+ | None => Stuck
+ | Some m2 => Next (nextinstr (rs #X30 <- (rs SP) #SP <- sp #X31 <- Vundef)) m2
+ end
+ | Pfreeframe sz pos =>
+ match Mem.loadv Mptr m (Val.offset_ptr rs#SP pos) with
+ | None => Stuck
+ | Some v =>
+ match rs SP with
+ | Vptr stk ofs =>
+ match Mem.free m stk 0 sz with
+ | None => Stuck
+ | Some m' => Next (nextinstr (rs#SP <- v #X31 <- Vundef)) m'
+ end
+ | _ => Stuck
+ end
+ end
+ | Plabel lbl =>
+ Next (nextinstr rs) m
+ | Ploadsymbol rd s ofs =>
+ Next (nextinstr (rs#rd <- (Genv.symbol_address ge s ofs))) m
+ | Ploadsymbol_high rd s ofs =>
+ Next (nextinstr (rs#rd <- (high_half ge s ofs))) m
+ | Ploadli rd i =>
+ Next (nextinstr (rs#X31 <- Vundef #rd <- (Vlong i))) m
+ | Ploadfi rd f =>
+ Next (nextinstr (rs#X31 <- Vundef #rd <- (Vfloat f))) m
+ | Ploadsi rd f =>
+ Next (nextinstr (rs#X31 <- Vundef #rd <- (Vsingle f))) m
+ | Pbtbl r tbl =>
+ match rs r with
+ | Vint n =>
+ match list_nth_z tbl (Int.unsigned n) with
+ | None => Stuck
+ | Some lbl => goto_label f lbl (rs#X5 <- Vundef #X31 <- Vundef) m
+ end
+ | _ => Stuck
+ end
+ | Pbuiltin ef args res =>
+ Stuck (**r treated specially below *)
+
+ (** The following instructions and directives are not generated directly by Asmgen,
+ so we do not model them. *)
+ | Pfence
+
+ | Pfmvxs _ _
+ | Pfmvxd _ _
+
+ | Pfmins _ _ _
+ | Pfmaxs _ _ _
+ | Pfsqrts _ _
+ | Pfmadds _ _ _ _
+ | Pfmsubs _ _ _ _
+ | Pfnmadds _ _ _ _
+ | Pfnmsubs _ _ _ _
+
+ | Pfmind _ _ _
+ | Pfmaxd _ _ _
+ | Pfsqrtd _ _
+ | Pfmaddd _ _ _ _
+ | Pfmsubd _ _ _ _
+ | Pfnmaddd _ _ _ _
+ | Pfnmsubd _ _ _ _
+ => Stuck
+ end.
+
+(** Translation of the LTL/Linear/Mach view of machine registers to
+ the RISC-V view. Note that no LTL register maps to [X31]. This
+ register is reserved as temporary, to be used by the generated RV32G
+ code. *)
+
+Definition preg_of (r: mreg) : preg :=
+ match r with
+ | R5 => X5 | R6 => X6 | R7 => X7
+ | R8 => X8 | R9 => X9 | R10 => X10 | R11 => X11
+ | R12 => X12 | R13 => X13 | R14 => X14 | R15 => X15
+ | R16 => X16 | R17 => X17 | R18 => X18 | R19 => X19
+ | R20 => X20 | R21 => X21 | R22 => X22 | R23 => X23
+ | R24 => X24 | R25 => X25 | R26 => X26 | R27 => X27
+ | R28 => X28 | R29 => X29 | R30 => X30
+
+ | Machregs.F0 => F0 | Machregs.F1 => F1 | Machregs.F2 => F2 | Machregs.F3 => F3
+ | Machregs.F4 => F4 | Machregs.F5 => F5 | Machregs.F6 => F6 | Machregs.F7 => F7
+ | Machregs.F8 => F8 | Machregs.F9 => F9 | Machregs.F10 => F10 | Machregs.F11 => F11
+ | Machregs.F12 => F12 | Machregs.F13 => F13 | Machregs.F14 => F14 | Machregs.F15 => F15
+ | Machregs.F16 => F16 | Machregs.F17 => F17 | Machregs.F18 => F18 | Machregs.F19 => F19
+ | Machregs.F20 => F20 | Machregs.F21 => F21 | Machregs.F22 => F22 | Machregs.F23 => F23
+ | Machregs.F24 => F24 | Machregs.F25 => F25 | Machregs.F26 => F26 | Machregs.F27 => F27
+ | Machregs.F28 => F28 | Machregs.F29 => F29 | Machregs.F30 => F30 | Machregs.F31 => F31
+ end.
+
+(** Extract the values of the arguments of an external call.
+ We exploit the calling conventions from module [Conventions], except that
+ we use RISC-V registers instead of locations. *)
+
+Inductive extcall_arg (rs: regset) (m: mem): loc -> val -> Prop :=
+ | extcall_arg_reg: forall r,
+ extcall_arg rs m (R r) (rs (preg_of r))
+ | extcall_arg_stack: forall ofs ty bofs v,
+ bofs = Stacklayout.fe_ofs_arg + 4 * ofs ->
+ Mem.loadv (chunk_of_type ty) m
+ (Val.offset_ptr rs#SP (Ptrofs.repr bofs)) = Some v ->
+ extcall_arg rs m (S Outgoing ofs ty) v.
+
+Inductive extcall_arg_pair (rs: regset) (m: mem): rpair loc -> val -> Prop :=
+ | extcall_arg_one: forall l v,
+ extcall_arg rs m l v ->
+ extcall_arg_pair rs m (One l) v
+ | extcall_arg_twolong: forall hi lo vhi vlo,
+ extcall_arg rs m hi vhi ->
+ extcall_arg rs m lo vlo ->
+ extcall_arg_pair rs m (Twolong hi lo) (Val.longofwords vhi vlo).
+
+Definition extcall_arguments
+ (rs: regset) (m: mem) (sg: signature) (args: list val) : Prop :=
+ list_forall2 (extcall_arg_pair rs m) (loc_arguments sg) args.
+
+Definition loc_external_result (sg: signature) : rpair preg :=
+ map_rpair preg_of (loc_result sg).
+
+(** Execution of the instruction at [rs PC]. *)
+
+Inductive state: Type :=
+ | State: regset -> mem -> state.
+
+Inductive step: state -> trace -> state -> Prop :=
+ | exec_step_internal:
+ forall b ofs f i rs m rs' m',
+ rs PC = Vptr b ofs ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ find_instr (Ptrofs.unsigned ofs) (fn_code f) = Some i ->
+ exec_instr f i rs m = Next rs' m' ->
+ step (State rs m) E0 (State rs' m')
+ | exec_step_builtin:
+ forall b ofs f ef args res rs m vargs t vres rs' m',
+ rs PC = Vptr b ofs ->
+ Genv.find_funct_ptr ge b = Some (Internal f) ->
+ find_instr (Ptrofs.unsigned ofs) f.(fn_code) = Some (Pbuiltin ef args res) ->
+ eval_builtin_args ge rs (rs SP) m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ rs' = nextinstr
+ (set_res res vres
+ (undef_regs (map preg_of (destroyed_by_builtin ef))
+ (rs#X31 <- Vundef))) ->
+ step (State rs m) t (State rs' m')
+ | exec_step_external:
+ forall b ef args res rs m t rs' m',
+ rs PC = Vptr b Ptrofs.zero ->
+ Genv.find_funct_ptr ge b = Some (External ef) ->
+ external_call ef ge args m t res m' ->
+ extcall_arguments rs m (ef_sig ef) args ->
+ rs' = (set_pair (loc_external_result (ef_sig ef) ) res rs)#PC <- (rs RA) ->
+ step (State rs m) t (State rs' m').
+
+End RELSEM.
+
+(** Execution of whole programs. *)
+
+Inductive initial_state (p: program): state -> Prop :=
+ | initial_state_intro: forall m0,
+ let ge := Genv.globalenv p in
+ let rs0 :=
+ (Pregmap.init Vundef)
+ # PC <- (Genv.symbol_address ge p.(prog_main) Ptrofs.zero)
+ # SP <- Vnullptr
+ # RA <- Vnullptr in
+ Genv.init_mem p = Some m0 ->
+ initial_state p (State rs0 m0).
+
+Inductive final_state: state -> int -> Prop :=
+ | final_state_intro: forall rs m r,
+ rs PC = Vnullptr ->
+ rs X10 = Vint r ->
+ final_state (State rs m) r.
+
+Definition semantics (p: program) :=
+ Semantics step (initial_state p) final_state (Genv.globalenv p).
+
+(** Determinacy of the [Asm] semantics. *)
+
+Remark extcall_arguments_determ:
+ forall rs m sg args1 args2,
+ extcall_arguments rs m sg args1 -> extcall_arguments rs m sg args2 -> args1 = args2.
+Proof.
+ intros until m.
+ assert (A: forall l v1 v2,
+ extcall_arg rs m l v1 -> extcall_arg rs m l v2 -> v1 = v2).
+ { intros. inv H; inv H0; congruence. }
+ assert (B: forall p v1 v2,
+ extcall_arg_pair rs m p v1 -> extcall_arg_pair rs m p v2 -> v1 = v2).
+ { intros. inv H; inv H0.
+ eapply A; eauto.
+ f_equal; eapply A; eauto. }
+ assert (C: forall ll vl1, list_forall2 (extcall_arg_pair rs m) ll vl1 ->
+ forall vl2, list_forall2 (extcall_arg_pair rs m) ll vl2 -> vl1 = vl2).
+ {
+ induction 1; intros vl2 EA; inv EA.
+ auto.
+ f_equal; eauto. }
+ intros. eapply C; eauto.
+Qed.
+
+Lemma semantics_determinate: forall p, determinate (semantics p).
+Proof.
+Ltac Equalities :=
+ match goal with
+ | [ H1: ?a = ?b, H2: ?a = ?c |- _ ] =>
+ rewrite H1 in H2; inv H2; Equalities
+ | _ => idtac
+ end.
+ intros; constructor; simpl; intros.
+- (* determ *)
+ inv H; inv H0; Equalities.
+ split. constructor. auto.
+ discriminate.
+ discriminate.
+ assert (vargs0 = vargs) by (eapply eval_builtin_args_determ; eauto). subst vargs0.
+ exploit external_call_determ. eexact H5. eexact H11. intros [A B].
+ split. auto. intros. destruct B; auto. subst. auto.
+ assert (args0 = args) by (eapply extcall_arguments_determ; eauto). subst args0.
+ exploit external_call_determ. eexact H3. eexact H8. intros [A B].
+ split. auto. intros. destruct B; auto. subst. auto.
+- (* trace length *)
+ red; intros. inv H; simpl.
+ omega.
+ eapply external_call_trace_length; eauto.
+ eapply external_call_trace_length; eauto.
+- (* initial states *)
+ inv H; inv H0. f_equal. congruence.
+- (* final no step *)
+ assert (NOTNULL: forall b ofs, Vnullptr <> Vptr b ofs).
+ { intros; unfold Vnullptr; destruct Archi.ptr64; congruence. }
+ inv H. unfold Vzero in H0. red; intros; red; intros.
+ inv H; rewrite H0 in *; eelim NOTNULL; eauto.
+- (* final states *)
+ inv H; inv H0. congruence.
+Qed.
+
+(** Classification functions for processor registers (used in Asmgenproof). *)
+
+Definition data_preg (r: preg) : bool :=
+ match r with
+ | IR RA => false
+ | IR X31 => false
+ | IR _ => true
+ | FR _ => true
+ | PC => false
+ end.
diff --git a/riscV/AsmToJSON.ml b/riscV/AsmToJSON.ml
new file mode 100644
index 00000000..3580b618
--- /dev/null
+++ b/riscV/AsmToJSON.ml
@@ -0,0 +1,18 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Bernhard Schommer, AbsInt Angewandte Informatik GmbH *)
+(* *)
+(* AbsInt Angewandte Informatik GmbH. All rights reserved. This file *)
+(* is distributed under the terms of the INRIA Non-Commercial *)
+(* License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(* Simple functions to serialize RISC-V Asm to JSON *)
+
+(* Dummy function *)
+
+let p_program oc prog =
+ ()
diff --git a/riscV/Asmexpand.ml b/riscV/Asmexpand.ml
new file mode 100644
index 00000000..02e573fc
--- /dev/null
+++ b/riscV/Asmexpand.ml
@@ -0,0 +1,605 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris-Rocquencourt *)
+(* Bernhard Schommer, AbsInt Angewandte Informatik GmbH *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(* Expanding built-ins and some pseudo-instructions by rewriting
+ of the RISC-V assembly code. *)
+
+open Asm
+open Asmexpandaux
+open AST
+open Camlcoq
+open Integers
+
+exception Error of string
+
+(* Useful constants and helper functions *)
+
+let _0 = Integers.Int.zero
+let _1 = Integers.Int.one
+let _2 = coqint_of_camlint 2l
+let _4 = coqint_of_camlint 4l
+let _8 = coqint_of_camlint 8l
+let _16 = coqint_of_camlint 16l
+let _m1 = coqint_of_camlint (-1l)
+
+let wordsize = if Archi.ptr64 then 8 else 4
+
+let align n a = (n + a - 1) land (-a)
+
+(* Emit instruction sequences that set or offset a register by a constant. *)
+
+let expand_loadimm32 dst n =
+ List.iter emit (Asmgen.loadimm32 dst n [])
+let expand_addptrofs dst src n =
+ List.iter emit (Asmgen.addptrofs dst src n [])
+let expand_storeind_ptr src base ofs =
+ List.iter emit (Asmgen.storeind_ptr src base ofs [])
+
+(* Built-ins. They come in two flavors:
+ - annotation statements: take their arguments in registers or stack
+ locations; generate no code;
+ - inlined by the compiler: take their arguments in arbitrary
+ registers.
+*)
+
+(* Fix-up code around calls to variadic functions. Floating-point arguments
+ residing in FP registers need to be moved to integer registers. *)
+
+let int_param_regs = [| X10; X11; X12; X13; X14; X15; X16; X17 |]
+let float_param_regs = [| F10; F11; F12; F13; F14; F15; F16; F17 |]
+
+let rec fixup_variadic_call pos tyl =
+ if pos < 8 then
+ match tyl with
+ | [] ->
+ ()
+ | (Tint | Tany32) :: tyl ->
+ fixup_variadic_call (pos + 1) tyl
+ | Tsingle :: tyl ->
+ let rs =float_param_regs.(pos)
+ and rd = int_param_regs.(pos) in
+ emit (Pfmvxs(rd, rs));
+ fixup_variadic_call (pos + 1) tyl
+ | Tlong :: tyl ->
+ let pos' = if Archi.ptr64 then pos + 1 else align pos 2 + 2 in
+ fixup_variadic_call pos' tyl
+ | (Tfloat | Tany64) :: tyl ->
+ if Archi.ptr64 then begin
+ let rs = float_param_regs.(pos)
+ and rd = int_param_regs.(pos) in
+ emit (Pfmvxd(rd, rs));
+ fixup_variadic_call (pos + 1) tyl
+ end else begin
+ let pos = align pos 2 in
+ if pos < 8 then begin
+ let rs = float_param_regs.(pos)
+ and rd1 = int_param_regs.(pos)
+ and rd2 = int_param_regs.(pos + 1) in
+ emit (Paddiw(X2, X X2, Integers.Int.neg _16));
+ emit (Pfsd(rs, X2, Ofsimm _0));
+ emit (Plw(rd1, X2, Ofsimm _0));
+ emit (Plw(rd2, X2, Ofsimm _4));
+ emit (Paddiw(X2, X X2, _16));
+ fixup_variadic_call (pos + 2) tyl
+ end
+ end
+
+let fixup_call sg =
+ if sg.sig_cc.cc_vararg then fixup_variadic_call 0 sg.sig_args
+
+(* Handling of annotations *)
+
+let expand_annot_val txt targ args res =
+ emit (Pbuiltin (EF_annot(txt,[targ]), args, BR_none));
+ match args, res with
+ | [BA(IR src)], BR(IR dst) ->
+ if dst <> src then emit (Pmv (dst, src))
+ | [BA(FR src)], BR(FR dst) ->
+ if dst <> src then emit (Pfmv (dst, src))
+ | _, _ ->
+ raise (Error "ill-formed __builtin_annot_val")
+
+(* Handling of memcpy *)
+
+(* Unaligned accesses are slow on RISC-V, so don't use them *)
+
+let offset_in_range ofs =
+ let ofs = Z.to_int64 ofs in -2048L <= ofs && ofs < 2048L
+
+let memcpy_small_arg sz arg tmp =
+ match arg with
+ | BA (IR r) ->
+ (r, _0)
+ | BA_addrstack ofs ->
+ if offset_in_range ofs
+ && offset_in_range (Ptrofs.add ofs (Ptrofs.repr (Z.of_uint sz)))
+ then (X2, ofs)
+ else begin expand_addptrofs tmp X2 ofs; (tmp, _0) end
+ | _ ->
+ assert false
+
+let expand_builtin_memcpy_small sz al src dst =
+ let (tsrc, tdst) =
+ if dst <> BA (IR X5) then (X5, X6) else (X6, X5) in
+ let (rsrc, osrc) = memcpy_small_arg sz src tsrc in
+ let (rdst, odst) = memcpy_small_arg sz dst tdst in
+ let rec copy osrc odst sz =
+ if sz >= 8 && al >= 8 then
+ begin
+ emit (Pfld (F0, rsrc, Ofsimm osrc));
+ emit (Pfsd (F0, rdst, Ofsimm odst));
+ copy (Ptrofs.add osrc _8) (Ptrofs.add odst _8) (sz - 8)
+ end
+ else if sz >= 4 && al >= 4 then
+ begin
+ emit (Plw (X31, rsrc, Ofsimm osrc));
+ emit (Psw (X31, rdst, Ofsimm odst));
+ copy (Ptrofs.add osrc _4) (Ptrofs.add odst _4) (sz - 4)
+ end
+ else if sz >= 2 && al >= 2 then
+ begin
+ emit (Plh (X31, rsrc, Ofsimm osrc));
+ emit (Psh (X31, rdst, Ofsimm odst));
+ copy (Ptrofs.add osrc _2) (Ptrofs.add odst _2) (sz - 2)
+ end
+ else if sz >= 1 then
+ begin
+ emit (Plb (X31, rsrc, Ofsimm osrc));
+ emit (Psb (X31, rdst, Ofsimm odst));
+ copy (Ptrofs.add osrc _1) (Ptrofs.add odst _1) (sz - 1)
+ end
+ in copy osrc odst sz
+
+let memcpy_big_arg sz arg tmp =
+ match arg with
+ | BA (IR r) -> if r <> tmp then emit (Pmv(tmp, r))
+ | BA_addrstack ofs ->
+ expand_addptrofs tmp X2 ofs
+ | _ ->
+ assert false
+
+let expand_builtin_memcpy_big sz al src dst =
+ assert (sz >= al);
+ assert (sz mod al = 0);
+ let (s, d) =
+ if dst <> BA (IR X5) then (X5, X6) else (X6, X5) in
+ memcpy_big_arg sz src s;
+ memcpy_big_arg sz dst d;
+ (* Use X7 as loop count, X1 and F0 as ld/st temporaries. *)
+ let (load, store, chunksize) =
+ if al >= 8 then
+ (Pfld (F0, s, Ofsimm _0), Pfsd (F0, d, Ofsimm _0), 8)
+ else if al >= 4 then
+ (Plw (X31, s, Ofsimm _0), Psw (X31, d, Ofsimm _0), 4)
+ else if al = 2 then
+ (Plh (X31, s, Ofsimm _0), Psh (X31, d, Ofsimm _0), 2)
+ else
+ (Plb (X31, s, Ofsimm _0), Psb (X31, d, Ofsimm _0), 1) in
+ expand_loadimm32 X7 (Z.of_uint (sz / chunksize));
+ let delta = Z.of_uint chunksize in
+ let lbl = new_label () in
+ emit (Plabel lbl);
+ emit load;
+ expand_addptrofs s s delta;
+ emit (Paddiw(X7, X X7, _m1));
+ emit store;
+ expand_addptrofs d d delta;
+ emit (Pbnew (X X7, X0, lbl))
+
+let expand_builtin_memcpy sz al args =
+ let (dst, src) =
+ match args with [d; s] -> (d, s) | _ -> assert false in
+ if sz <= 32
+ then expand_builtin_memcpy_small sz al src dst
+ else expand_builtin_memcpy_big sz al src dst
+
+(* Handling of volatile reads and writes *)
+
+let expand_builtin_vload_common chunk base ofs res =
+ match chunk, res with
+ | Mint8unsigned, BR(IR res) ->
+ emit (Plbu (res, base, Ofsimm ofs))
+ | Mint8signed, BR(IR res) ->
+ emit (Plb (res, base, Ofsimm ofs))
+ | Mint16unsigned, BR(IR res) ->
+ emit (Plhu (res, base, Ofsimm ofs))
+ | Mint16signed, BR(IR res) ->
+ emit (Plh (res, base, Ofsimm ofs))
+ | Mint32, BR(IR res) ->
+ emit (Plw (res, base, Ofsimm ofs))
+ | Mint64, BR(IR res) ->
+ emit (Pld (res, base, Ofsimm ofs))
+ | Mint64, BR_splitlong(BR(IR res1), BR(IR res2)) ->
+ let ofs' = Ptrofs.add ofs _4 in
+ if base <> res2 then begin
+ emit (Plw (res2, base, Ofsimm ofs));
+ emit (Plw (res1, base, Ofsimm ofs'))
+ end else begin
+ emit (Plw (res1, base, Ofsimm ofs'));
+ emit (Plw (res2, base, Ofsimm ofs))
+ end
+ | Mfloat32, BR(FR res) ->
+ emit (Pfls (res, base, Ofsimm ofs))
+ | Mfloat64, BR(FR res) ->
+ emit (Pfld (res, base, Ofsimm ofs))
+ | _ ->
+ assert false
+
+let expand_builtin_vload chunk args res =
+ match args with
+ | [BA(IR addr)] ->
+ expand_builtin_vload_common chunk addr _0 res
+ | [BA_addrstack ofs] ->
+ if offset_in_range (Z.add ofs (Memdata.size_chunk chunk)) then
+ expand_builtin_vload_common chunk X2 ofs res
+ else begin
+ expand_addptrofs X31 X2 ofs; (* X31 <- sp + ofs *)
+ expand_builtin_vload_common chunk X31 _0 res
+ end
+ | _ ->
+ assert false
+
+let expand_builtin_vstore_common chunk base ofs src =
+ match chunk, src with
+ | (Mint8signed | Mint8unsigned), BA(IR src) ->
+ emit (Psb (src, base, Ofsimm ofs))
+ | (Mint16signed | Mint16unsigned), BA(IR src) ->
+ emit (Psh (src, base, Ofsimm ofs))
+ | Mint32, BA(IR src) ->
+ emit (Psw (src, base, Ofsimm ofs))
+ | Mint64, BA(IR src) ->
+ emit (Psd (src, base, Ofsimm ofs))
+ | Mint64, BA_splitlong(BA(IR src1), BA(IR src2)) ->
+ let ofs' = Ptrofs.add ofs _4 in
+ emit (Psw (src2, base, Ofsimm ofs));
+ emit (Psw (src1, base, Ofsimm ofs'))
+ | Mfloat32, BA(FR src) ->
+ emit (Pfss (src, base, Ofsimm ofs))
+ | Mfloat64, BA(FR src) ->
+ emit (Pfsd (src, base, Ofsimm ofs))
+ | _ ->
+ assert false
+
+let expand_builtin_vstore chunk args =
+ match args with
+ | [BA(IR addr); src] ->
+ expand_builtin_vstore_common chunk addr _0 src
+ | [BA_addrstack ofs; src] ->
+ if offset_in_range (Z.add ofs (Memdata.size_chunk chunk)) then
+ expand_builtin_vstore_common chunk X2 ofs src
+ else begin
+ expand_addptrofs X31 X2 ofs; (* X31 <- sp + ofs *)
+ expand_builtin_vstore_common chunk X31 _0 src
+ end
+ | _ ->
+ assert false
+
+(* Handling of varargs *)
+
+(* Size in words of the arguments to a function. This includes both
+ arguments passed in registers and arguments passed on stack. *)
+
+let rec args_size sz = function
+ | [] -> sz
+ | (Tint | Tsingle | Tany32) :: l ->
+ args_size (sz + 1) l
+ | (Tlong | Tfloat | Tany64) :: l ->
+ args_size (if Archi.ptr64 then sz + 1 else align sz 2 + 2) l
+
+let arguments_size sg =
+ args_size 0 sg.sig_args
+
+let save_arguments first_reg base_ofs =
+ for i = first_reg to 7 do
+ expand_storeind_ptr
+ int_param_regs.(i)
+ X2
+ (Ptrofs.repr (Z.add base_ofs (Z.of_uint ((i - first_reg) * wordsize))))
+ done
+
+let vararg_start_ofs : Z.t option ref = ref None
+
+let expand_builtin_va_start r =
+ match !vararg_start_ofs with
+ | None ->
+ invalid_arg "Fatal error: va_start used in non-vararg function"
+ | Some ofs ->
+ expand_addptrofs X31 X2 (Ptrofs.repr ofs);
+ expand_storeind_ptr X31 r Ptrofs.zero
+
+(* Auxiliary for 64-bit integer arithmetic built-ins. They expand to
+ two instructions, one computing the low 32 bits of the result,
+ followed by another computing the high 32 bits. In cases where
+ the first instruction would overwrite arguments to the second
+ instruction, we must go through X31 to hold the low 32 bits of the result.
+*)
+
+let expand_int64_arith conflict rl fn =
+ if conflict then (fn X31; emit (Pmv(rl, X31))) else fn rl
+
+(* Byte swaps. There are no specific instructions, so we use standard,
+ not-very-efficient formulas. *)
+
+let expand_bswap16 d s =
+ (* d = (s & 0xFF) << 8 | (s >> 8) & 0xFF *)
+ emit (Pandiw(X31, X s, coqint_of_camlint 0xFFl));
+ emit (Pslliw(X31, X X31, _8));
+ emit (Psrliw(d, X s, _8));
+ emit (Pandiw(d, X d, coqint_of_camlint 0xFFl));
+ emit (Porw(d, X X31, X d))
+
+let expand_bswap32 d s =
+ (* d = (s << 24)
+ | (((s >> 8) & 0xFF) << 16)
+ | (((s >> 16) & 0xFF) << 8)
+ | (s >> 24) *)
+ emit (Pslliw(X1, X s, coqint_of_camlint 24l));
+ emit (Psrliw(X31, X s, _8));
+ emit (Pandiw(X31, X X31, coqint_of_camlint 0xFFl));
+ emit (Pslliw(X31, X X31, _16));
+ emit (Porw(X1, X X1, X X31));
+ emit (Psrliw(X31, X s, _16));
+ emit (Pandiw(X31, X X31, coqint_of_camlint 0xFFl));
+ emit (Pslliw(X31, X X31, _8));
+ emit (Porw(X1, X X1, X X31));
+ emit (Psrliw(X31, X s, coqint_of_camlint 24l));
+ emit (Porw(d, X X1, X X31))
+
+let expand_bswap64 d s =
+ (* d = s << 56
+ | (((s >> 8) & 0xFF) << 48)
+ | (((s >> 16) & 0xFF) << 40)
+ | (((s >> 24) & 0xFF) << 32)
+ | (((s >> 32) & 0xFF) << 24)
+ | (((s >> 40) & 0xFF) << 16)
+ | (((s >> 48) & 0xFF) << 8)
+ | s >> 56 *)
+ emit (Psllil(X1, X s, coqint_of_camlint 56l));
+ List.iter
+ (fun (n1, n2) ->
+ emit (Psrlil(X31, X s, coqint_of_camlint n1));
+ emit (Pandil(X31, X X31, coqint_of_camlint 0xFFl));
+ emit (Psllil(X31, X X31, coqint_of_camlint n2));
+ emit (Porl(X1, X X1, X X31)))
+ [(8l,48l); (16l,40l); (24l,32l); (32l,24l); (40l,16l); (48l,8l)];
+ emit (Psrlil(X31, X s, coqint_of_camlint 56l));
+ emit (Porl(d, X X1, X X31))
+
+(* Handling of compiler-inlined builtins *)
+
+let expand_builtin_inline name args res =
+ match name, args, res with
+ (* Synchronization *)
+ | "__builtin_membar", [], _ ->
+ ()
+ | "__builtin_fence", [], _ ->
+ emit Pfence
+ (* Vararg stuff *)
+ | "__builtin_va_start", [BA(IR a)], _ ->
+ expand_builtin_va_start a
+ (* Byte swaps *)
+ | "__builtin_bswap16", [BA(IR a1)], BR(IR res) ->
+ expand_bswap16 res a1
+ | ("__builtin_bswap"| "__builtin_bswap32"), [BA(IR a1)], BR(IR res) ->
+ expand_bswap32 res a1
+ | "__builtin_bswap64", [BA(IR a1)], BR(IR res) ->
+ expand_bswap64 res a1
+ | "__builtin_bswap64", [BA_splitlong(BA(IR ah), BA(IR al))],
+ BR_splitlong(BR(IR rh), BR(IR rl)) ->
+ assert (ah = X6 && al = X5 && rh = X5 && rl = X6);
+ expand_bswap32 X5 X5;
+ expand_bswap32 X6 X6
+ (* Float arithmetic *)
+ | "__builtin_fabs", [BA(FR a1)], BR(FR res) ->
+ emit (Pfabsd(res, a1))
+ | "__builtin_fsqrt", [BA(FR a1)], BR(FR res) ->
+ emit (Pfsqrtd(res, a1))
+ | "__builtin_fmadd", [BA(FR a1); BA(FR a2); BA(FR a3)], BR(FR res) ->
+ emit (Pfmaddd(res, a1, a2, a3))
+ | "__builtin_fmsub", [BA(FR a1); BA(FR a2); BA(FR a3)], BR(FR res) ->
+ emit (Pfmsubd(res, a1, a2, a3))
+ | "__builtin_fnmadd", [BA(FR a1); BA(FR a2); BA(FR a3)], BR(FR res) ->
+ emit (Pfnmaddd(res, a1, a2, a3))
+ | "__builtin_fnmsub", [BA(FR a1); BA(FR a2); BA(FR a3)], BR(FR res) ->
+ emit (Pfnmsubd(res, a1, a2, a3))
+ | "__builtin_fmin", [BA(FR a1); BA(FR a2)], BR(FR res) ->
+ emit (Pfmind(res, a1, a2))
+ | "__builtin_fmax", [BA(FR a1); BA(FR a2)], BR(FR res) ->
+ emit (Pfmaxd(res, a1, a2))
+ (* 64-bit integer arithmetic for a 32-bit platform *)
+ | "__builtin_negl", [BA_splitlong(BA(IR ah), BA(IR al))],
+ BR_splitlong(BR(IR rh), BR(IR rl)) ->
+ expand_int64_arith (rl = ah) rl
+ (fun rl ->
+ emit (Psltuw (X1, X0, X al));
+ emit (Psubw (rl, X0, X al));
+ emit (Psubw (rh, X0, X ah));
+ emit (Psubw (rh, X rh, X X1)))
+ | "__builtin_addl", [BA_splitlong(BA(IR ah), BA(IR al));
+ BA_splitlong(BA(IR bh), BA(IR bl))],
+ BR_splitlong(BR(IR rh), BR(IR rl)) ->
+ expand_int64_arith (rl = bl || rl = ah || rl = bh) rl
+ (fun rl ->
+ emit (Paddw (rl, X al, X bl));
+ emit (Psltuw (X1, X rl, X bl));
+ emit (Paddw (rh, X ah, X bh));
+ emit (Paddw (rh, X rh, X X1)))
+ | "__builtin_subl", [BA_splitlong(BA(IR ah), BA(IR al));
+ BA_splitlong(BA(IR bh), BA(IR bl))],
+ BR_splitlong(BR(IR rh), BR(IR rl)) ->
+ expand_int64_arith (rl = ah || rl = bh) rl
+ (fun rl ->
+ emit (Psltuw (X1, X al, X bl));
+ emit (Psubw (rl, X al, X bl));
+ emit (Psubw (rh, X ah, X bh));
+ emit (Psubw (rh, X rh, X X1)))
+ | "__builtin_mull", [BA(IR a); BA(IR b)],
+ BR_splitlong(BR(IR rh), BR(IR rl)) ->
+ expand_int64_arith (rl = a || rl = b) rl
+ (fun rl ->
+ emit (Pmulw (rl, X a, X b));
+ emit (Pmulhuw (rh, X a, X b)))
+
+ (* Catch-all *)
+ | _ ->
+ raise (Error ("unrecognized builtin " ^ name))
+
+(* Expansion of instructions *)
+
+let expand_instruction instr =
+ match instr with
+ | Pallocframe (sz, ofs) ->
+ let sg = get_current_function_sig() in
+ emit (Pmv (X30, X2));
+ if sg.sig_cc.cc_vararg then begin
+ let n = arguments_size sg in
+ let extra_sz = if n >= 8 then 0 else align 16 ((8 - n) * wordsize) in
+ let full_sz = Z.add sz (Z.of_uint extra_sz) in
+ expand_addptrofs X2 X2 (Ptrofs.repr (Z.neg full_sz));
+ expand_storeind_ptr X30 X2 ofs;
+ let va_ofs =
+ Z.add full_sz (Z.of_sint ((n - 8) * wordsize)) in
+ vararg_start_ofs := Some va_ofs;
+ save_arguments n va_ofs
+ end else begin
+ expand_addptrofs X2 X2 (Ptrofs.repr (Z.neg sz));
+ expand_storeind_ptr X30 X2 ofs;
+ vararg_start_ofs := None
+ end
+ | Pfreeframe (sz, ofs) ->
+ let sg = get_current_function_sig() in
+ let extra_sz =
+ if sg.sig_cc.cc_vararg then begin
+ let n = arguments_size sg in
+ if n >= 8 then 0 else align 16 ((8 - n) * wordsize)
+ end else 0 in
+ expand_addptrofs X2 X2 (Ptrofs.repr (Z.add sz (Z.of_uint extra_sz)))
+
+ | Pseqw(rd, rs1, rs2) ->
+ (* emulate based on the fact that x == 0 iff x <u 1 (unsigned cmp) *)
+ if rs2 = X0 then begin
+ emit (Psltiuw(rd, rs1, Int.one))
+ end else begin
+ emit (Pxorw(rd, rs1, rs2)); emit (Psltiuw(rd, X rd, Int.one))
+ end
+ | Psnew(rd, rs1, rs2) ->
+ (* emulate based on the fact that x != 0 iff 0 <u x (unsigned cmp) *)
+ if rs2 = X0 then begin
+ emit (Psltuw(rd, X0, rs1))
+ end else begin
+ emit (Pxorw(rd, rs1, rs2)); emit (Psltuw(rd, X0, X rd))
+ end
+ | Pseql(rd, rs1, rs2) ->
+ (* emulate based on the fact that x == 0 iff x <u 1 (unsigned cmp) *)
+ if rs2 = X0 then begin
+ emit (Psltiul(rd, rs1, Int64.one))
+ end else begin
+ emit (Pxorl(rd, rs1, rs2)); emit (Psltiul(rd, X rd, Int64.one))
+ end
+ | Psnel(rd, rs1, rs2) ->
+ (* emulate based on the fact that x != 0 iff 0 <u x (unsigned cmp) *)
+ if rs2 = X0 then begin
+ emit (Psltul(rd, X0, rs1))
+ end else begin
+ emit (Pxorl(rd, rs1, rs2)); emit (Psltul(rd, X0, X rd))
+ end
+ | Pcvtl2w(rd, rs) ->
+ assert Archi.ptr64;
+ emit (Paddiw(rd, rs, Int.zero)) (* 32-bit sign extension *)
+ | Pcvtw2l(r) ->
+ assert Archi.ptr64
+ (* no-operation because the 32-bit integer was kept sign extended already *)
+
+ | Pjal_r(r, sg) ->
+ fixup_call sg; emit instr
+ | Pjal_s(symb, sg) ->
+ fixup_call sg; emit instr
+ | Pj_r(r, sg) when r <> X1 ->
+ fixup_call sg; emit instr
+ | Pj_s(symb, sg) ->
+ fixup_call sg; emit instr
+
+ | Pbuiltin (ef,args,res) ->
+ begin match ef with
+ | EF_builtin (name,sg) ->
+ expand_builtin_inline (camlstring_of_coqstring name) args res
+ | EF_vload chunk ->
+ expand_builtin_vload chunk args res
+ | EF_vstore chunk ->
+ expand_builtin_vstore chunk args
+ | EF_annot_val (txt,targ) ->
+ expand_annot_val txt targ args res
+ | EF_memcpy(sz, al) ->
+ expand_builtin_memcpy (Z.to_int sz) (Z.to_int al) args
+ | EF_annot _ | EF_debug _ | EF_inline_asm _ ->
+ emit instr
+ | _ ->
+ assert false
+ end
+ | _ ->
+ emit instr
+
+(* NOTE: Dwarf register maps for RV32G are not yet specified
+ officially. This is just a placeholder. *)
+let int_reg_to_dwarf = function
+ | X1 -> 1 | X2 -> 2 | X3 -> 3
+ | X4 -> 4 | X5 -> 5 | X6 -> 6 | X7 -> 7
+ | X8 -> 8 | X9 -> 9 | X10 -> 10 | X11 -> 11
+ | X12 -> 12 | X13 -> 13 | X14 -> 14 | X15 -> 15
+ | X16 -> 16 | X17 -> 17 | X18 -> 18 | X19 -> 19
+ | X20 -> 20 | X21 -> 21 | X22 -> 22 | X23 -> 23
+ | X24 -> 24 | X25 -> 25 | X26 -> 26 | X27 -> 27
+ | X28 -> 28 | X29 -> 29 | X30 -> 30 | X31 -> 31
+
+let float_reg_to_dwarf = function
+ | F0 -> 32 | F1 -> 33 | F2 -> 34 | F3 -> 35
+ | F4 -> 36 | F5 -> 37 | F6 -> 38 | F7 -> 39
+ | F8 -> 40 | F9 -> 41 | F10 -> 42 | F11 -> 43
+ | F12 -> 44 | F13 -> 45 | F14 -> 46 | F15 -> 47
+ | F16 -> 48 | F17 -> 49 | F18 -> 50 | F19 -> 51
+ | F20 -> 52 | F21 -> 53 | F22 -> 54 | F23 -> 55
+ | F24 -> 56 | F25 -> 57 | F26 -> 58 | F27 -> 59
+ | F28 -> 60 | F29 -> 61 | F30 -> 62 | F31 -> 63
+
+let preg_to_dwarf = function
+ | IR r -> int_reg_to_dwarf r
+ | FR r -> float_reg_to_dwarf r
+ | _ -> assert false
+
+let expand_function id fn =
+ try
+ set_current_function fn;
+ if !Clflags.option_g then
+ expand_debug id (* sp= *) 2 preg_to_dwarf expand_instruction fn.fn_code
+ else
+ List.iter expand_instruction fn.fn_code;
+ Errors.OK (get_current_function ())
+ with Error s ->
+ Errors.Error (Errors.msg (coqstring_of_camlstring s))
+
+let expand_fundef id = function
+ | Internal f ->
+ begin match expand_function id f with
+ | Errors.OK tf -> Errors.OK (Internal tf)
+ | Errors.Error msg -> Errors.Error msg
+ end
+ | External ef ->
+ Errors.OK (External ef)
+
+let expand_program (p: Asm.program) : Asm.program Errors.res =
+ AST.transform_partial_program2 expand_fundef (fun id v -> Errors.OK v) p
diff --git a/riscV/Asmgen.v b/riscV/Asmgen.v
new file mode 100644
index 00000000..a704ed74
--- /dev/null
+++ b/riscV/Asmgen.v
@@ -0,0 +1,936 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Translation from Mach to RISC-V assembly language *)
+
+Require Archi.
+Require Import Coqlib Errors.
+Require Import AST Integers Floats Memdata.
+Require Import Op Locations Mach Asm.
+
+Local Open Scope string_scope.
+Local Open Scope error_monad_scope.
+
+(** The code generation functions take advantage of several
+ characteristics of the [Mach] code generated by earlier passes of the
+ compiler, mostly that argument and result registers are of the correct
+ types. These properties are true by construction, but it's easier to
+ recheck them during code generation and fail if they do not hold. *)
+
+(** Extracting integer or float registers. *)
+
+Definition ireg_of (r: mreg) : res ireg :=
+ match preg_of r with IR mr => OK mr | _ => Error(msg "Asmgen.ireg_of") end.
+
+Definition freg_of (r: mreg) : res freg :=
+ match preg_of r with FR mr => OK mr | _ => Error(msg "Asmgen.freg_of") end.
+
+(** Decomposition of 32-bit integer constants. They are split into either
+ small signed immediates that fit in 12-bits, or, if they do not fit,
+ into a (20-bit hi, 12-bit lo) pair where lo is sign-extended. *)
+
+Inductive immed32 : Type :=
+ | Imm32_single (imm: int)
+ | Imm32_pair (hi: int) (lo: int).
+
+Definition make_immed32 (val: int) :=
+ let lo := Int.sign_ext 12 val in
+ if Int.eq val lo
+ then Imm32_single val
+ else Imm32_pair (Int.shru (Int.sub val lo) (Int.repr 12)) lo.
+(*
+ let discr := Int.shr val (Int.repr 11) in
+ let hi := Int.shru val (Int.repr 12) in
+ if Int.eq discr Int.zero || Int.eq discr Int.mone
+ then Imm32_single val
+ else Imm32_pair (Int.add hi (Int.and discr Int.one)) (Int.sign_ext 12 val).
+*)
+
+(** Likewise, for 64-bit integer constants. *)
+
+Inductive immed64 : Type :=
+ | Imm64_single (imm: int64)
+ | Imm64_pair (hi: int64) (lo: int64)
+ | Imm64_large (imm: int64).
+
+Definition make_immed64 (val: int64) :=
+ let lo := Int64.sign_ext 12 val in
+ if Int64.eq val lo then Imm64_single lo else
+ let hi := Int64.zero_ext 20 (Int64.shru (Int64.sub val lo) (Int64.repr 12)) in
+ if Int64.eq val (Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo)
+ then Imm64_pair hi lo
+ else Imm64_large val.
+
+(** Smart constructors for arithmetic operations involving
+ a 32-bit or 64-bit integer constant. Depending on whether the
+ constant fits in 12 bits or not, one or several instructions
+ are generated as required to perform the operation
+ and prepended to the given instruction sequence [k]. *)
+
+Definition load_hilo32 (r: ireg) (hi lo: int) k :=
+ if Int.eq lo Int.zero then Pluiw r hi :: k
+ else Pluiw r hi :: Paddiw r r lo :: k.
+
+Definition loadimm32 (r: ireg) (n: int) (k: code) :=
+ match make_immed32 n with
+ | Imm32_single imm => Paddiw r X0 imm :: k
+ | Imm32_pair hi lo => load_hilo32 r hi lo k
+ end.
+
+Definition opimm32 (op: ireg -> ireg0 -> ireg0 -> instruction)
+ (opimm: ireg -> ireg0 -> int -> instruction)
+ (rd rs: ireg) (n: int) (k: code) :=
+ match make_immed32 n with
+ | Imm32_single imm => opimm rd rs imm :: k
+ | Imm32_pair hi lo => load_hilo32 X31 hi lo (op rd rs X31 :: k)
+ end.
+
+Definition addimm32 := opimm32 Paddw Paddiw.
+Definition andimm32 := opimm32 Pandw Pandiw.
+Definition orimm32 := opimm32 Porw Poriw.
+Definition xorimm32 := opimm32 Pxorw Pxoriw.
+Definition sltimm32 := opimm32 Psltw Psltiw.
+Definition sltuimm32 := opimm32 Psltuw Psltiuw.
+
+Definition load_hilo64 (r: ireg) (hi lo: int64) k :=
+ if Int64.eq lo Int64.zero then Pluil r hi :: k
+ else Pluil r hi :: Paddil r r lo :: k.
+
+Definition loadimm64 (r: ireg) (n: int64) (k: code) :=
+ match make_immed64 n with
+ | Imm64_single imm => Paddil r X0 imm :: k
+ | Imm64_pair hi lo => load_hilo64 r hi lo k
+ | Imm64_large imm => Ploadli r imm :: k
+ end.
+
+Definition opimm64 (op: ireg -> ireg0 -> ireg0 -> instruction)
+ (opimm: ireg -> ireg0 -> int64 -> instruction)
+ (rd rs: ireg) (n: int64) (k: code) :=
+ match make_immed64 n with
+ | Imm64_single imm => opimm rd rs imm :: k
+ | Imm64_pair hi lo => load_hilo64 X31 hi lo (op rd rs X31 :: k)
+ | Imm64_large imm => Ploadli X31 imm :: op rd rs X31 :: k
+ end.
+
+Definition addimm64 := opimm64 Paddl Paddil.
+Definition andimm64 := opimm64 Pandl Pandil.
+Definition orimm64 := opimm64 Porl Poril.
+Definition xorimm64 := opimm64 Pxorl Pxoril.
+Definition sltimm64 := opimm64 Psltl Psltil.
+Definition sltuimm64 := opimm64 Psltul Psltiul.
+
+Definition addptrofs (rd rs: ireg) (n: ptrofs) (k: code) :=
+ if Ptrofs.eq_dec n Ptrofs.zero then
+ Pmv rd rs :: k
+ else
+ if Archi.ptr64
+ then addimm64 rd rs (Ptrofs.to_int64 n) k
+ else addimm32 rd rs (Ptrofs.to_int n) k.
+
+(** Translation of conditional branches. *)
+
+Definition transl_cbranch_int32s (cmp: comparison) (r1 r2: ireg0) (lbl: label) :=
+ match cmp with
+ | Ceq => Pbeqw r1 r2 lbl
+ | Cne => Pbnew r1 r2 lbl
+ | Clt => Pbltw r1 r2 lbl
+ | Cle => Pbgew r2 r1 lbl
+ | Cgt => Pbltw r2 r1 lbl
+ | Cge => Pbgew r1 r2 lbl
+ end.
+
+Definition transl_cbranch_int32u (cmp: comparison) (r1 r2: ireg0) (lbl: label) :=
+ match cmp with
+ | Ceq => Pbeqw r1 r2 lbl
+ | Cne => Pbnew r1 r2 lbl
+ | Clt => Pbltuw r1 r2 lbl
+ | Cle => Pbgeuw r2 r1 lbl
+ | Cgt => Pbltuw r2 r1 lbl
+ | Cge => Pbgeuw r1 r2 lbl
+ end.
+
+Definition transl_cbranch_int64s (cmp: comparison) (r1 r2: ireg0) (lbl: label) :=
+ match cmp with
+ | Ceq => Pbeql r1 r2 lbl
+ | Cne => Pbnel r1 r2 lbl
+ | Clt => Pbltl r1 r2 lbl
+ | Cle => Pbgel r2 r1 lbl
+ | Cgt => Pbltl r2 r1 lbl
+ | Cge => Pbgel r1 r2 lbl
+ end.
+
+Definition transl_cbranch_int64u (cmp: comparison) (r1 r2: ireg0) (lbl: label) :=
+ match cmp with
+ | Ceq => Pbeql r1 r2 lbl
+ | Cne => Pbnel r1 r2 lbl
+ | Clt => Pbltul r1 r2 lbl
+ | Cle => Pbgeul r2 r1 lbl
+ | Cgt => Pbltul r2 r1 lbl
+ | Cge => Pbgeul r1 r2 lbl
+ end.
+
+Definition transl_cond_float (cmp: comparison) (rd: ireg) (fs1 fs2: freg) :=
+ match cmp with
+ | Ceq => (Pfeqd rd fs1 fs2, true)
+ | Cne => (Pfeqd rd fs1 fs2, false)
+ | Clt => (Pfltd rd fs1 fs2, true)
+ | Cle => (Pfled rd fs1 fs2, true)
+ | Cgt => (Pfltd rd fs2 fs1, true)
+ | Cge => (Pfled rd fs2 fs1, true)
+ end.
+
+Definition transl_cond_single (cmp: comparison) (rd: ireg) (fs1 fs2: freg) :=
+ match cmp with
+ | Ceq => (Pfeqs rd fs1 fs2, true)
+ | Cne => (Pfeqs rd fs1 fs2, false)
+ | Clt => (Pflts rd fs1 fs2, true)
+ | Cle => (Pfles rd fs1 fs2, true)
+ | Cgt => (Pflts rd fs2 fs1, true)
+ | Cge => (Pfles rd fs2 fs1, true)
+ end.
+
+Definition transl_cbranch
+ (cond: condition) (args: list mreg) (lbl: label) (k: code) :=
+ match cond, args with
+ | Ccomp c, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (transl_cbranch_int32s c r1 r2 lbl :: k)
+ | Ccompu c, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (transl_cbranch_int32u c r1 r2 lbl :: k)
+ | Ccompimm c n, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (if Int.eq n Int.zero then
+ transl_cbranch_int32s c r1 X0 lbl :: k
+ else
+ loadimm32 X31 n (transl_cbranch_int32s c r1 X31 lbl :: k))
+ | Ccompuimm c n, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (if Int.eq n Int.zero then
+ transl_cbranch_int32u c r1 X0 lbl :: k
+ else
+ loadimm32 X31 n (transl_cbranch_int32u c r1 X31 lbl :: k))
+ | Ccompl c, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (transl_cbranch_int64s c r1 r2 lbl :: k)
+ | Ccomplu c, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (transl_cbranch_int64u c r1 r2 lbl :: k)
+ | Ccomplimm c n, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (if Int64.eq n Int64.zero then
+ transl_cbranch_int64s c r1 X0 lbl :: k
+ else
+ loadimm64 X31 n (transl_cbranch_int64s c r1 X31 lbl :: k))
+ | Ccompluimm c n, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (if Int64.eq n Int64.zero then
+ transl_cbranch_int64u c r1 X0 lbl :: k
+ else
+ loadimm64 X31 n (transl_cbranch_int64u c r1 X31 lbl :: k))
+ | Ccompf c, f1 :: f2 :: nil =>
+ do r1 <- freg_of f1; do r2 <- freg_of f2;
+ let (insn, normal) := transl_cond_float c X31 r1 r2 in
+ OK (insn :: (if normal then Pbnew X31 X0 lbl else Pbeqw X31 X0 lbl) :: k)
+ | Cnotcompf c, f1 :: f2 :: nil =>
+ do r1 <- freg_of f1; do r2 <- freg_of f2;
+ let (insn, normal) := transl_cond_float c X31 r1 r2 in
+ OK (insn :: (if normal then Pbeqw X31 X0 lbl else Pbnew X31 X0 lbl) :: k)
+ | Ccompfs c, f1 :: f2 :: nil =>
+ do r1 <- freg_of f1; do r2 <- freg_of f2;
+ let (insn, normal) := transl_cond_single c X31 r1 r2 in
+ OK (insn :: (if normal then Pbnew X31 X0 lbl else Pbeqw X31 X0 lbl) :: k)
+ | Cnotcompfs c, f1 :: f2 :: nil =>
+ do r1 <- freg_of f1; do r2 <- freg_of f2;
+ let (insn, normal) := transl_cond_single c X31 r1 r2 in
+ OK (insn :: (if normal then Pbeqw X31 X0 lbl else Pbnew X31 X0 lbl) :: k)
+ | _, _ =>
+ Error(msg "Asmgen.transl_cond_branch")
+ end.
+
+(** Translation of a condition operator. The generated code sets the
+ [rd] target register to 0 or 1 depending on the truth value of the
+ condition. *)
+
+Definition transl_cond_int32s (cmp: comparison) (rd: ireg) (r1 r2: ireg0) (k: code) :=
+ match cmp with
+ | Ceq => Pseqw rd r1 r2 :: k
+ | Cne => Psnew rd r1 r2 :: k
+ | Clt => Psltw rd r1 r2 :: k
+ | Cle => Psltw rd r2 r1 :: Pxoriw rd rd Int.one :: k
+ | Cgt => Psltw rd r2 r1 :: k
+ | Cge => Psltw rd r1 r2 :: Pxoriw rd rd Int.one :: k
+ end.
+
+Definition transl_cond_int32u (cmp: comparison) (rd: ireg) (r1 r2: ireg0) (k: code) :=
+ match cmp with
+ | Ceq => Pseqw rd r1 r2 :: k
+ | Cne => Psnew rd r1 r2 :: k
+ | Clt => Psltuw rd r1 r2 :: k
+ | Cle => Psltuw rd r2 r1 :: Pxoriw rd rd Int.one :: k
+ | Cgt => Psltuw rd r2 r1 :: k
+ | Cge => Psltuw rd r1 r2 :: Pxoriw rd rd Int.one :: k
+ end.
+
+Definition transl_cond_int64s (cmp: comparison) (rd: ireg) (r1 r2: ireg0) (k: code) :=
+ match cmp with
+ | Ceq => Pseql rd r1 r2 :: k
+ | Cne => Psnel rd r1 r2 :: k
+ | Clt => Psltl rd r1 r2 :: k
+ | Cle => Psltl rd r2 r1 :: Pxoriw rd rd Int.one :: k
+ | Cgt => Psltl rd r2 r1 :: k
+ | Cge => Psltl rd r1 r2 :: Pxoriw rd rd Int.one :: k
+ end.
+
+Definition transl_cond_int64u (cmp: comparison) (rd: ireg) (r1 r2: ireg0) (k: code) :=
+ match cmp with
+ | Ceq => Pseql rd r1 r2 :: k
+ | Cne => Psnel rd r1 r2 :: k
+ | Clt => Psltul rd r1 r2 :: k
+ | Cle => Psltul rd r2 r1 :: Pxoriw rd rd Int.one :: k
+ | Cgt => Psltul rd r2 r1 :: k
+ | Cge => Psltul rd r1 r2 :: Pxoriw rd rd Int.one :: k
+ end.
+
+Definition transl_condimm_int32s (cmp: comparison) (rd: ireg) (r1: ireg) (n: int) (k: code) :=
+ if Int.eq n Int.zero then transl_cond_int32s cmp rd r1 X0 k else
+ match cmp with
+ | Ceq | Cne => xorimm32 rd r1 n (transl_cond_int32s cmp rd rd X0 k)
+ | Clt => sltimm32 rd r1 n k
+ | Cle => if Int.eq n (Int.repr Int.max_signed)
+ then loadimm32 rd Int.one k
+ else sltimm32 rd r1 (Int.add n Int.one) k
+ | _ => loadimm32 X31 n (transl_cond_int32s cmp rd r1 X31 k)
+ end.
+
+Definition transl_condimm_int32u (cmp: comparison) (rd: ireg) (r1: ireg) (n: int) (k: code) :=
+ if Int.eq n Int.zero then transl_cond_int32u cmp rd r1 X0 k else
+ match cmp with
+ | Clt => sltuimm32 rd r1 n k
+ | _ => loadimm32 X31 n (transl_cond_int32u cmp rd r1 X31 k)
+ end.
+
+Definition transl_condimm_int64s (cmp: comparison) (rd: ireg) (r1: ireg) (n: int64) (k: code) :=
+ if Int64.eq n Int64.zero then transl_cond_int64s cmp rd r1 X0 k else
+ match cmp with
+ | Ceq | Cne => xorimm64 rd r1 n (transl_cond_int64s cmp rd rd X0 k)
+ | Clt => sltimm64 rd r1 n k
+ | Cle => if Int64.eq n (Int64.repr Int64.max_signed)
+ then loadimm32 rd Int.one k
+ else sltimm64 rd r1 (Int64.add n Int64.one) k
+ | _ => loadimm64 X31 n (transl_cond_int64s cmp rd r1 X31 k)
+ end.
+
+Definition transl_condimm_int64u (cmp: comparison) (rd: ireg) (r1: ireg) (n: int64) (k: code) :=
+ if Int64.eq n Int64.zero then transl_cond_int64u cmp rd r1 X0 k else
+ match cmp with
+ | Clt => sltuimm64 rd r1 n k
+ | _ => loadimm64 X31 n (transl_cond_int64u cmp rd r1 X31 k)
+ end.
+
+Definition transl_cond_op
+ (cond: condition) (rd: ireg) (args: list mreg) (k: code) :=
+ match cond, args with
+ | Ccomp c, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (transl_cond_int32s c rd r1 r2 k)
+ | Ccompu c, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (transl_cond_int32u c rd r1 r2 k)
+ | Ccompimm c n, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (transl_condimm_int32s c rd r1 n k)
+ | Ccompuimm c n, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (transl_condimm_int32u c rd r1 n k)
+ | Ccompl c, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (transl_cond_int64s c rd r1 r2 k)
+ | Ccomplu c, a1 :: a2 :: nil =>
+ do r1 <- ireg_of a1; do r2 <- ireg_of a2;
+ OK (transl_cond_int64u c rd r1 r2 k)
+ | Ccomplimm c n, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (transl_condimm_int64s c rd r1 n k)
+ | Ccompluimm c n, a1 :: nil =>
+ do r1 <- ireg_of a1;
+ OK (transl_condimm_int64u c rd r1 n k)
+ | Ccompf c, f1 :: f2 :: nil =>
+ do r1 <- freg_of f1; do r2 <- freg_of f2;
+ let (insn, normal) := transl_cond_float c rd r1 r2 in
+ OK (insn :: if normal then k else Pxoriw rd rd Int.one :: k)
+ | Cnotcompf c, f1 :: f2 :: nil =>
+ do r1 <- freg_of f1; do r2 <- freg_of f2;
+ let (insn, normal) := transl_cond_float c rd r1 r2 in
+ OK (insn :: if normal then Pxoriw rd rd Int.one :: k else k)
+ | Ccompfs c, f1 :: f2 :: nil =>
+ do r1 <- freg_of f1; do r2 <- freg_of f2;
+ let (insn, normal) := transl_cond_single c rd r1 r2 in
+ OK (insn :: if normal then k else Pxoriw rd rd Int.one :: k)
+ | Cnotcompfs c, f1 :: f2 :: nil =>
+ do r1 <- freg_of f1; do r2 <- freg_of f2;
+ let (insn, normal) := transl_cond_single c rd r1 r2 in
+ OK (insn :: if normal then Pxoriw rd rd Int.one :: k else k)
+ | _, _ =>
+ Error(msg "Asmgen.transl_cond_op")
+ end.
+
+(** Translation of the arithmetic operation [r <- op(args)].
+ The corresponding instructions are prepended to [k]. *)
+
+Definition transl_op
+ (op: operation) (args: list mreg) (res: mreg) (k: code) :=
+ match op, args with
+ | Omove, a1 :: nil =>
+ match preg_of res, preg_of a1 with
+ | IR r, IR a => OK (Pmv r a :: k)
+ | FR r, FR a => OK (Pfmv r a :: k)
+ | _ , _ => Error(msg "Asmgen.Omove")
+ end
+ | Ointconst n, nil =>
+ do rd <- ireg_of res;
+ OK (loadimm32 rd n k)
+ | Olongconst n, nil =>
+ do rd <- ireg_of res;
+ OK (loadimm64 rd n k)
+ | Ofloatconst f, nil =>
+ do rd <- freg_of res;
+ OK (if Float.eq_dec f Float.zero
+ then Pfcvtdw rd X0 :: k
+ else Ploadfi rd f :: k)
+ | Osingleconst f, nil =>
+ do rd <- freg_of res;
+ OK (if Float32.eq_dec f Float32.zero
+ then Pfcvtsw rd X0 :: k
+ else Ploadsi rd f :: k)
+ | Oaddrsymbol s ofs, nil =>
+ do rd <- ireg_of res;
+ OK (if Archi.pic_code tt && negb (Ptrofs.eq ofs Ptrofs.zero)
+ then Ploadsymbol rd s Ptrofs.zero :: addptrofs rd rd ofs k
+ else Ploadsymbol rd s ofs :: k)
+ | Oaddrstack n, nil =>
+ do rd <- ireg_of res;
+ OK (addptrofs rd SP n k)
+
+ | Ocast8signed, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Pslliw rd rs (Int.repr 24) :: Psraiw rd rd (Int.repr 24) :: k)
+ | Ocast16signed, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Pslliw rd rs (Int.repr 16) :: Psraiw rd rd (Int.repr 16) :: k)
+ | Oadd, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Paddw rd rs1 rs2 :: k)
+ | Oaddimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (addimm32 rd rs n k)
+ | Oneg, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Psubw rd X0 rs :: k)
+ | Osub, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Psubw rd rs1 rs2 :: k)
+ | Omul, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pmulw rd rs1 rs2 :: k)
+ | Omulhs, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pmulhw rd rs1 rs2 :: k)
+ | Omulhu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pmulhuw rd rs1 rs2 :: k)
+ | Odiv, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pdivw rd rs1 rs2 :: k)
+ | Odivu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pdivuw rd rs1 rs2 :: k)
+ | Omod, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Premw rd rs1 rs2 :: k)
+ | Omodu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Premuw rd rs1 rs2 :: k)
+ | Oand, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pandw rd rs1 rs2 :: k)
+ | Oandimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (andimm32 rd rs n k)
+ | Oor, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Porw rd rs1 rs2 :: k)
+ | Oorimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (orimm32 rd rs n k)
+ | Oxor, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pxorw rd rs1 rs2 :: k)
+ | Oxorimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (xorimm32 rd rs n k)
+ | Oshl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Psllw rd rs1 rs2 :: k)
+ | Oshlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Pslliw rd rs n :: k)
+ | Oshr, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Psraw rd rs1 rs2 :: k)
+ | Oshrimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Psraiw rd rs n :: k)
+ | Oshru, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Psrlw rd rs1 rs2 :: k)
+ | Oshruimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Psrliw rd rs n :: k)
+ | Oshrximm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (if Int.eq n Int.zero then Pmv rd rs :: k else
+ Psraiw X31 rs (Int.repr 31) ::
+ Psrliw X31 X31 (Int.sub Int.iwordsize n) ::
+ Paddw X31 rs X31 ::
+ Psraiw rd X31 n :: k)
+
+ (* [Omakelong], [Ohighlong] should not occur *)
+ | Olowlong, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Pcvtl2w rd rs :: k)
+ | Ocast32signed, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ assertion (ireg_eq rd rs);
+ OK (Pcvtw2l rd :: k)
+ | Ocast32unsigned, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ assertion (ireg_eq rd rs);
+ OK (Pcvtw2l rd :: Psllil rd rd (Int.repr 32) :: Psrlil rd rd (Int.repr 32) :: k)
+ | Oaddl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Paddl rd rs1 rs2 :: k)
+ | Oaddlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (addimm64 rd rs n k)
+ | Onegl, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Psubl rd X0 rs :: k)
+ | Osubl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Psubl rd rs1 rs2 :: k)
+ | Omull, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pmull rd rs1 rs2 :: k)
+ | Omullhs, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pmulhl rd rs1 rs2 :: k)
+ | Omullhu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pmulhul rd rs1 rs2 :: k)
+ | Odivl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pdivl rd rs1 rs2 :: k)
+ | Odivlu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pdivul rd rs1 rs2 :: k)
+ | Omodl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Preml rd rs1 rs2 :: k)
+ | Omodlu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Premul rd rs1 rs2 :: k)
+ | Oandl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pandl rd rs1 rs2 :: k)
+ | Oandlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (andimm64 rd rs n k)
+ | Oorl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Porl rd rs1 rs2 :: k)
+ | Oorlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (orimm64 rd rs n k)
+ | Oxorl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pxorl rd rs1 rs2 :: k)
+ | Oxorlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (xorimm64 rd rs n k)
+ | Oshll, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Pslll rd rs1 rs2 :: k)
+ | Oshllimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Psllil rd rs n :: k)
+ | Oshrl, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Psral rd rs1 rs2 :: k)
+ | Oshrlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Psrail rd rs n :: k)
+ | Oshrlu, a1 :: a2 :: nil =>
+ do rd <- ireg_of res; do rs1 <- ireg_of a1; do rs2 <- ireg_of a2;
+ OK (Psrll rd rs1 rs2 :: k)
+ | Oshrluimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (Psrlil rd rs n :: k)
+ | Oshrxlimm n, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- ireg_of a1;
+ OK (if Int.eq n Int.zero then Pmv rd rs :: k else
+ Psrail X31 rs (Int.repr 63) ::
+ Psrlil X31 X31 (Int.sub Int64.iwordsize' n) ::
+ Paddl X31 rs X31 ::
+ Psrail rd X31 n :: k)
+
+ | Onegf, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfnegd rd rs :: k)
+ | Oabsf, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfabsd rd rs :: k)
+ | Oaddf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfaddd rd rs1 rs2 :: k)
+ | Osubf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfsubd rd rs1 rs2 :: k)
+ | Omulf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfmuld rd rs1 rs2 :: k)
+ | Odivf, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfdivd rd rs1 rs2 :: k)
+
+ | Onegfs, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfnegs rd rs :: k)
+ | Oabsfs, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfabss rd rs :: k)
+ | Oaddfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfadds rd rs1 rs2 :: k)
+ | Osubfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfsubs rd rs1 rs2 :: k)
+ | Omulfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfmuls rd rs1 rs2 :: k)
+ | Odivfs, a1 :: a2 :: nil =>
+ do rd <- freg_of res; do rs1 <- freg_of a1; do rs2 <- freg_of a2;
+ OK (Pfdivs rd rs1 rs2 :: k)
+
+ | Osingleoffloat, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfcvtsd rd rs :: k)
+ | Ofloatofsingle, a1 :: nil =>
+ do rd <- freg_of res; do rs <- freg_of a1;
+ OK (Pfcvtds rd rs :: k)
+
+ | Ointoffloat, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtwd rd rs :: k)
+ | Ointuoffloat, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtwud rd rs :: k)
+ | Ofloatofint, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfcvtdw rd rs :: k)
+ | Ofloatofintu, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfcvtdwu rd rs :: k)
+ | Ointofsingle, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtws rd rs :: k)
+ | Ointuofsingle, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtwus rd rs :: k)
+ | Osingleofint, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfcvtsw rd rs :: k)
+ | Osingleofintu, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfcvtswu rd rs :: k)
+
+ | Olongoffloat, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtld rd rs :: k)
+ | Olonguoffloat, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtlud rd rs :: k)
+ | Ofloatoflong, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfcvtdl rd rs :: k)
+ | Ofloatoflongu, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfcvtdlu rd rs :: k)
+ | Olongofsingle, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtls rd rs :: k)
+ | Olonguofsingle, a1 :: nil =>
+ do rd <- ireg_of res; do rs <- freg_of a1;
+ OK (Pfcvtlus rd rs :: k)
+ | Osingleoflong, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfcvtsl rd rs :: k)
+ | Osingleoflongu, a1 :: nil =>
+ do rd <- freg_of res; do rs <- ireg_of a1;
+ OK (Pfcvtslu rd rs :: k)
+
+ | Ocmp cmp, _ =>
+ do rd <- ireg_of res;
+ transl_cond_op cmp rd args k
+
+ | _, _ =>
+ Error(msg "Asmgen.transl_op")
+ end.
+
+(** Accessing data in the stack frame. *)
+
+Definition indexed_memory_access
+ (mk_instr: ireg -> offset -> instruction)
+ (base: ireg) (ofs: ptrofs) (k: code) :=
+ if Archi.ptr64 then
+ match make_immed64 (Ptrofs.to_int64 ofs) with
+ | Imm64_single imm =>
+ mk_instr base (Ofsimm (Ptrofs.of_int64 imm)) :: k
+ | Imm64_pair hi lo =>
+ Pluil X31 hi :: Paddl X31 base X31 :: mk_instr X31 (Ofsimm (Ptrofs.of_int64 lo)) :: k
+ | Imm64_large imm =>
+ Ploadli X31 imm :: Paddl X31 base X31 :: mk_instr X31 (Ofsimm Ptrofs.zero) :: k
+ end
+ else
+ match make_immed32 (Ptrofs.to_int ofs) with
+ | Imm32_single imm =>
+ mk_instr base (Ofsimm (Ptrofs.of_int imm)) :: k
+ | Imm32_pair hi lo =>
+ Pluiw X31 hi :: Paddw X31 base X31 :: mk_instr X31 (Ofsimm (Ptrofs.of_int lo)) :: k
+ end.
+
+Definition loadind (base: ireg) (ofs: ptrofs) (ty: typ) (dst: mreg) (k: code) :=
+ match ty, preg_of dst with
+ | Tint, IR rd => OK (indexed_memory_access (Plw rd) base ofs k)
+ | Tlong, IR rd => OK (indexed_memory_access (Pld rd) base ofs k)
+ | Tsingle, FR rd => OK (indexed_memory_access (Pfls rd) base ofs k)
+ | Tfloat, FR rd => OK (indexed_memory_access (Pfld rd) base ofs k)
+ | Tany32, IR rd => OK (indexed_memory_access (Plw_a rd) base ofs k)
+ | Tany64, IR rd => OK (indexed_memory_access (Pld_a rd) base ofs k)
+ | Tany64, FR rd => OK (indexed_memory_access (Pfld_a rd) base ofs k)
+ | _, _ => Error (msg "Asmgen.loadind")
+ end.
+
+Definition storeind (src: mreg) (base: ireg) (ofs: ptrofs) (ty: typ) (k: code) :=
+ match ty, preg_of src with
+ | Tint, IR rd => OK (indexed_memory_access (Psw rd) base ofs k)
+ | Tlong, IR rd => OK (indexed_memory_access (Psd rd) base ofs k)
+ | Tsingle, FR rd => OK (indexed_memory_access (Pfss rd) base ofs k)
+ | Tfloat, FR rd => OK (indexed_memory_access (Pfsd rd) base ofs k)
+ | Tany32, IR rd => OK (indexed_memory_access (Psw_a rd) base ofs k)
+ | Tany64, IR rd => OK (indexed_memory_access (Psd_a rd) base ofs k)
+ | Tany64, FR rd => OK (indexed_memory_access (Pfsd_a rd) base ofs k)
+ | _, _ => Error (msg "Asmgen.storeind")
+ end.
+
+Definition loadind_ptr (base: ireg) (ofs: ptrofs) (dst: ireg) (k: code) :=
+ indexed_memory_access (if Archi.ptr64 then Pld dst else Plw dst) base ofs k.
+
+Definition storeind_ptr (src: ireg) (base: ireg) (ofs: ptrofs) (k: code) :=
+ indexed_memory_access (if Archi.ptr64 then Psd src else Psw src) base ofs k.
+
+(** Translation of memory accesses: loads, and stores. *)
+
+Definition transl_memory_access
+ (mk_instr: ireg -> offset -> instruction)
+ (addr: addressing) (args: list mreg) (k: code) :=
+ match addr, args with
+ | Aindexed ofs, a1 :: nil =>
+ do rs <- ireg_of a1;
+ OK (indexed_memory_access mk_instr rs ofs k)
+ | Aglobal id ofs, nil =>
+ OK (Ploadsymbol_high X31 id ofs :: mk_instr X31 (Ofslow id ofs) :: k)
+ | Ainstack ofs, nil =>
+ OK (indexed_memory_access mk_instr SP ofs k)
+ | _, _ =>
+ Error(msg "Asmgen.transl_memory_access")
+ end.
+
+Definition transl_load (chunk: memory_chunk) (addr: addressing)
+ (args: list mreg) (dst: mreg) (k: code) :=
+ match chunk with
+ | Mint8signed =>
+ do r <- ireg_of dst;
+ transl_memory_access (Plb r) addr args k
+ | Mint8unsigned =>
+ do r <- ireg_of dst;
+ transl_memory_access (Plbu r) addr args k
+ | Mint16signed =>
+ do r <- ireg_of dst;
+ transl_memory_access (Plh r) addr args k
+ | Mint16unsigned =>
+ do r <- ireg_of dst;
+ transl_memory_access (Plhu r) addr args k
+ | Mint32 =>
+ do r <- ireg_of dst;
+ transl_memory_access (Plw r) addr args k
+ | Mint64 =>
+ do r <- ireg_of dst;
+ transl_memory_access (Pld r) addr args k
+ | Mfloat32 =>
+ do r <- freg_of dst;
+ transl_memory_access (Pfls r) addr args k
+ | Mfloat64 =>
+ do r <- freg_of dst;
+ transl_memory_access (Pfld r) addr args k
+ | _ =>
+ Error (msg "Asmgen.transl_load")
+ end.
+
+Definition transl_store (chunk: memory_chunk) (addr: addressing)
+ (args: list mreg) (src: mreg) (k: code) :=
+ match chunk with
+ | Mint8signed | Mint8unsigned =>
+ do r <- ireg_of src;
+ transl_memory_access (Psb r) addr args k
+ | Mint16signed | Mint16unsigned =>
+ do r <- ireg_of src;
+ transl_memory_access (Psh r) addr args k
+ | Mint32 =>
+ do r <- ireg_of src;
+ transl_memory_access (Psw r) addr args k
+ | Mint64 =>
+ do r <- ireg_of src;
+ transl_memory_access (Psd r) addr args k
+ | Mfloat32 =>
+ do r <- freg_of src;
+ transl_memory_access (Pfss r) addr args k
+ | Mfloat64 =>
+ do r <- freg_of src;
+ transl_memory_access (Pfsd r) addr args k
+ | _ =>
+ Error (msg "Asmgen.transl_store")
+ end.
+
+(** Function epilogue *)
+
+Definition make_epilogue (f: Mach.function) (k: code) :=
+ loadind_ptr SP f.(fn_retaddr_ofs) RA
+ (Pfreeframe f.(fn_stacksize) f.(fn_link_ofs) :: k).
+
+(** Translation of a Mach instruction. *)
+
+Definition transl_instr (f: Mach.function) (i: Mach.instruction)
+ (ep: bool) (k: code) :=
+ match i with
+ | Mgetstack ofs ty dst =>
+ loadind SP ofs ty dst k
+ | Msetstack src ofs ty =>
+ storeind src SP ofs ty k
+ | Mgetparam ofs ty dst =>
+ (* load via the frame pointer if it is valid *)
+ do c <- loadind X30 ofs ty dst k;
+ OK (if ep then c
+ else loadind_ptr SP f.(fn_link_ofs) X30 c)
+ | Mop op args res =>
+ transl_op op args res k
+ | Mload chunk addr args dst =>
+ transl_load chunk addr args dst k
+ | Mstore chunk addr args src =>
+ transl_store chunk addr args src k
+ | Mcall sig (inl r) =>
+ do r1 <- ireg_of r; OK (Pjal_r r1 sig :: k)
+ | Mcall sig (inr symb) =>
+ OK (Pjal_s symb sig :: k)
+ | Mtailcall sig (inl r) =>
+ do r1 <- ireg_of r;
+ OK (make_epilogue f (Pj_r r1 sig :: k))
+ | Mtailcall sig (inr symb) =>
+ OK (make_epilogue f (Pj_s symb sig :: k))
+ | Mbuiltin ef args res =>
+ OK (Pbuiltin ef (List.map (map_builtin_arg preg_of) args) (map_builtin_res preg_of res) :: k)
+ | Mlabel lbl =>
+ OK (Plabel lbl :: k)
+ | Mgoto lbl =>
+ OK (Pj_l lbl :: k)
+ | Mcond cond args lbl =>
+ transl_cbranch cond args lbl k
+ | Mjumptable arg tbl =>
+ do r <- ireg_of arg;
+ OK (Pbtbl r tbl :: k)
+ | Mreturn =>
+ OK (make_epilogue f (Pj_r RA f.(Mach.fn_sig) :: k))
+ end.
+
+(** Translation of a code sequence *)
+
+Definition it1_is_parent (before: bool) (i: Mach.instruction) : bool :=
+ match i with
+ | Msetstack src ofs ty => before
+ | Mgetparam ofs ty dst => negb (mreg_eq dst R30)
+ | Mop op args res => before && negb (mreg_eq res R30)
+ | _ => false
+ end.
+
+(** This is the naive definition that we no longer use because it
+ is not tail-recursive. It is kept as specification. *)
+
+Fixpoint transl_code (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
+ match il with
+ | nil => OK nil
+ | i1 :: il' =>
+ do k <- transl_code f il' (it1_is_parent it1p i1);
+ transl_instr f i1 it1p k
+ end.
+
+(** This is an equivalent definition in continuation-passing style
+ that runs in constant stack space. *)
+
+Fixpoint transl_code_rec (f: Mach.function) (il: list Mach.instruction)
+ (it1p: bool) (k: code -> res code) :=
+ match il with
+ | nil => k nil
+ | i1 :: il' =>
+ transl_code_rec f il' (it1_is_parent it1p i1)
+ (fun c1 => do c2 <- transl_instr f i1 it1p c1; k c2)
+ end.
+
+Definition transl_code' (f: Mach.function) (il: list Mach.instruction) (it1p: bool) :=
+ transl_code_rec f il it1p (fun c => OK c).
+
+(** Translation of a whole function. Note that we must check
+ that the generated code contains less than [2^32] instructions,
+ otherwise the offset part of the [PC] code pointer could wrap
+ around, leading to incorrect executions. *)
+
+Definition transl_function (f: Mach.function) :=
+ do c <- transl_code' f f.(Mach.fn_code) true;
+ OK (mkfunction f.(Mach.fn_sig)
+ (Pallocframe f.(fn_stacksize) f.(fn_link_ofs) ::
+ storeind_ptr RA SP f.(fn_retaddr_ofs) c)).
+
+Definition transf_function (f: Mach.function) : res Asm.function :=
+ do tf <- transl_function f;
+ if zlt Ptrofs.max_unsigned (list_length_z tf.(fn_code))
+ then Error (msg "code size exceeded")
+ else OK tf.
+
+Definition transf_fundef (f: Mach.fundef) : res Asm.fundef :=
+ transf_partial_fundef transf_function f.
+
+Definition transf_program (p: Mach.program) : res Asm.program :=
+ transform_partial_program transf_fundef p.
diff --git a/riscV/Asmgenproof.v b/riscV/Asmgenproof.v
new file mode 100644
index 00000000..da444a4b
--- /dev/null
+++ b/riscV/Asmgenproof.v
@@ -0,0 +1,1028 @@
+(* *********************************************************************)
+(* *)
+(* 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 Mach Conventions Asm.
+Require Import Asmgen Asmgenproof0 Asmgenproof1.
+
+Definition match_prog (p: Mach.program) (tp: Asm.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: Mach.program.
+Variable tprog: Asm.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 -> list_length_z tf.(fn_code) <= Ptrofs.max_unsigned.
+Proof.
+ intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0.
+ omega.
+Qed.
+
+Lemma exec_straight_exec:
+ forall fb f c ep tf tc c' rs m rs' m',
+ transl_code_at_pc ge (rs PC) fb f c ep tf tc ->
+ exec_straight tge tf tc rs m c' rs' m' ->
+ plus step tge (State rs m) E0 (State rs' m').
+Proof.
+ intros. inv H.
+ eapply exec_straight_steps_1; eauto.
+ eapply transf_function_no_overflow; eauto.
+ eapply functions_transl; eauto.
+Qed.
+
+Lemma exec_straight_at:
+ forall fb f c ep tf tc c' ep' tc' rs m rs' m',
+ transl_code_at_pc ge (rs PC) fb f c ep tf tc ->
+ transl_code f c' ep' = OK tc' ->
+ exec_straight tge tf tc rs m tc' rs' m' ->
+ transl_code_at_pc ge (rs' PC) fb f c' ep' tf tc'.
+Proof.
+ intros. inv H.
+ exploit exec_straight_steps_2; eauto.
+ eapply transf_function_no_overflow; eauto.
+ eapply functions_transl; eauto.
+ intros [ofs' [PC' CT']].
+ rewrite PC'. constructor; auto.
+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.
+
+Remark loadimm32_label:
+ forall r n k, tail_nolabel k (loadimm32 r n k).
+Proof.
+ intros; unfold loadimm32. destruct (make_immed32 n); TailNoLabel.
+ unfold load_hilo32. destruct (Int.eq lo Int.zero); TailNoLabel.
+Qed.
+Hint Resolve loadimm32_label: labels.
+
+Remark opimm32_label:
+ forall op opimm r1 r2 n k,
+ (forall r1 r2 r3, nolabel (op r1 r2 r3)) ->
+ (forall r1 r2 n, nolabel (opimm r1 r2 n)) ->
+ tail_nolabel k (opimm32 op opimm r1 r2 n k).
+Proof.
+ intros; unfold opimm32. destruct (make_immed32 n); TailNoLabel.
+ unfold load_hilo32. destruct (Int.eq lo Int.zero); TailNoLabel.
+Qed.
+Hint Resolve opimm32_label: labels.
+
+Remark loadimm64_label:
+ forall r n k, tail_nolabel k (loadimm64 r n k).
+Proof.
+ intros; unfold loadimm64. destruct (make_immed64 n); TailNoLabel.
+ unfold load_hilo64. destruct (Int64.eq lo Int64.zero); TailNoLabel.
+Qed.
+Hint Resolve loadimm64_label: labels.
+
+Remark opimm64_label:
+ forall op opimm r1 r2 n k,
+ (forall r1 r2 r3, nolabel (op r1 r2 r3)) ->
+ (forall r1 r2 n, nolabel (opimm r1 r2 n)) ->
+ tail_nolabel k (opimm64 op opimm r1 r2 n k).
+Proof.
+ intros; unfold opimm64. destruct (make_immed64 n); TailNoLabel.
+ unfold load_hilo64. destruct (Int64.eq lo Int64.zero); TailNoLabel.
+Qed.
+Hint Resolve opimm64_label: labels.
+
+Remark addptrofs_label:
+ forall r1 r2 n k, tail_nolabel k (addptrofs r1 r2 n k).
+Proof.
+ unfold addptrofs; intros. destruct (Ptrofs.eq_dec n Ptrofs.zero). TailNoLabel.
+ destruct Archi.ptr64. apply opimm64_label; TailNoLabel. apply opimm32_label; TailNoLabel.
+Qed.
+Hint Resolve addptrofs_label: labels.
+
+Remark transl_cond_float_nolabel:
+ forall c r1 r2 r3 insn normal,
+ transl_cond_float c r1 r2 r3 = (insn, normal) -> nolabel insn.
+Proof.
+ unfold transl_cond_float; intros. destruct c; inv H; exact I.
+Qed.
+
+Remark transl_cond_single_nolabel:
+ forall c r1 r2 r3 insn normal,
+ transl_cond_single c r1 r2 r3 = (insn, normal) -> nolabel insn.
+Proof.
+ unfold transl_cond_single; intros. destruct c; inv H; exact I.
+Qed.
+
+Remark transl_cbranch_label:
+ forall cond args lbl k c,
+ transl_cbranch cond args lbl k = OK c -> tail_nolabel k c.
+Proof.
+ intros. unfold transl_cbranch in H; destruct cond; TailNoLabel.
+- destruct c0; simpl; TailNoLabel.
+- destruct c0; simpl; TailNoLabel.
+- destruct (Int.eq n Int.zero).
+ destruct c0; simpl; TailNoLabel.
+ apply tail_nolabel_trans with (transl_cbranch_int32s c0 x X31 lbl :: k).
+ auto with labels. destruct c0; simpl; TailNoLabel.
+- destruct (Int.eq n Int.zero).
+ destruct c0; simpl; TailNoLabel.
+ apply tail_nolabel_trans with (transl_cbranch_int32u c0 x X31 lbl :: k).
+ auto with labels. destruct c0; simpl; TailNoLabel.
+- destruct c0; simpl; TailNoLabel.
+- destruct c0; simpl; TailNoLabel.
+- destruct (Int64.eq n Int64.zero).
+ destruct c0; simpl; TailNoLabel.
+ apply tail_nolabel_trans with (transl_cbranch_int64s c0 x X31 lbl :: k).
+ auto with labels. destruct c0; simpl; TailNoLabel.
+- destruct (Int64.eq n Int64.zero).
+ destruct c0; simpl; TailNoLabel.
+ apply tail_nolabel_trans with (transl_cbranch_int64u c0 x X31 lbl :: k).
+ auto with labels. destruct c0; simpl; TailNoLabel.
+- destruct (transl_cond_float c0 X31 x x0) as [insn normal] eqn:F; inv EQ2.
+ apply tail_nolabel_cons. eapply transl_cond_float_nolabel; eauto.
+ destruct normal; TailNoLabel.
+- destruct (transl_cond_float c0 X31 x x0) as [insn normal] eqn:F; inv EQ2.
+ apply tail_nolabel_cons. eapply transl_cond_float_nolabel; eauto.
+ destruct normal; TailNoLabel.
+- destruct (transl_cond_single c0 X31 x x0) as [insn normal] eqn:F; inv EQ2.
+ apply tail_nolabel_cons. eapply transl_cond_single_nolabel; eauto.
+ destruct normal; TailNoLabel.
+- destruct (transl_cond_single c0 X31 x x0) as [insn normal] eqn:F; inv EQ2.
+ apply tail_nolabel_cons. eapply transl_cond_single_nolabel; eauto.
+ destruct normal; TailNoLabel.
+Qed.
+
+Remark transl_cond_op_label:
+ forall cond args r k c,
+ transl_cond_op cond r args k = OK c -> tail_nolabel k c.
+Proof.
+ intros. unfold transl_cond_op in H; destruct cond; TailNoLabel.
+- destruct c0; simpl; TailNoLabel.
+- destruct c0; simpl; TailNoLabel.
+- unfold transl_condimm_int32s.
+ destruct (Int.eq n Int.zero).
++ destruct c0; simpl; TailNoLabel.
++ destruct c0; simpl.
+* eapply tail_nolabel_trans; [apply opimm32_label; intros; exact I | TailNoLabel].
+* eapply tail_nolabel_trans; [apply opimm32_label; intros; exact I | TailNoLabel].
+* apply opimm32_label; intros; exact I.
+* destruct (Int.eq n (Int.repr Int.max_signed)). apply loadimm32_label. apply opimm32_label; intros; exact I.
+* eapply tail_nolabel_trans. apply loadimm32_label. TailNoLabel.
+* eapply tail_nolabel_trans. apply loadimm32_label. TailNoLabel.
+- unfold transl_condimm_int32u.
+ destruct (Int.eq n Int.zero).
++ destruct c0; simpl; TailNoLabel.
++ destruct c0; simpl;
+ try (eapply tail_nolabel_trans; [apply loadimm32_label | TailNoLabel]).
+ apply opimm32_label; intros; exact I.
+- destruct c0; simpl; TailNoLabel.
+- destruct c0; simpl; TailNoLabel.
+- unfold transl_condimm_int64s.
+ destruct (Int64.eq n Int64.zero).
++ destruct c0; simpl; TailNoLabel.
++ destruct c0; simpl.
+* eapply tail_nolabel_trans; [apply opimm64_label; intros; exact I | TailNoLabel].
+* eapply tail_nolabel_trans; [apply opimm64_label; intros; exact I | TailNoLabel].
+* apply opimm64_label; intros; exact I.
+* destruct (Int64.eq n (Int64.repr Int64.max_signed)). apply loadimm32_label. apply opimm64_label; intros; exact I.
+* eapply tail_nolabel_trans. apply loadimm64_label. TailNoLabel.
+* eapply tail_nolabel_trans. apply loadimm64_label. TailNoLabel.
+- unfold transl_condimm_int64u.
+ destruct (Int64.eq n Int64.zero).
++ destruct c0; simpl; TailNoLabel.
++ destruct c0; simpl;
+ try (eapply tail_nolabel_trans; [apply loadimm64_label | TailNoLabel]).
+ apply opimm64_label; intros; exact I.
+- destruct (transl_cond_float c0 r x x0) as [insn normal] eqn:F; inv EQ2.
+ apply tail_nolabel_cons. eapply transl_cond_float_nolabel; eauto.
+ destruct normal; TailNoLabel.
+- destruct (transl_cond_float c0 r x x0) as [insn normal] eqn:F; inv EQ2.
+ apply tail_nolabel_cons. eapply transl_cond_float_nolabel; eauto.
+ destruct normal; TailNoLabel.
+- destruct (transl_cond_single c0 r x x0) as [insn normal] eqn:F; inv EQ2.
+ apply tail_nolabel_cons. eapply transl_cond_single_nolabel; eauto.
+ destruct normal; TailNoLabel.
+- destruct (transl_cond_single c0 r x x0) as [insn normal] eqn:F; inv EQ2.
+ apply tail_nolabel_cons. eapply transl_cond_single_nolabel; eauto.
+ destruct normal; TailNoLabel.
+Qed.
+
+Remark transl_op_label:
+ forall op args r k c,
+ transl_op op args r k = OK c -> tail_nolabel k c.
+Proof.
+Opaque Int.eq.
+ unfold transl_op; intros; destruct op; TailNoLabel.
+- destruct (preg_of r); try discriminate; destruct (preg_of m); inv H; TailNoLabel.
+- destruct (Float.eq_dec n Float.zero); TailNoLabel.
+- destruct (Float32.eq_dec n Float32.zero); TailNoLabel.
+- destruct (Archi.pic_code tt && negb (Ptrofs.eq ofs Ptrofs.zero)).
++ eapply tail_nolabel_trans; [|apply addptrofs_label]. TailNoLabel.
++ TailNoLabel.
+- apply opimm32_label; intros; exact I.
+- apply opimm32_label; intros; exact I.
+- apply opimm32_label; intros; exact I.
+- apply opimm32_label; intros; exact I.
+- destruct (Int.eq n Int.zero); TailNoLabel.
+- apply opimm64_label; intros; exact I.
+- apply opimm64_label; intros; exact I.
+- apply opimm64_label; intros; exact I.
+- apply opimm64_label; intros; exact I.
+- destruct (Int.eq n Int.zero); TailNoLabel.
+- eapply transl_cond_op_label; eauto.
+Qed.
+
+Remark indexed_memory_access_label:
+ forall (mk_instr: ireg -> offset -> instruction) base ofs k,
+ (forall r o, nolabel (mk_instr r o)) ->
+ tail_nolabel k (indexed_memory_access mk_instr base ofs k).
+Proof.
+ unfold indexed_memory_access; intros.
+ destruct Archi.ptr64.
+ destruct (make_immed64 (Ptrofs.to_int64 ofs)); TailNoLabel.
+ destruct (make_immed32 (Ptrofs.to_int ofs)); TailNoLabel.
+Qed.
+
+Remark loadind_label:
+ forall base ofs ty dst k c,
+ loadind base ofs ty dst k = OK c -> tail_nolabel k c.
+Proof.
+ unfold loadind; intros.
+ destruct ty, (preg_of dst); inv H; apply indexed_memory_access_label; intros; exact I.
+Qed.
+
+Remark storeind_label:
+ forall src base ofs ty k c,
+ storeind src base ofs ty k = OK c -> tail_nolabel k c.
+Proof.
+ unfold storeind; intros.
+ destruct ty, (preg_of src); inv H; apply indexed_memory_access_label; intros; exact I.
+Qed.
+
+Remark loadind_ptr_label:
+ forall base ofs dst k, tail_nolabel k (loadind_ptr base ofs dst k).
+Proof.
+ intros. apply indexed_memory_access_label. intros; destruct Archi.ptr64; exact I.
+Qed.
+
+Remark storeind_ptr_label:
+ forall src base ofs k, tail_nolabel k (storeind_ptr src base ofs k).
+Proof.
+ intros. apply indexed_memory_access_label. intros; destruct Archi.ptr64; exact I.
+Qed.
+
+Remark transl_memory_access_label:
+ forall (mk_instr: ireg -> offset -> instruction) addr args k c,
+ (forall r o, nolabel (mk_instr r o)) ->
+ transl_memory_access mk_instr addr args k = OK c ->
+ tail_nolabel k c.
+Proof.
+ unfold transl_memory_access; intros; destruct addr; TailNoLabel; apply indexed_memory_access_label; auto.
+Qed.
+
+Remark make_epilogue_label:
+ forall f k, tail_nolabel k (make_epilogue f k).
+Proof.
+ unfold make_epilogue; intros. eapply tail_nolabel_trans. apply loadind_ptr_label. TailNoLabel.
+Qed.
+
+Lemma transl_instr_label:
+ forall f i ep k c,
+ transl_instr f i ep k = OK c ->
+ match i with Mlabel lbl => c = Plabel lbl :: k | _ => tail_nolabel k c end.
+Proof.
+ unfold transl_instr; intros; destruct i; TailNoLabel.
+- eapply loadind_label; eauto.
+- eapply storeind_label; eauto.
+- destruct ep. eapply loadind_label; eauto.
+ eapply tail_nolabel_trans. apply loadind_ptr_label. eapply loadind_label; eauto.
+- eapply transl_op_label; eauto.
+- destruct m; monadInv H; eapply transl_memory_access_label; eauto; intros; exact I.
+- destruct m; monadInv H; eapply transl_memory_access_label; eauto; intros; exact I.
+- destruct s0; monadInv H; TailNoLabel.
+- destruct s0; monadInv H; (eapply tail_nolabel_trans; [eapply make_epilogue_label|TailNoLabel]).
+- eapply transl_cbranch_label; eauto.
+- eapply tail_nolabel_trans; [eapply make_epilogue_label|TailNoLabel].
+Qed.
+
+Lemma transl_instr_label':
+ forall lbl f i ep k c,
+ transl_instr f i ep k = OK c ->
+ find_label lbl c = if Mach.is_label lbl i then Some k else find_label lbl k.
+Proof.
+ intros. exploit transl_instr_label; eauto.
+ destruct i; try (intros [A B]; apply B).
+ intros. subst c. simpl. auto.
+Qed.
+
+Lemma transl_code_label:
+ forall lbl f c ep tc,
+ transl_code f c ep = OK tc ->
+ match Mach.find_label lbl c with
+ | None => find_label lbl tc = None
+ | Some c' => exists tc', find_label lbl tc = Some tc' /\ transl_code f c' false = OK tc'
+ end.
+Proof.
+ induction c; simpl; intros.
+ inv H. auto.
+ monadInv H. rewrite (transl_instr_label' lbl _ _ _ _ _ EQ0).
+ generalize (Mach.is_label_correct lbl a).
+ destruct (Mach.is_label lbl a); intros.
+ subst a. simpl in EQ. exists x; auto.
+ eapply IHc; eauto.
+Qed.
+
+Lemma transl_find_label:
+ forall lbl f tf,
+ transf_function f = OK tf ->
+ match Mach.find_label lbl f.(Mach.fn_code) with
+ | None => find_label lbl tf.(fn_code) = None
+ | Some c => exists tc, find_label lbl tf.(fn_code) = Some tc /\ transl_code f c false = OK tc
+ end.
+Proof.
+ intros. monadInv H. destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0.
+ monadInv EQ. rewrite transl_code'_transl_code in EQ0. unfold fn_code.
+ simpl. destruct (storeind_ptr_label X1 X2 (fn_retaddr_ofs f) x) as [A B]; rewrite B.
+ eapply transl_code_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 ->
+ Mach.find_label lbl f.(Mach.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. 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 f sg ros c, is_tail (Mcall sg ros :: c) f.(Mach.fn_code) ->
+ exists ra, return_address_offset f c ra.
+Proof.
+ intros. eapply Asmgenproof0.return_address_exists; eauto.
+- intros. exploit transl_instr_label; eauto.
+ destruct i; try (intros [A B]; apply A). intros. subst c0. repeat constructor.
+- intros. monadInv H0.
+ destruct (zlt Ptrofs.max_unsigned (list_length_z x.(fn_code))); inv EQ0. monadInv EQ.
+ rewrite transl_code'_transl_code in EQ0.
+ exists x; exists true; split; auto. unfold fn_code.
+ constructor. apply (storeind_ptr_label X1 X2 (fn_retaddr_ofs f0) x).
+- 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.
+*)
+
+Inductive match_states: Mach.state -> Asm.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#X30 = parent_sp s),
+ match_states (Mach.State s fb sp c ms m)
+ (Asm.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 (Mach.Callstate s fb ms m)
+ (Asm.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 (Mach.Returnstate s ms m)
+ (Asm.State rs m').
+
+Lemma exec_straight_steps:
+ forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2,
+ match_stack ge s ->
+ Mem.extends m2 m2' ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
+ (forall k c (TR: transl_instr f i ep k = OK c),
+ exists rs2,
+ exec_straight tge tf c rs1 m1' k rs2 m2'
+ /\ agree ms2 sp rs2
+ /\ (it1_is_parent ep i = true -> rs2#X30 = parent_sp s)) ->
+ exists st',
+ plus step tge (State rs1 m1') E0 st' /\
+ match_states (Mach.State s fb sp c ms2 m2) st'.
+Proof.
+ intros. inversion H2. subst. monadInv H7.
+ exploit H3; eauto. intros [rs2 [A [B C]]].
+ exists (State rs2 m2'); split.
+ eapply exec_straight_exec; eauto.
+ econstructor; eauto. eapply exec_straight_at; eauto.
+Qed.
+
+Lemma exec_straight_steps_goto:
+ forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2 lbl c',
+ match_stack ge s ->
+ Mem.extends m2 m2' ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ Mach.find_label lbl f.(Mach.fn_code) = Some c' ->
+ transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
+ it1_is_parent ep i = false ->
+ (forall k c (TR: transl_instr f i ep k = OK c),
+ exists jmp, exists k', exists rs2,
+ exec_straight tge tf c rs1 m1' (jmp :: k') rs2 m2'
+ /\ agree ms2 sp rs2
+ /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2') ->
+ exists st',
+ plus step tge (State rs1 m1') E0 st' /\
+ match_states (Mach.State s fb sp c' ms2 m2) st'.
+Proof.
+ intros. inversion H3. subst. monadInv H9.
+ exploit H5; eauto. intros [jmp [k' [rs2 [A [B C]]]]].
+ generalize (functions_transl _ _ _ H7 H8); intro FN.
+ generalize (transf_function_no_overflow _ _ H8); intro NOOV.
+ exploit exec_straight_steps_2; eauto.
+ intros [ofs' [PC2 CT2]].
+ exploit find_label_goto_label; eauto.
+ intros [tc' [rs3 [GOTO [AT' OTH]]]].
+ exists (State rs3 m2'); split.
+ eapply plus_right'.
+ eapply exec_straight_steps_1; eauto.
+ econstructor; eauto.
+ eapply find_instr_tail. eauto.
+ rewrite C. eexact GOTO.
+ traceEq.
+ econstructor; eauto.
+ apply agree_exten with rs2; auto with asmgen.
+ congruence.
+Qed.
+
+Lemma exec_straight_opt_steps_goto:
+ forall s fb f rs1 i c ep tf tc m1' m2 m2' sp ms2 lbl c',
+ match_stack ge s ->
+ Mem.extends m2 m2' ->
+ Genv.find_funct_ptr ge fb = Some (Internal f) ->
+ Mach.find_label lbl f.(Mach.fn_code) = Some c' ->
+ transl_code_at_pc ge (rs1 PC) fb f (i :: c) ep tf tc ->
+ it1_is_parent ep i = false ->
+ (forall k c (TR: transl_instr f i ep k = OK c),
+ exists jmp, exists k', exists rs2,
+ exec_straight_opt tge tf c rs1 m1' (jmp :: k') rs2 m2'
+ /\ agree ms2 sp rs2
+ /\ exec_instr tge tf jmp rs2 m2' = goto_label tf lbl rs2 m2') ->
+ exists st',
+ plus step tge (State rs1 m1') E0 st' /\
+ match_states (Mach.State s fb sp c' ms2 m2) st'.
+Proof.
+ intros. inversion H3. subst. monadInv H9.
+ exploit H5; eauto. intros [jmp [k' [rs2 [A [B C]]]]].
+ generalize (functions_transl _ _ _ H7 H8); intro FN.
+ generalize (transf_function_no_overflow _ _ H8); intro NOOV.
+ inv A.
+- exploit find_label_goto_label; eauto.
+ intros [tc' [rs3 [GOTO [AT' OTH]]]].
+ exists (State rs3 m2'); split.
+ apply plus_one. econstructor; eauto.
+ eapply find_instr_tail. eauto.
+ rewrite C. eexact GOTO.
+ econstructor; eauto.
+ apply agree_exten with rs2; auto with asmgen.
+ congruence.
+- exploit exec_straight_steps_2; eauto.
+ intros [ofs' [PC2 CT2]].
+ exploit find_label_goto_label; eauto.
+ intros [tc' [rs3 [GOTO [AT' OTH]]]].
+ exists (State rs3 m2'); split.
+ eapply plus_right'.
+ eapply exec_straight_steps_1; eauto.
+ econstructor; eauto.
+ eapply find_instr_tail. eauto.
+ rewrite C. eexact GOTO.
+ traceEq.
+ econstructor; eauto.
+ apply agree_exten with rs2; auto with asmgen.
+ congruence.
+Qed.
+
+(** 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. *)
+
+Definition measure (s: Mach.state) : nat :=
+ match s with
+ | Mach.State _ _ _ _ _ _ => 0%nat
+ | Mach.Callstate _ _ _ _ => 0%nat
+ | Mach.Returnstate _ _ _ => 1%nat
+ end.
+
+Remark preg_of_not_X30: forall r, negb (mreg_eq r R30) = true -> IR X30 <> preg_of r.
+Proof.
+ intros. change (IR X30) with (preg_of R30). red; intros.
+ exploit preg_of_injective; eauto. intros; subst r; discriminate.
+Qed.
+
+(** This is the simulation diagram. We prove it by case analysis on the Mach transition. *)
+
+Theorem step_simulation:
+ forall S1 t S2, Mach.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; inv MS.
+
+- (* Mlabel *)
+ left; eapply exec_straight_steps; eauto; intros.
+ monadInv TR. econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split. apply agree_nextinstr; auto. simpl; congruence.
+
+- (* Mgetstack *)
+ unfold load_stack in H.
+ exploit Mem.loadv_extends; eauto. intros [v' [A B]].
+ rewrite (sp_val _ _ _ AG) in A.
+ left; eapply exec_straight_steps; eauto. intros. simpl in TR.
+ exploit loadind_correct; eauto with asmgen. intros [rs' [P [Q R]]].
+ exists rs'; split. eauto.
+ split. eapply agree_set_mreg; eauto with asmgen. congruence.
+ simpl; congruence.
+
+- (* Msetstack *)
+ unfold store_stack in H.
+ assert (Val.lessdef (rs src) (rs0 (preg_of src))). eapply preg_val; eauto.
+ exploit Mem.storev_extends; eauto. intros [m2' [A B]].
+ left; eapply exec_straight_steps; eauto.
+ rewrite (sp_val _ _ _ AG) in A. intros. simpl in TR.
+ exploit storeind_correct; eauto with asmgen. intros [rs' [P Q]].
+ exists rs'; split. eauto.
+ split. eapply agree_undef_regs; eauto with asmgen.
+ simpl; intros. rewrite Q; auto with asmgen.
+
+- (* Mgetparam *)
+ assert (f0 = f) by congruence; subst f0.
+ unfold 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.
+ destruct ep.
+(* X30 contains parent *)
+ exploit loadind_correct. eexact EQ.
+ instantiate (2 := rs0). rewrite DXP; eauto. congruence.
+ intros [rs1 [P [Q R]]].
+ exists rs1; split. eauto.
+ split. eapply agree_set_mreg. eapply agree_set_mreg; eauto. congruence. auto with asmgen.
+ simpl; intros. rewrite R; auto with asmgen.
+ apply preg_of_not_X30; auto.
+(* GPR11 does not contain parent *)
+ rewrite chunk_of_Tptr in A.
+ exploit loadind_ptr_correct. eexact A. congruence. intros [rs1 [P [Q R]]].
+ exploit loadind_correct. eexact EQ. instantiate (2 := rs1). rewrite Q. eauto. congruence.
+ intros [rs2 [S [T U]]].
+ exists rs2; split. eapply exec_straight_trans; eauto.
+ split. eapply agree_set_mreg. eapply agree_set_mreg. eauto. eauto.
+ instantiate (1 := rs1#X30 <- (rs2#X30)). intros.
+ rewrite Pregmap.gso; auto with asmgen.
+ congruence.
+ intros. unfold Pregmap.set. destruct (PregEq.eq r' X30). congruence. auto with asmgen.
+ simpl; intros. rewrite U; auto with asmgen.
+ apply preg_of_not_X30; auto.
+
+- (* Mop *)
+ assert (eval_operation tge sp op (map rs args) m = Some v).
+ rewrite <- H. apply eval_operation_preserved. exact symbols_preserved.
+ exploit eval_operation_lessdef. eapply preg_vals; eauto. eauto. eexact H0.
+ intros [v' [A B]]. rewrite (sp_val _ _ _ AG) in A.
+ left; eapply exec_straight_steps; eauto; intros. simpl in TR.
+ exploit transl_op_correct; eauto. intros [rs2 [P [Q R]]].
+ exists rs2; split. eauto. split. auto.
+ apply agree_set_undef_mreg with rs0; auto.
+ apply Val.lessdef_trans with v'; auto.
+ simpl; intros. destruct (andb_prop _ _ H1); clear H1.
+ rewrite R; auto. apply preg_of_not_X30; auto.
+Local Transparent destroyed_by_op.
+ destruct op; simpl; auto; congruence.
+
+- (* Mload *)
+ assert (eval_addressing tge sp addr (map rs 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]].
+ left; eapply exec_straight_steps; eauto; intros. simpl in TR.
+ exploit transl_load_correct; eauto. intros [rs2 [P [Q R]]].
+ exists rs2; split. eauto.
+ split. eapply agree_set_undef_mreg; eauto. congruence.
+ intros; auto with asmgen.
+ simpl; congruence.
+
+- (* Mstore *)
+ assert (eval_addressing tge sp addr (map rs 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 (rs src) (rs0 (preg_of src))). eapply preg_val; eauto.
+ exploit Mem.storev_extends; eauto. intros [m2' [C D]].
+ left; eapply exec_straight_steps; eauto.
+ intros. simpl in TR. exploit transl_store_correct; eauto. intros [rs2 [P Q]].
+ exists rs2; split. eauto.
+ split. eapply agree_undef_regs; eauto with asmgen.
+ simpl; congruence.
+
+- (* Mcall *)
+ assert (f0 = f) by congruence. subst f0.
+ inv AT.
+ assert (NOOV: list_length_z tf.(fn_code) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto.
+ destruct ros as [rf|fid]; simpl in H; monadInv H5.
++ (* Indirect call *)
+ assert (rs rf = Vptr f' Ptrofs.zero).
+ destruct (rs rf); try discriminate.
+ revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence.
+ assert (rs0 x0 = Vptr f' Ptrofs.zero).
+ exploit ireg_val; eauto. rewrite H5; intros LD; inv LD; auto.
+ generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1.
+ assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x).
+ econstructor; eauto.
+ exploit return_address_offset_correct; eauto. intros; subst ra.
+ left; econstructor; split.
+ apply plus_one. eapply exec_step_internal. Simpl. rewrite <- H2; simpl; eauto.
+ eapply functions_transl; eauto. eapply find_instr_tail; eauto.
+ simpl. eauto.
+ econstructor; eauto.
+ econstructor; eauto.
+ eapply agree_sp_def; eauto.
+ simpl. eapply agree_exten; eauto. intros. Simpl.
+ Simpl. rewrite <- H2. auto.
++ (* Direct call *)
+ generalize (code_tail_next_int _ _ _ _ NOOV H6). intro CT1.
+ assert (TCA: transl_code_at_pc ge (Vptr fb (Ptrofs.add ofs Ptrofs.one)) fb f c false tf x).
+ econstructor; eauto.
+ exploit return_address_offset_correct; eauto. intros; subst ra.
+ left; econstructor; split.
+ apply plus_one. eapply exec_step_internal. eauto.
+ eapply functions_transl; eauto. eapply find_instr_tail; eauto.
+ simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. eauto.
+ econstructor; eauto.
+ econstructor; eauto.
+ eapply agree_sp_def; eauto.
+ simpl. eapply agree_exten; eauto. intros. Simpl.
+ Simpl. rewrite <- H2. auto.
+
+- (* Mtailcall *)
+ assert (f0 = f) by congruence. subst f0.
+ inversion AT; subst.
+ assert (NOOV: list_length_z tf.(fn_code) <= Ptrofs.max_unsigned).
+ eapply transf_function_no_overflow; eauto. exploit Mem.loadv_extends. eauto. eexact H1. auto. simpl. intros [parent' [A B]].
+ destruct ros as [rf|fid]; simpl in H; monadInv H7.
++ (* Indirect call *)
+ assert (rs rf = Vptr f' Ptrofs.zero).
+ destruct (rs rf); try discriminate.
+ revert H; predSpec Ptrofs.eq Ptrofs.eq_spec i Ptrofs.zero; intros; congruence.
+ assert (rs0 x0 = Vptr f' Ptrofs.zero).
+ exploit ireg_val; eauto. rewrite H7; intros LD; inv LD; auto.
+ exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
+ exploit exec_straight_steps_2; eauto using functions_transl.
+ intros (ofs' & P & Q).
+ left; econstructor; split.
+ (* execution *)
+ eapply plus_right'. eapply exec_straight_exec; eauto.
+ econstructor. eexact P. eapply functions_transl; eauto. eapply find_instr_tail. eexact Q.
+ simpl. reflexivity.
+ traceEq.
+ (* match states *)
+ econstructor; eauto.
+ apply agree_set_other; auto with asmgen.
+ Simpl. rewrite Z by (rewrite <- (ireg_of_eq _ _ EQ1); eauto with asmgen). assumption.
++ (* Direct call *)
+ exploit make_epilogue_correct; eauto. intros (rs1 & m1 & U & V & W & X & Y & Z).
+ exploit exec_straight_steps_2; eauto using functions_transl.
+ intros (ofs' & P & Q).
+ left; econstructor; split.
+ (* execution *)
+ eapply plus_right'. eapply exec_straight_exec; eauto.
+ econstructor. eexact P. eapply functions_transl; eauto. eapply find_instr_tail. eexact Q.
+ simpl. reflexivity.
+ traceEq.
+ (* match states *)
+ econstructor; eauto.
+ apply agree_set_other; auto with asmgen.
+ Simpl. unfold Genv.symbol_address. rewrite symbols_preserved. rewrite H. auto.
+
+- (* Mbuiltin *)
+ inv AT. monadInv H4.
+ exploit functions_transl; eauto. intro FN.
+ generalize (transf_function_no_overflow _ _ H3); intro NOOV.
+ exploit builtin_args_match; eauto. intros [vargs' [P Q]].
+ exploit external_call_mem_extends; eauto.
+ intros [vres' [m2' [A [B [C D]]]]].
+ left. econstructor; split. apply plus_one.
+ eapply exec_step_builtin. eauto. eauto.
+ eapply find_instr_tail; 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 := x).
+ unfold nextinstr. rewrite Pregmap.gss.
+ rewrite set_res_other. rewrite undef_regs_other_2. rewrite Pregmap.gso by congruence.
+ rewrite <- H1. simpl. econstructor; eauto.
+ eapply code_tail_next_int; eauto.
+ rewrite preg_notin_charact. intros. auto with asmgen.
+ auto with asmgen.
+ apply agree_nextinstr. eapply agree_set_res; auto.
+ eapply agree_undef_regs; eauto. intros. rewrite undef_regs_other_2; auto. apply Pregmap.gso; auto with asmgen.
+ congruence.
+
+- (* Mgoto *)
+ assert (f0 = f) by congruence. subst f0.
+ inv AT. monadInv H4.
+ exploit find_label_goto_label; eauto. intros [tc' [rs' [GOTO [AT2 INV]]]].
+ left; exists (State rs' m'); split.
+ apply plus_one. econstructor; eauto.
+ eapply functions_transl; eauto.
+ eapply find_instr_tail; eauto.
+ simpl; eauto.
+ econstructor; eauto.
+ eapply agree_exten; eauto with asmgen.
+ congruence.
+
+- (* Mcond true *)
+ assert (f0 = f) by congruence. subst f0.
+ exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC.
+ left; eapply exec_straight_opt_steps_goto; eauto.
+ intros. simpl in TR.
+ exploit transl_cbranch_correct_true; eauto. intros (rs' & jmp & A & B & C).
+ exists jmp; exists k; exists rs'.
+ split. eexact A.
+ split. apply agree_exten with rs0; auto with asmgen.
+ exact B.
+
+- (* Mcond false *)
+ exploit eval_condition_lessdef. eapply preg_vals; eauto. eauto. eauto. intros EC.
+ left; eapply exec_straight_steps; eauto. intros. simpl in TR.
+ exploit transl_cbranch_correct_false; eauto. intros (rs' & A & B).
+ exists rs'.
+ split. eexact A.
+ split. apply agree_exten with rs0; auto with asmgen.
+ simpl. congruence.
+
+- (* Mjumptable *)
+ assert (f0 = f) by congruence. subst f0.
+ inv AT. monadInv H6.
+ exploit functions_transl; eauto. intro FN.
+ generalize (transf_function_no_overflow _ _ H5); intro NOOV.
+ exploit find_label_goto_label. eauto. eauto.
+ instantiate (2 := rs0#X5 <- Vundef #X31 <- Vundef).
+ Simpl. eauto.
+ eauto.
+ intros [tc' [rs' [A [B C]]]].
+ exploit ireg_val; eauto. rewrite H. intros LD; inv LD.
+ left; econstructor; split.
+ apply plus_one. econstructor; eauto.
+ eapply find_instr_tail; eauto.
+ simpl. rewrite <- H9. unfold Mach.label in H0; unfold label; rewrite H0. eexact A.
+ econstructor; eauto.
+ eapply agree_undef_regs; eauto.
+ simpl. intros. rewrite C; auto with asmgen. Simpl.
+ congruence.
+
+- (* Mreturn *)
+ assert (f0 = f) by congruence. subst f0.
+ inversion AT; subst. simpl in H6; monadInv H6.
+ assert (NOOV: list_length_z tf.(fn_code) <= 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_steps_2; eauto using functions_transl.
+ intros (ofs' & P & Q).
+ left; econstructor; split.
+ (* execution *)
+ eapply plus_right'. eapply exec_straight_exec; eauto.
+ econstructor. eexact P. eapply functions_transl; eauto. eapply find_instr_tail. eexact Q.
+ simpl. reflexivity.
+ traceEq.
+ (* match states *)
+ econstructor; eauto.
+ apply agree_set_other; auto with asmgen.
+
+- (* internal function *)
+ exploit functions_translated; eauto. intros [tf [A B]]. monadInv B.
+ generalize EQ; intros EQ'. monadInv EQ'.
+ destruct (zlt Ptrofs.max_unsigned (list_length_z x0.(fn_code))); inversion EQ1. clear EQ1. subst x0.
+ unfold store_stack in *.
+ exploit Mem.alloc_extends. eauto. eauto. apply Zle_refl. apply Zle_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 := Pallocframe (fn_stacksize f) (fn_link_ofs f) ::
+ storeind_ptr RA SP (fn_retaddr_ofs f) x0) in *.
+ set (tf := {| fn_sig := Mach.fn_sig f; fn_code := tfbody |}) in *.
+ set (rs2 := nextinstr (rs0#X30 <- (parent_sp s) #SP <- sp #X31 <- Vundef)).
+ exploit (storeind_ptr_correct tge tf SP (fn_retaddr_ofs f) RA x0 rs2 m2').
+ rewrite chunk_of_Tptr in P. change (rs2 X1) with (rs0 X1). rewrite ATLR.
+ change (rs2 X2) with sp. eexact P.
+ congruence. congruence.
+ intros (rs3 & U & V).
+ assert (EXEC_PROLOGUE:
+ exec_straight tge tf
+ tf.(fn_code) rs0 m'
+ x0 rs3 m3').
+ { change (fn_code tf) with tfbody; unfold tfbody.
+ apply exec_straight_step with rs2 m2'.
+ unfold exec_instr. rewrite C. fold sp.
+ rewrite <- (sp_val _ _ _ AG). rewrite chunk_of_Tptr in F. rewrite F. reflexivity.
+ reflexivity.
+ eexact U. }
+ exploit exec_straight_steps_2; eauto using functions_transl. omega. constructor.
+ intros (ofs' & X & Y).
+ left; exists (State rs3 m3'); split.
+ eapply exec_straight_steps_1; eauto. omega. constructor.
+ econstructor; eauto.
+ rewrite X; econstructor; eauto.
+ apply agree_exten with rs2; eauto with asmgen.
+ unfold rs2.
+ apply agree_nextinstr. 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. rewrite V by auto with asmgen. reflexivity.
+
+- (* external function *)
+ 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.
+
+- (* return *)
+ inv STACKS. simpl in *.
+ right. split. omega. split. auto.
+ rewrite <- ATPC in H5.
+ econstructor; eauto. congruence.
+Qed.
+
+Lemma transf_initial_states:
+ forall st1, Mach.initial_state prog st1 ->
+ exists st2, Asm.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 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 -> Mach.final_state st1 r -> Asm.final_state st2 r.
+Proof.
+ intros. inv H0. inv H. constructor. assumption.
+ compute in H1. inv H1.
+ generalize (preg_val _ _ _ R10 AG). rewrite H2. intros LD; inv LD. auto.
+Qed.
+
+Theorem transf_program_correct:
+ forward_simulation (Mach.semantics return_address_offset prog) (Asm.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/riscV/Asmgenproof1.v b/riscV/Asmgenproof1.v
new file mode 100644
index 00000000..7f070c12
--- /dev/null
+++ b/riscV/Asmgenproof1.v
@@ -0,0 +1,1411 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+Require Import Coqlib Errors Maps.
+Require Import AST Integers Floats Values Memory Globalenvs.
+Require Import Op Locations Mach Conventions.
+Require Import Asm Asmgen Asmgenproof0.
+
+(** Decomposition of integer constants. *)
+
+Lemma make_immed32_sound:
+ forall n,
+ match make_immed32 n with
+ | Imm32_single imm => n = imm
+ | Imm32_pair hi lo => n = Int.add (Int.shl hi (Int.repr 12)) lo
+ end.
+Proof.
+ intros; unfold make_immed32. set (lo := Int.sign_ext 12 n).
+ predSpec Int.eq Int.eq_spec n lo.
+- auto.
+- set (m := Int.sub n lo).
+ assert (A: Int.eqmod (two_p 12) (Int.unsigned lo) (Int.unsigned n)) by (apply Int.eqmod_sign_ext'; compute; auto).
+ assert (B: Int.eqmod (two_p 12) (Int.unsigned n - Int.unsigned lo) 0).
+ { replace 0 with (Int.unsigned n - Int.unsigned n) by omega.
+ auto using Int.eqmod_sub, Int.eqmod_refl. }
+ assert (C: Int.eqmod (two_p 12) (Int.unsigned m) 0).
+ { apply Int.eqmod_trans with (Int.unsigned n - Int.unsigned lo); auto.
+ apply Int.eqmod_divides with Int.modulus. apply Int.eqm_sym; apply Int.eqm_unsigned_repr.
+ exists (two_p (32-12)); auto. }
+ assert (D: Int.modu m (Int.repr 4096) = Int.zero).
+ { apply Int.eqmod_mod_eq in C. unfold Int.modu.
+ change (Int.unsigned (Int.repr 4096)) with (two_p 12). rewrite C.
+ reflexivity.
+ apply two_p_gt_ZERO; omega. }
+ rewrite <- (Int.divu_pow2 m (Int.repr 4096) (Int.repr 12)) by auto.
+ rewrite Int.shl_mul_two_p.
+ change (two_p (Int.unsigned (Int.repr 12))) with 4096.
+ replace (Int.mul (Int.divu m (Int.repr 4096)) (Int.repr 4096)) with m.
+ unfold m. rewrite Int.sub_add_opp. rewrite Int.add_assoc. rewrite <- (Int.add_commut lo).
+ rewrite Int.add_neg_zero. rewrite Int.add_zero. auto.
+ rewrite (Int.modu_divu_Euclid m (Int.repr 4096)) at 1 by (vm_compute; congruence).
+ rewrite D. apply Int.add_zero.
+Qed.
+
+Lemma make_immed64_sound:
+ forall n,
+ match make_immed64 n with
+ | Imm64_single imm => n = imm
+ | Imm64_pair hi lo => n = Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo
+ | Imm64_large imm => n = imm
+ end.
+Proof.
+ intros; unfold make_immed64. set (lo := Int64.sign_ext 12 n).
+ predSpec Int64.eq Int64.eq_spec n lo.
+- auto.
+- set (m := Int64.sub n lo).
+ set (p := Int64.zero_ext 20 (Int64.shru m (Int64.repr 12))).
+ predSpec Int64.eq Int64.eq_spec n (Int64.add (Int64.sign_ext 32 (Int64.shl p (Int64.repr 12))) lo).
+ auto.
+ auto.
+Qed.
+
+(** Properties of registers *)
+
+Lemma ireg_of_not_X31:
+ forall m r, ireg_of m = OK r -> IR r <> IR X31.
+Proof.
+ intros. erewrite <- ireg_of_eq; eauto with asmgen.
+Qed.
+
+Lemma ireg_of_not_X31':
+ forall m r, ireg_of m = OK r -> r <> X31.
+Proof.
+ intros. apply ireg_of_not_X31 in H. congruence.
+Qed.
+
+Hint Resolve ireg_of_not_X31 ireg_of_not_X31': asmgen.
+
+(** Useful simplification tactic *)
+
+Ltac Simplif :=
+ ((rewrite nextinstr_inv by eauto with asmgen)
+ || (rewrite nextinstr_inv1 by eauto with asmgen)
+ || (rewrite Pregmap.gss)
+ || (rewrite nextinstr_pc)
+ || (rewrite Pregmap.gso by eauto with asmgen)); auto with asmgen.
+
+Ltac Simpl := repeat Simplif.
+
+(** * Correctness of RISC-V constructor functions *)
+
+Section CONSTRUCTORS.
+
+Variable ge: genv.
+Variable fn: function.
+
+(** 32-bit integer constants and arithmetic *)
+
+Lemma load_hilo32_correct:
+ forall rd hi lo k rs m,
+ exists rs',
+ exec_straight ge fn (load_hilo32 rd hi lo k) rs m k rs' m
+ /\ rs'#rd = Vint (Int.add (Int.shl hi (Int.repr 12)) lo)
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold load_hilo32; intros.
+ predSpec Int.eq Int.eq_spec lo Int.zero.
+- subst lo. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. rewrite Int.add_zero. Simpl.
+ intros; Simpl.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split. Simpl.
+ intros; Simpl.
+Qed.
+
+Lemma loadimm32_correct:
+ forall rd n k rs m,
+ exists rs',
+ exec_straight ge fn (loadimm32 rd n k) rs m k rs' m
+ /\ rs'#rd = Vint n
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold loadimm32; intros. generalize (make_immed32_sound n); intros E.
+ destruct (make_immed32 n).
+- subst imm. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. rewrite Int.add_zero_l; Simpl.
+ intros; Simpl.
+- rewrite E. apply load_hilo32_correct.
+Qed.
+
+Lemma opimm32_correct:
+ forall (op: ireg -> ireg0 -> ireg0 -> instruction)
+ (opi: ireg -> ireg0 -> int -> instruction)
+ (sem: val -> val -> val) m,
+ (forall d s1 s2 rs,
+ exec_instr ge fn (op d s1 s2) rs m = Next (nextinstr (rs#d <- (sem rs##s1 rs##s2))) m) ->
+ (forall d s n rs,
+ exec_instr ge fn (opi d s n) rs m = Next (nextinstr (rs#d <- (sem rs##s (Vint n)))) m) ->
+ forall rd r1 n k rs,
+ r1 <> X31 ->
+ exists rs',
+ exec_straight ge fn (opimm32 op opi rd r1 n k) rs m k rs' m
+ /\ rs'#rd = sem rs##r1 (Vint n)
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. unfold opimm32. generalize (make_immed32_sound n); intros E.
+ destruct (make_immed32 n).
+- subst imm. econstructor; split.
+ apply exec_straight_one. rewrite H0. simpl; eauto. auto.
+ split. Simpl. intros; Simpl.
+- destruct (load_hilo32_correct X31 hi lo (op rd r1 X31 :: k) rs m)
+ as (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one.
+ rewrite H; eauto. auto.
+ split. Simpl. simpl. rewrite B, C, E. auto. congruence. congruence.
+ intros; Simpl.
+Qed.
+
+(** 64-bit integer constants and arithmetic *)
+
+Lemma load_hilo64_correct:
+ forall rd hi lo k rs m,
+ exists rs',
+ exec_straight ge fn (load_hilo64 rd hi lo k) rs m k rs' m
+ /\ rs'#rd = Vlong (Int64.add (Int64.sign_ext 32 (Int64.shl hi (Int64.repr 12))) lo)
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ unfold load_hilo64; intros.
+ predSpec Int64.eq Int64.eq_spec lo Int64.zero.
+- subst lo. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. rewrite Int64.add_zero. Simpl.
+ intros; Simpl.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split. Simpl.
+ intros; Simpl.
+Qed.
+
+Lemma loadimm64_correct:
+ forall rd n k rs m,
+ exists rs',
+ exec_straight ge fn (loadimm64 rd n k) rs m k rs' m
+ /\ rs'#rd = Vlong n
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ unfold loadimm64; intros. generalize (make_immed64_sound n); intros E.
+ destruct (make_immed64 n).
+- subst imm. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. rewrite Int64.add_zero_l; Simpl.
+ intros; Simpl.
+- exploit load_hilo64_correct; eauto. intros (rs' & A & B & C).
+ rewrite E. exists rs'; eauto.
+- subst imm. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl.
+ intros; Simpl.
+Qed.
+
+Lemma opimm64_correct:
+ forall (op: ireg -> ireg0 -> ireg0 -> instruction)
+ (opi: ireg -> ireg0 -> int64 -> instruction)
+ (sem: val -> val -> val) m,
+ (forall d s1 s2 rs,
+ exec_instr ge fn (op d s1 s2) rs m = Next (nextinstr (rs#d <- (sem rs###s1 rs###s2))) m) ->
+ (forall d s n rs,
+ exec_instr ge fn (opi d s n) rs m = Next (nextinstr (rs#d <- (sem rs###s (Vlong n)))) m) ->
+ forall rd r1 n k rs,
+ r1 <> X31 ->
+ exists rs',
+ exec_straight ge fn (opimm64 op opi rd r1 n k) rs m k rs' m
+ /\ rs'#rd = sem rs##r1 (Vlong n)
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. unfold opimm64. generalize (make_immed64_sound n); intros E.
+ destruct (make_immed64 n).
+- subst imm. econstructor; split.
+ apply exec_straight_one. rewrite H0. simpl; eauto. auto.
+ split. Simpl. intros; Simpl.
+- destruct (load_hilo64_correct X31 hi lo (op rd r1 X31 :: k) rs m)
+ as (rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_trans. eexact A. apply exec_straight_one.
+ rewrite H; eauto. auto.
+ split. Simpl. simpl. rewrite B, C, E. auto. congruence. congruence.
+ intros; Simpl.
+- subst imm. econstructor; split.
+ eapply exec_straight_two. simpl; eauto. rewrite H. simpl; eauto. auto. auto.
+ split. Simpl. intros; Simpl.
+Qed.
+
+(** Add offset to pointer *)
+
+Lemma addptrofs_correct:
+ forall rd r1 n k rs m,
+ r1 <> X31 ->
+ exists rs',
+ exec_straight ge fn (addptrofs rd r1 n k) rs m k rs' m
+ /\ Val.lessdef (Val.offset_ptr rs#r1 n) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ unfold addptrofs; intros.
+ destruct (Ptrofs.eq_dec n Ptrofs.zero).
+- subst n. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split. Simpl. destruct (rs r1); simpl; auto. rewrite Ptrofs.add_zero; auto.
+ intros; Simpl.
+- destruct Archi.ptr64 eqn:SF.
++ unfold addimm64.
+ exploit (opimm64_correct Paddl Paddil Val.addl); eauto. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split; auto.
+ rewrite B. simpl. destruct (rs r1); simpl; auto. rewrite SF.
+ rewrite Ptrofs.of_int64_to_int64 by auto. auto.
++ unfold addimm32.
+ exploit (opimm32_correct Paddw Paddiw Val.add); eauto. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split; auto.
+ rewrite B. simpl. destruct (rs r1); simpl; auto. rewrite SF.
+ rewrite Ptrofs.of_int_to_int by auto. auto.
+Qed.
+
+Lemma addptrofs_correct_2:
+ forall rd r1 n k (rs: regset) m b ofs,
+ r1 <> X31 -> rs#r1 = Vptr b ofs ->
+ exists rs',
+ exec_straight ge fn (addptrofs rd r1 n k) rs m k rs' m
+ /\ rs'#rd = Vptr b (Ptrofs.add ofs n)
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. exploit (addptrofs_correct rd r1 n); eauto. intros (rs' & A & B & C).
+ exists rs'; intuition eauto.
+ rewrite H0 in B. inv B. auto.
+Qed.
+
+(** Translation of conditional branches *)
+
+Lemma transl_cbranch_int32s_correct:
+ forall cmp r1 r2 lbl (rs: regset) m b,
+ Val.cmp_bool cmp rs##r1 rs##r2 = Some b ->
+ exec_instr ge fn (transl_cbranch_int32s cmp r1 r2 lbl) rs m =
+ eval_branch fn lbl rs m (Some b).
+Proof.
+ intros. destruct cmp; simpl; rewrite ? H.
+- destruct rs##r1; simpl in H; try discriminate. destruct rs##r2; inv H.
+ simpl; auto.
+- destruct rs##r1; simpl in H; try discriminate. destruct rs##r2; inv H.
+ simpl; auto.
+- auto.
+- rewrite <- Val.swap_cmp_bool. simpl. rewrite H; auto.
+- rewrite <- Val.swap_cmp_bool. simpl. rewrite H; auto.
+- auto.
+Qed.
+
+Lemma transl_cbranch_int32u_correct:
+ forall cmp r1 r2 lbl (rs: regset) m b,
+ Val.cmpu_bool (Mem.valid_pointer m) cmp rs##r1 rs##r2 = Some b ->
+ exec_instr ge fn (transl_cbranch_int32u cmp r1 r2 lbl) rs m =
+ eval_branch fn lbl rs m (Some b).
+Proof.
+ intros. destruct cmp; simpl; rewrite ? H; auto.
+- rewrite <- Val.swap_cmpu_bool. simpl. rewrite H; auto.
+- rewrite <- Val.swap_cmpu_bool. simpl. rewrite H; auto.
+Qed.
+
+Lemma transl_cbranch_int64s_correct:
+ forall cmp r1 r2 lbl (rs: regset) m b,
+ Val.cmpl_bool cmp rs###r1 rs###r2 = Some b ->
+ exec_instr ge fn (transl_cbranch_int64s cmp r1 r2 lbl) rs m =
+ eval_branch fn lbl rs m (Some b).
+Proof.
+ intros. destruct cmp; simpl; rewrite ? H.
+- destruct rs###r1; simpl in H; try discriminate. destruct rs###r2; inv H.
+ simpl; auto.
+- destruct rs###r1; simpl in H; try discriminate. destruct rs###r2; inv H.
+ simpl; auto.
+- auto.
+- rewrite <- Val.swap_cmpl_bool. simpl. rewrite H; auto.
+- rewrite <- Val.swap_cmpl_bool. simpl. rewrite H; auto.
+- auto.
+Qed.
+
+Lemma transl_cbranch_int64u_correct:
+ forall cmp r1 r2 lbl (rs: regset) m b,
+ Val.cmplu_bool (Mem.valid_pointer m) cmp rs###r1 rs###r2 = Some b ->
+ exec_instr ge fn (transl_cbranch_int64u cmp r1 r2 lbl) rs m =
+ eval_branch fn lbl rs m (Some b).
+Proof.
+ intros. destruct cmp; simpl; rewrite ? H; auto.
+- rewrite <- Val.swap_cmplu_bool. simpl. rewrite H; auto.
+- rewrite <- Val.swap_cmplu_bool. simpl. rewrite H; auto.
+Qed.
+
+Lemma transl_cond_float_correct:
+ forall (rs: regset) m cmp rd r1 r2 insn normal v,
+ transl_cond_float cmp rd r1 r2 = (insn, normal) ->
+ v = (if normal then Val.cmpf cmp rs#r1 rs#r2 else Val.notbool (Val.cmpf cmp rs#r1 rs#r2)) ->
+ exec_instr ge fn insn rs m = Next (nextinstr (rs#rd <- v)) m.
+Proof.
+ intros. destruct cmp; simpl in H; inv H; auto.
+- rewrite Val.negate_cmpf_eq. auto.
+- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpf, Val.cmpf_bool.
+ rewrite <- Float.cmp_swap. auto.
+- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpf, Val.cmpf_bool.
+ rewrite <- Float.cmp_swap. auto.
+Qed.
+
+Lemma transl_cond_single_correct:
+ forall (rs: regset) m cmp rd r1 r2 insn normal v,
+ transl_cond_single cmp rd r1 r2 = (insn, normal) ->
+ v = (if normal then Val.cmpfs cmp rs#r1 rs#r2 else Val.notbool (Val.cmpfs cmp rs#r1 rs#r2)) ->
+ exec_instr ge fn insn rs m = Next (nextinstr (rs#rd <- v)) m.
+Proof.
+ intros. destruct cmp; simpl in H; inv H; auto.
+- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpfs, Val.cmpfs_bool.
+ rewrite Float32.cmp_ne_eq. destruct (Float32.cmp Ceq f0 f); auto.
+- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpfs, Val.cmpfs_bool.
+ rewrite <- Float32.cmp_swap. auto.
+- simpl. f_equal. f_equal. f_equal. destruct (rs r2), (rs r1); auto. unfold Val.cmpfs, Val.cmpfs_bool.
+ rewrite <- Float32.cmp_swap. auto.
+Qed.
+
+Remark branch_on_X31:
+ forall normal lbl (rs: regset) m b,
+ rs#X31 = Val.of_bool (eqb normal b) ->
+ exec_instr ge fn (if normal then Pbnew X31 X0 lbl else Pbeqw X31 X0 lbl) rs m =
+ eval_branch fn lbl rs m (Some b).
+Proof.
+ intros. destruct normal; simpl; rewrite H; simpl; destruct b; reflexivity.
+Qed.
+
+Ltac ArgsInv :=
+ repeat (match goal with
+ | [ H: Error _ = OK _ |- _ ] => discriminate
+ | [ H: match ?args with nil => _ | _ :: _ => _ end = OK _ |- _ ] => destruct args
+ | [ H: bind _ _ = OK _ |- _ ] => monadInv H
+ | [ H: match _ with left _ => _ | right _ => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv
+ | [ H: match _ with true => _ | false => assertion_failed end = OK _ |- _ ] => monadInv H; ArgsInv
+ end);
+ subst;
+ repeat (match goal with
+ | [ H: ireg_of _ = OK _ |- _ ] => simpl in *; rewrite (ireg_of_eq _ _ H) in *
+ | [ H: freg_of _ = OK _ |- _ ] => simpl in *; rewrite (freg_of_eq _ _ H) in *
+ end).
+
+Inductive exec_straight_opt: code -> regset -> mem -> code -> regset -> mem -> Prop :=
+ | exec_straight_opt_refl: forall c rs m,
+ exec_straight_opt c rs m c rs m
+ | exec_straight_opt_intro: forall c1 rs1 m1 c2 rs2 m2,
+ exec_straight ge fn c1 rs1 m1 c2 rs2 m2 ->
+ exec_straight_opt c1 rs1 m1 c2 rs2 m2.
+
+Remark exec_straight_opt_right:
+ forall c3 rs3 m3 c1 rs1 m1 c2 rs2 m2,
+ exec_straight_opt c1 rs1 m1 c2 rs2 m2 ->
+ exec_straight ge fn c2 rs2 m2 c3 rs3 m3 ->
+ exec_straight ge fn c1 rs1 m1 c3 rs3 m3.
+Proof.
+ destruct 1; intros. auto. eapply exec_straight_trans; eauto.
+Qed.
+
+Lemma transl_cbranch_correct_1:
+ forall cond args lbl k c m ms b sp rs m',
+ transl_cbranch cond args lbl k = OK c ->
+ eval_condition cond (List.map ms args) m = Some b ->
+ agree ms sp rs ->
+ Mem.extends m m' ->
+ exists rs', exists insn,
+ exec_straight_opt c rs m' (insn :: k) rs' m'
+ /\ exec_instr ge fn insn rs' m' = eval_branch fn lbl rs' m' (Some b)
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros until m'; intros TRANSL EVAL AG MEXT.
+ set (vl' := map rs (map preg_of args)).
+ assert (EVAL': eval_condition cond vl' m' = Some b).
+ { apply eval_condition_lessdef with (map ms args) m; auto. eapply preg_vals; eauto. }
+ clear EVAL MEXT AG.
+ destruct cond; simpl in TRANSL; ArgsInv.
+- exists rs, (transl_cbranch_int32s c0 x x0 lbl).
+ intuition auto. constructor. apply transl_cbranch_int32s_correct; auto.
+- exists rs, (transl_cbranch_int32u c0 x x0 lbl).
+ intuition auto. constructor. apply transl_cbranch_int32u_correct; auto.
+- predSpec Int.eq Int.eq_spec n Int.zero.
++ subst n. exists rs, (transl_cbranch_int32s c0 x X0 lbl).
+ intuition auto. constructor. apply transl_cbranch_int32s_correct; auto.
++ exploit (loadimm32_correct X31 n); eauto. intros (rs' & A & B & C).
+ exists rs', (transl_cbranch_int32s c0 x X31 lbl).
+ split. constructor; eexact A. split; auto.
+ apply transl_cbranch_int32s_correct; auto.
+ simpl; rewrite B, C; eauto with asmgen.
+- predSpec Int.eq Int.eq_spec n Int.zero.
++ subst n. exists rs, (transl_cbranch_int32u c0 x X0 lbl).
+ intuition auto. constructor. apply transl_cbranch_int32u_correct; auto.
++ exploit (loadimm32_correct X31 n); eauto. intros (rs' & A & B & C).
+ exists rs', (transl_cbranch_int32u c0 x X31 lbl).
+ split. constructor; eexact A. split; auto.
+ apply transl_cbranch_int32u_correct; auto.
+ simpl; rewrite B, C; eauto with asmgen.
+- exists rs, (transl_cbranch_int64s c0 x x0 lbl).
+ intuition auto. constructor. apply transl_cbranch_int64s_correct; auto.
+- exists rs, (transl_cbranch_int64u c0 x x0 lbl).
+ intuition auto. constructor. apply transl_cbranch_int64u_correct; auto.
+- predSpec Int64.eq Int64.eq_spec n Int64.zero.
++ subst n. exists rs, (transl_cbranch_int64s c0 x X0 lbl).
+ intuition auto. constructor. apply transl_cbranch_int64s_correct; auto.
++ exploit (loadimm64_correct X31 n); eauto. intros (rs' & A & B & C).
+ exists rs', (transl_cbranch_int64s c0 x X31 lbl).
+ split. constructor; eexact A. split; auto.
+ apply transl_cbranch_int64s_correct; auto.
+ simpl; rewrite B, C; eauto with asmgen.
+- predSpec Int64.eq Int64.eq_spec n Int64.zero.
++ subst n. exists rs, (transl_cbranch_int64u c0 x X0 lbl).
+ intuition auto. constructor. apply transl_cbranch_int64u_correct; auto.
++ exploit (loadimm64_correct X31 n); eauto. intros (rs' & A & B & C).
+ exists rs', (transl_cbranch_int64u c0 x X31 lbl).
+ split. constructor; eexact A. split; auto.
+ apply transl_cbranch_int64u_correct; auto.
+ simpl; rewrite B, C; eauto with asmgen.
+- destruct (transl_cond_float c0 X31 x x0) as [insn normal] eqn:TC; inv EQ2.
+ set (v := if normal then Val.cmpf c0 rs#x rs#x0 else Val.notbool (Val.cmpf c0 rs#x rs#x0)).
+ assert (V: v = Val.of_bool (eqb normal b)).
+ { unfold v, Val.cmpf. rewrite EVAL'. destruct normal, b; reflexivity. }
+ econstructor; econstructor.
+ split. constructor. apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto.
+ split. rewrite V; destruct normal, b; reflexivity.
+ intros; Simpl.
+- destruct (transl_cond_float c0 X31 x x0) as [insn normal] eqn:TC; inv EQ2.
+ assert (EVAL'': Val.cmpf_bool c0 (rs x) (rs x0) = Some (negb b)).
+ { destruct (Val.cmpf_bool c0 (rs x) (rs x0)) as [[]|]; inv EVAL'; auto. }
+ set (v := if normal then Val.cmpf c0 rs#x rs#x0 else Val.notbool (Val.cmpf c0 rs#x rs#x0)).
+ assert (V: v = Val.of_bool (xorb normal b)).
+ { unfold v, Val.cmpf. rewrite EVAL''. destruct normal, b; reflexivity. }
+ econstructor; econstructor.
+ split. constructor. apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto.
+ split. rewrite V; destruct normal, b; reflexivity.
+ intros; Simpl.
+- destruct (transl_cond_single c0 X31 x x0) as [insn normal] eqn:TC; inv EQ2.
+ set (v := if normal then Val.cmpfs c0 rs#x rs#x0 else Val.notbool (Val.cmpfs c0 rs#x rs#x0)).
+ assert (V: v = Val.of_bool (eqb normal b)).
+ { unfold v, Val.cmpfs. rewrite EVAL'. destruct normal, b; reflexivity. }
+ econstructor; econstructor.
+ split. constructor. apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto.
+ split. rewrite V; destruct normal, b; reflexivity.
+ intros; Simpl.
+- destruct (transl_cond_single c0 X31 x x0) as [insn normal] eqn:TC; inv EQ2.
+ assert (EVAL'': Val.cmpfs_bool c0 (rs x) (rs x0) = Some (negb b)).
+ { destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) as [[]|]; inv EVAL'; auto. }
+ set (v := if normal then Val.cmpfs c0 rs#x rs#x0 else Val.notbool (Val.cmpfs c0 rs#x rs#x0)).
+ assert (V: v = Val.of_bool (xorb normal b)).
+ { unfold v, Val.cmpfs. rewrite EVAL''. destruct normal, b; reflexivity. }
+ econstructor; econstructor.
+ split. constructor. apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto.
+ split. rewrite V; destruct normal, b; reflexivity.
+ intros; Simpl.
+Qed.
+
+Lemma transl_cbranch_correct_true:
+ forall cond args lbl k c m ms sp rs m',
+ transl_cbranch cond args lbl k = OK c ->
+ eval_condition cond (List.map ms args) m = Some true ->
+ agree ms sp rs ->
+ Mem.extends m m' ->
+ exists rs', exists insn,
+ exec_straight_opt c rs m' (insn :: k) rs' m'
+ /\ exec_instr ge fn insn rs' m' = goto_label fn lbl rs' m'
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. eapply transl_cbranch_correct_1 with (b := true); eauto.
+Qed.
+
+Lemma transl_cbranch_correct_false:
+ forall cond args lbl k c m ms sp rs m',
+ transl_cbranch cond args lbl k = OK c ->
+ eval_condition cond (List.map ms args) m = Some false ->
+ agree ms sp rs ->
+ Mem.extends m m' ->
+ exists rs',
+ exec_straight ge fn c rs m' k rs' m'
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. exploit transl_cbranch_correct_1; eauto. simpl.
+ intros (rs' & insn & A & B & C).
+ exists (nextinstr rs').
+ split. eapply exec_straight_opt_right; eauto. apply exec_straight_one; auto.
+ intros; Simpl.
+Qed.
+
+(** Translation of condition operators *)
+
+Lemma transl_cond_int32s_correct:
+ forall cmp rd r1 r2 k rs m,
+ exists rs',
+ exec_straight ge fn (transl_cond_int32s cmp rd r1 r2 k) rs m k rs' m
+ /\ Val.lessdef (Val.cmp cmp rs##r1 rs##r2) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros. destruct cmp; simpl.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl. destruct (rs##r1); auto. destruct (rs##r2); auto.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl. destruct (rs##r1); auto. destruct (rs##r2); auto.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold Val.cmp. rewrite <- Val.swap_cmp_bool.
+ simpl. rewrite (Val.negate_cmp_bool Clt).
+ destruct (Val.cmp_bool Clt rs##r2 rs##r1) as [[]|]; auto.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl. unfold Val.cmp. rewrite <- Val.swap_cmp_bool. auto.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold Val.cmp. rewrite (Val.negate_cmp_bool Clt).
+ destruct (Val.cmp_bool Clt rs##r1 rs##r2) as [[]|]; auto.
+Qed.
+
+Lemma transl_cond_int32u_correct:
+ forall cmp rd r1 r2 k rs m,
+ exists rs',
+ exec_straight ge fn (transl_cond_int32u cmp rd r1 r2 k) rs m k rs' m
+ /\ rs'#rd = Val.cmpu (Mem.valid_pointer m) cmp rs##r1 rs##r2
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros. destruct cmp; simpl.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold Val.cmpu. rewrite <- Val.swap_cmpu_bool.
+ simpl. rewrite (Val.negate_cmpu_bool (Mem.valid_pointer m) Cle).
+ destruct (Val.cmpu_bool (Mem.valid_pointer m) Cle rs##r1 rs##r2) as [[]|]; auto.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl. unfold Val.cmpu. rewrite <- Val.swap_cmpu_bool. auto.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold Val.cmpu. rewrite (Val.negate_cmpu_bool (Mem.valid_pointer m) Clt).
+ destruct (Val.cmpu_bool (Mem.valid_pointer m) Clt rs##r1 rs##r2) as [[]|]; auto.
+Qed.
+
+Lemma transl_cond_int64s_correct:
+ forall cmp rd r1 r2 k rs m,
+ exists rs',
+ exec_straight ge fn (transl_cond_int64s cmp rd r1 r2 k) rs m k rs' m
+ /\ Val.lessdef (Val.maketotal (Val.cmpl cmp rs###r1 rs###r2)) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros. destruct cmp; simpl.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl. destruct (rs###r1); auto. destruct (rs###r2); auto.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl. destruct (rs###r1); auto. destruct (rs###r2); auto.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold Val.cmpl. rewrite <- Val.swap_cmpl_bool.
+ simpl. rewrite (Val.negate_cmpl_bool Clt).
+ destruct (Val.cmpl_bool Clt rs###r2 rs###r1) as [[]|]; auto.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl. unfold Val.cmpl. rewrite <- Val.swap_cmpl_bool. auto.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold Val.cmpl. rewrite (Val.negate_cmpl_bool Clt).
+ destruct (Val.cmpl_bool Clt rs###r1 rs###r2) as [[]|]; auto.
+Qed.
+
+Lemma transl_cond_int64u_correct:
+ forall cmp rd r1 r2 k rs m,
+ exists rs',
+ exec_straight ge fn (transl_cond_int64u cmp rd r1 r2 k) rs m k rs' m
+ /\ rs'#rd = Val.maketotal (Val.cmplu (Mem.valid_pointer m) cmp rs###r1 rs###r2)
+ /\ forall r, r <> PC -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros. destruct cmp; simpl.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold Val.cmplu. rewrite <- Val.swap_cmplu_bool.
+ simpl. rewrite (Val.negate_cmplu_bool (Mem.valid_pointer m) Cle).
+ destruct (Val.cmplu_bool (Mem.valid_pointer m) Cle rs###r1 rs###r2) as [[]|]; auto.
+- econstructor; split. apply exec_straight_one; [simpl; eauto|auto].
+ split; intros; Simpl. unfold Val.cmplu. rewrite <- Val.swap_cmplu_bool. auto.
+- econstructor; split.
+ eapply exec_straight_two. simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold Val.cmplu. rewrite (Val.negate_cmplu_bool (Mem.valid_pointer m) Clt).
+ destruct (Val.cmplu_bool (Mem.valid_pointer m) Clt rs###r1 rs###r2) as [[]|]; auto.
+Qed.
+
+Lemma transl_condimm_int32s_correct:
+ forall cmp rd r1 n k rs m,
+ r1 <> X31 ->
+ exists rs',
+ exec_straight ge fn (transl_condimm_int32s cmp rd r1 n k) rs m k rs' m
+ /\ Val.lessdef (Val.cmp cmp rs#r1 (Vint n)) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. unfold transl_condimm_int32s.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+- subst n. exploit transl_cond_int32s_correct. intros (rs' & A & B & C).
+ exists rs'; eauto.
+- assert (DFL:
+ exists rs',
+ exec_straight ge fn (loadimm32 X31 n (transl_cond_int32s cmp rd r1 X31 k)) rs m k rs' m
+ /\ Val.lessdef (Val.cmp cmp rs#r1 (Vint n)) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r).
+ { exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1).
+ exploit transl_cond_int32s_correct; eauto. intros (rs2 & A2 & B2 & C2).
+ exists rs2; split.
+ eapply exec_straight_trans. eexact A1. eexact A2.
+ split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. auto.
+ intros; transitivity (rs1 r); auto. }
+ destruct cmp.
++ unfold xorimm32.
+ exploit (opimm32_correct Pxorw Pxoriw Val.xor); eauto. intros (rs1 & A1 & B1 & C1).
+ exploit transl_cond_int32s_correct; eauto. intros (rs2 & A2 & B2 & C2).
+ exists rs2; split.
+ eapply exec_straight_trans. eexact A1. eexact A2.
+ split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto.
+ unfold Val.cmp in B2; simpl in B2; rewrite Int.xor_is_zero in B2. exact B2.
+ intros; transitivity (rs1 r); auto.
++ unfold xorimm32.
+ exploit (opimm32_correct Pxorw Pxoriw Val.xor); eauto. intros (rs1 & A1 & B1 & C1).
+ exploit transl_cond_int32s_correct; eauto. intros (rs2 & A2 & B2 & C2).
+ exists rs2; split.
+ eapply exec_straight_trans. eexact A1. eexact A2.
+ split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto.
+ unfold Val.cmp in B2; simpl in B2; rewrite Int.xor_is_zero in B2. exact B2.
+ intros; transitivity (rs1 r); auto.
++ exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1).
+ exists rs1; split. eexact A1. split; auto. rewrite B1; auto.
++ predSpec Int.eq Int.eq_spec n (Int.repr Int.max_signed).
+* subst n. exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1).
+ exists rs1; split. eexact A1. split; auto.
+ unfold Val.cmp; destruct (rs#r1); simpl; auto. rewrite B1.
+ unfold Int.lt. rewrite zlt_false. auto.
+ change (Int.signed (Int.repr Int.max_signed)) with Int.max_signed.
+ generalize (Int.signed_range i); omega.
+* exploit (opimm32_correct Psltw Psltiw (Val.cmp Clt)); eauto. intros (rs1 & A1 & B1 & C1).
+ exists rs1; split. eexact A1. split; auto.
+ rewrite B1. unfold Val.cmp; simpl; destruct (rs#r1); simpl; auto.
+ unfold Int.lt. replace (Int.signed (Int.add n Int.one)) with (Int.signed n + 1).
+ destruct (zlt (Int.signed n) (Int.signed i)).
+ rewrite zlt_false by omega. auto.
+ rewrite zlt_true by omega. auto.
+ rewrite Int.add_signed. symmetry; apply Int.signed_repr.
+ assert (Int.signed n <> Int.max_signed).
+ { red; intros E. elim H1. rewrite <- (Int.repr_signed n). rewrite E. auto. }
+ generalize (Int.signed_range n); omega.
++ apply DFL.
++ apply DFL.
+Qed.
+
+Lemma transl_condimm_int32u_correct:
+ forall cmp rd r1 n k rs m,
+ r1 <> X31 ->
+ exists rs',
+ exec_straight ge fn (transl_condimm_int32u cmp rd r1 n k) rs m k rs' m
+ /\ Val.lessdef (Val.cmpu (Mem.valid_pointer m) cmp rs#r1 (Vint n)) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. unfold transl_condimm_int32u.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+- subst n. exploit transl_cond_int32u_correct. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split; auto. rewrite B; auto.
+- assert (DFL:
+ exists rs',
+ exec_straight ge fn (loadimm32 X31 n (transl_cond_int32u cmp rd r1 X31 k)) rs m k rs' m
+ /\ Val.lessdef (Val.cmpu (Mem.valid_pointer m) cmp rs#r1 (Vint n)) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r).
+ { exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1).
+ exploit transl_cond_int32u_correct; eauto. intros (rs2 & A2 & B2 & C2).
+ exists rs2; split.
+ eapply exec_straight_trans. eexact A1. eexact A2.
+ split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. rewrite B2; auto.
+ intros; transitivity (rs1 r); auto. }
+ destruct cmp.
++ apply DFL.
++ apply DFL.
++ exploit (opimm32_correct Psltuw Psltiuw (Val.cmpu (Mem.valid_pointer m) Clt) m); eauto.
+ intros (rs1 & A1 & B1 & C1).
+ exists rs1; split. eexact A1. split; auto. rewrite B1; auto.
++ apply DFL.
++ apply DFL.
++ apply DFL.
+Qed.
+
+Lemma transl_condimm_int64s_correct:
+ forall cmp rd r1 n k rs m,
+ r1 <> X31 ->
+ exists rs',
+ exec_straight ge fn (transl_condimm_int64s cmp rd r1 n k) rs m k rs' m
+ /\ Val.lessdef (Val.maketotal (Val.cmpl cmp rs#r1 (Vlong n))) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. unfold transl_condimm_int64s.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero.
+- subst n. exploit transl_cond_int64s_correct. intros (rs' & A & B & C).
+ exists rs'; eauto.
+- assert (DFL:
+ exists rs',
+ exec_straight ge fn (loadimm64 X31 n (transl_cond_int64s cmp rd r1 X31 k)) rs m k rs' m
+ /\ Val.lessdef (Val.maketotal (Val.cmpl cmp rs#r1 (Vlong n))) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r).
+ { exploit loadimm64_correct; eauto. intros (rs1 & A1 & B1 & C1).
+ exploit transl_cond_int64s_correct; eauto. intros (rs2 & A2 & B2 & C2).
+ exists rs2; split.
+ eapply exec_straight_trans. eexact A1. eexact A2.
+ split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. auto.
+ intros; transitivity (rs1 r); auto. }
+ destruct cmp.
++ unfold xorimm64.
+ exploit (opimm64_correct Pxorl Pxoril Val.xorl); eauto. intros (rs1 & A1 & B1 & C1).
+ exploit transl_cond_int64s_correct; eauto. intros (rs2 & A2 & B2 & C2).
+ exists rs2; split.
+ eapply exec_straight_trans. eexact A1. eexact A2.
+ split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto.
+ unfold Val.cmpl in B2; simpl in B2; rewrite Int64.xor_is_zero in B2. exact B2.
+ intros; transitivity (rs1 r); auto.
++ unfold xorimm64.
+ exploit (opimm64_correct Pxorl Pxoril Val.xorl); eauto. intros (rs1 & A1 & B1 & C1).
+ exploit transl_cond_int64s_correct; eauto. intros (rs2 & A2 & B2 & C2).
+ exists rs2; split.
+ eapply exec_straight_trans. eexact A1. eexact A2.
+ split. simpl in B2; rewrite B1 in B2; simpl in B2. destruct (rs#r1); auto.
+ unfold Val.cmpl in B2; simpl in B2; rewrite Int64.xor_is_zero in B2. exact B2.
+ intros; transitivity (rs1 r); auto.
++ exploit (opimm64_correct Psltl Psltil (fun v1 v2 => Val.maketotal (Val.cmpl Clt v1 v2))); eauto. intros (rs1 & A1 & B1 & C1).
+ exists rs1; split. eexact A1. split; auto. rewrite B1; auto.
++ predSpec Int64.eq Int64.eq_spec n (Int64.repr Int64.max_signed).
+* subst n. exploit loadimm32_correct; eauto. intros (rs1 & A1 & B1 & C1).
+ exists rs1; split. eexact A1. split; auto.
+ unfold Val.cmpl; destruct (rs#r1); simpl; auto. rewrite B1.
+ unfold Int64.lt. rewrite zlt_false. auto.
+ change (Int64.signed (Int64.repr Int64.max_signed)) with Int64.max_signed.
+ generalize (Int64.signed_range i); omega.
+* exploit (opimm64_correct Psltl Psltil (fun v1 v2 => Val.maketotal (Val.cmpl Clt v1 v2))); eauto. intros (rs1 & A1 & B1 & C1).
+ exists rs1; split. eexact A1. split; auto.
+ rewrite B1. unfold Val.cmpl; simpl; destruct (rs#r1); simpl; auto.
+ unfold Int64.lt. replace (Int64.signed (Int64.add n Int64.one)) with (Int64.signed n + 1).
+ destruct (zlt (Int64.signed n) (Int64.signed i)).
+ rewrite zlt_false by omega. auto.
+ rewrite zlt_true by omega. auto.
+ rewrite Int64.add_signed. symmetry; apply Int64.signed_repr.
+ assert (Int64.signed n <> Int64.max_signed).
+ { red; intros E. elim H1. rewrite <- (Int64.repr_signed n). rewrite E. auto. }
+ generalize (Int64.signed_range n); omega.
++ apply DFL.
++ apply DFL.
+Qed.
+
+Lemma transl_condimm_int64u_correct:
+ forall cmp rd r1 n k rs m,
+ r1 <> X31 ->
+ exists rs',
+ exec_straight ge fn (transl_condimm_int64u cmp rd r1 n k) rs m k rs' m
+ /\ Val.lessdef (Val.maketotal (Val.cmplu (Mem.valid_pointer m) cmp rs#r1 (Vlong n))) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. unfold transl_condimm_int64u.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero.
+- subst n. exploit transl_cond_int64u_correct. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. split; auto. rewrite B; auto.
+- assert (DFL:
+ exists rs',
+ exec_straight ge fn (loadimm64 X31 n (transl_cond_int64u cmp rd r1 X31 k)) rs m k rs' m
+ /\ Val.lessdef (Val.maketotal (Val.cmplu (Mem.valid_pointer m) cmp rs#r1 (Vlong n))) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r).
+ { exploit loadimm64_correct; eauto. intros (rs1 & A1 & B1 & C1).
+ exploit transl_cond_int64u_correct; eauto. intros (rs2 & A2 & B2 & C2).
+ exists rs2; split.
+ eapply exec_straight_trans. eexact A1. eexact A2.
+ split. simpl in B2. rewrite B1, C1 in B2 by auto with asmgen. rewrite B2; auto.
+ intros; transitivity (rs1 r); auto. }
+ destruct cmp.
++ apply DFL.
++ apply DFL.
++ exploit (opimm64_correct Psltul Psltiul (fun v1 v2 => Val.maketotal (Val.cmplu (Mem.valid_pointer m) Clt v1 v2)) m); eauto.
+ intros (rs1 & A1 & B1 & C1).
+ exists rs1; split. eexact A1. split; auto. rewrite B1; auto.
++ apply DFL.
++ apply DFL.
++ apply DFL.
+Qed.
+
+Lemma transl_cond_op_correct:
+ forall cond rd args k c rs m,
+ transl_cond_op cond rd args k = OK c ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ Val.lessdef (Val.of_optbool (eval_condition cond (map rs (map preg_of args)) m)) rs'#rd
+ /\ forall r, r <> PC -> r <> rd -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ assert (MKTOT: forall ob, Val.of_optbool ob = Val.maketotal (option_map Val.of_bool ob)).
+ { destruct ob as [[]|]; reflexivity. }
+ intros until m; intros TR.
+ destruct cond; simpl in TR; ArgsInv.
++ (* cmp *)
+ exploit transl_cond_int32s_correct; eauto. intros (rs' & A & B & C). exists rs'; eauto.
++ (* cmpu *)
+ exploit transl_cond_int32u_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; repeat split; eauto. rewrite B; auto.
++ (* cmpimm *)
+ apply transl_condimm_int32s_correct; eauto with asmgen.
++ (* cmpuimm *)
+ apply transl_condimm_int32u_correct; eauto with asmgen.
++ (* cmpl *)
+ exploit transl_cond_int64s_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; repeat split; eauto. rewrite MKTOT; eauto.
++ (* cmplu *)
+ exploit transl_cond_int64u_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; repeat split; eauto. rewrite B, MKTOT; eauto.
++ (* cmplimm *)
+ exploit transl_condimm_int64s_correct; eauto. instantiate (1 := x); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; repeat split; eauto. rewrite MKTOT; eauto.
++ (* cmpluimm *)
+ exploit transl_condimm_int64u_correct; eauto. instantiate (1 := x); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; repeat split; eauto. rewrite MKTOT; eauto.
++ (* cmpf *)
+ destruct (transl_cond_float c0 rd x x0) as [insn normal] eqn:TR.
+ fold (Val.cmpf c0 (rs x) (rs x0)).
+ set (v := Val.cmpf c0 (rs x) (rs x0)).
+ destruct normal; inv EQ2.
+* econstructor; split.
+ apply exec_straight_one. eapply transl_cond_float_correct with (v := v); eauto. auto.
+ split; intros; Simpl.
+* econstructor; split.
+ eapply exec_straight_two.
+ eapply transl_cond_float_correct with (v := Val.notbool v); eauto.
+ simpl; reflexivity.
+ auto. auto.
+ split; intros; Simpl. unfold v, Val.cmpf. destruct (Val.cmpf_bool c0 (rs x) (rs x0)) as [[]|]; auto.
++ (* notcmpf *)
+ destruct (transl_cond_float c0 rd x x0) as [insn normal] eqn:TR.
+ rewrite Val.notbool_negb_3. fold (Val.cmpf c0 (rs x) (rs x0)).
+ set (v := Val.cmpf c0 (rs x) (rs x0)).
+ destruct normal; inv EQ2.
+* econstructor; split.
+ eapply exec_straight_two.
+ eapply transl_cond_float_correct with (v := v); eauto.
+ simpl; reflexivity.
+ auto. auto.
+ split; intros; Simpl. unfold v, Val.cmpf. destruct (Val.cmpf_bool c0 (rs x) (rs x0)) as [[]|]; auto.
+* econstructor; split.
+ apply exec_straight_one. eapply transl_cond_float_correct with (v := Val.notbool v); eauto. auto.
+ split; intros; Simpl.
++ (* cmpfs *)
+ destruct (transl_cond_single c0 rd x x0) as [insn normal] eqn:TR.
+ fold (Val.cmpfs c0 (rs x) (rs x0)).
+ set (v := Val.cmpfs c0 (rs x) (rs x0)).
+ destruct normal; inv EQ2.
+* econstructor; split.
+ apply exec_straight_one. eapply transl_cond_single_correct with (v := v); eauto. auto.
+ split; intros; Simpl.
+* econstructor; split.
+ eapply exec_straight_two.
+ eapply transl_cond_single_correct with (v := Val.notbool v); eauto.
+ simpl; reflexivity.
+ auto. auto.
+ split; intros; Simpl. unfold v, Val.cmpfs. destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) as [[]|]; auto.
++ (* notcmpfs *)
+ destruct (transl_cond_single c0 rd x x0) as [insn normal] eqn:TR.
+ rewrite Val.notbool_negb_3. fold (Val.cmpfs c0 (rs x) (rs x0)).
+ set (v := Val.cmpfs c0 (rs x) (rs x0)).
+ destruct normal; inv EQ2.
+* econstructor; split.
+ eapply exec_straight_two.
+ eapply transl_cond_single_correct with (v := v); eauto.
+ simpl; reflexivity.
+ auto. auto.
+ split; intros; Simpl. unfold v, Val.cmpfs. destruct (Val.cmpfs_bool c0 (rs x) (rs x0)) as [[]|]; auto.
+* econstructor; split.
+ apply exec_straight_one. eapply transl_cond_single_correct with (v := Val.notbool v); eauto. auto.
+ split; intros; Simpl.
+Qed.
+
+(** Some arithmetic properties. *)
+
+Remark cast32unsigned_from_cast32signed:
+ forall i, Int64.repr (Int.unsigned i) = Int64.zero_ext 32 (Int64.repr (Int.signed i)).
+Proof.
+ intros. apply Int64.same_bits_eq; intros.
+ rewrite Int64.bits_zero_ext, !Int64.testbit_repr by tauto.
+ rewrite Int.bits_signed by tauto. fold (Int.testbit i i0).
+ change Int.zwordsize with 32.
+ destruct (zlt i0 32). auto. apply Int.bits_above. auto.
+Qed.
+
+(* Translation of arithmetic operations *)
+
+Ltac SimplEval H :=
+ match type of H with
+ | Some _ = None _ => discriminate
+ | Some _ = Some _ => inv H
+ | ?a = Some ?b => let A := fresh in assert (A: Val.maketotal a = b) by (rewrite H; reflexivity)
+end.
+
+Ltac TranslOpSimpl :=
+ econstructor; split;
+ [ apply exec_straight_one; [simpl; eauto | reflexivity]
+ | split; [ apply Val.lessdef_same; Simpl; fail | intros; Simpl; fail ] ].
+
+Lemma transl_op_correct:
+ forall op args res k (rs: regset) m v c,
+ transl_op op args res k = OK c ->
+ eval_operation ge (rs#SP) op (map rs (map preg_of args)) m = Some v ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ Val.lessdef v rs'#(preg_of res)
+ /\ forall r, data_preg r = true -> r <> preg_of res -> preg_notin r (destroyed_by_op op) -> rs' r = rs r.
+Proof.
+ assert (SAME: forall v1 v2, v1 = v2 -> Val.lessdef v2 v1). { intros; subst; auto. }
+Opaque Int.eq.
+ intros until c; intros TR EV.
+ unfold transl_op in TR; destruct op; ArgsInv; simpl in EV; SimplEval EV; try TranslOpSimpl.
+- (* move *)
+ destruct (preg_of res), (preg_of m0); inv TR; TranslOpSimpl.
+- (* intconst *)
+ exploit loadimm32_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* longconst *)
+ exploit loadimm64_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* floatconst *)
+ destruct (Float.eq_dec n Float.zero).
++ subst n. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
++ econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
+- (* singleconst *)
+ destruct (Float32.eq_dec n Float32.zero).
++ subst n. econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
++ econstructor; split.
+ apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
+- (* addrsymbol *)
+ destruct (Archi.pic_code tt && negb (Ptrofs.eq ofs Ptrofs.zero)).
++ set (rs1 := nextinstr (rs#x <- (Genv.symbol_address ge id Ptrofs.zero))).
+ exploit (addptrofs_correct x x ofs k rs1 m); eauto with asmgen.
+ intros (rs2 & A & B & C).
+ exists rs2; split.
+ apply exec_straight_step with rs1 m; auto.
+ split. replace ofs with (Ptrofs.add Ptrofs.zero ofs) by (apply Ptrofs.add_zero_l).
+ rewrite Genv.shift_symbol_address.
+ replace (rs1 x) with (Genv.symbol_address ge id Ptrofs.zero) in B by (unfold rs1; Simpl).
+ exact B.
+ intros. rewrite C by eauto with asmgen. unfold rs1; Simpl.
++ TranslOpSimpl.
+- (* stackoffset *)
+ exploit addptrofs_correct. instantiate (1 := X2); auto with asmgen. intros (rs' & A & B & C).
+ exists rs'; split; eauto. auto with asmgen.
+- (* cast8signed *)
+ econstructor; split.
+ eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto.
+ split; intros; Simpl.
+ assert (A: Int.ltu (Int.repr 24) Int.iwordsize = true) by auto.
+ destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A.
+ apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity.
+- (* cast16signed *)
+ econstructor; split.
+ eapply exec_straight_two. simpl;eauto. simpl;eauto. auto. auto.
+ split; intros; Simpl.
+ assert (A: Int.ltu (Int.repr 16) Int.iwordsize = true) by auto.
+ destruct (rs x0); auto; simpl. rewrite A; simpl. rewrite A.
+ apply Val.lessdef_same. f_equal. apply Int.sign_ext_shr_shl. split; reflexivity.
+- (* addimm *)
+ exploit (opimm32_correct Paddw Paddiw Val.add); auto. instantiate (1 := x0); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* andimm *)
+ exploit (opimm32_correct Pandw Pandiw Val.and); auto. instantiate (1 := x0); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* orimm *)
+ exploit (opimm32_correct Porw Poriw Val.or); auto. instantiate (1 := x0); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* xorimm *)
+ exploit (opimm32_correct Pxorw Pxoriw Val.xor); auto. instantiate (1 := x0); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* shrximm *)
+ clear H. exploit Val.shrx_shr_2; eauto. intros E; subst v; clear EV.
+ destruct (Int.eq n Int.zero).
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
++ change (Int.repr 32) with Int.iwordsize. set (n' := Int.sub Int.iwordsize n).
+ econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl.
+- (* longofintu *)
+ econstructor; split.
+ eapply exec_straight_three. simpl; eauto. simpl; eauto. simpl; eauto. auto. auto. auto.
+ split; intros; Simpl. destruct (rs x0); auto. simpl.
+ assert (A: Int.ltu (Int.repr 32) Int64.iwordsize' = true) by auto.
+ rewrite A; simpl. rewrite A. apply Val.lessdef_same. f_equal.
+ rewrite cast32unsigned_from_cast32signed. apply Int64.zero_ext_shru_shl. compute; auto.
+- (* addlimm *)
+ exploit (opimm64_correct Paddl Paddil Val.addl); auto. instantiate (1 := x0); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* andimm *)
+ exploit (opimm64_correct Pandl Pandil Val.andl); auto. instantiate (1 := x0); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* orimm *)
+ exploit (opimm64_correct Porl Poril Val.orl); auto. instantiate (1 := x0); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* xorimm *)
+ exploit (opimm64_correct Pxorl Pxoril Val.xorl); auto. instantiate (1 := x0); eauto with asmgen.
+ intros (rs' & A & B & C).
+ exists rs'; split; eauto. rewrite B; auto with asmgen.
+- (* shrxlimm *)
+ clear H. exploit Val.shrxl_shrl_2; eauto. intros E; subst v; clear EV.
+ destruct (Int.eq n Int.zero).
++ econstructor; split. apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl.
++ change (Int.repr 64) with Int64.iwordsize'. set (n' := Int.sub Int64.iwordsize' n).
+ econstructor; split.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ eapply exec_straight_step. simpl; reflexivity. auto.
+ apply exec_straight_one. simpl; reflexivity. auto.
+ split; intros; Simpl.
+- (* cond *)
+ exploit transl_cond_op_correct; eauto. intros (rs' & A & B & C).
+ exists rs'; split. eexact A. eauto with asmgen.
+Qed.
+
+(** Memory accesses *)
+
+Lemma indexed_memory_access_correct:
+ forall mk_instr base ofs k rs m,
+ base <> X31 ->
+ exists base' ofs' rs',
+ exec_straight_opt (indexed_memory_access mk_instr base ofs k) rs m
+ (mk_instr base' ofs' :: k) rs' m
+ /\ Val.offset_ptr rs'#base' (eval_offset ge ofs') = Val.offset_ptr rs#base ofs
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ unfold indexed_memory_access; intros.
+ destruct Archi.ptr64 eqn:SF.
+- generalize (make_immed64_sound (Ptrofs.to_int64 ofs)); intros EQ.
+ destruct (make_immed64 (Ptrofs.to_int64 ofs)).
++ econstructor; econstructor; econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. subst imm. rewrite Ptrofs.of_int64_to_int64 by auto. auto.
++ econstructor; econstructor; econstructor; split.
+ constructor. eapply exec_straight_two.
+ simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. destruct (rs base); auto; simpl. rewrite SF. simpl.
+ rewrite Ptrofs.add_assoc. f_equal. f_equal.
+ rewrite <- (Ptrofs.of_int64_to_int64 SF ofs). rewrite EQ.
+ symmetry; auto with ptrofs.
++ econstructor; econstructor; econstructor; split.
+ constructor. eapply exec_straight_two.
+ simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. unfold eval_offset. destruct (rs base); auto; simpl. rewrite SF. simpl.
+ rewrite Ptrofs.add_zero. subst imm. rewrite Ptrofs.of_int64_to_int64 by auto. auto.
+- generalize (make_immed32_sound (Ptrofs.to_int ofs)); intros EQ.
+ destruct (make_immed32 (Ptrofs.to_int ofs)).
++ econstructor; econstructor; econstructor; split.
+ apply exec_straight_opt_refl.
+ split; auto. simpl. subst imm. rewrite Ptrofs.of_int_to_int by auto. auto.
++ econstructor; econstructor; econstructor; split.
+ constructor. eapply exec_straight_two.
+ simpl; eauto. simpl; eauto. auto. auto.
+ split; intros; Simpl. destruct (rs base); auto; simpl. rewrite SF. simpl.
+ rewrite Ptrofs.add_assoc. f_equal. f_equal.
+ rewrite <- (Ptrofs.of_int_to_int SF ofs). rewrite EQ.
+ symmetry; auto with ptrofs.
+Qed.
+
+Lemma indexed_load_access_correct:
+ forall chunk (mk_instr: ireg -> offset -> instruction) rd m,
+ (forall base ofs rs,
+ exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m rd base ofs) ->
+ forall (base: ireg) ofs k (rs: regset) v,
+ Mem.loadv chunk m (Val.offset_ptr rs#base ofs) = Some v ->
+ base <> X31 -> rd <> PC ->
+ exists rs',
+ exec_straight ge fn (indexed_memory_access mk_instr base ofs k) rs m k rs' m
+ /\ rs'#rd = v
+ /\ forall r, r <> PC -> r <> X31 -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros until m; intros EXEC; intros until v; intros LOAD NOT31 NOTPC.
+ exploit indexed_memory_access_correct; eauto.
+ intros (base' & ofs' & rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite EXEC.
+ unfold exec_load. rewrite B, LOAD. eauto. Simpl.
+ split; intros; Simpl.
+Qed.
+
+Lemma indexed_store_access_correct:
+ forall chunk (mk_instr: ireg -> offset -> instruction) r1 m,
+ (forall base ofs rs,
+ exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk rs m r1 base ofs) ->
+ forall (base: ireg) ofs k (rs: regset) m',
+ Mem.storev chunk m (Val.offset_ptr rs#base ofs) (rs#r1) = Some m' ->
+ base <> X31 -> r1 <> X31 -> r1 <> PC ->
+ exists rs',
+ exec_straight ge fn (indexed_memory_access mk_instr base ofs k) rs m k rs' m'
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros until m; intros EXEC; intros until m'; intros STORE NOT31 NOT31' NOTPC.
+ exploit indexed_memory_access_correct; eauto.
+ intros (base' & ofs' & rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A. apply exec_straight_one. rewrite EXEC.
+ unfold exec_store. rewrite B, C, STORE by auto. eauto. auto.
+ intros; Simpl.
+Qed.
+
+Lemma loadind_correct:
+ forall (base: ireg) ofs ty dst k c (rs: regset) m v,
+ loadind base ofs ty dst k = OK c ->
+ Mem.loadv (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) = Some v ->
+ base <> X31 ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ rs'#(preg_of dst) = v
+ /\ forall r, r <> PC -> r <> X31 -> r <> preg_of dst -> rs'#r = rs#r.
+Proof.
+ intros until v; intros TR LOAD NOT31.
+ assert (A: exists mk_instr,
+ c = indexed_memory_access mk_instr base ofs k
+ /\ forall base' ofs' rs',
+ exec_instr ge fn (mk_instr base' ofs') rs' m =
+ exec_load ge (chunk_of_type ty) rs' m (preg_of dst) base' ofs').
+ { unfold loadind in TR. destruct ty, (preg_of dst); inv TR; econstructor; split; eauto. }
+ destruct A as (mk_instr & B & C). subst c.
+ eapply indexed_load_access_correct; eauto with asmgen.
+Qed.
+
+Lemma storeind_correct:
+ forall (base: ireg) ofs ty src k c (rs: regset) m m',
+ storeind src base ofs ty k = OK c ->
+ Mem.storev (chunk_of_type ty) m (Val.offset_ptr rs#base ofs) rs#(preg_of src) = Some m' ->
+ base <> X31 ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m'
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros until m'; intros TR STORE NOT31.
+ assert (A: exists mk_instr,
+ c = indexed_memory_access mk_instr base ofs k
+ /\ forall base' ofs' rs',
+ exec_instr ge fn (mk_instr base' ofs') rs' m =
+ exec_store ge (chunk_of_type ty) rs' m (preg_of src) base' ofs').
+ { unfold storeind in TR. destruct ty, (preg_of src); inv TR; econstructor; split; eauto. }
+ destruct A as (mk_instr & B & C). subst c.
+ eapply indexed_store_access_correct; eauto with asmgen.
+Qed.
+
+Lemma loadind_ptr_correct:
+ forall (base: ireg) ofs (dst: ireg) k (rs: regset) m v,
+ Mem.loadv Mptr m (Val.offset_ptr rs#base ofs) = Some v ->
+ base <> X31 ->
+ exists rs',
+ exec_straight ge fn (loadind_ptr base ofs dst k) rs m k rs' m
+ /\ rs'#dst = v
+ /\ forall r, r <> PC -> r <> X31 -> r <> dst -> rs'#r = rs#r.
+Proof.
+ intros. eapply indexed_load_access_correct; eauto with asmgen.
+ intros. unfold Mptr. destruct Archi.ptr64; auto.
+Qed.
+
+Lemma storeind_ptr_correct:
+ forall (base: ireg) ofs (src: ireg) k (rs: regset) m m',
+ Mem.storev Mptr m (Val.offset_ptr rs#base ofs) rs#src = Some m' ->
+ base <> X31 -> src <> X31 ->
+ exists rs',
+ exec_straight ge fn (storeind_ptr src base ofs k) rs m k rs' m'
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros. eapply indexed_store_access_correct with (r1 := src); eauto with asmgen.
+ intros. unfold Mptr. destruct Archi.ptr64; auto.
+Qed.
+
+Lemma transl_memory_access_correct:
+ forall mk_instr addr args k c (rs: regset) m v,
+ transl_memory_access mk_instr addr args k = OK c ->
+ eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v ->
+ exists base ofs rs',
+ exec_straight_opt c rs m (mk_instr base ofs :: k) rs' m
+ /\ Val.offset_ptr rs'#base (eval_offset ge ofs) = v
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros until v; intros TR EV.
+ unfold transl_memory_access in TR; destruct addr; ArgsInv.
+- (* indexed *)
+ inv EV. apply indexed_memory_access_correct; eauto with asmgen.
+- (* global *)
+ simpl in EV. inv EV. inv TR. econstructor; econstructor; econstructor; split.
+ constructor. apply exec_straight_one. simpl; eauto. auto.
+ split; intros; Simpl. unfold eval_offset. apply low_high_half.
+- (* stack *)
+ inv TR. inv EV. apply indexed_memory_access_correct; eauto with asmgen.
+Qed.
+
+Lemma transl_load_access_correct:
+ forall chunk (mk_instr: ireg -> offset -> instruction) addr args k c rd (rs: regset) m v v',
+ (forall base ofs rs,
+ exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m rd base ofs) ->
+ transl_memory_access mk_instr addr args k = OK c ->
+ eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v ->
+ Mem.loadv chunk m v = Some v' ->
+ rd <> PC ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ rs'#rd = v'
+ /\ forall r, r <> PC -> r <> X31 -> r <> rd -> rs'#r = rs#r.
+Proof.
+ intros until v'; intros INSTR TR EV LOAD NOTPC.
+ exploit transl_memory_access_correct; eauto.
+ intros (base & ofs & rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A. apply exec_straight_one.
+ rewrite INSTR. unfold exec_load. rewrite B, LOAD. reflexivity. Simpl.
+ split; intros; Simpl.
+Qed.
+
+Lemma transl_store_access_correct:
+ forall chunk (mk_instr: ireg -> offset -> instruction) addr args k c r1 (rs: regset) m v m',
+ (forall base ofs rs,
+ exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk rs m r1 base ofs) ->
+ transl_memory_access mk_instr addr args k = OK c ->
+ eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some v ->
+ Mem.storev chunk m v rs#r1 = Some m' ->
+ r1 <> PC -> r1 <> X31 ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m'
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros until m'; intros INSTR TR EV STORE NOTPC NOT31.
+ exploit transl_memory_access_correct; eauto.
+ intros (base & ofs & rs' & A & B & C).
+ econstructor; split.
+ eapply exec_straight_opt_right. eexact A. apply exec_straight_one.
+ rewrite INSTR. unfold exec_store. rewrite B, C, STORE by auto. reflexivity. auto.
+ intros; Simpl.
+Qed.
+
+Lemma transl_load_correct:
+ forall chunk addr args dst k c (rs: regset) m a v,
+ transl_load chunk addr args dst k = OK c ->
+ eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a ->
+ Mem.loadv chunk m a = Some v ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m
+ /\ rs'#(preg_of dst) = v
+ /\ forall r, r <> PC -> r <> X31 -> r <> preg_of dst -> rs'#r = rs#r.
+Proof.
+ intros until v; intros TR EV LOAD.
+ assert (A: exists mk_instr,
+ transl_memory_access mk_instr addr args k = OK c
+ /\ forall base ofs rs,
+ exec_instr ge fn (mk_instr base ofs) rs m = exec_load ge chunk rs m (preg_of dst) base ofs).
+ { unfold transl_load in TR; destruct chunk; ArgsInv; econstructor; (split; [eassumption|auto]). }
+ destruct A as (mk_instr & B & C).
+ eapply transl_load_access_correct; eauto with asmgen.
+Qed.
+
+Lemma transl_store_correct:
+ forall chunk addr args src k c (rs: regset) m a m',
+ transl_store chunk addr args src k = OK c ->
+ eval_addressing ge rs#SP addr (map rs (map preg_of args)) = Some a ->
+ Mem.storev chunk m a rs#(preg_of src) = Some m' ->
+ exists rs',
+ exec_straight ge fn c rs m k rs' m'
+ /\ forall r, r <> PC -> r <> X31 -> rs'#r = rs#r.
+Proof.
+ intros until m'; intros TR EV STORE.
+ assert (A: exists mk_instr chunk',
+ transl_memory_access mk_instr addr args k = OK c
+ /\ (forall base ofs rs,
+ exec_instr ge fn (mk_instr base ofs) rs m = exec_store ge chunk' rs m (preg_of src) base ofs)
+ /\ Mem.storev chunk m a rs#(preg_of src) = Mem.storev chunk' m a rs#(preg_of src)).
+ { unfold transl_store in TR; destruct chunk; ArgsInv;
+ (econstructor; econstructor; split; [eassumption | split; [ intros; simpl; reflexivity | auto]]).
+ destruct a; auto. apply Mem.store_signed_unsigned_8.
+ destruct a; auto. apply Mem.store_signed_unsigned_16.
+ }
+ destruct A as (mk_instr & chunk' & B & C & D).
+ rewrite D in STORE; clear D.
+ eapply transl_store_access_correct; eauto with asmgen.
+Qed.
+
+(** Function epilogues *)
+
+Lemma make_epilogue_correct:
+ forall ge0 f m stk soff cs m' ms rs k tm,
+ load_stack m (Vptr stk soff) Tptr f.(fn_link_ofs) = Some (parent_sp cs) ->
+ load_stack m (Vptr stk soff) Tptr f.(fn_retaddr_ofs) = Some (parent_ra cs) ->
+ Mem.free m stk 0 f.(fn_stacksize) = Some m' ->
+ agree ms (Vptr stk soff) rs ->
+ Mem.extends m tm ->
+ match_stack ge0 cs ->
+ exists rs', exists tm',
+ exec_straight ge fn (make_epilogue f k) rs tm k rs' tm'
+ /\ agree ms (parent_sp cs) rs'
+ /\ Mem.extends m' tm'
+ /\ rs'#RA = parent_ra cs
+ /\ rs'#SP = parent_sp cs
+ /\ (forall r, r <> PC -> r <> RA -> r <> SP -> r <> X31 -> rs'#r = rs#r).
+Proof.
+ intros until tm; intros LP LRA FREE AG MEXT MCS.
+ exploit Mem.loadv_extends. eauto. eexact LP. auto. simpl. intros (parent' & LP' & LDP').
+ exploit Mem.loadv_extends. eauto. eexact LRA. auto. simpl. intros (ra' & LRA' & LDRA').
+ exploit lessdef_parent_sp; eauto. intros EQ; subst parent'; clear LDP'.
+ exploit lessdef_parent_ra; eauto. intros EQ; subst ra'; clear LDRA'.
+ exploit Mem.free_parallel_extends; eauto. intros (tm' & FREE' & MEXT').
+ unfold make_epilogue.
+ rewrite chunk_of_Tptr in *.
+ exploit (loadind_ptr_correct SP (fn_retaddr_ofs f) RA (Pfreeframe (fn_stacksize f) (fn_link_ofs f) :: k) rs tm).
+ rewrite <- (sp_val _ _ _ AG). simpl. eexact LRA'. congruence.
+ intros (rs1 & A1 & B1 & C1).
+ econstructor; econstructor; split.
+ eapply exec_straight_trans. eexact A1. apply exec_straight_one. simpl.
+ rewrite (C1 X2) by auto with asmgen. rewrite <- (sp_val _ _ _ AG). simpl; rewrite LP'.
+ rewrite FREE'. eauto. auto.
+ split. apply agree_nextinstr. apply agree_set_other; auto with asmgen.
+ apply agree_change_sp with (Vptr stk soff).
+ apply agree_exten with rs; auto. intros; apply C1; auto with asmgen.
+ eapply parent_sp_def; eauto.
+ split. auto.
+ split. Simpl.
+ split. Simpl.
+ intros. Simpl.
+Qed.
+
+End CONSTRUCTORS.
+
+
+
diff --git a/riscV/CBuiltins.ml b/riscV/CBuiltins.ml
new file mode 100644
index 00000000..385f9d13
--- /dev/null
+++ b/riscV/CBuiltins.ml
@@ -0,0 +1,69 @@
+(* *********************************************************************)
+(* *)
+(* 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 GNU General Public License as published by *)
+(* the Free Software Foundation, either version 2 of the License, or *)
+(* (at your option) any later version. This file is also distributed *)
+(* under the terms of the INRIA Non-Commercial License Agreement. *)
+(* *)
+(* *********************************************************************)
+
+(* Processor-dependent builtin C functions *)
+
+open C
+
+let builtins = {
+ Builtins.typedefs = [
+ "__builtin_va_list", TPtr(TVoid [], [])
+ ];
+ Builtins.functions = [
+ (* Synchronization *)
+ "__builtin_fence",
+ (TVoid [], [], false);
+ (* Integer arithmetic *)
+ "__builtin_bswap",
+ (TInt(IUInt, []), [TInt(IUInt, [])], false);
+ "__builtin_bswap64",
+ (TInt(IULongLong, []), [TInt(IULongLong, [])], false);
+ "__builtin_bswap32",
+ (TInt(IUInt, []), [TInt(IUInt, [])], false);
+ "__builtin_bswap16",
+ (TInt(IUShort, []), [TInt(IUShort, [])], false);
+ (* Float arithmetic *)
+ "__builtin_fmadd",
+ (TFloat(FDouble, []),
+ [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])],
+ false);
+ "__builtin_fmsub",
+ (TFloat(FDouble, []),
+ [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])],
+ false);
+ "__builtin_fnmadd",
+ (TFloat(FDouble, []),
+ [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])],
+ false);
+ "__builtin_fnmsub",
+ (TFloat(FDouble, []),
+ [TFloat(FDouble, []); TFloat(FDouble, []); TFloat(FDouble, [])],
+ false);
+ "__builtin_fsqrt",
+ (TFloat(FDouble, []), [TFloat(FDouble, [])], false);
+ "__builtin_fmax",
+ (TFloat(FDouble, []), [TFloat(FDouble, []); TFloat(FDouble, [])], false);
+ "__builtin_fmin",
+ (TFloat(FDouble, []), [TFloat(FDouble, []); TFloat(FDouble, [])], false);
+ ]
+}
+
+let va_list_type = TPtr(TVoid [], []) (* to check! *)
+let size_va_list = if Archi.ptr64 then 8 else 4
+let va_list_scalar = true
+
+(* Expand memory references inside extended asm statements. Used in C2C. *)
+
+let asm_mem_argument arg = Printf.sprintf "0(%s)" arg
diff --git a/riscV/CombineOp.v b/riscV/CombineOp.v
new file mode 100644
index 00000000..6236f38f
--- /dev/null
+++ b/riscV/CombineOp.v
@@ -0,0 +1,138 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Recognition of combined operations, addressing modes and conditions
+ during the [CSE] phase. *)
+
+Require Import Coqlib.
+Require Import AST.
+Require Import Integers.
+Require Import Op.
+Require Import CSEdomain.
+
+Section COMBINE.
+
+Variable get: valnum -> option rhs.
+
+Function combine_compimm_ne_0 (x: valnum) : option(condition * list valnum) :=
+ match get x with
+ | Some(Op (Ocmp c) ys) => Some (c, ys)
+ | _ => None
+ end.
+
+Function combine_compimm_eq_0 (x: valnum) : option(condition * list valnum) :=
+ match get x with
+ | Some(Op (Ocmp c) ys) => Some (negate_condition c, ys)
+ | _ => None
+ end.
+
+Function combine_compimm_eq_1 (x: valnum) : option(condition * list valnum) :=
+ match get x with
+ | Some(Op (Ocmp c) ys) => Some (c, ys)
+ | _ => None
+ end.
+
+Function combine_compimm_ne_1 (x: valnum) : option(condition * list valnum) :=
+ match get x with
+ | Some(Op (Ocmp c) ys) => Some (negate_condition c, ys)
+ | _ => None
+ end.
+
+Function combine_cond (cond: condition) (args: list valnum) : option(condition * list valnum) :=
+ match cond, args with
+ | Ccompimm Cne n, x::nil =>
+ if Int.eq_dec n Int.zero then combine_compimm_ne_0 x
+ else if Int.eq_dec n Int.one then combine_compimm_ne_1 x
+ else None
+ | Ccompimm Ceq n, x::nil =>
+ if Int.eq_dec n Int.zero then combine_compimm_eq_0 x
+ else if Int.eq_dec n Int.one then combine_compimm_eq_1 x
+ else None
+ | Ccompuimm Cne n, x::nil =>
+ if Int.eq_dec n Int.zero then combine_compimm_ne_0 x
+ else if Int.eq_dec n Int.one then combine_compimm_ne_1 x
+ else None
+ | Ccompuimm Ceq n, x::nil =>
+ if Int.eq_dec n Int.zero then combine_compimm_eq_0 x
+ else if Int.eq_dec n Int.one then combine_compimm_eq_1 x
+ else None
+ | _, _ => None
+ end.
+
+Function combine_addr (addr: addressing) (args: list valnum) : option(addressing * list valnum) :=
+ match addr, args with
+ | Aindexed n, x::nil =>
+ match get x with
+ | Some(Op (Oaddimm m) ys) =>
+ if Archi.ptr64 then None else Some(Aindexed (Ptrofs.add (Ptrofs.of_int m) n), ys)
+ | Some(Op (Oaddlimm m) ys) =>
+ if Archi.ptr64 then Some(Aindexed (Ptrofs.add (Ptrofs.of_int64 m) n), ys) else None
+ | _ => None
+ end
+ | _, _ => None
+ end.
+
+Function combine_op (op: operation) (args: list valnum) : option(operation * list valnum) :=
+ match op, args with
+ | Oaddimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oaddimm m) ys) => Some(Oaddimm (Int.add m n), ys)
+ | _ => None
+ end
+ | Oandimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oandimm m) ys) =>
+ Some(let p := Int.and m n in
+ if Int.eq p m then (Omove, x :: nil) else (Oandimm p, ys))
+ | _ => None
+ end
+ | Oorimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oorimm m) ys) => Some(Oorimm (Int.or m n), ys)
+ | _ => None
+ end
+ | Oxorimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oxorimm m) ys) => Some(Oxorimm (Int.xor m n), ys)
+ | _ => None
+ end
+ | Oaddlimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oaddlimm m) ys) => Some(Oaddlimm (Int64.add m n), ys)
+ | _ => None
+ end
+ | Oandlimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oandlimm m) ys) =>
+ Some(let p := Int64.and m n in
+ if Int64.eq p m then (Omove, x :: nil) else (Oandlimm p, ys))
+ | _ => None
+ end
+ | Oorlimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oorlimm m) ys) => Some(Oorlimm (Int64.or m n), ys)
+ | _ => None
+ end
+ | Oxorlimm n, x :: nil =>
+ match get x with
+ | Some(Op (Oxorlimm m) ys) => Some(Oxorlimm (Int64.xor m n), ys)
+ | _ => None
+ end
+ | Ocmp cond, _ =>
+ match combine_cond cond args with
+ | Some(cond', args') => Some(Ocmp cond', args')
+ | None => None
+ end
+ | _, _ => None
+ end.
+
+End COMBINE.
diff --git a/riscV/CombineOpproof.v b/riscV/CombineOpproof.v
new file mode 100644
index 00000000..6843da29
--- /dev/null
+++ b/riscV/CombineOpproof.v
@@ -0,0 +1,172 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Recognition of combined operations, addressing modes and conditions
+ during the [CSE] phase. *)
+
+Require Import Coqlib.
+Require Import AST.
+Require Import Integers.
+Require Import Values.
+Require Import Memory.
+Require Import Op.
+Require Import Registers.
+Require Import RTL.
+Require Import CSEdomain.
+Require Import CombineOp.
+
+Section COMBINE.
+
+Variable ge: genv.
+Variable sp: val.
+Variable m: mem.
+Variable get: valnum -> option rhs.
+Variable valu: valnum -> val.
+Hypothesis get_sound: forall v rhs, get v = Some rhs -> rhs_eval_to valu ge sp m rhs (valu v).
+
+Lemma get_op_sound:
+ forall v op vl, get v = Some (Op op vl) -> eval_operation ge sp op (map valu vl) m = Some (valu v).
+Proof.
+ intros. exploit get_sound; eauto. intros REV; inv REV; auto.
+Qed.
+
+Ltac UseGetSound :=
+ match goal with
+ | [ H: get _ = Some _ |- _ ] =>
+ let x := fresh "EQ" in (generalize (get_op_sound _ _ _ H); intros x; simpl in x; FuncInv)
+ end.
+
+Lemma combine_compimm_ne_0_sound:
+ forall x cond args,
+ combine_compimm_ne_0 get x = Some(cond, args) ->
+ eval_condition cond (map valu args) m = Val.cmp_bool Cne (valu x) (Vint Int.zero) /\
+ eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Cne (valu x) (Vint Int.zero).
+Proof.
+ intros until args. functional induction (combine_compimm_ne_0 get x); intros EQ; inv EQ.
+ (* of cmp *)
+ UseGetSound. rewrite <- H.
+ destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto.
+Qed.
+
+Lemma combine_compimm_eq_0_sound:
+ forall x cond args,
+ combine_compimm_eq_0 get x = Some(cond, args) ->
+ eval_condition cond (map valu args) m = Val.cmp_bool Ceq (valu x) (Vint Int.zero) /\
+ eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Ceq (valu x) (Vint Int.zero).
+Proof.
+ intros until args. functional induction (combine_compimm_eq_0 get x); intros EQ; inv EQ.
+ (* of cmp *)
+ UseGetSound. rewrite <- H.
+ rewrite eval_negate_condition.
+ destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto.
+Qed.
+
+Lemma combine_compimm_eq_1_sound:
+ forall x cond args,
+ combine_compimm_eq_1 get x = Some(cond, args) ->
+ eval_condition cond (map valu args) m = Val.cmp_bool Ceq (valu x) (Vint Int.one) /\
+ eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Ceq (valu x) (Vint Int.one).
+Proof.
+ intros until args. functional induction (combine_compimm_eq_1 get x); intros EQ; inv EQ.
+ (* of cmp *)
+ UseGetSound. rewrite <- H.
+ destruct (eval_condition cond (map valu args) m); simpl; auto. destruct b; auto.
+Qed.
+
+Lemma combine_compimm_ne_1_sound:
+ forall x cond args,
+ combine_compimm_ne_1 get x = Some(cond, args) ->
+ eval_condition cond (map valu args) m = Val.cmp_bool Cne (valu x) (Vint Int.one) /\
+ eval_condition cond (map valu args) m = Val.cmpu_bool (Mem.valid_pointer m) Cne (valu x) (Vint Int.one).
+Proof.
+ intros until args. functional induction (combine_compimm_ne_1 get x); intros EQ; inv EQ.
+ (* of cmp *)
+ UseGetSound. rewrite <- H.
+ rewrite eval_negate_condition.
+ destruct (eval_condition c (map valu args) m); simpl; auto. destruct b; auto.
+Qed.
+
+Theorem combine_cond_sound:
+ forall cond args cond' args',
+ combine_cond get cond args = Some(cond', args') ->
+ eval_condition cond' (map valu args') m = eval_condition cond (map valu args) m.
+Proof.
+ intros. functional inversion H; subst.
+ (* compimm ne zero *)
+ - simpl; eapply combine_compimm_ne_0_sound; eauto.
+ (* compimm ne one *)
+ - simpl; eapply combine_compimm_ne_1_sound; eauto.
+ (* compimm eq zero *)
+ - simpl; eapply combine_compimm_eq_0_sound; eauto.
+ (* compimm eq one *)
+ - simpl; eapply combine_compimm_eq_1_sound; eauto.
+ (* compuimm ne zero *)
+ - simpl; eapply combine_compimm_ne_0_sound; eauto.
+ (* compuimm ne one *)
+ - simpl; eapply combine_compimm_ne_1_sound; eauto.
+ (* compuimm eq zero *)
+ - simpl; eapply combine_compimm_eq_0_sound; eauto.
+ (* compuimm eq one *)
+ - simpl; eapply combine_compimm_eq_1_sound; eauto.
+Qed.
+
+Theorem combine_addr_sound:
+ forall addr args addr' args',
+ combine_addr get addr args = Some(addr', args') ->
+ eval_addressing ge sp addr' (map valu args') = eval_addressing ge sp addr (map valu args).
+Proof.
+ intros. functional inversion H; subst.
+- (* indexed - addimm *)
+ UseGetSound. simpl. rewrite <- H0. destruct v; auto. simpl; rewrite H7; simpl.
+ rewrite Ptrofs.add_assoc. auto.
+- (* indexed - addimml *)
+ UseGetSound. simpl. rewrite <- H0. destruct v; auto. simpl; rewrite H7; simpl.
+ rewrite Ptrofs.add_assoc. auto.
+Qed.
+
+Theorem combine_op_sound:
+ forall op args op' args',
+ combine_op get op args = Some(op', args') ->
+ eval_operation ge sp op' (map valu args') m = eval_operation ge sp op (map valu args) m.
+Proof.
+ intros. functional inversion H; subst.
+ (* addimm - addimm *)
+ - UseGetSound. FuncInv. simpl.
+ rewrite <- H0. rewrite Val.add_assoc. auto.
+ (* andimm - andimm *)
+ - UseGetSound; simpl.
+ generalize (Int.eq_spec p m0); rewrite H7; intros.
+ rewrite <- H0. rewrite Val.and_assoc. simpl. fold p. rewrite H1. auto.
+ - UseGetSound; simpl.
+ rewrite <- H0. rewrite Val.and_assoc. auto.
+ (* orimm - orimm *)
+ - UseGetSound. simpl. rewrite <- H0. rewrite Val.or_assoc. auto.
+ (* xorimm - xorimm *)
+ - UseGetSound. simpl. rewrite <- H0. rewrite Val.xor_assoc. auto.
+ (* addlimm - addlimm *)
+ - UseGetSound. FuncInv. simpl.
+ rewrite <- H0. rewrite Val.addl_assoc. auto.
+ (* andlimm - andlimm *)
+ - UseGetSound; simpl.
+ generalize (Int64.eq_spec p m0); rewrite H7; intros.
+ rewrite <- H0. rewrite Val.andl_assoc. simpl. fold p. rewrite H1. auto.
+ - UseGetSound; simpl.
+ rewrite <- H0. rewrite Val.andl_assoc. auto.
+ (* orlimm - orlimm *)
+ - UseGetSound. simpl. rewrite <- H0. rewrite Val.orl_assoc. auto.
+ (* xorlimm - xorlimm *)
+ - UseGetSound. simpl. rewrite <- H0. rewrite Val.xorl_assoc. auto.
+ (* cmp *)
+ - simpl. decEq; decEq. eapply combine_cond_sound; eauto.
+Qed.
+
+End COMBINE.
diff --git a/riscV/ConstpropOp.vp b/riscV/ConstpropOp.vp
new file mode 100644
index 00000000..37241816
--- /dev/null
+++ b/riscV/ConstpropOp.vp
@@ -0,0 +1,291 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris *)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Strength reduction for operators and conditions.
+ This is the machine-dependent part of [Constprop]. *)
+
+Require Archi.
+Require Import Coqlib Compopts.
+Require Import AST Integers Floats.
+Require Import Op Registers.
+Require Import ValueDomain.
+
+(** * Converting known values to constants *)
+
+Definition const_for_result (a: aval) : option operation :=
+ match a with
+ | I n => Some(Ointconst n)
+ | L n => if Archi.ptr64 then Some(Olongconst n) else None
+ | F n => if Compopts.generate_float_constants tt then Some(Ofloatconst n) else None
+ | FS n => if Compopts.generate_float_constants tt then Some(Osingleconst n) else None
+ | Ptr(Gl id ofs) => Some(Oaddrsymbol id ofs)
+ | Ptr(Stk ofs) => Some(Oaddrstack ofs)
+ | _ => None
+ end.
+
+(** * Operator strength reduction *)
+
+(** We now define auxiliary functions for strength reduction of
+ operators and addressing modes: replacing an operator with a cheaper
+ one if some of its arguments are statically known. These are again
+ large pattern-matchings expressed in indirect style. *)
+
+Nondetfunction cond_strength_reduction
+ (cond: condition) (args: list reg) (vl: list aval) :=
+ match cond, args, vl with
+ | Ccomp c, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
+ (Ccompimm (swap_comparison c) n1, r2 :: nil)
+ | Ccomp c, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
+ (Ccompimm c n2, r1 :: nil)
+ | Ccompu c, r1 :: r2 :: nil, I n1 :: v2 :: nil =>
+ (Ccompuimm (swap_comparison c) n1, r2 :: nil)
+ | Ccompu c, r1 :: r2 :: nil, v1 :: I n2 :: nil =>
+ (Ccompuimm c n2, r1 :: nil)
+ | Ccompl c, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
+ (Ccomplimm (swap_comparison c) n1, r2 :: nil)
+ | Ccompl c, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
+ (Ccomplimm c n2, r1 :: nil)
+ | Ccomplu c, r1 :: r2 :: nil, L n1 :: v2 :: nil =>
+ (Ccompluimm (swap_comparison c) n1, r2 :: nil)
+ | Ccomplu c, r1 :: r2 :: nil, v1 :: L n2 :: nil =>
+ (Ccompluimm c n2, r1 :: nil)
+ | _, _, _ =>
+ (cond, args)
+ end.
+
+Definition make_cmp_base (c: condition) (args: list reg) (vl: list aval) :=
+ let (c', args') := cond_strength_reduction c args vl in (Ocmp c', args').
+
+Nondetfunction make_cmp (c: condition) (args: list reg) (vl: list aval) :=
+ match c, args, vl with
+ | Ccompimm Ceq n, r1 :: nil, v1 :: nil =>
+ if Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1) then (Omove, r1 :: nil)
+ else if Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1) then (Oxorimm Int.one, r1 :: nil)
+ else make_cmp_base c args vl
+ | Ccompimm Cne n, r1 :: nil, v1 :: nil =>
+ if Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1) then (Omove, r1 :: nil)
+ else if Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1) then (Oxorimm Int.one, r1 :: nil)
+ else make_cmp_base c args vl
+ | _, _, _ =>
+ make_cmp_base c args vl
+ end.
+
+Definition make_addimm (n: int) (r: reg) :=
+ if Int.eq n Int.zero
+ then (Omove, r :: nil)
+ else (Oaddimm n, r :: nil).
+
+Definition make_shlimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int.iwordsize then (Oshlimm n, r1 :: nil)
+ else (Oshl, r1 :: r2 :: nil).
+
+Definition make_shrimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int.iwordsize then (Oshrimm n, r1 :: nil)
+ else (Oshr, r1 :: r2 :: nil).
+
+Definition make_shruimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int.iwordsize then (Oshruimm n, r1 :: nil)
+ else (Oshru, r1 :: r2 :: nil).
+
+Definition make_mulimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then
+ (Ointconst Int.zero, nil)
+ else if Int.eq n Int.one then
+ (Omove, r1 :: nil)
+ else
+ match Int.is_power2 n with
+ | Some l => (Oshlimm l, r1 :: nil)
+ | None => (Omul, r1 :: r2 :: nil)
+ end.
+
+Definition make_andimm (n: int) (r: reg) (a: aval) :=
+ if Int.eq n Int.zero then (Ointconst Int.zero, nil)
+ else if Int.eq n Int.mone then (Omove, r :: nil)
+ else if match a with Uns _ m => Int.eq (Int.zero_ext m (Int.not n)) Int.zero
+ | _ => false end
+ then (Omove, r :: nil)
+ else (Oandimm n, r :: nil).
+
+Definition make_orimm (n: int) (r: reg) :=
+ if Int.eq n Int.zero then (Omove, r :: nil)
+ else if Int.eq n Int.mone then (Ointconst Int.mone, nil)
+ else (Oorimm n, r :: nil).
+
+Definition make_xorimm (n: int) (r: reg) :=
+ if Int.eq n Int.zero then (Omove, r :: nil)
+ else (Oxorimm n, r :: nil).
+
+Definition make_divimm n (r1 r2: reg) :=
+ match Int.is_power2 n with
+ | Some l => if Int.ltu l (Int.repr 31)
+ then (Oshrximm l, r1 :: nil)
+ else (Odiv, r1 :: r2 :: nil)
+ | None => (Odiv, r1 :: r2 :: nil)
+ end.
+
+Definition make_divuimm n (r1 r2: reg) :=
+ match Int.is_power2 n with
+ | Some l => (Oshruimm l, r1 :: nil)
+ | None => (Odivu, r1 :: r2 :: nil)
+ end.
+
+Definition make_moduimm n (r1 r2: reg) :=
+ match Int.is_power2 n with
+ | Some l => (Oandimm (Int.sub n Int.one), r1 :: nil)
+ | None => (Omodu, r1 :: r2 :: nil)
+ end.
+
+Definition make_addlimm (n: int64) (r: reg) :=
+ if Int64.eq n Int64.zero
+ then (Omove, r :: nil)
+ else (Oaddlimm n, r :: nil).
+
+Definition make_shllimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int64.iwordsize' then (Oshllimm n, r1 :: nil)
+ else (Oshll, r1 :: r2 :: nil).
+
+Definition make_shrlimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int64.iwordsize' then (Oshrlimm n, r1 :: nil)
+ else (Oshrl, r1 :: r2 :: nil).
+
+Definition make_shrluimm (n: int) (r1 r2: reg) :=
+ if Int.eq n Int.zero then (Omove, r1 :: nil)
+ else if Int.ltu n Int64.iwordsize' then (Oshrluimm n, r1 :: nil)
+ else (Oshrlu, r1 :: r2 :: nil).
+
+Definition make_mullimm (n: int64) (r1 r2: reg) :=
+ if Int64.eq n Int64.zero then
+ (Olongconst Int64.zero, nil)
+ else if Int64.eq n Int64.one then
+ (Omove, r1 :: nil)
+ else
+ match Int64.is_power2' n with
+ | Some l => (Oshllimm l, r1 :: nil)
+ | None => (Omull, r1 :: r2 :: nil)
+ end.
+
+Definition make_andlimm (n: int64) (r: reg) (a: aval) :=
+ if Int64.eq n Int64.zero then (Olongconst Int64.zero, nil)
+ else if Int64.eq n Int64.mone then (Omove, r :: nil)
+ else (Oandlimm n, r :: nil).
+
+Definition make_orlimm (n: int64) (r: reg) :=
+ if Int64.eq n Int64.zero then (Omove, r :: nil)
+ else if Int64.eq n Int64.mone then (Olongconst Int64.mone, nil)
+ else (Oorlimm n, r :: nil).
+
+Definition make_xorlimm (n: int64) (r: reg) :=
+ if Int64.eq n Int64.zero then (Omove, r :: nil)
+ else (Oxorlimm n, r :: nil).
+
+Definition make_divlimm n (r1 r2: reg) :=
+ match Int64.is_power2' n with
+ | Some l => if Int.ltu l (Int.repr 63)
+ then (Oshrxlimm l, r1 :: nil)
+ else (Odivl, r1 :: r2 :: nil)
+ | None => (Odivl, r1 :: r2 :: nil)
+ end.
+
+Definition make_divluimm n (r1 r2: reg) :=
+ match Int64.is_power2' n with
+ | Some l => (Oshrluimm l, r1 :: nil)
+ | None => (Odivlu, r1 :: r2 :: nil)
+ end.
+
+Definition make_modluimm n (r1 r2: reg) :=
+ match Int64.is_power2 n with
+ | Some l => (Oandlimm (Int64.sub n Int64.one), r1 :: nil)
+ | None => (Omodlu, r1 :: r2 :: nil)
+ end.
+
+Definition make_mulfimm (n: float) (r r1 r2: reg) :=
+ if Float.eq_dec n (Float.of_int (Int.repr 2))
+ then (Oaddf, r :: r :: nil)
+ else (Omulf, r1 :: r2 :: nil).
+
+Definition make_mulfsimm (n: float32) (r r1 r2: reg) :=
+ if Float32.eq_dec n (Float32.of_int (Int.repr 2))
+ then (Oaddfs, r :: r :: nil)
+ else (Omulfs, r1 :: r2 :: nil).
+
+Definition make_cast8signed (r: reg) (a: aval) :=
+ if vincl a (Sgn Ptop 8) then (Omove, r :: nil) else (Ocast8signed, r :: nil).
+Definition make_cast16signed (r: reg) (a: aval) :=
+ if vincl a (Sgn Ptop 16) then (Omove, r :: nil) else (Ocast16signed, r :: nil).
+
+Nondetfunction op_strength_reduction
+ (op: operation) (args: list reg) (vl: list aval) :=
+ match op, args, vl with
+ | Ocast8signed, r1 :: nil, v1 :: nil => make_cast8signed r1 v1
+ | Ocast16signed, r1 :: nil, v1 :: nil => make_cast16signed r1 v1
+ | Oadd, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_addimm n1 r2
+ | Oadd, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_addimm n2 r1
+ | Osub, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_addimm (Int.neg n2) r1
+ | Omul, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_mulimm n1 r2 r1
+ | Omul, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_mulimm n2 r1 r2
+ | Odiv, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_divimm n2 r1 r2
+ | Odivu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_divuimm n2 r1 r2
+ | Omodu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_moduimm n2 r1 r2
+ | Oand, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_andimm n1 r2 v2
+ | Oand, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_andimm n2 r1 v1
+ | Oandimm n, r1 :: nil, v1 :: nil => make_andimm n r1 v1
+ | Oor, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_orimm n1 r2
+ | Oor, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_orimm n2 r1
+ | Oxor, r1 :: r2 :: nil, I n1 :: v2 :: nil => make_xorimm n1 r2
+ | Oxor, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_xorimm n2 r1
+ | Oshl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shlimm n2 r1 r2
+ | Oshr, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrimm n2 r1 r2
+ | Oshru, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shruimm n2 r1 r2
+ | Oaddl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_addlimm n1 r2
+ | Oaddl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_addlimm n2 r1
+ | Osubl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_addlimm (Int64.neg n2) r1
+ | Omull, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_mullimm n1 r2 r1
+ | Omull, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_mullimm n2 r1 r2
+ | Odivl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_divlimm n2 r1 r2
+ | Odivlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_divluimm n2 r1 r2
+ | Omodlu, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_modluimm n2 r1 r2
+ | Oandl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_andlimm n1 r2 v2
+ | Oandl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_andlimm n2 r1 v1
+ | Oandlimm n, r1 :: nil, v1 :: nil => make_andlimm n r1 v1
+ | Oorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_orlimm n1 r2
+ | Oorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_orlimm n2 r1
+ | Oxorl, r1 :: r2 :: nil, L n1 :: v2 :: nil => make_xorlimm n1 r2
+ | Oxorl, r1 :: r2 :: nil, v1 :: L n2 :: nil => make_xorlimm n2 r1
+ | Oshll, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shllimm n2 r1 r2
+ | Oshrl, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrlimm n2 r1 r2
+ | Oshrlu, r1 :: r2 :: nil, v1 :: I n2 :: nil => make_shrluimm n2 r1 r2
+ | Ocmp c, args, vl => make_cmp c args vl
+ | Omulf, r1 :: r2 :: nil, v1 :: F n2 :: nil => make_mulfimm n2 r1 r1 r2
+ | Omulf, r1 :: r2 :: nil, F n1 :: v2 :: nil => make_mulfimm n1 r2 r1 r2
+ | Omulfs, r1 :: r2 :: nil, v1 :: FS n2 :: nil => make_mulfsimm n2 r1 r1 r2
+ | Omulfs, r1 :: r2 :: nil, FS n1 :: v2 :: nil => make_mulfsimm n1 r2 r1 r2
+ | _, _, _ => (op, args)
+ end.
+
+Nondetfunction addr_strength_reduction
+ (addr: addressing) (args: list reg) (vl: list aval) :=
+ match addr, args, vl with
+ | Aindexed n, r1 :: nil, Ptr(Gl symb n1) :: nil =>
+ if Archi.pic_code tt
+ then (addr, args)
+ else (Aglobal symb (Ptrofs.add n1 n), nil)
+ | Aindexed n, r1 :: nil, Ptr(Stk n1) :: nil =>
+ (Ainstack (Ptrofs.add n1 n), nil)
+ | _, _, _ =>
+ (addr, args)
+ end.
+
diff --git a/riscV/ConstpropOpproof.v b/riscV/ConstpropOpproof.v
new file mode 100644
index 00000000..f2e2b95e
--- /dev/null
+++ b/riscV/ConstpropOpproof.v
@@ -0,0 +1,715 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris *)
+(* *)
+(* 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 operator strength reduction. *)
+
+Require Import Coqlib Compopts.
+Require Import Integers Floats Values Memory Globalenvs Events.
+Require Import Op Registers RTL ValueDomain.
+Require Import ConstpropOp.
+
+Section STRENGTH_REDUCTION.
+
+Variable bc: block_classification.
+Variable ge: genv.
+Hypothesis GENV: genv_match bc ge.
+Variable sp: block.
+Hypothesis STACK: bc sp = BCstack.
+Variable ae: AE.t.
+Variable e: regset.
+Variable m: mem.
+Hypothesis MATCH: ematch bc e ae.
+
+Lemma match_G:
+ forall r id ofs,
+ AE.get r ae = Ptr(Gl id ofs) -> Val.lessdef e#r (Genv.symbol_address ge id ofs).
+Proof.
+ intros. apply vmatch_ptr_gl with bc; auto. rewrite <- H. apply MATCH.
+Qed.
+
+Lemma match_S:
+ forall r ofs,
+ AE.get r ae = Ptr(Stk ofs) -> Val.lessdef e#r (Vptr sp ofs).
+Proof.
+ intros. apply vmatch_ptr_stk with bc; auto. rewrite <- H. apply MATCH.
+Qed.
+
+Ltac InvApproxRegs :=
+ match goal with
+ | [ H: _ :: _ = _ :: _ |- _ ] =>
+ injection H; clear H; intros; InvApproxRegs
+ | [ H: ?v = AE.get ?r ae |- _ ] =>
+ generalize (MATCH r); rewrite <- H; clear H; intro; InvApproxRegs
+ | _ => idtac
+ end.
+
+Ltac SimplVM :=
+ match goal with
+ | [ H: vmatch _ ?v (I ?n) |- _ ] =>
+ let E := fresh in
+ assert (E: v = Vint n) by (inversion H; auto);
+ rewrite E in *; clear H; SimplVM
+ | [ H: vmatch _ ?v (L ?n) |- _ ] =>
+ let E := fresh in
+ assert (E: v = Vlong n) by (inversion H; auto);
+ rewrite E in *; clear H; SimplVM
+ | [ H: vmatch _ ?v (F ?n) |- _ ] =>
+ let E := fresh in
+ assert (E: v = Vfloat n) by (inversion H; auto);
+ rewrite E in *; clear H; SimplVM
+ | [ H: vmatch _ ?v (FS ?n) |- _ ] =>
+ let E := fresh in
+ assert (E: v = Vsingle n) by (inversion H; auto);
+ rewrite E in *; clear H; SimplVM
+ | [ H: vmatch _ ?v (Ptr(Gl ?id ?ofs)) |- _ ] =>
+ let E := fresh in
+ assert (E: Val.lessdef v (Genv.symbol_address ge id ofs)) by (eapply vmatch_ptr_gl; eauto);
+ clear H; SimplVM
+ | [ H: vmatch _ ?v (Ptr(Stk ?ofs)) |- _ ] =>
+ let E := fresh in
+ assert (E: Val.lessdef v (Vptr sp ofs)) by (eapply vmatch_ptr_stk; eauto);
+ clear H; SimplVM
+ | _ => idtac
+ end.
+
+Lemma const_for_result_correct:
+ forall a op v,
+ const_for_result a = Some op ->
+ vmatch bc v a ->
+ exists v', eval_operation ge (Vptr sp Ptrofs.zero) op nil m = Some v' /\ Val.lessdef v v'.
+Proof.
+ unfold const_for_result. generalize Archi.ptr64; intros ptr64; intros.
+ destruct a; inv H; SimplVM.
+- (* integer *)
+ exists (Vint n); auto.
+- (* long *)
+ destruct ptr64; inv H2. exists (Vlong n); auto.
+- (* float *)
+ destruct (Compopts.generate_float_constants tt); inv H2. exists (Vfloat f); auto.
+- (* single *)
+ destruct (Compopts.generate_float_constants tt); inv H2. exists (Vsingle f); auto.
+- (* pointer *)
+ destruct p; try discriminate; SimplVM.
+ + (* global *)
+ inv H2. exists (Genv.symbol_address ge id ofs); auto.
+ + (* stack *)
+ inv H2. exists (Vptr sp ofs); split; auto. simpl. rewrite Ptrofs.add_zero_l; auto.
+Qed.
+
+Lemma cond_strength_reduction_correct:
+ forall cond args vl,
+ vl = map (fun r => AE.get r ae) args ->
+ let (cond', args') := cond_strength_reduction cond args vl in
+ eval_condition cond' e##args' m = eval_condition cond e##args m.
+Proof.
+ intros until vl. unfold cond_strength_reduction.
+ case (cond_strength_reduction_match cond args vl); simpl; intros; InvApproxRegs; SimplVM.
+- apply Val.swap_cmp_bool.
+- auto.
+- apply Val.swap_cmpu_bool.
+- auto.
+- apply Val.swap_cmpl_bool.
+- auto.
+- apply Val.swap_cmplu_bool.
+- auto.
+- auto.
+Qed.
+
+Lemma make_cmp_base_correct:
+ forall c args vl,
+ vl = map (fun r => AE.get r ae) args ->
+ let (op', args') := make_cmp_base c args vl in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op' e##args' m = Some v
+ /\ Val.lessdef (Val.of_optbool (eval_condition c e##args m)) v.
+Proof.
+ intros. unfold make_cmp_base.
+ generalize (cond_strength_reduction_correct c args vl H).
+ destruct (cond_strength_reduction c args vl) as [c' args']. intros EQ.
+ econstructor; split. simpl; eauto. rewrite EQ. auto.
+Qed.
+
+Lemma make_cmp_correct:
+ forall c args vl,
+ vl = map (fun r => AE.get r ae) args ->
+ let (op', args') := make_cmp c args vl in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op' e##args' m = Some v
+ /\ Val.lessdef (Val.of_optbool (eval_condition c e##args m)) v.
+Proof.
+ intros c args vl.
+ assert (Y: forall r, vincl (AE.get r ae) (Uns Ptop 1) = true ->
+ e#r = Vundef \/ e#r = Vint Int.zero \/ e#r = Vint Int.one).
+ { intros. apply vmatch_Uns_1 with bc Ptop. eapply vmatch_ge. eapply vincl_ge; eauto. apply MATCH. }
+ unfold make_cmp. case (make_cmp_match c args vl); intros.
+- destruct (Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1)) eqn:E1.
+ simpl in H; inv H. InvBooleans. subst n.
+ exists (e#r1); split; auto. simpl.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
+ destruct (Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1)) eqn:E0.
+ simpl in H; inv H. InvBooleans. subst n.
+ exists (Val.xor e#r1 (Vint Int.one)); split; auto. simpl.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
+ apply make_cmp_base_correct; auto.
+- destruct (Int.eq_dec n Int.zero && vincl v1 (Uns Ptop 1)) eqn:E0.
+ simpl in H; inv H. InvBooleans. subst n.
+ exists (e#r1); split; auto. simpl.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
+ destruct (Int.eq_dec n Int.one && vincl v1 (Uns Ptop 1)) eqn:E1.
+ simpl in H; inv H. InvBooleans. subst n.
+ exists (Val.xor e#r1 (Vint Int.one)); split; auto. simpl.
+ exploit Y; eauto. intros [A | [A | A]]; rewrite A; simpl; auto.
+ apply make_cmp_base_correct; auto.
+- apply make_cmp_base_correct; auto.
+Qed.
+
+Lemma make_addimm_correct:
+ forall n r,
+ let (op, args) := make_addimm n r in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.add e#r (Vint n)) v.
+Proof.
+ intros. unfold make_addimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros.
+ subst. exists (e#r); split; auto.
+ destruct (e#r); simpl; auto; rewrite ?Int.add_zero, ?Ptrofs.add_zero; auto.
+ destruct Archi.ptr64; auto.
+ exists (Val.add e#r (Vint n)); split; auto.
+Qed.
+
+Lemma make_shlimm_correct:
+ forall n r1 r2,
+ e#r2 = Vint n ->
+ let (op, args) := make_shlimm n r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shl e#r1 (Vint n)) v.
+Proof.
+ intros; unfold make_shlimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
+ exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shl_zero. auto.
+ destruct (Int.ltu n Int.iwordsize).
+ econstructor; split. simpl. eauto. auto.
+ econstructor; split. simpl. eauto. rewrite H; auto.
+Qed.
+
+Lemma make_shrimm_correct:
+ forall n r1 r2,
+ e#r2 = Vint n ->
+ let (op, args) := make_shrimm n r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shr e#r1 (Vint n)) v.
+Proof.
+ intros; unfold make_shrimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
+ exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shr_zero. auto.
+ destruct (Int.ltu n Int.iwordsize).
+ econstructor; split. simpl. eauto. auto.
+ econstructor; split. simpl. eauto. rewrite H; auto.
+Qed.
+
+Lemma make_shruimm_correct:
+ forall n r1 r2,
+ e#r2 = Vint n ->
+ let (op, args) := make_shruimm n r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shru e#r1 (Vint n)) v.
+Proof.
+ intros; unfold make_shruimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
+ exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.shru_zero. auto.
+ destruct (Int.ltu n Int.iwordsize).
+ econstructor; split. simpl. eauto. auto.
+ econstructor; split. simpl. eauto. rewrite H; auto.
+Qed.
+
+Lemma make_mulimm_correct:
+ forall n r1 r2,
+ e#r2 = Vint n ->
+ let (op, args) := make_mulimm n r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mul e#r1 (Vint n)) v.
+Proof.
+ intros; unfold make_mulimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
+ exists (Vint Int.zero); split; auto. destruct (e#r1); simpl; auto. rewrite Int.mul_zero; auto.
+ predSpec Int.eq Int.eq_spec n Int.one; intros. subst.
+ exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int.mul_one; auto.
+ destruct (Int.is_power2 n) eqn:?; intros.
+ rewrite (Val.mul_pow2 e#r1 _ _ Heqo). econstructor; split. simpl; eauto. auto.
+ econstructor; split; eauto. simpl. rewrite H; auto.
+Qed.
+
+Lemma make_divimm_correct:
+ forall n r1 r2 v,
+ Val.divs e#r1 e#r2 = Some v ->
+ e#r2 = Vint n ->
+ let (op, args) := make_divimm n r1 r2 in
+ exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
+Proof.
+ intros; unfold make_divimm.
+ destruct (Int.is_power2 n) eqn:?.
+ destruct (Int.ltu i (Int.repr 31)) eqn:?.
+ exists v; split; auto. simpl. eapply Val.divs_pow2; eauto. congruence.
+ exists v; auto.
+ exists v; auto.
+Qed.
+
+Lemma make_divuimm_correct:
+ forall n r1 r2 v,
+ Val.divu e#r1 e#r2 = Some v ->
+ e#r2 = Vint n ->
+ let (op, args) := make_divuimm n r1 r2 in
+ exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
+Proof.
+ intros; unfold make_divuimm.
+ destruct (Int.is_power2 n) eqn:?.
+ econstructor; split. simpl; eauto.
+ rewrite H0 in H. erewrite Val.divu_pow2 by eauto. auto.
+ exists v; auto.
+Qed.
+
+Lemma make_moduimm_correct:
+ forall n r1 r2 v,
+ Val.modu e#r1 e#r2 = Some v ->
+ e#r2 = Vint n ->
+ let (op, args) := make_moduimm n r1 r2 in
+ exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
+Proof.
+ intros; unfold make_moduimm.
+ destruct (Int.is_power2 n) eqn:?.
+ exists v; split; auto. simpl. decEq. eapply Val.modu_pow2; eauto. congruence.
+ exists v; auto.
+Qed.
+
+Lemma make_andimm_correct:
+ forall n r x,
+ vmatch bc e#r x ->
+ let (op, args) := make_andimm n r x in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.and e#r (Vint n)) v.
+Proof.
+ intros; unfold make_andimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros.
+ subst n. exists (Vint Int.zero); split; auto. destruct (e#r); simpl; auto. rewrite Int.and_zero; auto.
+ predSpec Int.eq Int.eq_spec n Int.mone; intros.
+ subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.and_mone; auto.
+ destruct (match x with Uns _ k => Int.eq (Int.zero_ext k (Int.not n)) Int.zero
+ | _ => false end) eqn:UNS.
+ destruct x; try congruence.
+ exists (e#r); split; auto.
+ inv H; auto. simpl. replace (Int.and i n) with i; auto.
+ generalize (Int.eq_spec (Int.zero_ext n0 (Int.not n)) Int.zero); rewrite UNS; intro EQ.
+ Int.bit_solve. destruct (zlt i0 n0).
+ replace (Int.testbit n i0) with (negb (Int.testbit Int.zero i0)).
+ rewrite Int.bits_zero. simpl. rewrite andb_true_r. auto.
+ rewrite <- EQ. rewrite Int.bits_zero_ext by omega. rewrite zlt_true by auto.
+ rewrite Int.bits_not by auto. apply negb_involutive.
+ rewrite H6 by auto. auto.
+ econstructor; split; eauto. auto.
+Qed.
+
+Lemma make_orimm_correct:
+ forall n r,
+ let (op, args) := make_orimm n r in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.or e#r (Vint n)) v.
+Proof.
+ intros; unfold make_orimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros.
+ subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.or_zero; auto.
+ predSpec Int.eq Int.eq_spec n Int.mone; intros.
+ subst n. exists (Vint Int.mone); split; auto. destruct (e#r); simpl; auto. rewrite Int.or_mone; auto.
+ econstructor; split; eauto. auto.
+Qed.
+
+Lemma make_xorimm_correct:
+ forall n r,
+ let (op, args) := make_xorimm n r in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.xor e#r (Vint n)) v.
+Proof.
+ intros; unfold make_xorimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros.
+ subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int.xor_zero; auto.
+ predSpec Int.eq Int.eq_spec n Int.mone; intros.
+ subst n. exists (Val.notint e#r); split; auto.
+ econstructor; split; eauto. auto.
+Qed.
+
+Lemma make_addlimm_correct:
+ forall n r,
+ let (op, args) := make_addlimm n r in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.addl e#r (Vlong n)) v.
+Proof.
+ intros. unfold make_addlimm.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
+ subst. exists (e#r); split; auto.
+ destruct (e#r); simpl; auto; rewrite ? Int64.add_zero, ? Ptrofs.add_zero; auto.
+ destruct Archi.ptr64; auto.
+ exists (Val.addl e#r (Vlong n)); split; auto.
+Qed.
+
+Lemma make_shllimm_correct:
+ forall n r1 r2,
+ e#r2 = Vint n ->
+ let (op, args) := make_shllimm n r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shll e#r1 (Vint n)) v.
+Proof.
+ intros; unfold make_shllimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
+ exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
+ unfold Int64.shl'. rewrite Z.shiftl_0_r, Int64.repr_unsigned. auto.
+ destruct (Int.ltu n Int64.iwordsize').
+ econstructor; split. simpl. eauto. auto.
+ econstructor; split. simpl. eauto. rewrite H; auto.
+Qed.
+
+Lemma make_shrlimm_correct:
+ forall n r1 r2,
+ e#r2 = Vint n ->
+ let (op, args) := make_shrlimm n r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shrl e#r1 (Vint n)) v.
+Proof.
+ intros; unfold make_shrlimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
+ exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
+ unfold Int64.shr'. rewrite Z.shiftr_0_r, Int64.repr_signed. auto.
+ destruct (Int.ltu n Int64.iwordsize').
+ econstructor; split. simpl. eauto. auto.
+ econstructor; split. simpl. eauto. rewrite H; auto.
+Qed.
+
+Lemma make_shrluimm_correct:
+ forall n r1 r2,
+ e#r2 = Vint n ->
+ let (op, args) := make_shrluimm n r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.shrlu e#r1 (Vint n)) v.
+Proof.
+ intros; unfold make_shrluimm.
+ predSpec Int.eq Int.eq_spec n Int.zero; intros. subst.
+ exists (e#r1); split; auto. destruct (e#r1); simpl; auto.
+ unfold Int64.shru'. rewrite Z.shiftr_0_r, Int64.repr_unsigned. auto.
+ destruct (Int.ltu n Int64.iwordsize').
+ econstructor; split. simpl. eauto. auto.
+ econstructor; split. simpl. eauto. rewrite H; auto.
+Qed.
+
+Lemma make_mullimm_correct:
+ forall n r1 r2,
+ e#r2 = Vlong n ->
+ let (op, args) := make_mullimm n r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mull e#r1 (Vlong n)) v.
+Proof.
+ intros; unfold make_mullimm.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero; intros. subst.
+ exists (Vlong Int64.zero); split; auto. destruct (e#r1); simpl; auto. rewrite Int64.mul_zero; auto.
+ predSpec Int64.eq Int64.eq_spec n Int64.one; intros. subst.
+ exists (e#r1); split; auto. destruct (e#r1); simpl; auto. rewrite Int64.mul_one; auto.
+ destruct (Int64.is_power2' n) eqn:?; intros.
+ exists (Val.shll e#r1 (Vint i)); split; auto.
+ destruct (e#r1); simpl; auto.
+ erewrite Int64.is_power2'_range by eauto.
+ erewrite Int64.mul_pow2' by eauto. auto.
+ econstructor; split; eauto. simpl; rewrite H; auto.
+Qed.
+
+Lemma make_divlimm_correct:
+ forall n r1 r2 v,
+ Val.divls e#r1 e#r2 = Some v ->
+ e#r2 = Vlong n ->
+ let (op, args) := make_divlimm n r1 r2 in
+ exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
+Proof.
+ intros; unfold make_divlimm.
+ destruct (Int64.is_power2' n) eqn:?. destruct (Int.ltu i (Int.repr 63)) eqn:?.
+ rewrite H0 in H. econstructor; split. simpl; eauto. eapply Val.divls_pow2; eauto. auto.
+ exists v; auto.
+ exists v; auto.
+Qed.
+
+Lemma make_divluimm_correct:
+ forall n r1 r2 v,
+ Val.divlu e#r1 e#r2 = Some v ->
+ e#r2 = Vlong n ->
+ let (op, args) := make_divluimm n r1 r2 in
+ exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
+Proof.
+ intros; unfold make_divluimm.
+ destruct (Int64.is_power2' n) eqn:?.
+ econstructor; split. simpl; eauto.
+ rewrite H0 in H. destruct (e#r1); inv H. destruct (Int64.eq n Int64.zero); inv H2.
+ simpl.
+ erewrite Int64.is_power2'_range by eauto.
+ erewrite Int64.divu_pow2' by eauto. auto.
+ exists v; auto.
+Qed.
+
+Lemma make_modluimm_correct:
+ forall n r1 r2 v,
+ Val.modlu e#r1 e#r2 = Some v ->
+ e#r2 = Vlong n ->
+ let (op, args) := make_modluimm n r1 r2 in
+ exists w, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some w /\ Val.lessdef v w.
+Proof.
+ intros; unfold make_modluimm.
+ destruct (Int64.is_power2 n) eqn:?.
+ exists v; split; auto. simpl. decEq.
+ rewrite H0 in H. destruct (e#r1); inv H. destruct (Int64.eq n Int64.zero); inv H2.
+ simpl. erewrite Int64.modu_and by eauto. auto.
+ exists v; auto.
+Qed.
+
+Lemma make_andlimm_correct:
+ forall n r x,
+ let (op, args) := make_andlimm n r x in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.andl e#r (Vlong n)) v.
+Proof.
+ intros; unfold make_andlimm.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
+ subst n. exists (Vlong Int64.zero); split; auto. destruct (e#r); simpl; auto. rewrite Int64.and_zero; auto.
+ predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
+ subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.and_mone; auto.
+ econstructor; split; eauto. auto.
+Qed.
+
+Lemma make_orlimm_correct:
+ forall n r,
+ let (op, args) := make_orlimm n r in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.orl e#r (Vlong n)) v.
+Proof.
+ intros; unfold make_orlimm.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
+ subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.or_zero; auto.
+ predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
+ subst n. exists (Vlong Int64.mone); split; auto. destruct (e#r); simpl; auto. rewrite Int64.or_mone; auto.
+ econstructor; split; eauto. auto.
+Qed.
+
+Lemma make_xorlimm_correct:
+ forall n r,
+ let (op, args) := make_xorlimm n r in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.xorl e#r (Vlong n)) v.
+Proof.
+ intros; unfold make_xorlimm.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero; intros.
+ subst n. exists (e#r); split; auto. destruct (e#r); simpl; auto. rewrite Int64.xor_zero; auto.
+ predSpec Int64.eq Int64.eq_spec n Int64.mone; intros.
+ subst n. exists (Val.notl e#r); split; auto.
+ econstructor; split; eauto. auto.
+Qed.
+
+Lemma make_mulfimm_correct:
+ forall n r1 r2,
+ e#r2 = Vfloat n ->
+ let (op, args) := make_mulfimm n r1 r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mulf e#r1 e#r2) v.
+Proof.
+ intros; unfold make_mulfimm.
+ destruct (Float.eq_dec n (Float.of_int (Int.repr 2))); intros.
+ simpl. econstructor; split. eauto. rewrite H; subst n.
+ destruct (e#r1); simpl; auto. rewrite Float.mul2_add; auto.
+ simpl. econstructor; split; eauto.
+Qed.
+
+Lemma make_mulfimm_correct_2:
+ forall n r1 r2,
+ e#r1 = Vfloat n ->
+ let (op, args) := make_mulfimm n r2 r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mulf e#r1 e#r2) v.
+Proof.
+ intros; unfold make_mulfimm.
+ destruct (Float.eq_dec n (Float.of_int (Int.repr 2))); intros.
+ simpl. econstructor; split. eauto. rewrite H; subst n.
+ destruct (e#r2); simpl; auto. rewrite Float.mul2_add; auto.
+ rewrite Float.mul_commut; auto.
+ simpl. econstructor; split; eauto.
+Qed.
+
+Lemma make_mulfsimm_correct:
+ forall n r1 r2,
+ e#r2 = Vsingle n ->
+ let (op, args) := make_mulfsimm n r1 r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mulfs e#r1 e#r2) v.
+Proof.
+ intros; unfold make_mulfsimm.
+ destruct (Float32.eq_dec n (Float32.of_int (Int.repr 2))); intros.
+ simpl. econstructor; split. eauto. rewrite H; subst n.
+ destruct (e#r1); simpl; auto. rewrite Float32.mul2_add; auto.
+ simpl. econstructor; split; eauto.
+Qed.
+
+Lemma make_mulfsimm_correct_2:
+ forall n r1 r2,
+ e#r1 = Vsingle n ->
+ let (op, args) := make_mulfsimm n r2 r1 r2 in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.mulfs e#r1 e#r2) v.
+Proof.
+ intros; unfold make_mulfsimm.
+ destruct (Float32.eq_dec n (Float32.of_int (Int.repr 2))); intros.
+ simpl. econstructor; split. eauto. rewrite H; subst n.
+ destruct (e#r2); simpl; auto. rewrite Float32.mul2_add; auto.
+ rewrite Float32.mul_commut; auto.
+ simpl. econstructor; split; eauto.
+Qed.
+
+Lemma make_cast8signed_correct:
+ forall r x,
+ vmatch bc e#r x ->
+ let (op, args) := make_cast8signed r x in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.sign_ext 8 e#r) v.
+Proof.
+ intros; unfold make_cast8signed. destruct (vincl x (Sgn Ptop 8)) eqn:INCL.
+ exists e#r; split; auto.
+ assert (V: vmatch bc e#r (Sgn Ptop 8)).
+ { eapply vmatch_ge; eauto. apply vincl_ge; auto. }
+ inv V; simpl; auto. rewrite is_sgn_sign_ext in H4 by auto. rewrite H4; auto.
+ econstructor; split; simpl; eauto.
+Qed.
+
+Lemma make_cast16signed_correct:
+ forall r x,
+ vmatch bc e#r x ->
+ let (op, args) := make_cast16signed r x in
+ exists v, eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v /\ Val.lessdef (Val.sign_ext 16 e#r) v.
+Proof.
+ intros; unfold make_cast16signed. destruct (vincl x (Sgn Ptop 16)) eqn:INCL.
+ exists e#r; split; auto.
+ assert (V: vmatch bc e#r (Sgn Ptop 16)).
+ { eapply vmatch_ge; eauto. apply vincl_ge; auto. }
+ inv V; simpl; auto. rewrite is_sgn_sign_ext in H4 by auto. rewrite H4; auto.
+ econstructor; split; simpl; eauto.
+Qed.
+
+Lemma op_strength_reduction_correct:
+ forall op args vl v,
+ vl = map (fun r => AE.get r ae) args ->
+ eval_operation ge (Vptr sp Ptrofs.zero) op e##args m = Some v ->
+ let (op', args') := op_strength_reduction op args vl in
+ exists w, eval_operation ge (Vptr sp Ptrofs.zero) op' e##args' m = Some w /\ Val.lessdef v w.
+Proof.
+ intros until v; unfold op_strength_reduction;
+ case (op_strength_reduction_match op args vl); simpl; intros.
+- (* cast8signed *)
+ InvApproxRegs; SimplVM; inv H0. apply make_cast8signed_correct; auto.
+- (* cast16signed *)
+ InvApproxRegs; SimplVM; inv H0. apply make_cast16signed_correct; auto.
+- (* add 1 *)
+ rewrite Val.add_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_addimm_correct; auto.
+- (* add 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_addimm_correct; auto.
+- (* sub *)
+ InvApproxRegs; SimplVM; inv H0. rewrite Val.sub_add_opp. apply make_addimm_correct; auto.
+- (* mul 1 *)
+ rewrite Val.mul_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_mulimm_correct; auto.
+- (* mul 2*)
+ InvApproxRegs; SimplVM; inv H0. apply make_mulimm_correct; auto.
+- (* divs *)
+ assert (e#r2 = Vint n2). clear H0. InvApproxRegs; SimplVM; auto.
+ apply make_divimm_correct; auto.
+- (* divu *)
+ assert (e#r2 = Vint n2). clear H0. InvApproxRegs; SimplVM; auto.
+ apply make_divuimm_correct; auto.
+- (* modu *)
+ assert (e#r2 = Vint n2). clear H0. InvApproxRegs; SimplVM; auto.
+ apply make_moduimm_correct; auto.
+- (* and 1 *)
+ rewrite Val.and_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_andimm_correct; auto.
+- (* and 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_andimm_correct; auto.
+- (* andimm *)
+ inv H; inv H0. apply make_andimm_correct; auto.
+- (* or 1 *)
+ rewrite Val.or_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_orimm_correct; auto.
+- (* or 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_orimm_correct; auto.
+- (* xor 1 *)
+ rewrite Val.xor_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_xorimm_correct; auto.
+- (* xor 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_xorimm_correct; auto.
+- (* shl *)
+ InvApproxRegs; SimplVM; inv H0. apply make_shlimm_correct; auto.
+- (* shr *)
+ InvApproxRegs; SimplVM; inv H0. apply make_shrimm_correct; auto.
+- (* shru *)
+ InvApproxRegs; SimplVM; inv H0. apply make_shruimm_correct; auto.
+- (* addl 1 *)
+ rewrite Val.addl_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_addlimm_correct; auto.
+- (* addl 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_addlimm_correct; auto.
+- (* subl *)
+ InvApproxRegs; SimplVM; inv H0.
+ replace (Val.subl e#r1 (Vlong n2)) with (Val.addl e#r1 (Vlong (Int64.neg n2))).
+ apply make_addlimm_correct; auto.
+ unfold Val.addl, Val.subl. destruct Archi.ptr64 eqn:SF, e#r1; auto.
+ rewrite Int64.sub_add_opp; auto.
+ rewrite Ptrofs.sub_add_opp. do 2 f_equal. auto with ptrofs.
+ rewrite Int64.sub_add_opp; auto.
+- (* mull 1 *)
+ rewrite Val.mull_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_mullimm_correct; auto.
+- (* mull 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_mullimm_correct; auto.
+- (* divl *)
+ assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
+ apply make_divlimm_correct; auto.
+- (* divlu *)
+ assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
+ apply make_divluimm_correct; auto.
+- (* modlu *)
+ assert (e#r2 = Vlong n2). clear H0. InvApproxRegs; SimplVM; auto.
+ apply make_modluimm_correct; auto.
+- (* andl 1 *)
+ rewrite Val.andl_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_andlimm_correct; auto.
+- (* andl 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_andlimm_correct; auto.
+- (* andlimm *)
+ inv H; inv H0. apply make_andlimm_correct; auto.
+- (* orl 1 *)
+ rewrite Val.orl_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_orlimm_correct; auto.
+- (* orl 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_orlimm_correct; auto.
+- (* xorl 1 *)
+ rewrite Val.xorl_commut in H0. InvApproxRegs; SimplVM; inv H0. apply make_xorlimm_correct; auto.
+- (* xorl 2 *)
+ InvApproxRegs; SimplVM; inv H0. apply make_xorlimm_correct; auto.
+- (* shll *)
+ InvApproxRegs; SimplVM; inv H0. apply make_shllimm_correct; auto.
+- (* shrl *)
+ InvApproxRegs; SimplVM; inv H0. apply make_shrlimm_correct; auto.
+- (* shrlu *)
+ InvApproxRegs; SimplVM; inv H0. apply make_shrluimm_correct; auto.
+- (* cond *)
+ inv H0. apply make_cmp_correct; auto.
+- (* mulf 1 *)
+ InvApproxRegs; SimplVM; inv H0. rewrite <- H2. apply make_mulfimm_correct; auto.
+- (* mulf 2 *)
+ InvApproxRegs; SimplVM; inv H0. fold (Val.mulf (Vfloat n1) e#r2).
+ rewrite <- H2. apply make_mulfimm_correct_2; auto.
+- (* mulfs 1 *)
+ InvApproxRegs; SimplVM; inv H0. rewrite <- H2. apply make_mulfsimm_correct; auto.
+- (* mulfs 2 *)
+ InvApproxRegs; SimplVM; inv H0. fold (Val.mulfs (Vsingle n1) e#r2).
+ rewrite <- H2. apply make_mulfsimm_correct_2; auto.
+- (* default *)
+ exists v; auto.
+Qed.
+
+Lemma addr_strength_reduction_correct:
+ forall addr args vl res,
+ vl = map (fun r => AE.get r ae) args ->
+ eval_addressing ge (Vptr sp Ptrofs.zero) addr e##args = Some res ->
+ let (addr', args') := addr_strength_reduction addr args vl in
+ exists res', eval_addressing ge (Vptr sp Ptrofs.zero) addr' e##args' = Some res' /\ Val.lessdef res res'.
+Proof.
+ intros until res. unfold addr_strength_reduction.
+ destruct (addr_strength_reduction_match addr args vl); simpl;
+ intros VL EA; InvApproxRegs; SimplVM; try (inv EA).
+- destruct (Archi.pic_code tt).
++ exists (Val.offset_ptr e#r1 n); auto.
++ simpl. rewrite Genv.shift_symbol_address. econstructor; split; eauto.
+ inv H0; simpl; auto.
+- rewrite Ptrofs.add_zero_l. econstructor; split; eauto.
+ change (Vptr sp (Ptrofs.add n1 n)) with (Val.offset_ptr (Vptr sp n1) n).
+ inv H0; simpl; auto.
+- exists res; auto.
+Qed.
+
+End STRENGTH_REDUCTION.
diff --git a/riscV/Conventions1.v b/riscV/Conventions1.v
new file mode 100644
index 00000000..922f878a
--- /dev/null
+++ b/riscV/Conventions1.v
@@ -0,0 +1,436 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Function calling conventions and other conventions regarding the use of
+ machine registers and stack slots. *)
+
+Require Import Coqlib Decidableplus.
+Require Import AST Machregs Locations.
+
+(** * Classification of machine registers *)
+
+(** Machine registers (type [mreg] in module [Locations]) are divided in
+ the following groups:
+- Callee-save registers, whose value is preserved across a function call.
+- Caller-save registers that can be modified during a function call.
+
+ We follow the RISC-V application binary interface (ABI) in our choice
+ of callee- and caller-save registers.
+*)
+
+Definition is_callee_save (r: mreg) : bool :=
+ match r with
+ | R5 | R6 | R7 => false
+ | R8 | R9 => true
+ | R10 | R11 | R12 | R13 | R14 | R15 | R16 | R17 => false
+ | R18 | R19 | R20 | R21 | R22 | R23 | R24 | R25 | R26 | R27 => true
+ | R28 | R29 | R30 => false
+ | F0 | F1 | F2 | F3 | F4 | F5 | F6 | F7 => false
+ | F8 | F9 => true
+ | F10 | F11 | F12 | F13 | F14 | F15 | F16 | F17 => false
+ | F18 | F19 | F20 | F21 | F22 | F23 | F24 | F25 | F26 | F27 => true
+ | F28 | F29 | F30 | F31 => false
+ end.
+
+Definition int_caller_save_regs :=
+ R5 :: R6 :: R7 ::
+ R10 :: R11 :: R12 :: R13 :: R14 :: R15 :: R16 :: R17 ::
+ R28 :: R29 :: R30 ::
+ nil.
+
+Definition float_caller_save_regs :=
+ F0 :: F1 :: F2 :: F3 :: F4 :: F5 :: F6 :: F7 ::
+ F10 :: F11 :: F12 :: F13 :: F14 :: F15 :: F16 :: F17 ::
+ F28 :: F29 :: F30 :: F31 ::
+ nil.
+
+Definition int_callee_save_regs :=
+ R8 :: R9 ::
+ R18 :: R19 :: R20 :: R21 :: R22 :: R23 :: R24 :: R25 :: R26 :: R27 ::
+ nil.
+
+Definition float_callee_save_regs :=
+ F8 :: F9 ::
+ F18 :: F19 :: F20 :: F21 :: F22 :: F23 :: F24 :: F25 :: F26 :: F27 ::
+ nil.
+
+Definition destroyed_at_call :=
+ List.filter (fun r => negb (is_callee_save r)) all_mregs.
+
+Definition dummy_int_reg := R6. (**r Used in [Coloring]. *)
+Definition dummy_float_reg := F0 . (**r Used in [Coloring]. *)
+
+Definition is_float_reg (r: mreg) :=
+ match r with
+ | R5 | R6 | R7 | R8 | R9 | R10 | R11
+ | R12 | R13 | R14 | R15 | R16 | R17 | R18 | R19
+ | R20 | R21 | R22 | R23 | R24 | R25 | R26 | R27
+ | R28 | R29 | R30 => false
+
+ | F0 | F1 | F2 | F3 | F4 | F5 | F6 | F7
+ | F8 | F9 | F10 | F11 | F12 | F13 | F14 | F15
+ | F16 | F17 | F18 | F19 | F20 | F21 | F22 | F23
+ | F24 | F25 | F26 | F27 | F28 | F29 | F30 | F31 => true
+ end.
+
+(** * Function calling conventions *)
+
+(** The functions in this section determine the locations (machine registers
+ and stack slots) used to communicate arguments and results between the
+ caller and the callee during function calls. These locations are functions
+ of the signature of the function and of the call instruction.
+ Agreement between the caller and the callee on the locations to use
+ is guaranteed by our dynamic semantics for Cminor and RTL, which demand
+ that the signature of the call instruction is identical to that of the
+ called function.
+
+ Calling conventions are largely arbitrary: they must respect the properties
+ proved in this section (such as no overlapping between the locations
+ of function arguments), but this leaves much liberty in choosing actual
+ locations. To ensure binary interoperability of code generated by our
+ compiler with libraries compiled by another compiler, we
+ implement the standard RISC-V conventions. *)
+
+(** ** Location of function result *)
+
+(** The result value of a function is passed back to the caller in
+ registers [R10] or [F10] or [R10,R11], depending on the type of the
+ returned value. We treat a function without result as a function
+ with one integer result. *)
+
+Definition loc_result (s: signature) : rpair mreg :=
+ match s.(sig_res) with
+ | None => One R10
+ | Some (Tint | Tany32) => One R10
+ | Some (Tfloat | Tsingle | Tany64) => One F10
+ | Some Tlong => if Archi.ptr64 then One R10 else Twolong R11 R10
+ end.
+
+(** The result registers have types compatible with that given in the signature. *)
+
+Lemma loc_result_type:
+ forall sig,
+ subtype (proj_sig_res sig) (typ_rpair mreg_type (loc_result sig)) = true.
+Proof.
+ intros. unfold proj_sig_res, loc_result, mreg_type;
+ destruct (sig_res sig) as [[]|]; auto; destruct Archi.ptr64; auto.
+Qed.
+
+(** The result locations are caller-save registers *)
+
+Lemma loc_result_caller_save:
+ forall (s: signature),
+ forall_rpair (fun r => is_callee_save r = false) (loc_result s).
+Proof.
+ intros. unfold loc_result, is_callee_save;
+ destruct (sig_res s) as [[]|]; simpl; auto; destruct Archi.ptr64; simpl; auto.
+Qed.
+
+(** If the result is in a pair of registers, those registers are distinct and have type [Tint] at least. *)
+
+Lemma loc_result_pair:
+ forall sg,
+ match loc_result sg with
+ | One _ => True
+ | Twolong r1 r2 =>
+ r1 <> r2 /\ sg.(sig_res) = Some Tlong
+ /\ subtype Tint (mreg_type r1) = true /\ subtype Tint (mreg_type r2) = true
+ /\ Archi.splitlong = true
+ end.
+Proof.
+ intros. change Archi.splitlong with (negb Archi.ptr64).
+ unfold loc_result; destruct (sig_res sg) as [[]|]; auto.
+ unfold mreg_type; destruct Archi.ptr64; auto.
+ split; auto. congruence.
+Qed.
+
+(** The location of the result depends only on the result part of the signature *)
+
+Lemma loc_result_exten:
+ forall s1 s2, s1.(sig_res) = s2.(sig_res) -> loc_result s1 = loc_result s2.
+Proof.
+ intros. unfold loc_result. rewrite H; auto.
+Qed.
+
+(** ** Location of function arguments *)
+
+(** The RISC-V ABI states the following convention for passing arguments
+ to a function:
+
+- Arguments are passed in registers when possible.
+
+- Up to eight integer registers (ai: int_param_regs) and up to eight
+ floating-point registers (fai: float_param_regs) are used for this
+ purpose.
+
+- If the arguments to a function are conceptualized as fields of a C
+ struct, each with pointer alignment, the argument registers are a
+ shadow of the first eight pointer-words of that struct. If argument
+ i < 8 is a floating-point type, it is passed in floating-point
+ register fa_i; otherwise, it is passed in integer register a_i.
+
+- When primitive arguments twice the size of a pointer-word are passed
+ on the stack, they are naturally aligned. When they are passed in the
+ integer registers, they reside in an aligned even-odd register pair,
+ with the even register holding the least-significant bits.
+
+- Floating-point arguments to variadic functions (except those that
+ are explicitly named in the parameter list) are passed in integer
+ registers.
+
+- The portion of the conceptual struct that is not passed in argument
+ registers is passed on the stack. The stack pointer sp points to the
+ first argument not passed in a register.
+
+The bit about variadic functions doesn't quite fit CompCert's model.
+We do our best by passing the FP arguments in registers, as usual,
+and reserving the corresponding integer registers, so that fixup
+code can be introduced in the Asmexpand pass.
+*)
+
+Definition int_param_regs :=
+ R10 :: R11 :: R12 :: R13 :: R14 :: R15 :: R16 :: R17 :: nil.
+Definition float_param_regs :=
+ F10 :: F11 :: F12 :: F13 :: F14 :: F15 :: F16 :: F17 :: nil.
+
+Definition one_arg (regs: list mreg) (rn: Z) (ofs: Z) (ty: typ)
+ (rec: Z -> Z -> list (rpair loc)) :=
+ match list_nth_z regs rn with
+ | Some r =>
+ One(R r) :: rec (rn + 1) ofs
+ | None =>
+ let ofs := align ofs (typealign ty) in
+ One(S Outgoing ofs ty) :: rec rn (ofs + (if Archi.ptr64 then 2 else 1))
+ end.
+
+Definition two_args (regs: list mreg) (rn: Z) (ofs: Z)
+ (rec: Z -> Z -> list (rpair loc)) :=
+ let rn := align rn 2 in
+ match list_nth_z regs rn, list_nth_z regs (rn + 1) with
+ | Some r1, Some r2 =>
+ Twolong (R r2) (R r1) :: rec (rn + 2) ofs
+ | _, _ =>
+ let ofs := align ofs 2 in
+ Twolong (S Outgoing (ofs + 1) Tint) (S Outgoing ofs Tint) ::
+ rec rn (ofs + 2)
+ end.
+
+Definition hybrid_arg (regs: list mreg) (rn: Z) (ofs: Z) (ty: typ)
+ (rec: Z -> Z -> list (rpair loc)) :=
+ let rn := align rn 2 in
+ match list_nth_z regs rn with
+ | Some r =>
+ One (R r) :: rec (rn + 2) ofs
+ | None =>
+ let ofs := align ofs 2 in
+ One (S Outgoing ofs ty) :: rec rn (ofs + 2)
+ end.
+
+Fixpoint loc_arguments_rec (va: bool)
+ (tyl: list typ) (r ofs: Z) {struct tyl} : list (rpair loc) :=
+ match tyl with
+ | nil => nil
+ | (Tint | Tany32) as ty :: tys =>
+ one_arg int_param_regs r ofs ty (loc_arguments_rec va tys)
+ | Tsingle as ty :: tys =>
+ one_arg float_param_regs r ofs ty (loc_arguments_rec va tys)
+ | Tlong as ty :: tys =>
+ if Archi.ptr64
+ then one_arg int_param_regs r ofs ty (loc_arguments_rec va tys)
+ else two_args int_param_regs r ofs (loc_arguments_rec va tys)
+ | (Tfloat | Tany64) as ty :: tys =>
+ if va && negb Archi.ptr64
+ then hybrid_arg float_param_regs r ofs ty (loc_arguments_rec va tys)
+ else one_arg float_param_regs r ofs ty (loc_arguments_rec va tys)
+ end.
+
+(** [loc_arguments s] returns the list of locations where to store arguments
+ when calling a function with signature [s]. *)
+
+Definition loc_arguments (s: signature) : list (rpair loc) :=
+ loc_arguments_rec s.(sig_cc).(cc_vararg) s.(sig_args) 0 0.
+
+(** [size_arguments s] returns the number of [Outgoing] slots used
+ to call a function with signature [s]. *)
+
+Definition max_outgoing_1 (accu: Z) (l: loc) : Z :=
+ match l with
+ | S Outgoing ofs ty => Z.max accu (ofs + typesize ty)
+ | _ => accu
+ end.
+
+Definition max_outgoing_2 (accu: Z) (rl: rpair loc) : Z :=
+ match rl with
+ | One l => max_outgoing_1 accu l
+ | Twolong l1 l2 => max_outgoing_1 (max_outgoing_1 accu l1) l2
+ end.
+
+Definition size_arguments (s: signature) : Z :=
+ List.fold_left max_outgoing_2 (loc_arguments s) 0.
+
+(** Argument locations are either non-temporary registers or [Outgoing]
+ stack slots at nonnegative offsets. *)
+
+Definition loc_argument_acceptable (l: loc) : Prop :=
+ match l with
+ | R r => is_callee_save r = false
+ | S Outgoing ofs ty => ofs >= 0 /\ (typealign ty | ofs)
+ | _ => False
+ end.
+
+Lemma loc_arguments_rec_charact:
+ forall va tyl rn ofs p,
+ ofs >= 0 ->
+ In p (loc_arguments_rec va tyl rn ofs) -> forall_rpair loc_argument_acceptable p.
+Proof.
+ set (OK := fun (l: list (rpair loc)) =>
+ forall p, In p l -> forall_rpair loc_argument_acceptable p).
+ set (OKF := fun (f: Z -> Z -> list (rpair loc)) =>
+ forall rn ofs, ofs >= 0 -> OK (f rn ofs)).
+ set (OKREGS := fun (l: list mreg) => forall r, In r l -> is_callee_save r = false).
+ assert (AL: forall ofs ty, ofs >= 0 -> align ofs (typealign ty) >= 0).
+ { intros.
+ assert (ofs <= align ofs (typealign ty)) by (apply align_le; apply typealign_pos).
+ omega. }
+ assert (SK: (if Archi.ptr64 then 2 else 1) > 0).
+ { destruct Archi.ptr64; omega. }
+ assert (A: forall regs rn ofs ty f,
+ OKREGS regs -> OKF f -> ofs >= 0 -> OK (one_arg regs rn ofs ty f)).
+ { intros until f; intros OR OF OO; red; unfold one_arg; intros.
+ destruct (list_nth_z regs rn) as [r|] eqn:NTH; destruct H.
+ - subst p; simpl. apply OR. eapply list_nth_z_in; eauto.
+ - eapply OF; eauto.
+ - subst p; simpl. auto using align_divides, typealign_pos.
+ - eapply OF; [idtac|eauto]. generalize (AL ofs ty OO); omega.
+ }
+ assert (B: forall regs rn ofs f,
+ OKREGS regs -> OKF f -> ofs >= 0 -> OK (two_args regs rn ofs f)).
+ { intros until f; intros OR OF OO; unfold two_args.
+ set (rn' := align rn 2).
+ set (ofs' := align ofs 2).
+ assert (OO': ofs' >= 0) by (apply (AL ofs Tlong); auto).
+ assert (DFL: OK (Twolong (S Outgoing (ofs' + 1) Tint) (S Outgoing ofs' Tint)
+ :: f rn' (ofs' + 2))).
+ { red; simpl; intros. destruct H.
+ - subst p; simpl.
+ repeat split; auto using Z.divide_1_l. omega.
+ - eapply OF; [idtac|eauto]. omega.
+ }
+ destruct (list_nth_z regs rn') as [r1|] eqn:NTH1;
+ destruct (list_nth_z regs (rn' + 1)) as [r2|] eqn:NTH2;
+ try apply DFL.
+ red; simpl; intros; destruct H.
+ - subst p; simpl. split; apply OR; eauto using list_nth_z_in.
+ - eapply OF; [idtac|eauto]. auto.
+ }
+ assert (C: forall regs rn ofs ty f,
+ OKREGS regs -> OKF f -> ofs >= 0 -> typealign ty = 1 -> OK (hybrid_arg regs rn ofs ty f)).
+ { intros until f; intros OR OF OO OTY; unfold hybrid_arg; red; intros.
+ set (rn' := align rn 2) in *.
+ destruct (list_nth_z regs rn') as [r|] eqn:NTH; destruct H.
+ - subst p; simpl. apply OR. eapply list_nth_z_in; eauto.
+ - eapply OF; eauto.
+ - subst p; simpl. rewrite OTY. split. apply (AL ofs Tlong OO). apply Z.divide_1_l.
+ - eapply OF; [idtac|eauto]. generalize (AL ofs Tlong OO); simpl; omega.
+ }
+ assert (D: OKREGS int_param_regs).
+ { red. decide_goal. }
+ assert (E: OKREGS float_param_regs).
+ { red. decide_goal. }
+
+ cut (forall va tyl rn ofs, ofs >= 0 -> OK (loc_arguments_rec va tyl rn ofs)).
+ unfold OK. eauto.
+ induction tyl as [ | ty1 tyl]; intros until ofs; intros OO; simpl.
+- red; simpl; tauto.
+- destruct ty1.
++ (* int *) apply A; auto.
++ (* float *)
+ destruct (va && negb Archi.ptr64).
+ apply C; auto.
+ apply A; auto.
++ (* long *)
+ destruct Archi.ptr64.
+ apply A; auto.
+ apply B; auto.
++ (* single *)
+ apply A; auto.
++ (* any32 *)
+ apply A; auto.
++ (* any64 *)
+ destruct (va && negb Archi.ptr64).
+ apply C; auto.
+ apply A; auto.
+Qed.
+
+Lemma loc_arguments_acceptable:
+ forall (s: signature) (p: rpair loc),
+ In p (loc_arguments s) -> forall_rpair loc_argument_acceptable p.
+Proof.
+ unfold loc_arguments; intros. eapply loc_arguments_rec_charact; eauto. omega.
+Qed.
+
+(** The offsets of [Outgoing] arguments are below [size_arguments s]. *)
+
+Remark fold_max_outgoing_above:
+ forall l n, fold_left max_outgoing_2 l n >= n.
+Proof.
+ assert (A: forall n l, max_outgoing_1 n l >= n).
+ { intros; unfold max_outgoing_1. destruct l as [_ | []]; xomega. }
+ induction l; simpl; intros.
+ - omega.
+ - eapply Zge_trans. eauto.
+ destruct a; simpl. apply A. eapply Zge_trans; eauto.
+Qed.
+
+Lemma size_arguments_above:
+ forall s, size_arguments s >= 0.
+Proof.
+ intros. apply fold_max_outgoing_above.
+Qed.
+
+Lemma loc_arguments_bounded:
+ forall (s: signature) (ofs: Z) (ty: typ),
+ In (S Outgoing ofs ty) (regs_of_rpairs (loc_arguments s)) ->
+ ofs + typesize ty <= size_arguments s.
+Proof.
+ intros until ty.
+ assert (A: forall n l, n <= max_outgoing_1 n l).
+ { intros; unfold max_outgoing_1. destruct l as [_ | []]; xomega. }
+ assert (B: forall p n,
+ In (S Outgoing ofs ty) (regs_of_rpair p) ->
+ ofs + typesize ty <= max_outgoing_2 n p).
+ { intros. destruct p; simpl in H; intuition; subst; simpl.
+ - xomega.
+ - eapply Zle_trans. 2: apply A. xomega.
+ - xomega. }
+ assert (C: forall l n,
+ In (S Outgoing ofs ty) (regs_of_rpairs l) ->
+ ofs + typesize ty <= fold_left max_outgoing_2 l n).
+ { induction l; simpl; intros.
+ - contradiction.
+ - rewrite in_app_iff in H. destruct H.
+ + eapply Zle_trans. eapply B; eauto. apply Zge_le. apply fold_max_outgoing_above.
+ + apply IHl; auto.
+ }
+ apply C.
+Qed.
+
+Lemma loc_arguments_main:
+ loc_arguments signature_main = nil.
+Proof.
+ reflexivity.
+Qed.
diff --git a/riscV/Machregs.v b/riscV/Machregs.v
new file mode 100644
index 00000000..e286bbad
--- /dev/null
+++ b/riscV/Machregs.v
@@ -0,0 +1,253 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+Require Import String.
+Require Import Coqlib.
+Require Import Decidableplus.
+Require Import Maps.
+Require Import AST.
+Require Import Integers.
+Require Import Op.
+
+(** ** Machine registers *)
+
+(** The following type defines the machine registers that can be referenced
+ as locations. These include:
+- Integer registers that can be allocated to RTL pseudo-registers ([Rxx]).
+- Floating-point registers that can be allocated to RTL pseudo-registers
+ ([Fxx]).
+
+ The type [mreg] does not include reserved machine registers such as
+ the zero register (x0), the link register (x1), the stack pointer
+ (x2), the global pointer (x3), and the thread pointer (x4).
+ Finally, register x31 is reserved for use as a temporary by the
+ assembly-code generator [Asmgen].
+*)
+
+Inductive mreg: Type :=
+ (** Allocatable integer regs. *)
+ | R5: mreg | R6: mreg | R7: mreg
+ | R8: mreg | R9: mreg | R10: mreg | R11: mreg
+ | R12: mreg | R13: mreg | R14: mreg | R15: mreg
+ | R16: mreg | R17: mreg | R18: mreg | R19: mreg
+ | R20: mreg | R21: mreg | R22: mreg | R23: mreg
+ | R24: mreg | R25: mreg | R26: mreg | R27: mreg
+ | R28: mreg | R29: mreg | R30: mreg
+ (** Allocatable double-precision float regs *)
+ | F0: mreg | F1: mreg | F2: mreg | F3: mreg
+ | F4: mreg | F5: mreg | F6: mreg | F7: mreg
+ | F8: mreg | F9: mreg | F10: mreg | F11: mreg
+ | F12: mreg | F13: mreg | F14: mreg | F15: mreg
+ | F16: mreg | F17: mreg | F18: mreg | F19: mreg
+ | F20: mreg | F21: mreg | F22: mreg | F23: mreg
+ | F24: mreg | F25: mreg | F26: mreg | F27: mreg
+ | F28: mreg | F29: mreg | F30: mreg | F31: mreg.
+
+Lemma mreg_eq: forall (r1 r2: mreg), {r1 = r2} + {r1 <> r2}.
+Proof. decide equality. Defined.
+Global Opaque mreg_eq.
+
+Definition all_mregs :=
+ R5 :: R6 :: R7 :: R8 :: R9 :: R10 :: R11 :: R12 :: R13 :: R14 :: R15
+ :: R16 :: R17 :: R18 :: R19 :: R20 :: R21 :: R22 :: R23
+ :: R24 :: R25 :: R26 :: R27 :: R28 :: R29 :: R30
+ :: F0 :: F1 :: F2 :: F3 :: F4 :: F5 :: F6 :: F7
+ :: F8 :: F9 :: F10 :: F11 :: F12 :: F13 :: F14 :: F15
+ :: F16 :: F17 :: F18 :: F19 :: F20 :: F21 :: F22 :: F23
+ :: F24 :: F25 :: F26 :: F27 :: F28 :: F29 :: F30 :: F31
+ :: nil.
+
+Lemma all_mregs_complete:
+ forall (r: mreg), In r all_mregs.
+Proof.
+ assert (forall r, proj_sumbool (In_dec mreg_eq r all_mregs) = true) by (destruct r; reflexivity).
+ intros. specialize (H r). InvBooleans. auto.
+Qed.
+
+Instance Decidable_eq_mreg : forall (x y: mreg), Decidable (eq x y) := Decidable_eq mreg_eq.
+
+Instance Finite_mreg : Finite mreg := {
+ Finite_elements := all_mregs;
+ Finite_elements_spec := all_mregs_complete
+}.
+
+Definition mreg_type (r: mreg): typ :=
+ match r with
+ | R5 | R6 | R7 | R8 | R9 | R10 | R11
+ | R12 | R13 | R14 | R15 | R16 | R17 | R18 | R19
+ | R20 | R21 | R22 | R23 | R24 | R25 | R26 | R27
+ | R28 | R29 | R30 => if Archi.ptr64 then Tany64 else Tany32
+
+ | F0 | F1 | F2 | F3 | F4 | F5 | F6 | F7
+ | F8 | F9 | F10 | F11 | F12 | F13 | F14 | F15
+ | F16 | F17 | F18 | F19 | F20 | F21 | F22 | F23
+ | F24 | F25 | F26 | F27 | F28 | F29 | F30 | F31 => Tany64
+ end.
+
+Open Scope positive_scope.
+
+Module IndexedMreg <: INDEXED_TYPE.
+ Definition t := mreg.
+ Definition eq := mreg_eq.
+ Definition index (r: mreg): positive :=
+ match r with
+ | R5 => 1 | R6 => 2 | R7 => 3
+ | R8 => 4 | R9 => 5 | R10 => 6 | R11 => 7
+ | R12 => 8 | R13 => 9 | R14 => 10 | R15 => 11
+ | R16 => 12 | R17 => 13 | R18 => 14 | R19 => 15
+ | R20 => 16 | R21 => 17 | R22 => 18 | R23 => 19
+ | R24 => 20 | R25 => 21 | R26 => 22 | R27 => 23
+ | R28 => 24 | R29 => 25 | R30 => 26
+
+ | F0 => 28 | F1 => 29 | F2 => 30 | F3 => 31
+ | F4 => 32 | F5 => 33 | F6 => 34 | F7 => 35
+ | F8 => 36 | F9 => 37 | F10 => 38 | F11 => 39
+ | F12 => 40 | F13 => 41 | F14 => 42 | F15 => 43
+ | F16 => 44 | F17 => 45 | F18 => 46 | F19 => 47
+ | F20 => 48 | F21 => 49 | F22 => 50 | F23 => 51
+ | F24 => 52 | F25 => 53 | F26 => 54 | F27 => 55
+ | F28 => 56 | F29 => 57 | F30 => 58 | F31 => 59
+ end.
+ Lemma index_inj:
+ forall r1 r2, index r1 = index r2 -> r1 = r2.
+ Proof.
+ decide_goal.
+ Qed.
+End IndexedMreg.
+
+Definition is_stack_reg (r: mreg) : bool := false.
+
+(** ** Names of registers *)
+
+Local Open Scope string_scope.
+
+Definition register_names :=
+ ("X5", R5) :: ("X6", R6) :: ("X7", R7) ::
+ ("X8", R8) :: ("X9", R9) :: ("X10", R10) :: ("X11", R11) ::
+ ("X12", R12) :: ("X13", R13) :: ("X14", R14) :: ("X15", R15) ::
+ ("X16", R16) :: ("X17", R17) :: ("X18", R18) :: ("X19", R19) ::
+ ("X20", R20) :: ("X21", R21) :: ("X22", R22) :: ("X23", R23) ::
+ ("X24", R24) :: ("X25", R25) :: ("X26", R26) :: ("X27", R27) ::
+ ("X28", R28) :: ("X29", R29) :: ("X30", R30) ::
+
+ ("F0", F0) :: ("F1", F1) :: ("F2", F2) :: ("F3", F3) ::
+ ("F4", F4) :: ("F5", F5) :: ("F6", F6) :: ("F7", F7) ::
+ ("F8", F8) :: ("F9", F9) :: ("F10", F10) :: ("F11", F11) ::
+ ("F12", F12) :: ("F13", F13) :: ("F14", F14) :: ("F15", F15) ::
+ ("F16", F16) :: ("F17", F17) :: ("F18", F18) :: ("F19", F19) ::
+ ("F20", F20) :: ("F21", F21) :: ("F22", F22) :: ("F23", F23) ::
+ ("F24", F24) :: ("F25", F25) :: ("F26", F26) :: ("F27", F27) ::
+ ("F27", F27) :: ("F28", F28) :: ("F29", F29) :: ("F30", F30) ::
+ nil.
+
+Definition register_by_name (s: string) : option mreg :=
+ let fix assoc (l: list (string * mreg)) : option mreg :=
+ match l with
+ | nil => None
+ | (s1, r1) :: l' => if string_dec s s1 then Some r1 else assoc l'
+ end
+ in assoc register_names.
+
+(** ** Destroyed registers, preferred registers *)
+
+Definition destroyed_by_op (op: operation): list mreg :=
+ match op with
+ | Ointoffloat | Ointuoffloat | Ointofsingle | Ointuofsingle
+ | Olongoffloat | Olonguoffloat | Olongofsingle | Olonguofsingle
+ => F6 :: nil
+ | _ => nil
+ end.
+
+Definition destroyed_by_load (chunk: memory_chunk) (addr: addressing): list mreg := nil.
+
+Definition destroyed_by_store (chunk: memory_chunk) (addr: addressing): list mreg := nil.
+
+Definition destroyed_by_cond (cond: condition): list mreg := nil.
+
+Definition destroyed_by_jumptable: list mreg := R5 :: nil.
+
+Fixpoint destroyed_by_clobber (cl: list string): list mreg :=
+ match cl with
+ | nil => nil
+ | c1 :: cl =>
+ match register_by_name c1 with
+ | Some r => r :: destroyed_by_clobber cl
+ | None => destroyed_by_clobber cl
+ end
+ end.
+
+Definition destroyed_by_builtin (ef: external_function): list mreg :=
+ match ef with
+ | EF_inline_asm txt sg clob => destroyed_by_clobber clob
+ | EF_memcpy sz al => R5 :: R6 :: R7 :: F0 :: nil
+ | _ => nil
+ end.
+
+Definition destroyed_by_setstack (ty: typ): list mreg := nil.
+
+Definition destroyed_at_function_entry: list mreg := R30 :: nil.
+
+Definition temp_for_parent_frame: mreg := R30.
+
+Definition destroyed_at_indirect_call: list mreg :=
+ R10 :: R11 :: R12 :: R13 :: R14 :: R15 :: R16 :: R17 :: nil.
+
+Definition mregs_for_operation (op: operation): list (option mreg) * option mreg := (nil, None).
+
+Definition mregs_for_builtin (ef: external_function): list (option mreg) * list(option mreg) :=
+ match ef with
+ | EF_builtin name sg =>
+ if (negb Archi.ptr64) && string_dec name "__builtin_bswap64" then
+ (Some R6 :: Some R5 :: nil, Some R5 :: Some R6 :: nil)
+ else
+ (nil, nil)
+ | _ =>
+ (nil, nil)
+ end.
+
+Global Opaque
+ destroyed_by_op destroyed_by_load destroyed_by_store
+ destroyed_by_cond destroyed_by_jumptable destroyed_by_builtin
+ destroyed_by_setstack destroyed_at_function_entry temp_for_parent_frame
+ mregs_for_operation mregs_for_builtin.
+
+(** Two-address operations. Return [true] if the first argument and
+ the result must be in the same location *and* are unconstrained
+ by [mregs_for_operation]. There are two: the pseudo [Ocast32signed],
+ because it expands to a no-op owing to the representation of 32-bit
+ integers as their 64-bit sign extension; and [Ocast32unsigned],
+ because it builds on the same magic no-op. *)
+
+Definition two_address_op (op: operation) : bool :=
+ match op with
+ | Ocast32signed | Ocast32unsigned => true
+ | _ => false
+ end.
+
+(** Constraints on constant propagation for builtins *)
+
+Definition builtin_constraints (ef: external_function) :
+ list builtin_arg_constraint :=
+ match ef with
+ | EF_builtin id sg => nil
+ | EF_vload _ => OK_addrstack :: nil
+ | EF_vstore _ => OK_addrstack :: OK_default :: nil
+ | EF_memcpy _ _ => OK_addrstack :: OK_addrstack :: nil
+ | EF_annot txt targs => map (fun _ => OK_all) targs
+ | EF_debug kind txt targs => map (fun _ => OK_all) targs
+ | _ => nil
+ end.
diff --git a/riscV/Machregsaux.ml b/riscV/Machregsaux.ml
new file mode 100644
index 00000000..473e0602
--- /dev/null
+++ b/riscV/Machregsaux.ml
@@ -0,0 +1,33 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Auxiliary functions on machine registers *)
+
+open Camlcoq
+open Machregs
+
+let register_names : (mreg, string) Hashtbl.t = Hashtbl.create 31
+
+let _ =
+ List.iter
+ (fun (s, r) -> Hashtbl.add register_names r (camlstring_of_coqstring s))
+ Machregs.register_names
+
+let is_scratch_register r = false
+
+let name_of_register r =
+ try Some (Hashtbl.find register_names r) with Not_found -> None
+
+let register_by_name s =
+ Machregs.register_by_name (coqstring_uppercase_ascii_of_camlstring s)
+
+let can_reserve_register r = Conventions1.is_callee_save r
diff --git a/riscV/Machregsaux.mli b/riscV/Machregsaux.mli
new file mode 100644
index 00000000..9404568d
--- /dev/null
+++ b/riscV/Machregsaux.mli
@@ -0,0 +1,18 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Auxiliary functions on machine registers *)
+
+val name_of_register: Machregs.mreg -> string option
+val register_by_name: string -> Machregs.mreg option
+val is_scratch_register: string -> bool
+val can_reserve_register: Machregs.mreg -> bool
diff --git a/riscV/NeedOp.v b/riscV/NeedOp.v
new file mode 100644
index 00000000..117bbcb4
--- /dev/null
+++ b/riscV/NeedOp.v
@@ -0,0 +1,173 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+Require Import Coqlib.
+Require Import AST Integers Floats.
+Require Import Values Memory Globalenvs.
+Require Import Op RTL.
+Require Import NeedDomain.
+
+(** Neededness analysis for RISC-V operators *)
+
+Definition op1 (nv: nval) := nv :: nil.
+Definition op2 (nv: nval) := nv :: nv :: nil.
+
+Definition needs_of_condition (cond: condition): list nval := nil.
+
+Definition needs_of_operation (op: operation) (nv: nval): list nval :=
+ match op with
+ | Omove => op1 nv
+ | Ointconst n => nil
+ | Olongconst n => nil
+ | Ofloatconst n => nil
+ | Osingleconst n => nil
+ | Oaddrsymbol id ofs => nil
+ | Oaddrstack ofs => nil
+ | Ocast8signed => op1 (sign_ext 8 nv)
+ | Ocast16signed => op1 (sign_ext 16 nv)
+ | Oadd => op2 (modarith nv)
+ | Oaddimm n => op1 (modarith nv)
+ | Oneg => op1 (modarith nv)
+ | Osub => op2 (default nv)
+ | Omul => op2 (modarith nv)
+ | Omulhs | Omulhu | Odiv | Odivu | Omod | Omodu => op2 (default nv)
+ | Oand => op2 (bitwise nv)
+ | Oandimm n => op1 (andimm nv n)
+ | Oor => op2 (bitwise nv)
+ | Oorimm n => op1 (orimm nv n)
+ | Oxor => op2 (bitwise nv)
+ | Oxorimm n => op1 (bitwise nv)
+ | Oshl | Oshr | Oshru => op2 (default nv)
+ | Oshlimm n => op1 (shlimm nv n)
+ | Oshrimm n => op1 (shrimm nv n)
+ | Oshruimm n => op1 (shruimm nv n)
+ | Oshrximm n => op1 (default nv)
+ | Omakelong => op2 (default nv)
+ | Olowlong | Ohighlong => op1 (default nv)
+ | Ocast32signed => op1 (default nv)
+ | Ocast32unsigned => op1 (default nv)
+ | Oaddl => op2 (default nv)
+ | Oaddlimm n => op1 (default nv)
+ | Onegl => op1 (default nv)
+ | Osubl => op2 (default nv)
+ | Omull => op2 (default nv)
+ | Omullhs | Omullhu | Odivl | Odivlu | Omodl | Omodlu => op2 (default nv)
+ | Oandl => op2 (default nv)
+ | Oandlimm n => op1 (default nv)
+ | Oorl => op2 (default nv)
+ | Oorlimm n => op1 (default nv)
+ | Oxorl => op2 (default nv)
+ | Oxorlimm n => op1 (default nv)
+ | Oshll | Oshrl | Oshrlu => op2 (default nv)
+ | Oshllimm n => op1 (default nv)
+ | Oshrlimm n => op1 (default nv)
+ | Oshrluimm n => op1 (default nv)
+ | Oshrxlimm n => op1 (default nv)
+ | Onegf | Oabsf => op1 (default nv)
+ | Oaddf | Osubf | Omulf | Odivf => op2 (default nv)
+ | Onegfs | Oabsfs => op1 (default nv)
+ | Oaddfs | Osubfs | Omulfs | Odivfs => op2 (default nv)
+ | Ofloatofsingle | Osingleoffloat => op1 (default nv)
+ | Ointoffloat | Ointuoffloat | Ofloatofint | Ofloatofintu => op1 (default nv)
+ | Olongoffloat | Olonguoffloat | Ofloatoflong | Ofloatoflongu => op1 (default nv)
+ | Ointofsingle | Ointuofsingle | Osingleofint | Osingleofintu => op1 (default nv)
+ | Olongofsingle | Olonguofsingle | Osingleoflong | Osingleoflongu => op1 (default nv)
+ | Ocmp c => needs_of_condition c
+ end.
+
+Definition operation_is_redundant (op: operation) (nv: nval): bool :=
+ match op with
+ | Ocast8signed => sign_ext_redundant 8 nv
+ | Ocast16signed => sign_ext_redundant 16 nv
+ | Oandimm n => andimm_redundant nv n
+ | Oorimm n => orimm_redundant nv n
+ | _ => false
+ end.
+
+Ltac InvAgree :=
+ match goal with
+ | [H: vagree_list nil _ _ |- _ ] => inv H; InvAgree
+ | [H: vagree_list (_::_) _ _ |- _ ] => inv H; InvAgree
+ | _ => idtac
+ end.
+
+Ltac TrivialExists :=
+ match goal with
+ | [ |- exists v, Some ?x = Some v /\ _ ] => exists x; split; auto
+ | _ => idtac
+ end.
+
+Section SOUNDNESS.
+
+Variable ge: genv.
+Variable sp: block.
+Variables m m': mem.
+Hypothesis PERM: forall b ofs k p, Mem.perm m b ofs k p -> Mem.perm m' b ofs k p.
+
+Lemma needs_of_condition_sound:
+ forall cond args b args',
+ eval_condition cond args m = Some b ->
+ vagree_list args args' (needs_of_condition cond) ->
+ eval_condition cond args' m' = Some b.
+Proof.
+ intros. unfold needs_of_condition in H0.
+ eapply default_needs_of_condition_sound; eauto.
+Qed.
+
+Lemma needs_of_operation_sound:
+ forall op args v nv args',
+ eval_operation ge (Vptr sp Ptrofs.zero) op args m = Some v ->
+ vagree_list args args' (needs_of_operation op nv) ->
+ nv <> Nothing ->
+ exists v',
+ eval_operation ge (Vptr sp Ptrofs.zero) op args' m' = Some v'
+ /\ vagree v v' nv.
+Proof.
+ unfold needs_of_operation; intros; destruct op; try (eapply default_needs_of_operation_sound; eauto; fail);
+ simpl in *; FuncInv; InvAgree; TrivialExists.
+- apply sign_ext_sound; auto. compute; auto.
+- apply sign_ext_sound; auto. compute; auto.
+- apply add_sound; auto.
+- apply add_sound; auto with na.
+- apply neg_sound; auto.
+- apply mul_sound; auto.
+- apply and_sound; auto.
+- apply andimm_sound; auto.
+- apply or_sound; auto.
+- apply orimm_sound; auto.
+- apply xor_sound; auto.
+- apply xor_sound; auto with na.
+- apply shlimm_sound; auto.
+- apply shrimm_sound; auto.
+- apply shruimm_sound; auto.
+Qed.
+
+Lemma operation_is_redundant_sound:
+ forall op nv arg1 args v arg1' args',
+ operation_is_redundant op nv = true ->
+ eval_operation ge (Vptr sp Ptrofs.zero) op (arg1 :: args) m = Some v ->
+ vagree_list (arg1 :: args) (arg1' :: args') (needs_of_operation op nv) ->
+ vagree v arg1' nv.
+Proof.
+ intros. destruct op; simpl in *; try discriminate; inv H1; FuncInv; subst.
+- apply sign_ext_redundant_sound; auto. omega.
+- apply sign_ext_redundant_sound; auto. omega.
+- apply andimm_redundant_sound; auto.
+- apply orimm_redundant_sound; auto.
+Qed.
+
+End SOUNDNESS.
diff --git a/riscV/Op.v b/riscV/Op.v
new file mode 100644
index 00000000..ce90ebee
--- /dev/null
+++ b/riscV/Op.v
@@ -0,0 +1,1340 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Operators and addressing modes. The abstract syntax and dynamic
+ semantics for the CminorSel, RTL, LTL and Mach languages depend on the
+ following types, defined in this library:
+- [condition]: boolean conditions for conditional branches;
+- [operation]: arithmetic and logical operations;
+- [addressing]: addressing modes for load and store operations.
+
+ These types are processor-specific and correspond roughly to what the
+ processor can compute in one instruction. In other terms, these
+ types reflect the state of the program after instruction selection.
+ For a processor-independent set of operations, see the abstract
+ syntax and dynamic semantics of the Cminor language.
+*)
+
+Require Import BoolEqual Coqlib.
+Require Import AST Integers Floats.
+Require Import Values Memory Globalenvs Events.
+
+Set Implicit Arguments.
+
+(** Conditions (boolean-valued operators). *)
+
+Inductive condition : Type :=
+ | Ccomp (c: comparison) (**r signed integer comparison *)
+ | Ccompu (c: comparison) (**r unsigned integer comparison *)
+ | Ccompimm (c: comparison) (n: int) (**r signed integer comparison with a constant *)
+ | Ccompuimm (c: comparison) (n: int) (**r unsigned integer comparison with a constant *)
+ | Ccompl (c: comparison) (**r signed 64-bit integer comparison *)
+ | Ccomplu (c: comparison) (**r unsigned 64-bit integer comparison *)
+ | Ccomplimm (c: comparison) (n: int64) (**r signed 64-bit integer comparison with a constant *)
+ | Ccompluimm (c: comparison) (n: int64) (**r unsigned 64-bit integer comparison with a constant *)
+ | Ccompf (c: comparison) (**r 64-bit floating-point comparison *)
+ | Cnotcompf (c: comparison) (**r negation of a floating-point comparison *)
+ | Ccompfs (c: comparison) (**r 32-bit floating-point comparison *)
+ | Cnotcompfs (c: comparison). (**r negation of a floating-point comparison *)
+
+(** Arithmetic and logical operations. In the descriptions, [rd] is the
+ result of the operation and [r1], [r2], etc, are the arguments. *)
+
+Inductive operation : Type :=
+ | Omove (**r [rd = r1] *)
+ | Ointconst (n: int) (**r [rd] is set to the given integer constant *)
+ | Olongconst (n: int64) (**r [rd] is set to the given integer constant *)
+ | Ofloatconst (n: float) (**r [rd] is set to the given float constant *)
+ | Osingleconst (n: float32)(**r [rd] is set to the given float constant *)
+ | Oaddrsymbol (id: ident) (ofs: ptrofs) (**r [rd] is set to the address of the symbol plus the given offset *)
+ | Oaddrstack (ofs: ptrofs) (**r [rd] is set to the stack pointer plus the given offset *)
+(*c 32-bit integer arithmetic: *)
+ | Ocast8signed (**r [rd] is 8-bit sign extension of [r1] *)
+ | Ocast16signed (**r [rd] is 16-bit sign extension of [r1] *)
+ | Oadd (**r [rd = r1 + r2] *)
+ | Oaddimm (n: int) (**r [rd = r1 + n] *)
+ | Oneg (**r [rd = - r1] *)
+ | Osub (**r [rd = r1 - r2] *)
+ | Omul (**r [rd = r1 * r2] *)
+ | Omulhs (**r [rd = high part of r1 * r2, signed] *)
+ | Omulhu (**r [rd = high part of r1 * r2, unsigned] *)
+ | Odiv (**r [rd = r1 / r2] (signed) *)
+ | Odivu (**r [rd = r1 / r2] (unsigned) *)
+ | Omod (**r [rd = r1 % r2] (signed) *)
+ | Omodu (**r [rd = r1 % r2] (unsigned) *)
+ | Oand (**r [rd = r1 & r2] *)
+ | Oandimm (n: int) (**r [rd = r1 & n] *)
+ | Oor (**r [rd = r1 | r2] *)
+ | Oorimm (n: int) (**r [rd = r1 | n] *)
+ | Oxor (**r [rd = r1 ^ r2] *)
+ | Oxorimm (n: int) (**r [rd = r1 ^ n] *)
+ | Oshl (**r [rd = r1 << r2] *)
+ | Oshlimm (n: int) (**r [rd = r1 << n] *)
+ | Oshr (**r [rd = r1 >> r2] (signed) *)
+ | Oshrimm (n: int) (**r [rd = r1 >> n] (signed) *)
+ | Oshru (**r [rd = r1 >> r2] (unsigned) *)
+ | Oshruimm (n: int) (**r [rd = r1 >> n] (unsigned) *)
+ | Oshrximm (n: int) (**r [rd = r1 / 2^n] (signed) *)
+(*c 64-bit integer arithmetic: *)
+ | Omakelong (**r [rd = r1 << 32 | r2] *)
+ | Olowlong (**r [rd = low-word(r1)] *)
+ | Ohighlong (**r [rd = high-word(r1)] *)
+ | Ocast32signed (**r [rd] is 32-bit sign extension of [r1] *)
+ | Ocast32unsigned (**r [rd] is 32-bit zero extension of [r1] *)
+ | Oaddl (**r [rd = r1 + r2] *)
+ | Oaddlimm (n: int64) (**r [rd = r1 + n] *)
+ | Onegl (**r [rd = - r1] *)
+ | Osubl (**r [rd = r1 - r2] *)
+ | Omull (**r [rd = r1 * r2] *)
+ | Omullhs (**r [rd = high part of r1 * r2, signed] *)
+ | Omullhu (**r [rd = high part of r1 * r2, unsigned] *)
+ | Odivl (**r [rd = r1 / r2] (signed) *)
+ | Odivlu (**r [rd = r1 / r2] (unsigned) *)
+ | Omodl (**r [rd = r1 % r2] (signed) *)
+ | Omodlu (**r [rd = r1 % r2] (unsigned) *)
+ | Oandl (**r [rd = r1 & r2] *)
+ | Oandlimm (n: int64) (**r [rd = r1 & n] *)
+ | Oorl (**r [rd = r1 | r2] *)
+ | Oorlimm (n: int64) (**r [rd = r1 | n] *)
+ | Oxorl (**r [rd = r1 ^ r2] *)
+ | Oxorlimm (n: int64) (**r [rd = r1 ^ n] *)
+ | Oshll (**r [rd = r1 << r2] *)
+ | Oshllimm (n: int) (**r [rd = r1 << n] *)
+ | Oshrl (**r [rd = r1 >> r2] (signed) *)
+ | Oshrlimm (n: int) (**r [rd = r1 >> n] (signed) *)
+ | Oshrlu (**r [rd = r1 >> r2] (unsigned) *)
+ | Oshrluimm (n: int) (**r [rd = r1 >> n] (unsigned) *)
+ | Oshrxlimm (n: int) (**r [rd = r1 / 2^n] (signed) *)
+(*c Floating-point arithmetic: *)
+ | Onegf (**r [rd = - r1] *)
+ | Oabsf (**r [rd = abs(r1)] *)
+ | Oaddf (**r [rd = r1 + r2] *)
+ | Osubf (**r [rd = r1 - r2] *)
+ | Omulf (**r [rd = r1 * r2] *)
+ | Odivf (**r [rd = r1 / r2] *)
+ | Onegfs (**r [rd = - r1] *)
+ | Oabsfs (**r [rd = abs(r1)] *)
+ | Oaddfs (**r [rd = r1 + r2] *)
+ | Osubfs (**r [rd = r1 - r2] *)
+ | Omulfs (**r [rd = r1 * r2] *)
+ | Odivfs (**r [rd = r1 / r2] *)
+ | Osingleoffloat (**r [rd] is [r1] truncated to single-precision float *)
+ | Ofloatofsingle (**r [rd] is [r1] extended to double-precision float *)
+(*c Conversions between int and float: *)
+ | Ointoffloat (**r [rd = signed_int_of_float64(r1)] *)
+ | Ointuoffloat (**r [rd = unsigned_int_of_float64(r1)] *)
+ | Ofloatofint (**r [rd = float64_of_signed_int(r1)] *)
+ | Ofloatofintu (**r [rd = float64_of_unsigned_int(r1)] *)
+ | Ointofsingle (**r [rd = signed_int_of_float32(r1)] *)
+ | Ointuofsingle (**r [rd = unsigned_int_of_float32(r1)] *)
+ | Osingleofint (**r [rd = float32_of_signed_int(r1)] *)
+ | Osingleofintu (**r [rd = float32_of_unsigned_int(r1)] *)
+ | Olongoffloat (**r [rd = signed_long_of_float64(r1)] *)
+ | Olonguoffloat (**r [rd = unsigned_long_of_float64(r1)] *)
+ | Ofloatoflong (**r [rd = float64_of_signed_long(r1)] *)
+ | Ofloatoflongu (**r [rd = float64_of_unsigned_long(r1)] *)
+ | Olongofsingle (**r [rd = signed_long_of_float32(r1)] *)
+ | Olonguofsingle (**r [rd = unsigned_long_of_float32(r1)] *)
+ | Osingleoflong (**r [rd = float32_of_signed_long(r1)] *)
+ | Osingleoflongu (**r [rd = float32_of_unsigned_int(r1)] *)
+(*c Boolean tests: *)
+ | Ocmp (cond: condition). (**r [rd = 1] if condition holds, [rd = 0] otherwise. *)
+
+(** Addressing modes. [r1], [r2], etc, are the arguments to the
+ addressing. *)
+
+Inductive addressing: Type :=
+ | Aindexed: ptrofs -> addressing (**r Address is [r1 + offset] *)
+ | Aglobal: ident -> ptrofs -> addressing (**r Address is global plus offset *)
+ | Ainstack: ptrofs -> addressing. (**r Address is [stack_pointer + offset] *)
+
+(** Comparison functions (used in modules [CSE] and [Allocation]). *)
+
+Definition eq_condition (x y: condition) : {x=y} + {x<>y}.
+Proof.
+ generalize Int.eq_dec Int64.eq_dec; intro.
+ assert (forall (x y: comparison), {x=y}+{x<>y}). decide equality.
+ decide equality.
+Defined.
+
+Definition eq_addressing (x y: addressing) : {x=y} + {x<>y}.
+Proof.
+ generalize ident_eq Ptrofs.eq_dec; intros.
+ decide equality.
+Defined.
+
+Definition eq_operation: forall (x y: operation), {x=y} + {x<>y}.
+Proof.
+ generalize Int.eq_dec Int64.eq_dec Ptrofs.eq_dec Float.eq_dec Float32.eq_dec ident_eq eq_condition; intros.
+ decide equality.
+Defined.
+
+(* Alternate definition:
+Definition beq_operation: forall (x y: operation), bool.
+Proof.
+ generalize Int.eq_dec Int64.eq_dec Ptrofs.eq_dec Float.eq_dec Float32.eq_dec ident_eq eq_condition; boolean_equality.
+Defined.
+
+Definition eq_operation: forall (x y: operation), {x=y} + {x<>y}.
+Proof.
+ decidable_equality_from beq_operation.
+Defined.
+*)
+
+Global Opaque eq_condition eq_addressing eq_operation.
+
+(** * Evaluation functions *)
+
+(** Evaluation of conditions, operators and addressing modes applied
+ to lists of values. Return [None] when the computation can trigger an
+ error, e.g. integer division by zero. [eval_condition] returns a boolean,
+ [eval_operation] and [eval_addressing] return a value. *)
+
+Definition eval_condition (cond: condition) (vl: list val) (m: mem): option bool :=
+ match cond, vl with
+ | Ccomp c, v1 :: v2 :: nil => Val.cmp_bool c v1 v2
+ | Ccompu c, v1 :: v2 :: nil => Val.cmpu_bool (Mem.valid_pointer m) c v1 v2
+ | Ccompimm c n, v1 :: nil => Val.cmp_bool c v1 (Vint n)
+ | Ccompuimm c n, v1 :: nil => Val.cmpu_bool (Mem.valid_pointer m) c v1 (Vint n)
+ | Ccompl c, v1 :: v2 :: nil => Val.cmpl_bool c v1 v2
+ | Ccomplu c, v1 :: v2 :: nil => Val.cmplu_bool (Mem.valid_pointer m) c v1 v2
+ | Ccomplimm c n, v1 :: nil => Val.cmpl_bool c v1 (Vlong n)
+ | Ccompluimm c n, v1 :: nil => Val.cmplu_bool (Mem.valid_pointer m) c v1 (Vlong n)
+ | Ccompf c, v1 :: v2 :: nil => Val.cmpf_bool c v1 v2
+ | Cnotcompf c, v1 :: v2 :: nil => option_map negb (Val.cmpf_bool c v1 v2)
+ | Ccompfs c, v1 :: v2 :: nil => Val.cmpfs_bool c v1 v2
+ | Cnotcompfs c, v1 :: v2 :: nil => option_map negb (Val.cmpfs_bool c v1 v2)
+ | _, _ => None
+ end.
+
+Definition eval_operation
+ (F V: Type) (genv: Genv.t F V) (sp: val)
+ (op: operation) (vl: list val) (m: mem): option val :=
+ match op, vl with
+ | Omove, v1::nil => Some v1
+ | Ointconst n, nil => Some (Vint n)
+ | Olongconst n, nil => Some (Vlong n)
+ | Ofloatconst n, nil => Some (Vfloat n)
+ | Osingleconst n, nil => Some (Vsingle n)
+ | Oaddrsymbol s ofs, nil => Some (Genv.symbol_address genv s ofs)
+ | Oaddrstack ofs, nil => Some (Val.offset_ptr sp ofs)
+ | Ocast8signed, v1 :: nil => Some (Val.sign_ext 8 v1)
+ | Ocast16signed, v1 :: nil => Some (Val.sign_ext 16 v1)
+ | Oadd, v1 :: v2 :: nil => Some (Val.add v1 v2)
+ | Oaddimm n, v1 :: nil => Some (Val.add v1 (Vint n))
+ | Oneg, v1 :: nil => Some (Val.neg v1)
+ | Osub, v1 :: v2 :: nil => Some (Val.sub v1 v2)
+ | Omul, v1 :: v2 :: nil => Some (Val.mul v1 v2)
+ | Omulhs, v1::v2::nil => Some (Val.mulhs v1 v2)
+ | Omulhu, v1::v2::nil => Some (Val.mulhu v1 v2)
+ | Odiv, v1 :: v2 :: nil => Val.divs v1 v2
+ | Odivu, v1 :: v2 :: nil => Val.divu v1 v2
+ | Omod, v1 :: v2 :: nil => Val.mods v1 v2
+ | Omodu, v1 :: v2 :: nil => Val.modu v1 v2
+ | Oand, v1 :: v2 :: nil => Some (Val.and v1 v2)
+ | Oandimm n, v1 :: nil => Some (Val.and v1 (Vint n))
+ | Oor, v1 :: v2 :: nil => Some (Val.or v1 v2)
+ | Oorimm n, v1 :: nil => Some (Val.or v1 (Vint n))
+ | Oxor, v1 :: v2 :: nil => Some (Val.xor v1 v2)
+ | Oxorimm n, v1 :: nil => Some (Val.xor v1 (Vint n))
+ | Oshl, v1 :: v2 :: nil => Some (Val.shl v1 v2)
+ | Oshlimm n, v1 :: nil => Some (Val.shl v1 (Vint n))
+ | Oshr, v1 :: v2 :: nil => Some (Val.shr v1 v2)
+ | Oshrimm n, v1 :: nil => Some (Val.shr v1 (Vint n))
+ | Oshru, v1 :: v2 :: nil => Some (Val.shru v1 v2)
+ | Oshruimm n, v1 :: nil => Some (Val.shru v1 (Vint n))
+ | Oshrximm n, v1::nil => Val.shrx v1 (Vint n)
+ | Omakelong, v1::v2::nil => Some (Val.longofwords v1 v2)
+ | Olowlong, v1::nil => Some (Val.loword v1)
+ | Ohighlong, v1::nil => Some (Val.hiword v1)
+ | Ocast32signed, v1 :: nil => Some (Val.longofint v1)
+ | Ocast32unsigned, v1 :: nil => Some (Val.longofintu v1)
+ | Oaddl, v1 :: v2 :: nil => Some (Val.addl v1 v2)
+ | Oaddlimm n, v1::nil => Some (Val.addl v1 (Vlong n))
+ | Onegl, v1::nil => Some (Val.negl v1)
+ | Osubl, v1::v2::nil => Some (Val.subl v1 v2)
+ | Omull, v1::v2::nil => Some (Val.mull v1 v2)
+ | Omullhs, v1::v2::nil => Some (Val.mullhs v1 v2)
+ | Omullhu, v1::v2::nil => Some (Val.mullhu v1 v2)
+ | Odivl, v1::v2::nil => Val.divls v1 v2
+ | Odivlu, v1::v2::nil => Val.divlu v1 v2
+ | Omodl, v1::v2::nil => Val.modls v1 v2
+ | Omodlu, v1::v2::nil => Val.modlu v1 v2
+ | Oandl, v1::v2::nil => Some(Val.andl v1 v2)
+ | Oandlimm n, v1::nil => Some (Val.andl v1 (Vlong n))
+ | Oorl, v1::v2::nil => Some(Val.orl v1 v2)
+ | Oorlimm n, v1::nil => Some (Val.orl v1 (Vlong n))
+ | Oxorl, v1::v2::nil => Some(Val.xorl v1 v2)
+ | Oxorlimm n, v1::nil => Some (Val.xorl v1 (Vlong n))
+ | Oshll, v1::v2::nil => Some (Val.shll v1 v2)
+ | Oshllimm n, v1::nil => Some (Val.shll v1 (Vint n))
+ | Oshrl, v1::v2::nil => Some (Val.shrl v1 v2)
+ | Oshrlimm n, v1::nil => Some (Val.shrl v1 (Vint n))
+ | Oshrlu, v1::v2::nil => Some (Val.shrlu v1 v2)
+ | Oshrluimm n, v1::nil => Some (Val.shrlu v1 (Vint n))
+ | Oshrxlimm n, v1::nil => Val.shrxl v1 (Vint n)
+ | Onegf, v1::nil => Some (Val.negf v1)
+ | Oabsf, v1::nil => Some (Val.absf v1)
+ | Oaddf, v1::v2::nil => Some (Val.addf v1 v2)
+ | Osubf, v1::v2::nil => Some (Val.subf v1 v2)
+ | Omulf, v1::v2::nil => Some (Val.mulf v1 v2)
+ | Odivf, v1::v2::nil => Some (Val.divf v1 v2)
+ | Onegfs, v1::nil => Some (Val.negfs v1)
+ | Oabsfs, v1::nil => Some (Val.absfs v1)
+ | Oaddfs, v1::v2::nil => Some (Val.addfs v1 v2)
+ | Osubfs, v1::v2::nil => Some (Val.subfs v1 v2)
+ | Omulfs, v1::v2::nil => Some (Val.mulfs v1 v2)
+ | Odivfs, v1::v2::nil => Some (Val.divfs v1 v2)
+ | Osingleoffloat, v1::nil => Some (Val.singleoffloat v1)
+ | Ofloatofsingle, v1::nil => Some (Val.floatofsingle v1)
+ | Ointoffloat, v1::nil => Val.intoffloat v1
+ | Ointuoffloat, v1::nil => Val.intuoffloat v1
+ | Ofloatofint, v1::nil => Val.floatofint v1
+ | Ofloatofintu, v1::nil => Val.floatofintu v1
+ | Ointofsingle, v1::nil => Val.intofsingle v1
+ | Ointuofsingle, v1::nil => Val.intuofsingle v1
+ | Osingleofint, v1::nil => Val.singleofint v1
+ | Osingleofintu, v1::nil => Val.singleofintu v1
+ | Olongoffloat, v1::nil => Val.longoffloat v1
+ | Olonguoffloat, v1::nil => Val.longuoffloat v1
+ | Ofloatoflong, v1::nil => Val.floatoflong v1
+ | Ofloatoflongu, v1::nil => Val.floatoflongu v1
+ | Olongofsingle, v1::nil => Val.longofsingle v1
+ | Olonguofsingle, v1::nil => Val.longuofsingle v1
+ | Osingleoflong, v1::nil => Val.singleoflong v1
+ | Osingleoflongu, v1::nil => Val.singleoflongu v1
+ | Ocmp c, _ => Some (Val.of_optbool (eval_condition c vl m))
+ | _, _ => None
+ end.
+
+Definition eval_addressing
+ (F V: Type) (genv: Genv.t F V) (sp: val)
+ (addr: addressing) (vl: list val) : option val :=
+ match addr, vl with
+ | Aindexed n, v1 :: nil => Some (Val.offset_ptr v1 n)
+ | Aglobal s ofs, nil => Some (Genv.symbol_address genv s ofs)
+ | Ainstack n, nil => Some (Val.offset_ptr sp n)
+ | _, _ => None
+ end.
+
+Remark eval_addressing_Ainstack:
+ forall (F V: Type) (genv: Genv.t F V) sp ofs,
+ eval_addressing genv sp (Ainstack ofs) nil = Some (Val.offset_ptr sp ofs).
+Proof.
+ intros. reflexivity.
+Qed.
+
+Remark eval_addressing_Ainstack_inv:
+ forall (F V: Type) (genv: Genv.t F V) sp ofs vl v,
+ eval_addressing genv sp (Ainstack ofs) vl = Some v -> vl = nil /\ v = Val.offset_ptr sp ofs.
+Proof.
+ unfold eval_addressing; intros; destruct vl; inv H; auto.
+Qed.
+
+Ltac FuncInv :=
+ match goal with
+ | H: (match ?x with nil => _ | _ :: _ => _ end = Some _) |- _ =>
+ destruct x; simpl in H; FuncInv
+ | H: (match ?v with Vundef => _ | Vint _ => _ | Vfloat _ => _ | Vptr _ _ => _ end = Some _) |- _ =>
+ destruct v; simpl in H; FuncInv
+ | H: (if Archi.ptr64 then _ else _) = Some _ |- _ =>
+ destruct Archi.ptr64 eqn:?; FuncInv
+ | H: (Some _ = Some _) |- _ =>
+ injection H; intros; clear H; FuncInv
+ | H: (None = Some _) |- _ =>
+ discriminate H
+ | _ =>
+ idtac
+ end.
+
+(** * Static typing of conditions, operators and addressing modes. *)
+
+Definition type_of_condition (c: condition) : list typ :=
+ match c with
+ | Ccomp _ => Tint :: Tint :: nil
+ | Ccompu _ => Tint :: Tint :: nil
+ | Ccompimm _ _ => Tint :: nil
+ | Ccompuimm _ _ => Tint :: nil
+ | Ccompl _ => Tlong :: Tlong :: nil
+ | Ccomplu _ => Tlong :: Tlong :: nil
+ | Ccomplimm _ _ => Tlong :: nil
+ | Ccompluimm _ _ => Tlong :: nil
+ | Ccompf _ => Tfloat :: Tfloat :: nil
+ | Cnotcompf _ => Tfloat :: Tfloat :: nil
+ | Ccompfs _ => Tsingle :: Tsingle :: nil
+ | Cnotcompfs _ => Tsingle :: Tsingle :: nil
+ end.
+
+Definition type_of_operation (op: operation) : list typ * typ :=
+ match op with
+ | Omove => (nil, Tint) (* treated specially *)
+ | Ointconst _ => (nil, Tint)
+ | Olongconst _ => (nil, Tlong)
+ | Ofloatconst f => (nil, Tfloat)
+ | Osingleconst f => (nil, Tsingle)
+ | Oaddrsymbol _ _ => (nil, Tptr)
+ | Oaddrstack _ => (nil, Tptr)
+ | Ocast8signed => (Tint :: nil, Tint)
+ | Ocast16signed => (Tint :: nil, Tint)
+ | Oadd => (Tint :: Tint :: nil, Tint)
+ | Oaddimm _ => (Tint :: nil, Tint)
+ | Oneg => (Tint :: nil, Tint)
+ | Osub => (Tint :: Tint :: nil, Tint)
+ | Omul => (Tint :: Tint :: nil, Tint)
+ | Omulhs => (Tint :: Tint :: nil, Tint)
+ | Omulhu => (Tint :: Tint :: nil, Tint)
+ | Odiv => (Tint :: Tint :: nil, Tint)
+ | Odivu => (Tint :: Tint :: nil, Tint)
+ | Omod => (Tint :: Tint :: nil, Tint)
+ | Omodu => (Tint :: Tint :: nil, Tint)
+ | Oand => (Tint :: Tint :: nil, Tint)
+ | Oandimm _ => (Tint :: nil, Tint)
+ | Oor => (Tint :: Tint :: nil, Tint)
+ | Oorimm _ => (Tint :: nil, Tint)
+ | Oxor => (Tint :: Tint :: nil, Tint)
+ | Oxorimm _ => (Tint :: nil, Tint)
+ | Oshl => (Tint :: Tint :: nil, Tint)
+ | Oshlimm _ => (Tint :: nil, Tint)
+ | Oshr => (Tint :: Tint :: nil, Tint)
+ | Oshrimm _ => (Tint :: nil, Tint)
+ | Oshru => (Tint :: Tint :: nil, Tint)
+ | Oshruimm _ => (Tint :: nil, Tint)
+ | Oshrximm _ => (Tint :: nil, Tint)
+ | Omakelong => (Tint :: Tint :: nil, Tlong)
+ | Olowlong => (Tlong :: nil, Tint)
+ | Ohighlong => (Tlong :: nil, Tint)
+ | Ocast32signed => (Tint :: nil, Tlong)
+ | Ocast32unsigned => (Tint :: nil, Tlong)
+ | Oaddl => (Tlong :: Tlong :: nil, Tlong)
+ | Oaddlimm _ => (Tlong :: nil, Tlong)
+ | Onegl => (Tlong :: nil, Tlong)
+ | Osubl => (Tlong :: Tlong :: nil, Tlong)
+ | Omull => (Tlong :: Tlong :: nil, Tlong)
+ | Omullhs => (Tlong :: Tlong :: nil, Tlong)
+ | Omullhu => (Tlong :: Tlong :: nil, Tlong)
+ | Odivl => (Tlong :: Tlong :: nil, Tlong)
+ | Odivlu => (Tlong :: Tlong :: nil, Tlong)
+ | Omodl => (Tlong :: Tlong :: nil, Tlong)
+ | Omodlu => (Tlong :: Tlong :: nil, Tlong)
+ | Oandl => (Tlong :: Tlong :: nil, Tlong)
+ | Oandlimm _ => (Tlong :: nil, Tlong)
+ | Oorl => (Tlong :: Tlong :: nil, Tlong)
+ | Oorlimm _ => (Tlong :: nil, Tlong)
+ | Oxorl => (Tlong :: Tlong :: nil, Tlong)
+ | Oxorlimm _ => (Tlong :: nil, Tlong)
+ | Oshll => (Tlong :: Tint :: nil, Tlong)
+ | Oshllimm _ => (Tlong :: nil, Tlong)
+ | Oshrl => (Tlong :: Tint :: nil, Tlong)
+ | Oshrlimm _ => (Tlong :: nil, Tlong)
+ | Oshrlu => (Tlong :: Tint :: nil, Tlong)
+ | Oshrluimm _ => (Tlong :: nil, Tlong)
+ | Oshrxlimm _ => (Tlong :: nil, Tlong)
+ | Onegf => (Tfloat :: nil, Tfloat)
+ | Oabsf => (Tfloat :: nil, Tfloat)
+ | Oaddf => (Tfloat :: Tfloat :: nil, Tfloat)
+ | Osubf => (Tfloat :: Tfloat :: nil, Tfloat)
+ | Omulf => (Tfloat :: Tfloat :: nil, Tfloat)
+ | Odivf => (Tfloat :: Tfloat :: nil, Tfloat)
+ | Onegfs => (Tsingle :: nil, Tsingle)
+ | Oabsfs => (Tsingle :: nil, Tsingle)
+ | Oaddfs => (Tsingle :: Tsingle :: nil, Tsingle)
+ | Osubfs => (Tsingle :: Tsingle :: nil, Tsingle)
+ | Omulfs => (Tsingle :: Tsingle :: nil, Tsingle)
+ | Odivfs => (Tsingle :: Tsingle :: nil, Tsingle)
+ | Osingleoffloat => (Tfloat :: nil, Tsingle)
+ | Ofloatofsingle => (Tsingle :: nil, Tfloat)
+ | Ointoffloat => (Tfloat :: nil, Tint)
+ | Ointuoffloat => (Tfloat :: nil, Tint)
+ | Ofloatofint => (Tint :: nil, Tfloat)
+ | Ofloatofintu => (Tint :: nil, Tfloat)
+ | Ointofsingle => (Tsingle :: nil, Tint)
+ | Ointuofsingle => (Tsingle :: nil, Tint)
+ | Osingleofint => (Tint :: nil, Tsingle)
+ | Osingleofintu => (Tint :: nil, Tsingle)
+ | Olongoffloat => (Tfloat :: nil, Tlong)
+ | Olonguoffloat => (Tfloat :: nil, Tlong)
+ | Ofloatoflong => (Tlong :: nil, Tfloat)
+ | Ofloatoflongu => (Tlong :: nil, Tfloat)
+ | Olongofsingle => (Tsingle :: nil, Tlong)
+ | Olonguofsingle => (Tsingle :: nil, Tlong)
+ | Osingleoflong => (Tlong :: nil, Tsingle)
+ | Osingleoflongu => (Tlong :: nil, Tsingle)
+ | Ocmp c => (type_of_condition c, Tint)
+ end.
+
+Definition type_of_addressing (addr: addressing) : list typ :=
+ match addr with
+ | Aindexed _ => Tptr :: nil
+ | Aglobal _ _ => nil
+ | Ainstack _ => nil
+ end.
+
+(** Weak type soundness results for [eval_operation]:
+ the result values, when defined, are always of the type predicted
+ by [type_of_operation]. *)
+
+Section SOUNDNESS.
+
+Variable A V: Type.
+Variable genv: Genv.t A V.
+
+Remark type_add:
+ forall v1 v2, Val.has_type (Val.add v1 v2) Tint.
+Proof.
+ intros. unfold Val.has_type, Val.add. destruct Archi.ptr64, v1, v2; auto.
+Qed.
+
+Remark type_addl:
+ forall v1 v2, Val.has_type (Val.addl v1 v2) Tlong.
+Proof.
+ intros. unfold Val.has_type, Val.addl. destruct Archi.ptr64, v1, v2; auto.
+Qed.
+
+Lemma type_of_operation_sound:
+ forall op vl sp v m,
+ op <> Omove ->
+ eval_operation genv sp op vl m = Some v ->
+ Val.has_type v (snd (type_of_operation op)).
+Proof with (try exact I; try reflexivity; auto using Val.Vptr_has_type).
+ intros.
+ destruct op; simpl; simpl in H0; FuncInv; subst; simpl.
+ (* move *)
+ - congruence.
+ (* intconst, longconst, floatconst, singleconst *)
+ - exact I.
+ - exact I.
+ - exact I.
+ - exact I.
+ (* addrsymbol *)
+ - unfold Genv.symbol_address. destruct (Genv.find_symbol genv id)...
+ (* addrstack *)
+ - destruct sp... apply Val.Vptr_has_type.
+ (* castsigned *)
+ - destruct v0...
+ - destruct v0...
+ (* add, addimm *)
+ - apply type_add.
+ - apply type_add.
+ (* neg, sub *)
+ - destruct v0...
+ - unfold Val.sub. destruct v0; destruct v1...
+ unfold Val.has_type; destruct Archi.ptr64...
+ destruct Archi.ptr64... destruct (eq_block b b0)...
+ (* mul, mulhs, mulhu *)
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ (* div, divu *)
+ - destruct v0; destruct v1; simpl in *; inv H0.
+ destruct (Int.eq i0 Int.zero || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2...
+ - destruct v0; destruct v1; simpl in *; inv H0.
+ destruct (Int.eq i0 Int.zero); inv H2...
+ (* mod, modu *)
+ - destruct v0; destruct v1; simpl in *; inv H0.
+ destruct (Int.eq i0 Int.zero || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2...
+ - destruct v0; destruct v1; simpl in *; inv H0.
+ destruct (Int.eq i0 Int.zero); inv H2...
+ (* and, andimm *)
+ - destruct v0; destruct v1...
+ - destruct v0...
+ (* or, orimm *)
+ - destruct v0; destruct v1...
+ - destruct v0...
+ (* xor, xorimm *)
+ - destruct v0; destruct v1...
+ - destruct v0...
+ (* shl, shlimm *)
+ - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
+ - destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
+ (* shr, shrimm *)
+ - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
+ - destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
+ (* shru, shruimm *)
+ - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int.iwordsize)...
+ - destruct v0; simpl... destruct (Int.ltu n Int.iwordsize)...
+ (* shrx *)
+ - destruct v0; simpl in H0; try discriminate. destruct (Int.ltu n (Int.repr 31)); inv H0...
+ (* makelong, lowlong, highlong *)
+ - destruct v0; destruct v1...
+ - destruct v0...
+ - destruct v0...
+ (* cast32 *)
+ - destruct v0...
+ - destruct v0...
+ (* addl, addlimm *)
+ - apply type_addl.
+ - apply type_addl.
+ (* negl, subl *)
+ - destruct v0...
+ - unfold Val.subl. destruct v0; destruct v1...
+ unfold Val.has_type; destruct Archi.ptr64...
+ destruct Archi.ptr64... destruct (eq_block b b0)...
+ (* mull, mullhs, mullhu *)
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ (* divl, divlu *)
+ - destruct v0; destruct v1; simpl in *; inv H0.
+ destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2...
+ - destruct v0; destruct v1; simpl in *; inv H0.
+ destruct (Int64.eq i0 Int64.zero); inv H2...
+ (* modl, modlu *)
+ - destruct v0; destruct v1; simpl in *; inv H0.
+ destruct (Int64.eq i0 Int64.zero || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2...
+ - destruct v0; destruct v1; simpl in *; inv H0.
+ destruct (Int64.eq i0 Int64.zero); inv H2...
+ (* andl, andlimm *)
+ - destruct v0; destruct v1...
+ - destruct v0...
+ (* orl, orlimm *)
+ - destruct v0; destruct v1...
+ - destruct v0...
+ (* xorl, xorlimm *)
+ - destruct v0; destruct v1...
+ - destruct v0...
+ (* shll, shllimm *)
+ - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
+ - destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
+ (* shr, shrimm *)
+ - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
+ - destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
+ (* shru, shruimm *)
+ - destruct v0; destruct v1; simpl... destruct (Int.ltu i0 Int64.iwordsize')...
+ - destruct v0; simpl... destruct (Int.ltu n Int64.iwordsize')...
+ (* shrxl *)
+ - destruct v0; simpl in H0; try discriminate. destruct (Int.ltu n (Int.repr 63)); inv H0...
+ (* negf, absf *)
+ - destruct v0...
+ - destruct v0...
+ (* addf, subf *)
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ (* mulf, divf *)
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ (* negfs, absfs *)
+ - destruct v0...
+ - destruct v0...
+ (* addfs, subfs *)
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ (* mulfs, divfs *)
+ - destruct v0; destruct v1...
+ - destruct v0; destruct v1...
+ (* singleoffloat, floatofsingle *)
+ - destruct v0...
+ - destruct v0...
+ (* intoffloat, intuoffloat *)
+ - destruct v0; simpl in H0; inv H0. destruct (Float.to_int f); inv H2...
+ - destruct v0; simpl in H0; inv H0. destruct (Float.to_intu f); inv H2...
+ (* floatofint, floatofintu *)
+ - destruct v0; simpl in H0; inv H0...
+ - destruct v0; simpl in H0; inv H0...
+ (* intofsingle, intuofsingle *)
+ - destruct v0; simpl in H0; inv H0. destruct (Float32.to_int f); inv H2...
+ - destruct v0; simpl in H0; inv H0. destruct (Float32.to_intu f); inv H2...
+ (* singleofint, singleofintu *)
+ - destruct v0; simpl in H0; inv H0...
+ - destruct v0; simpl in H0; inv H0...
+ (* longoffloat, longuoffloat *)
+ - destruct v0; simpl in H0; inv H0. destruct (Float.to_long f); inv H2...
+ - destruct v0; simpl in H0; inv H0. destruct (Float.to_longu f); inv H2...
+ (* floatoflong, floatoflongu *)
+ - destruct v0; simpl in H0; inv H0...
+ - destruct v0; simpl in H0; inv H0...
+ (* longofsingle, longuofsingle *)
+ - destruct v0; simpl in H0; inv H0. destruct (Float32.to_long f); inv H2...
+ - destruct v0; simpl in H0; inv H0. destruct (Float32.to_longu f); inv H2...
+ (* singleoflong, singleoflongu *)
+ - destruct v0; simpl in H0; inv H0...
+ - destruct v0; simpl in H0; inv H0...
+ (* cmp *)
+ - destruct (eval_condition cond vl m)... destruct b...
+Qed.
+
+End SOUNDNESS.
+
+(** * Manipulating and transforming operations *)
+
+(** Recognition of move operations. *)
+
+Definition is_move_operation
+ (A: Type) (op: operation) (args: list A) : option A :=
+ match op, args with
+ | Omove, arg :: nil => Some arg
+ | _, _ => None
+ end.
+
+Lemma is_move_operation_correct:
+ forall (A: Type) (op: operation) (args: list A) (a: A),
+ is_move_operation op args = Some a ->
+ op = Omove /\ args = a :: nil.
+Proof.
+ intros until a. unfold is_move_operation; destruct op;
+ try (intros; discriminate).
+ destruct args. intros; discriminate.
+ destruct args. intros. intuition congruence.
+ intros; discriminate.
+Qed.
+
+(** [negate_condition cond] returns a condition that is logically
+ equivalent to the negation of [cond]. *)
+
+Definition negate_condition (cond: condition): condition :=
+ match cond with
+ | Ccomp c => Ccomp(negate_comparison c)
+ | Ccompu c => Ccompu(negate_comparison c)
+ | Ccompimm c n => Ccompimm (negate_comparison c) n
+ | Ccompuimm c n => Ccompuimm (negate_comparison c) n
+ | Ccompl c => Ccompl(negate_comparison c)
+ | Ccomplu c => Ccomplu(negate_comparison c)
+ | Ccomplimm c n => Ccomplimm (negate_comparison c) n
+ | Ccompluimm c n => Ccompluimm (negate_comparison c) n
+ | Ccompf c => Cnotcompf c
+ | Cnotcompf c => Ccompf c
+ | Ccompfs c => Cnotcompfs c
+ | Cnotcompfs c => Ccompfs c
+ end.
+
+Lemma eval_negate_condition:
+ forall cond vl m,
+ eval_condition (negate_condition cond) vl m = option_map negb (eval_condition cond vl m).
+Proof.
+ intros. destruct cond; simpl.
+ repeat (destruct vl; auto). apply Val.negate_cmp_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmpu_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmp_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmpu_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmpl_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmplu_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmpl_bool.
+ repeat (destruct vl; auto). apply Val.negate_cmplu_bool.
+ repeat (destruct vl; auto).
+ repeat (destruct vl; auto). destruct (Val.cmpf_bool c v v0) as [[]|]; auto.
+ repeat (destruct vl; auto).
+ repeat (destruct vl; auto). destruct (Val.cmpfs_bool c v v0) as [[]|]; auto.
+Qed.
+
+(** Shifting stack-relative references. This is used in [Stacking]. *)
+
+Definition shift_stack_addressing (delta: Z) (addr: addressing) :=
+ match addr with
+ | Ainstack ofs => Ainstack (Ptrofs.add ofs (Ptrofs.repr delta))
+ | _ => addr
+ end.
+
+Definition shift_stack_operation (delta: Z) (op: operation) :=
+ match op with
+ | Oaddrstack ofs => Oaddrstack (Ptrofs.add ofs (Ptrofs.repr delta))
+ | _ => op
+ end.
+
+Lemma type_shift_stack_addressing:
+ forall delta addr, type_of_addressing (shift_stack_addressing delta addr) = type_of_addressing addr.
+Proof.
+ intros. destruct addr; auto.
+Qed.
+
+Lemma type_shift_stack_operation:
+ forall delta op, type_of_operation (shift_stack_operation delta op) = type_of_operation op.
+Proof.
+ intros. destruct op; auto.
+Qed.
+
+Lemma eval_shift_stack_addressing:
+ forall F V (ge: Genv.t F V) sp addr vl delta,
+ eval_addressing ge (Vptr sp Ptrofs.zero) (shift_stack_addressing delta addr) vl =
+ eval_addressing ge (Vptr sp (Ptrofs.repr delta)) addr vl.
+Proof.
+ intros. destruct addr; simpl; auto. destruct vl; auto.
+ rewrite Ptrofs.add_zero_l, Ptrofs.add_commut; auto.
+Qed.
+
+Lemma eval_shift_stack_operation:
+ forall F V (ge: Genv.t F V) sp op vl m delta,
+ eval_operation ge (Vptr sp Ptrofs.zero) (shift_stack_operation delta op) vl m =
+ eval_operation ge (Vptr sp (Ptrofs.repr delta)) op vl m.
+Proof.
+ intros. destruct op; simpl; auto. destruct vl; auto.
+ rewrite Ptrofs.add_zero_l, Ptrofs.add_commut; auto.
+Qed.
+
+(** Offset an addressing mode [addr] by a quantity [delta], so that
+ it designates the pointer [delta] bytes past the pointer designated
+ by [addr]. May be undefined, in which case [None] is returned. *)
+
+Definition offset_addressing (addr: addressing) (delta: Z) : option addressing :=
+ match addr with
+ | Aindexed n => Some(Aindexed (Ptrofs.add n (Ptrofs.repr delta)))
+ | Aglobal id n => Some(Aglobal id (Ptrofs.add n (Ptrofs.repr delta)))
+ | Ainstack n => Some(Ainstack (Ptrofs.add n (Ptrofs.repr delta)))
+ end.
+
+Lemma eval_offset_addressing:
+ forall (F V: Type) (ge: Genv.t F V) sp addr args delta addr' v,
+ offset_addressing addr delta = Some addr' ->
+ eval_addressing ge sp addr args = Some v ->
+ Archi.ptr64 = false ->
+ eval_addressing ge sp addr' args = Some(Val.add v (Vint (Int.repr delta))).
+Proof.
+ intros.
+ assert (A: forall x n,
+ Val.offset_ptr x (Ptrofs.add n (Ptrofs.repr delta)) =
+ Val.add (Val.offset_ptr x n) (Vint (Int.repr delta))).
+ { intros; destruct x; simpl; auto. rewrite H1.
+ rewrite Ptrofs.add_assoc. f_equal; f_equal; f_equal. symmetry; auto with ptrofs. }
+ destruct addr; simpl in H; inv H; simpl in *; FuncInv; subst.
+- rewrite A; auto.
+- unfold Genv.symbol_address. destruct (Genv.find_symbol ge i); auto.
+ simpl. rewrite H1. f_equal; f_equal; f_equal. symmetry; auto with ptrofs.
+- rewrite A; auto.
+Qed.
+
+(** Operations that are so cheap to recompute that CSE should not factor them out. *)
+
+Definition is_trivial_op (op: operation) : bool :=
+ match op with
+ | Omove => true
+ | Ointconst n => Int.eq (Int.sign_ext 12 n) n
+ | Olongconst n => Int64.eq (Int64.sign_ext 12 n) n
+ | Oaddrstack _ => true
+ | _ => false
+ end.
+
+(** Operations that depend on the memory state. *)
+
+Definition op_depends_on_memory (op: operation) : bool :=
+ match op with
+ | Ocmp (Ccompu _) => negb Archi.ptr64
+ | Ocmp (Ccompuimm _ _) => negb Archi.ptr64
+ | Ocmp (Ccomplu _) => Archi.ptr64
+ | Ocmp (Ccompluimm _ _) => Archi.ptr64
+ | _ => false
+ end.
+
+Lemma op_depends_on_memory_correct:
+ forall (F V: Type) (ge: Genv.t F V) sp op args m1 m2,
+ op_depends_on_memory op = false ->
+ eval_operation ge sp op args m1 = eval_operation ge sp op args m2.
+Proof.
+ intros until m2. destruct op; simpl; try congruence.
+ destruct cond; simpl; intros SF; auto; rewrite ? negb_false_iff in SF;
+ unfold Val.cmpu_bool, Val.cmplu_bool; rewrite SF; reflexivity.
+Qed.
+
+(** Global variables mentioned in an operation or addressing mode *)
+
+Definition globals_addressing (addr: addressing) : list ident :=
+ match addr with
+ | Aglobal s ofs => s :: nil
+ | _ => nil
+ end.
+
+Definition globals_operation (op: operation) : list ident :=
+ match op with
+ | Oaddrsymbol s ofs => s :: nil
+ | _ => nil
+ end.
+
+(** * Invariance and compatibility properties. *)
+
+(** [eval_operation] and [eval_addressing] depend on a global environment
+ for resolving references to global symbols. We show that they give
+ the same results if a global environment is replaced by another that
+ assigns the same addresses to the same symbols. *)
+
+Section GENV_TRANSF.
+
+Variable F1 F2 V1 V2: Type.
+Variable ge1: Genv.t F1 V1.
+Variable ge2: Genv.t F2 V2.
+Hypothesis agree_on_symbols:
+ forall (s: ident), Genv.find_symbol ge2 s = Genv.find_symbol ge1 s.
+
+Lemma eval_addressing_preserved:
+ forall sp addr vl,
+ eval_addressing ge2 sp addr vl = eval_addressing ge1 sp addr vl.
+Proof.
+ intros.
+ unfold eval_addressing; destruct addr; auto. destruct vl; auto.
+ unfold Genv.symbol_address. rewrite agree_on_symbols; auto.
+Qed.
+
+Lemma eval_operation_preserved:
+ forall sp op vl m,
+ eval_operation ge2 sp op vl m = eval_operation ge1 sp op vl m.
+Proof.
+ intros.
+ unfold eval_operation; destruct op; auto. destruct vl; auto.
+ unfold Genv.symbol_address. rewrite agree_on_symbols; auto.
+Qed.
+
+End GENV_TRANSF.
+
+(** Compatibility of the evaluation functions with value injections. *)
+
+Section EVAL_COMPAT.
+
+Variable F1 F2 V1 V2: Type.
+Variable ge1: Genv.t F1 V1.
+Variable ge2: Genv.t F2 V2.
+Variable f: meminj.
+
+Variable m1: mem.
+Variable m2: mem.
+
+Hypothesis valid_pointer_inj:
+ forall b1 ofs b2 delta,
+ f b1 = Some(b2, delta) ->
+ Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ Mem.valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
+
+Hypothesis weak_valid_pointer_inj:
+ forall b1 ofs b2 delta,
+ f b1 = Some(b2, delta) ->
+ Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ Mem.weak_valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
+
+Hypothesis weak_valid_pointer_no_overflow:
+ forall b1 ofs b2 delta,
+ f b1 = Some(b2, delta) ->
+ Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ 0 <= Ptrofs.unsigned ofs + Ptrofs.unsigned (Ptrofs.repr delta) <= Ptrofs.max_unsigned.
+
+Hypothesis valid_different_pointers_inj:
+ forall b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
+ b1 <> b2 ->
+ Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs1) = true ->
+ Mem.valid_pointer m1 b2 (Ptrofs.unsigned ofs2) = true ->
+ f b1 = Some (b1', delta1) ->
+ f b2 = Some (b2', delta2) ->
+ b1' <> b2' \/
+ Ptrofs.unsigned (Ptrofs.add ofs1 (Ptrofs.repr delta1)) <> Ptrofs.unsigned (Ptrofs.add ofs2 (Ptrofs.repr delta2)).
+
+Ltac InvInject :=
+ match goal with
+ | [ H: Val.inject _ (Vint _) _ |- _ ] =>
+ inv H; InvInject
+ | [ H: Val.inject _ (Vfloat _) _ |- _ ] =>
+ inv H; InvInject
+ | [ H: Val.inject _ (Vptr _ _) _ |- _ ] =>
+ inv H; InvInject
+ | [ H: Val.inject_list _ nil _ |- _ ] =>
+ inv H; InvInject
+ | [ H: Val.inject_list _ (_ :: _) _ |- _ ] =>
+ inv H; InvInject
+ | _ => idtac
+ end.
+
+Lemma eval_condition_inj:
+ forall cond vl1 vl2 b,
+ Val.inject_list f vl1 vl2 ->
+ eval_condition cond vl1 m1 = Some b ->
+ eval_condition cond vl2 m2 = Some b.
+Proof.
+ intros. destruct cond; simpl in H0; FuncInv; InvInject; simpl; auto.
+- inv H3; inv H2; simpl in H0; inv H0; auto.
+- eauto 3 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
+- inv H3; simpl in H0; inv H0; auto.
+- eauto 3 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
+- inv H3; inv H2; simpl in H0; inv H0; auto.
+- eauto 3 using Val.cmplu_bool_inject, Mem.valid_pointer_implies.
+- inv H3; simpl in H0; inv H0; auto.
+- eauto 3 using Val.cmplu_bool_inject, Mem.valid_pointer_implies.
+- inv H3; inv H2; simpl in H0; inv H0; auto.
+- inv H3; inv H2; simpl in H0; inv H0; auto.
+- inv H3; inv H2; simpl in H0; inv H0; auto.
+- inv H3; inv H2; simpl in H0; inv H0; auto.
+Qed.
+
+Ltac TrivialExists :=
+ match goal with
+ | [ |- exists v2, Some ?v1 = Some v2 /\ Val.inject _ _ v2 ] =>
+ exists v1; split; auto
+ | _ => idtac
+ end.
+
+Lemma eval_operation_inj:
+ forall op sp1 vl1 sp2 vl2 v1,
+ (forall id ofs,
+ In id (globals_operation op) ->
+ Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
+ Val.inject f sp1 sp2 ->
+ Val.inject_list f vl1 vl2 ->
+ eval_operation ge1 sp1 op vl1 m1 = Some v1 ->
+ exists v2, eval_operation ge2 sp2 op vl2 m2 = Some v2 /\ Val.inject f v1 v2.
+Proof.
+ intros until v1; intros GL; intros. destruct op; simpl in H1; simpl; FuncInv; InvInject; TrivialExists.
+ (* addrsymbol *)
+ - apply GL; simpl; auto.
+ (* addrstack *)
+ - apply Val.offset_ptr_inject; auto.
+ (* castsigned *)
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
+ (* add, addimm *)
+ - apply Val.add_inject; auto.
+ - apply Val.add_inject; auto.
+ (* neg, sub *)
+ - inv H4; simpl; auto.
+ - apply Val.sub_inject; auto.
+ (* mul, mulhs, mulhu *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ (* div, divu *)
+ - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ destruct (Int.eq i0 Int.zero
+ || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2.
+ TrivialExists.
+ - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ destruct (Int.eq i0 Int.zero); inv H2. TrivialExists.
+ (* mod, modu *)
+ - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ destruct (Int.eq i0 Int.zero
+ || Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H2.
+ TrivialExists.
+ - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ destruct (Int.eq i0 Int.zero); inv H2. TrivialExists.
+ (* and, andimm *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; simpl; auto.
+ (* or, orimm *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; simpl; auto.
+ (* xor, xorimm *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; simpl; auto.
+ (* shl, shlimm *)
+ - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
+ - inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
+ (* shr, shrimm *)
+ - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
+ - inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
+ (* shru, shruimm *)
+ - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int.iwordsize); auto.
+ - inv H4; simpl; auto. destruct (Int.ltu n Int.iwordsize); auto.
+ (* shrx *)
+ - inv H4; simpl in H1; try discriminate. simpl.
+ destruct (Int.ltu n (Int.repr 31)); inv H1. TrivialExists.
+ (* makelong, highlong, lowlong *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
+ (* cast32 *)
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
+ (* addl, addlimm *)
+ - apply Val.addl_inject; auto.
+ - apply Val.addl_inject; auto.
+ (* negl, subl *)
+ - inv H4; simpl; auto.
+ - apply Val.subl_inject; auto.
+ (* mull, mullhs, mullhu *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ (* divl, divlu *)
+ - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ destruct (Int64.eq i0 Int64.zero
+ || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2.
+ TrivialExists.
+ - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ destruct (Int64.eq i0 Int64.zero); inv H2. TrivialExists.
+ (* modl, modlu *)
+ - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ destruct (Int64.eq i0 Int64.zero
+ || Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H2.
+ TrivialExists.
+ - inv H4; inv H3; simpl in H1; inv H1. simpl.
+ destruct (Int64.eq i0 Int64.zero); inv H2. TrivialExists.
+ (* andl, andlimm *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; simpl; auto.
+ (* orl, orlimm *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; simpl; auto.
+ (* xorl, xorlimm *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; simpl; auto.
+ (* shll, shllimm *)
+ - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
+ - inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
+ (* shr, shrimm *)
+ - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
+ - inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
+ (* shru, shruimm *)
+ - inv H4; inv H2; simpl; auto. destruct (Int.ltu i0 Int64.iwordsize'); auto.
+ - inv H4; simpl; auto. destruct (Int.ltu n Int64.iwordsize'); auto.
+ (* shrx *)
+ - inv H4; simpl in H1; try discriminate. simpl.
+ destruct (Int.ltu n (Int.repr 63)); inv H1. TrivialExists.
+ (* negf, absf *)
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
+ (* addf, subf *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ (* mulf, divf *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ (* negfs, absfs *)
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
+ (* addfs, subfs *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ (* mulfs, divfs *)
+ - inv H4; inv H2; simpl; auto.
+ - inv H4; inv H2; simpl; auto.
+ (* singleoffloat, floatofsingle *)
+ - inv H4; simpl; auto.
+ - inv H4; simpl; auto.
+ (* intoffloat, intuoffloat *)
+ - inv H4; simpl in H1; inv H1. simpl. destruct (Float.to_int f0); simpl in H2; inv H2.
+ exists (Vint i); auto.
+ - inv H4; simpl in H1; inv H1. simpl. destruct (Float.to_intu f0); simpl in H2; inv H2.
+ exists (Vint i); auto.
+ (* floatofint, floatofintu *)
+ - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ (* intofsingle, intuofsingle *)
+ - inv H4; simpl in H1; inv H1. simpl. destruct (Float32.to_int f0); simpl in H2; inv H2.
+ exists (Vint i); auto.
+ - inv H4; simpl in H1; inv H1. simpl. destruct (Float32.to_intu f0); simpl in H2; inv H2.
+ exists (Vint i); auto.
+ (* singleofint, singleofintu *)
+ - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ (* longoffloat, longuoffloat *)
+ - inv H4; simpl in H1; inv H1. simpl. destruct (Float.to_long f0); simpl in H2; inv H2.
+ exists (Vlong i); auto.
+ - inv H4; simpl in H1; inv H1. simpl. destruct (Float.to_longu f0); simpl in H2; inv H2.
+ exists (Vlong i); auto.
+ (* floatoflong, floatoflongu *)
+ - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ (* longofsingle, longuofsingle *)
+ - inv H4; simpl in H1; inv H1. simpl. destruct (Float32.to_long f0); simpl in H2; inv H2.
+ exists (Vlong i); auto.
+ - inv H4; simpl in H1; inv H1. simpl. destruct (Float32.to_longu f0); simpl in H2; inv H2.
+ exists (Vlong i); auto.
+ (* singleoflong, singleoflongu *)
+ - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ - inv H4; simpl in H1; inv H1. simpl. TrivialExists.
+ (* cmp *)
+ - subst v1. destruct (eval_condition cond vl1 m1) eqn:?.
+ exploit eval_condition_inj; eauto. intros EQ; rewrite EQ.
+ destruct b; simpl; constructor.
+ simpl; constructor.
+Qed.
+
+Lemma eval_addressing_inj:
+ forall addr sp1 vl1 sp2 vl2 v1,
+ (forall id ofs,
+ In id (globals_addressing addr) ->
+ Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
+ Val.inject f sp1 sp2 ->
+ Val.inject_list f vl1 vl2 ->
+ eval_addressing ge1 sp1 addr vl1 = Some v1 ->
+ exists v2, eval_addressing ge2 sp2 addr vl2 = Some v2 /\ Val.inject f v1 v2.
+Proof.
+ intros. destruct addr; simpl in H2; simpl; FuncInv; InvInject; TrivialExists.
+ apply Val.offset_ptr_inject; auto.
+ apply H; simpl; auto.
+ apply Val.offset_ptr_inject; auto.
+Qed.
+
+End EVAL_COMPAT.
+
+(** Compatibility of the evaluation functions with the ``is less defined'' relation over values. *)
+
+Section EVAL_LESSDEF.
+
+Variable F V: Type.
+Variable genv: Genv.t F V.
+
+Remark valid_pointer_extends:
+ forall m1 m2, Mem.extends m1 m2 ->
+ forall b1 ofs b2 delta,
+ Some(b1, 0) = Some(b2, delta) ->
+ Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ Mem.valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
+Proof.
+ intros. inv H0. rewrite Ptrofs.add_zero. eapply Mem.valid_pointer_extends; eauto.
+Qed.
+
+Remark weak_valid_pointer_extends:
+ forall m1 m2, Mem.extends m1 m2 ->
+ forall b1 ofs b2 delta,
+ Some(b1, 0) = Some(b2, delta) ->
+ Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ Mem.weak_valid_pointer m2 b2 (Ptrofs.unsigned (Ptrofs.add ofs (Ptrofs.repr delta))) = true.
+Proof.
+ intros. inv H0. rewrite Ptrofs.add_zero. eapply Mem.weak_valid_pointer_extends; eauto.
+Qed.
+
+Remark weak_valid_pointer_no_overflow_extends:
+ forall m1 b1 ofs b2 delta,
+ Some(b1, 0) = Some(b2, delta) ->
+ Mem.weak_valid_pointer m1 b1 (Ptrofs.unsigned ofs) = true ->
+ 0 <= Ptrofs.unsigned ofs + Ptrofs.unsigned (Ptrofs.repr delta) <= Ptrofs.max_unsigned.
+Proof.
+ intros. inv H. rewrite Zplus_0_r. apply Ptrofs.unsigned_range_2.
+Qed.
+
+Remark valid_different_pointers_extends:
+ forall m1 b1 ofs1 b2 ofs2 b1' delta1 b2' delta2,
+ b1 <> b2 ->
+ Mem.valid_pointer m1 b1 (Ptrofs.unsigned ofs1) = true ->
+ Mem.valid_pointer m1 b2 (Ptrofs.unsigned ofs2) = true ->
+ Some(b1, 0) = Some (b1', delta1) ->
+ Some(b2, 0) = Some (b2', delta2) ->
+ b1' <> b2' \/
+ Ptrofs.unsigned(Ptrofs.add ofs1 (Ptrofs.repr delta1)) <> Ptrofs.unsigned(Ptrofs.add ofs2 (Ptrofs.repr delta2)).
+Proof.
+ intros. inv H2; inv H3. auto.
+Qed.
+
+Lemma eval_condition_lessdef:
+ forall cond vl1 vl2 b m1 m2,
+ Val.lessdef_list vl1 vl2 ->
+ Mem.extends m1 m2 ->
+ eval_condition cond vl1 m1 = Some b ->
+ eval_condition cond vl2 m2 = Some b.
+Proof.
+ intros. eapply eval_condition_inj with (f := fun b => Some(b, 0)) (m1 := m1).
+ apply valid_pointer_extends; auto.
+ apply weak_valid_pointer_extends; auto.
+ apply weak_valid_pointer_no_overflow_extends.
+ apply valid_different_pointers_extends; auto.
+ rewrite <- val_inject_list_lessdef. eauto. auto.
+Qed.
+
+Lemma eval_operation_lessdef:
+ forall sp op vl1 vl2 v1 m1 m2,
+ Val.lessdef_list vl1 vl2 ->
+ Mem.extends m1 m2 ->
+ eval_operation genv sp op vl1 m1 = Some v1 ->
+ exists v2, eval_operation genv sp op vl2 m2 = Some v2 /\ Val.lessdef v1 v2.
+Proof.
+ intros. rewrite val_inject_list_lessdef in H.
+ assert (exists v2 : val,
+ eval_operation genv sp op vl2 m2 = Some v2
+ /\ Val.inject (fun b => Some(b, 0)) v1 v2).
+ eapply eval_operation_inj with (m1 := m1) (sp1 := sp).
+ apply valid_pointer_extends; auto.
+ apply weak_valid_pointer_extends; auto.
+ apply weak_valid_pointer_no_overflow_extends.
+ apply valid_different_pointers_extends; auto.
+ intros. apply val_inject_lessdef. auto.
+ apply val_inject_lessdef; auto.
+ eauto.
+ auto.
+ destruct H2 as [v2 [A B]]. exists v2; split; auto. rewrite val_inject_lessdef; auto.
+Qed.
+
+Lemma eval_addressing_lessdef:
+ forall sp addr vl1 vl2 v1,
+ Val.lessdef_list vl1 vl2 ->
+ eval_addressing genv sp addr vl1 = Some v1 ->
+ exists v2, eval_addressing genv sp addr vl2 = Some v2 /\ Val.lessdef v1 v2.
+Proof.
+ intros. rewrite val_inject_list_lessdef in H.
+ assert (exists v2 : val,
+ eval_addressing genv sp addr vl2 = Some v2
+ /\ Val.inject (fun b => Some(b, 0)) v1 v2).
+ eapply eval_addressing_inj with (sp1 := sp).
+ intros. rewrite <- val_inject_lessdef; auto.
+ rewrite <- val_inject_lessdef; auto.
+ eauto. auto.
+ destruct H1 as [v2 [A B]]. exists v2; split; auto. rewrite val_inject_lessdef; auto.
+Qed.
+
+End EVAL_LESSDEF.
+
+(** Compatibility of the evaluation functions with memory injections. *)
+
+Section EVAL_INJECT.
+
+Variable F V: Type.
+Variable genv: Genv.t F V.
+Variable f: meminj.
+Hypothesis globals: meminj_preserves_globals genv f.
+Variable sp1: block.
+Variable sp2: block.
+Variable delta: Z.
+Hypothesis sp_inj: f sp1 = Some(sp2, delta).
+
+Remark symbol_address_inject:
+ forall id ofs, Val.inject f (Genv.symbol_address genv id ofs) (Genv.symbol_address genv id ofs).
+Proof.
+ intros. unfold Genv.symbol_address. destruct (Genv.find_symbol genv id) eqn:?; auto.
+ exploit (proj1 globals); eauto. intros.
+ econstructor; eauto. rewrite Ptrofs.add_zero; auto.
+Qed.
+
+Lemma eval_condition_inject:
+ forall cond vl1 vl2 b m1 m2,
+ Val.inject_list f vl1 vl2 ->
+ Mem.inject f m1 m2 ->
+ eval_condition cond vl1 m1 = Some b ->
+ eval_condition cond vl2 m2 = Some b.
+Proof.
+ intros. eapply eval_condition_inj with (f := f) (m1 := m1); eauto.
+ intros; eapply Mem.valid_pointer_inject_val; eauto.
+ intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
+ intros; eapply Mem.weak_valid_pointer_inject_no_overflow; eauto.
+ intros; eapply Mem.different_pointers_inject; eauto.
+Qed.
+
+Lemma eval_addressing_inject:
+ forall addr vl1 vl2 v1,
+ Val.inject_list f vl1 vl2 ->
+ eval_addressing genv (Vptr sp1 Ptrofs.zero) addr vl1 = Some v1 ->
+ exists v2,
+ eval_addressing genv (Vptr sp2 Ptrofs.zero) (shift_stack_addressing delta addr) vl2 = Some v2
+ /\ Val.inject f v1 v2.
+Proof.
+ intros.
+ rewrite eval_shift_stack_addressing.
+ eapply eval_addressing_inj with (sp1 := Vptr sp1 Ptrofs.zero); eauto.
+ intros. apply symbol_address_inject.
+ econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
+Qed.
+
+Lemma eval_operation_inject:
+ forall op vl1 vl2 v1 m1 m2,
+ Val.inject_list f vl1 vl2 ->
+ Mem.inject f m1 m2 ->
+ eval_operation genv (Vptr sp1 Ptrofs.zero) op vl1 m1 = Some v1 ->
+ exists v2,
+ eval_operation genv (Vptr sp2 Ptrofs.zero) (shift_stack_operation delta op) vl2 m2 = Some v2
+ /\ Val.inject f v1 v2.
+Proof.
+ intros.
+ rewrite eval_shift_stack_operation. simpl.
+ eapply eval_operation_inj with (sp1 := Vptr sp1 Ptrofs.zero) (m1 := m1); eauto.
+ intros; eapply Mem.valid_pointer_inject_val; eauto.
+ intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
+ intros; eapply Mem.weak_valid_pointer_inject_no_overflow; eauto.
+ intros; eapply Mem.different_pointers_inject; eauto.
+ intros. apply symbol_address_inject.
+ econstructor; eauto. rewrite Ptrofs.add_zero_l; auto.
+Qed.
+
+End EVAL_INJECT.
diff --git a/riscV/PrintOp.ml b/riscV/PrintOp.ml
new file mode 100644
index 00000000..9ec474b3
--- /dev/null
+++ b/riscV/PrintOp.ml
@@ -0,0 +1,166 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Pretty-printing of operators, conditions, addressing modes *)
+
+open Printf
+open Camlcoq
+open Integers
+open Op
+
+let comparison_name = function
+ | Ceq -> "=="
+ | Cne -> "!="
+ | Clt -> "<"
+ | Cle -> "<="
+ | Cgt -> ">"
+ | Cge -> ">="
+
+let print_condition reg pp = function
+ | (Ccomp c, [r1;r2]) ->
+ fprintf pp "%a %ss %a" reg r1 (comparison_name c) reg r2
+ | (Ccompu c, [r1;r2]) ->
+ fprintf pp "%a %su %a" reg r1 (comparison_name c) reg r2
+ | (Ccompimm(c, n), [r1]) ->
+ fprintf pp "%a %ss %ld" reg r1 (comparison_name c) (camlint_of_coqint n)
+ | (Ccompuimm(c, n), [r1]) ->
+ fprintf pp "%a %su %ld" reg r1 (comparison_name c) (camlint_of_coqint n)
+ | (Ccompf c, [r1;r2]) ->
+ fprintf pp "%a %sf %a" reg r1 (comparison_name c) reg r2
+ | (Ccompl c, [r1;r2]) ->
+ fprintf pp "%a %sls %a" reg r1 (comparison_name c) reg r2
+ | (Ccomplu c, [r1;r2]) ->
+ fprintf pp "%a %slu %a" reg r1 (comparison_name c) reg r2
+ | (Ccomplimm(c, n), [r1]) ->
+ fprintf pp "%a %sls %Ld" reg r1 (comparison_name c) (camlint64_of_coqint n)
+ | (Ccompluimm(c, n), [r1]) ->
+ fprintf pp "%a %slu %Lu" reg r1 (comparison_name c) (camlint64_of_coqint n)
+ | (Cnotcompf c, [r1;r2]) ->
+ fprintf pp "%a not(%sf) %a" reg r1 (comparison_name c) reg r2
+ | (Ccompfs c, [r1;r2]) ->
+ fprintf pp "%a %sfs %a" reg r1 (comparison_name c) reg r2
+ | (Cnotcompfs c, [r1;r2]) ->
+ fprintf pp "%a not(%sfs) %a" reg r1 (comparison_name c) reg r2
+ | _ ->
+ fprintf pp "<bad condition>"
+
+let print_operation reg pp = function
+ | Omove, [r1] -> reg pp r1
+ | Ointconst n, [] -> fprintf pp "%ld" (camlint_of_coqint n)
+ | Olongconst n, [] -> fprintf pp "%LdL" (camlint64_of_coqint n)
+ | Ofloatconst n, [] -> fprintf pp "%F" (camlfloat_of_coqfloat n)
+ | Osingleconst n, [] -> fprintf pp "%Ff" (camlfloat_of_coqfloat32 n)
+ | Oaddrsymbol(id, ofs), [] ->
+ fprintf pp "\"%s\" + %Ld" (extern_atom id) (camlint64_of_ptrofs ofs)
+ | Oaddrstack ofs, [] ->
+ fprintf pp "stack(%Ld)" (camlint64_of_ptrofs ofs)
+ | Ocast8signed, [r1] -> fprintf pp "int8signed(%a)" reg r1
+ | Ocast16signed, [r1] -> fprintf pp "int16signed(%a)" reg r1
+ | Oadd, [r1;r2] -> fprintf pp "%a + %a" reg r1 reg r2
+ | Oaddimm n, [r1] -> fprintf pp "%a + %ld" reg r1 (camlint_of_coqint n)
+ | Oneg, [r1] -> fprintf pp "-(%a)" reg r1
+ | Osub, [r1;r2] -> fprintf pp "%a - %a" reg r1 reg r2
+ | Omul, [r1;r2] -> fprintf pp "%a * %a" reg r1 reg r2
+ | Omulhs, [r1;r2] -> fprintf pp "%a *hs %a" reg r1 reg r2
+ | Omulhu, [r1;r2] -> fprintf pp "%a *hu %a" reg r1 reg r2
+ | Odiv, [r1;r2] -> fprintf pp "%a /s %a" reg r1 reg r2
+ | Odivu, [r1;r2] -> fprintf pp "%a /u %a" reg r1 reg r2
+ | Omod, [r1;r2] -> fprintf pp "%a %%s %a" reg r1 reg r2
+ | Omodu, [r1;r2] -> fprintf pp "%a %%u %a" reg r1 reg r2
+ | Oand, [r1;r2] -> fprintf pp "%a & %a" reg r1 reg r2
+ | Oandimm n, [r1] -> fprintf pp "%a & %ld" reg r1 (camlint_of_coqint n)
+ | Oor, [r1;r2] -> fprintf pp "%a | %a" reg r1 reg r2
+ | Oorimm n, [r1] -> fprintf pp "%a | %ld" reg r1 (camlint_of_coqint n)
+ | Oxor, [r1;r2] -> fprintf pp "%a ^ %a" reg r1 reg r2
+ | Oxorimm n, [r1] -> fprintf pp "%a ^ %ld" reg r1 (camlint_of_coqint n)
+ | Oshl, [r1;r2] -> fprintf pp "%a << %a" reg r1 reg r2
+ | Oshlimm n, [r1] -> fprintf pp "%a << %ld" reg r1 (camlint_of_coqint n)
+ | Oshr, [r1;r2] -> fprintf pp "%a >>s %a" reg r1 reg r2
+ | Oshrimm n, [r1] -> fprintf pp "%a >>s %ld" reg r1 (camlint_of_coqint n)
+ | Oshru, [r1;r2] -> fprintf pp "%a >>u %a" reg r1 reg r2
+ | Oshruimm n, [r1] -> fprintf pp "%a >>u %ld" reg r1 (camlint_of_coqint n)
+ | Oshrximm n, [r1] -> fprintf pp "%a >>x %ld" reg r1 (camlint_of_coqint n)
+
+ | Omakelong, [r1;r2] -> fprintf pp "makelong(%a,%a)" reg r1 reg r2
+ | Olowlong, [r1] -> fprintf pp "lowlong(%a)" reg r1
+ | Ohighlong, [r1] -> fprintf pp "highlong(%a)" reg r1
+ | Ocast32signed, [r1] -> fprintf pp "long32signed(%a)" reg r1
+ | Ocast32unsigned, [r1] -> fprintf pp "long32unsigned(%a)" reg r1
+ | Oaddl, [r1;r2] -> fprintf pp "%a +l %a" reg r1 reg r2
+ | Oaddlimm n, [r1] -> fprintf pp "%a +l %Ld" reg r1 (camlint64_of_coqint n)
+ | Onegl, [r1] -> fprintf pp "-l (%a)" reg r1
+ | Osubl, [r1;r2] -> fprintf pp "%a -l %a" reg r1 reg r2
+ | Omull, [r1;r2] -> fprintf pp "%a *l %a" reg r1 reg r2
+ | Omullhs, [r1;r2] -> fprintf pp "%a *lhs %a" reg r1 reg r2
+ | Omullhu, [r1;r2] -> fprintf pp "%a *lhu %a" reg r1 reg r2
+ | Odivl, [r1;r2] -> fprintf pp "%a /ls %a" reg r1 reg r2
+ | Odivlu, [r1;r2] -> fprintf pp "%a /lu %a" reg r1 reg r2
+ | Omodl, [r1;r2] -> fprintf pp "%a %%ls %a" reg r1 reg r2
+ | Omodlu, [r1;r2] -> fprintf pp "%a %%lu %a" reg r1 reg r2
+ | Oandl, [r1;r2] -> fprintf pp "%a &l %a" reg r1 reg r2
+ | Oandlimm n, [r1] -> fprintf pp "%a &l %Ld" reg r1 (camlint64_of_coqint n)
+ | Oorl, [r1;r2] -> fprintf pp "%a |l %a" reg r1 reg r2
+ | Oorlimm n, [r1] -> fprintf pp "%a |l %Ld" reg r1 (camlint64_of_coqint n)
+ | Oxorl, [r1;r2] -> fprintf pp "%a ^l %a" reg r1 reg r2
+ | Oxorlimm n, [r1] -> fprintf pp "%a ^l %Ld" reg r1 (camlint64_of_coqint n)
+ | Oshll, [r1;r2] -> fprintf pp "%a <<l %a" reg r1 reg r2
+ | Oshllimm n, [r1] -> fprintf pp "%a <<l %Ld" reg r1 (camlint64_of_coqint n)
+ | Oshrl, [r1;r2] -> fprintf pp "%a >>ls %a" reg r1 reg r2
+ | Oshrlimm n, [r1] -> fprintf pp "%a >>ls %ld" reg r1 (camlint_of_coqint n)
+ | Oshrlu, [r1;r2] -> fprintf pp "%a >>lu %a" reg r1 reg r2
+ | Oshrluimm n, [r1] -> fprintf pp "%a >>lu %ld" reg r1 (camlint_of_coqint n)
+ | Oshrxlimm n, [r1] -> fprintf pp "%a >>lx %ld" reg r1 (camlint_of_coqint n)
+
+ | Onegf, [r1] -> fprintf pp "negf(%a)" reg r1
+ | Oabsf, [r1] -> fprintf pp "absf(%a)" reg r1
+ | Oaddf, [r1;r2] -> fprintf pp "%a +f %a" reg r1 reg r2
+ | Osubf, [r1;r2] -> fprintf pp "%a -f %a" reg r1 reg r2
+ | Omulf, [r1;r2] -> fprintf pp "%a *f %a" reg r1 reg r2
+ | Odivf, [r1;r2] -> fprintf pp "%a /f %a" reg r1 reg r2
+ | Onegfs, [r1] -> fprintf pp "negfs(%a)" reg r1
+ | Oabsfs, [r1] -> fprintf pp "absfs(%a)" reg r1
+ | Oaddfs, [r1;r2] -> fprintf pp "%a +fs %a" reg r1 reg r2
+ | Osubfs, [r1;r2] -> fprintf pp "%a -fs %a" reg r1 reg r2
+ | Omulfs, [r1;r2] -> fprintf pp "%a *fs %a" reg r1 reg r2
+ | Odivfs, [r1;r2] -> fprintf pp "%a /fs %a" reg r1 reg r2
+ | Osingleoffloat, [r1] -> fprintf pp "singleoffloat(%a)" reg r1
+ | Ofloatofsingle, [r1] -> fprintf pp "floatofsingle(%a)" reg r1
+ | Ointoffloat, [r1] -> fprintf pp "intoffloat(%a)" reg r1
+ | Ointuoffloat, [r1] -> fprintf pp "intuoffloat(%a)" reg r1
+ | Ofloatofint, [r1] -> fprintf pp "floatofint(%a)" reg r1
+ | Ofloatofintu, [r1] -> fprintf pp "floatofintu(%a)" reg r1
+ | Olongoffloat, [r1] -> fprintf pp "longoffloat(%a)" reg r1
+ | Olonguoffloat, [r1] -> fprintf pp "longuoffloat(%a)" reg r1
+ | Ofloatoflong, [r1] -> fprintf pp "floatoflong(%a)" reg r1
+ | Ofloatoflongu, [r1] -> fprintf pp "floatoflongu(%a)" reg r1
+ | Ointofsingle, [r1] -> fprintf pp "intofsingle(%a)" reg r1
+ | Ointuofsingle, [r1] -> fprintf pp "intuofsingle(%a)" reg r1
+ | Osingleofint, [r1] -> fprintf pp "singleofint(%a)" reg r1
+ | Osingleofintu, [r1] -> fprintf pp "singleofintu(%a)" reg r1
+ | Olongofsingle, [r1] -> fprintf pp "longofsingle(%a)" reg r1
+ | Olonguofsingle, [r1] -> fprintf pp "longuofsingle(%a)" reg r1
+ | Osingleoflong, [r1] -> fprintf pp "singleoflong(%a)" reg r1
+ | Osingleoflongu, [r1] -> fprintf pp "singleoflongu(%a)" reg r1
+ | Ocmp c, args -> print_condition reg pp (c, args)
+ | _ -> fprintf pp "<bad operator>"
+
+let print_addressing reg pp = function
+ | Aindexed n, [r1] -> fprintf pp "%a + %Ld" reg r1 (camlint64_of_ptrofs n)
+ | Aglobal(id, ofs), [] ->
+ fprintf pp "\"%s\" + %Ld" (extern_atom id) (camlint64_of_ptrofs ofs)
+ | Ainstack ofs, [] -> fprintf pp "stack(%Ld)" (camlint64_of_ptrofs ofs)
+ | _ -> fprintf pp "<bad addressing>"
diff --git a/riscV/SelectLong.vp b/riscV/SelectLong.vp
new file mode 100644
index 00000000..b3e07bf5
--- /dev/null
+++ b/riscV/SelectLong.vp
@@ -0,0 +1,364 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Instruction selection for 64-bit integer operations *)
+
+Require Import Coqlib.
+Require Import Compopts.
+Require Import AST Integers Floats.
+Require Import Op CminorSel.
+Require Import SelectOp SplitLong.
+
+Local Open Scope cminorsel_scope.
+Local Open Scope string_scope.
+
+Section SELECT.
+
+Context {hf: helper_functions}.
+
+Definition longconst (n: int64) : expr :=
+ if Archi.splitlong then SplitLong.longconst n else Eop (Olongconst n) Enil.
+
+Definition is_longconst (e: expr) :=
+ if Archi.splitlong then SplitLong.is_longconst e else
+ match e with
+ | Eop (Olongconst n) Enil => Some n
+ | _ => None
+ end.
+
+Definition intoflong (e: expr) :=
+ if Archi.splitlong then SplitLong.intoflong e else
+ match is_longconst e with
+ | Some n => Eop (Ointconst (Int.repr (Int64.unsigned n))) Enil
+ | None => Eop Olowlong (e ::: Enil)
+ end.
+
+Definition longofint (e: expr) :=
+ if Archi.splitlong then SplitLong.longofint e else
+ match is_intconst e with
+ | Some n => longconst (Int64.repr (Int.signed n))
+ | None => Eop Ocast32signed (e ::: Enil)
+ end.
+
+Definition longofintu (e: expr) :=
+ if Archi.splitlong then SplitLong.longofintu e else
+ match is_intconst e with
+ | Some n => longconst (Int64.repr (Int.unsigned n))
+ | None => Eop Ocast32unsigned (e ::: Enil)
+ end.
+
+(** ** Integer addition and pointer addition *)
+
+Nondetfunction addlimm (n: int64) (e: expr) :=
+ if Int64.eq n Int64.zero then e else
+ match e with
+ | Eop (Olongconst m) Enil => longconst (Int64.add n m)
+ | Eop (Oaddrsymbol s m) Enil => Eop (Oaddrsymbol s (Ptrofs.add (Ptrofs.of_int64 n) m)) Enil
+ | Eop (Oaddrstack m) Enil => Eop (Oaddrstack (Ptrofs.add (Ptrofs.of_int64 n) m)) Enil
+ | Eop (Oaddlimm m) (t ::: Enil) => Eop (Oaddlimm(Int64.add n m)) (t ::: Enil)
+ | _ => Eop (Oaddlimm n) (e ::: Enil)
+ end.
+
+Nondetfunction addl (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.addl e1 e2 else
+ match e1, e2 with
+ | Eop (Olongconst n1) Enil, t2 => addlimm n1 t2
+ | t1, Eop (Olongconst n2) Enil => addlimm n2 t1
+ | Eop (Oaddlimm n1) (t1:::Enil), Eop (Oaddlimm n2) (t2:::Enil) =>
+ addlimm (Int64.add n1 n2) (Eop Oaddl (t1:::t2:::Enil))
+ | Eop (Oaddlimm n1) (t1:::Enil), Eop (Oaddrstack n2) Enil =>
+ Eop Oaddl (Eop (Oaddrstack (Ptrofs.add (Ptrofs.of_int64 n1) n2)) Enil ::: t1 ::: Enil)
+ | Eop (Oaddrstack n1) Enil, Eop (Oaddlimm n2) (t2:::Enil) =>
+ Eop Oaddl (Eop (Oaddrstack (Ptrofs.add n1 (Ptrofs.of_int64 n2))) Enil ::: t2 ::: Enil)
+ | Eop (Oaddlimm n1) (t1:::Enil), t2 =>
+ addlimm n1 (Eop Oaddl (t1:::t2:::Enil))
+ | t1, Eop (Oaddlimm n2) (t2:::Enil) =>
+ addlimm n2 (Eop Oaddl (t1:::t2:::Enil))
+ | _, _ => Eop Oaddl (e1:::e2:::Enil)
+ end.
+
+(** ** Integer and pointer subtraction *)
+
+Nondetfunction subl (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.subl e1 e2 else
+ match e1, e2 with
+ | t1, Eop (Olongconst n2) Enil =>
+ addlimm (Int64.neg n2) t1
+ | Eop (Oaddlimm n1) (t1:::Enil), Eop (Oaddlimm n2) (t2:::Enil) =>
+ addlimm (Int64.sub n1 n2) (Eop Osubl (t1:::t2:::Enil))
+ | Eop (Oaddlimm n1) (t1:::Enil), t2 =>
+ addlimm n1 (Eop Osubl (t1:::t2:::Enil))
+ | t1, Eop (Oaddlimm n2) (t2:::Enil) =>
+ addlimm (Int64.neg n2) (Eop Osubl (t1:::t2:::Enil))
+ | _, _ => Eop Osubl (e1:::e2:::Enil)
+ end.
+
+Definition negl (e: expr) :=
+ if Archi.splitlong then SplitLong.negl e else
+ match is_longconst e with
+ | Some n => longconst (Int64.neg n)
+ | None => Eop Onegl (e ::: Enil)
+ end.
+
+(** ** Immediate shifts *)
+
+Nondetfunction shllimm (e1: expr) (n: int) :=
+ if Archi.splitlong then SplitLong.shllimm e1 n else
+ if Int.eq n Int.zero then
+ e1
+ else if negb (Int.ltu n Int64.iwordsize') then
+ Eop Oshll (e1 ::: Eop (Ointconst n) Enil ::: Enil)
+ else match e1 with
+ | Eop (Olongconst n1) Enil =>
+ longconst (Int64.shl' n1 n)
+ | Eop (Oshllimm n1) (t1:::Enil) =>
+ if Int.ltu (Int.add n n1) Int64.iwordsize'
+ then Eop (Oshllimm (Int.add n n1)) (t1:::Enil)
+ else Eop (Oshllimm n) (e1:::Enil)
+ | _ =>
+ Eop (Oshllimm n) (e1:::Enil)
+ end.
+
+Nondetfunction shrluimm (e1: expr) (n: int) :=
+ if Archi.splitlong then SplitLong.shrluimm e1 n else
+ if Int.eq n Int.zero then e1 else
+ if negb (Int.ltu n Int64.iwordsize') then
+ Eop Oshrlu (e1:::Eop (Ointconst n) Enil:::Enil)
+ else
+ match e1 with
+ | Eop (Olongconst n1) Enil =>
+ longconst (Int64.shru' n1 n)
+ | Eop (Oshrluimm n1) (t1:::Enil) =>
+ if Int.ltu (Int.add n n1) Int64.iwordsize'
+ then Eop (Oshrluimm (Int.add n n1)) (t1:::Enil)
+ else Eop (Oshrluimm n) (e1:::Enil)
+ | _ =>
+ Eop (Oshrluimm n) (e1:::Enil)
+ end.
+
+Nondetfunction shrlimm (e1: expr) (n: int) :=
+ if Archi.splitlong then SplitLong.shrlimm e1 n else
+ if Int.eq n Int.zero then e1 else
+ if negb (Int.ltu n Int64.iwordsize') then
+ Eop Oshrl (e1:::Eop (Ointconst n) Enil:::Enil)
+ else
+ match e1 with
+ | Eop (Olongconst n1) Enil =>
+ longconst (Int64.shr' n1 n)
+ | Eop (Oshrlimm n1) (t1:::Enil) =>
+ if Int.ltu (Int.add n n1) Int64.iwordsize'
+ then Eop (Oshrlimm (Int.add n n1)) (t1:::Enil)
+ else Eop (Oshrlimm n) (e1:::Enil)
+ | _ =>
+ Eop (Oshrlimm n) (e1:::Enil)
+ end.
+
+(** ** General shifts *)
+
+Definition shll (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.shll e1 e2 else
+ match is_intconst e2 with
+ | Some n2 => shllimm e1 n2
+ | None => Eop Oshll (e1:::e2:::Enil)
+ end.
+
+Definition shrl (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.shrl e1 e2 else
+ match is_intconst e2 with
+ | Some n2 => shrlimm e1 n2
+ | None => Eop Oshrl (e1:::e2:::Enil)
+ end.
+
+Definition shrlu (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.shrlu e1 e2 else
+ match is_intconst e2 with
+ | Some n2 => shrluimm e1 n2
+ | _ => Eop Oshrlu (e1:::e2:::Enil)
+ end.
+
+(** ** Integer multiply *)
+
+Definition mullimm_base (n1: int64) (e2: expr) :=
+ match Int64.one_bits' n1 with
+ | i :: nil =>
+ shllimm e2 i
+ | i :: j :: nil =>
+ Elet e2 (addl (shllimm (Eletvar 0) i) (shllimm (Eletvar 0) j))
+ | _ =>
+ Eop Omull (e2 ::: longconst n1 ::: Enil)
+ end.
+
+Nondetfunction mullimm (n1: int64) (e2: expr) :=
+ if Archi.splitlong then SplitLong.mullimm n1 e2
+ else if Int64.eq n1 Int64.zero then longconst Int64.zero
+ else if Int64.eq n1 Int64.one then e2
+ else match e2 with
+ | Eop (Olongconst n2) Enil => longconst (Int64.mul n1 n2)
+ | Eop (Oaddlimm n2) (t2:::Enil) => addlimm (Int64.mul n1 n2) (mullimm_base n1 t2)
+ | _ => mullimm_base n1 e2
+ end.
+
+Nondetfunction mull (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.mull e1 e2 else
+ match e1, e2 with
+ | Eop (Olongconst n1) Enil, t2 => mullimm n1 t2
+ | t1, Eop (Olongconst n2) Enil => mullimm n2 t1
+ | _, _ => Eop Omull (e1:::e2:::Enil)
+ end.
+
+Definition mullhu (e1: expr) (n2: int64) :=
+ if Archi.splitlong then SplitLong.mullhu e1 n2 else
+ Eop Omullhu (e1 ::: longconst n2 ::: Enil).
+
+Definition mullhs (e1: expr) (n2: int64) :=
+ if Archi.splitlong then SplitLong.mullhs e1 n2 else
+ Eop Omullhs (e1 ::: longconst n2 ::: Enil).
+
+(** ** Bitwise and, or, xor *)
+
+Nondetfunction andlimm (n1: int64) (e2: expr) :=
+ if Int64.eq n1 Int64.zero then longconst Int64.zero else
+ if Int64.eq n1 Int64.mone then e2 else
+ match e2 with
+ | Eop (Olongconst n2) Enil =>
+ longconst (Int64.and n1 n2)
+ | Eop (Oandlimm n2) (t2:::Enil) =>
+ Eop (Oandlimm (Int64.and n1 n2)) (t2:::Enil)
+ | _ =>
+ Eop (Oandlimm n1) (e2:::Enil)
+ end.
+
+Nondetfunction andl (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.andl e1 e2 else
+ match e1, e2 with
+ | Eop (Olongconst n1) Enil, t2 => andlimm n1 t2
+ | t1, Eop (Olongconst n2) Enil => andlimm n2 t1
+ | _, _ => Eop Oandl (e1:::e2:::Enil)
+ end.
+
+Nondetfunction orlimm (n1: int64) (e2: expr) :=
+ if Int64.eq n1 Int64.zero then e2 else
+ if Int64.eq n1 Int64.mone then longconst Int64.mone else
+ match e2 with
+ | Eop (Olongconst n2) Enil => longconst (Int64.or n1 n2)
+ | Eop (Oorlimm n2) (t2:::Enil) => Eop (Oorlimm (Int64.or n1 n2)) (t2:::Enil)
+ | _ => Eop (Oorlimm n1) (e2:::Enil)
+ end.
+
+Nondetfunction orl (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.orl e1 e2 else
+ match e1, e2 with
+ | Eop (Olongconst n1) Enil, t2 => orlimm n1 t2
+ | t1, Eop (Olongconst n2) Enil => orlimm n2 t1
+ | _, _ => Eop Oorl (e1:::e2:::Enil)
+ end.
+
+Nondetfunction xorlimm (n1: int64) (e2: expr) :=
+ if Int64.eq n1 Int64.zero then e2 else
+ match e2 with
+ | Eop (Olongconst n2) Enil => longconst (Int64.xor n1 n2)
+ | Eop (Oxorlimm n2) (t2:::Enil) =>
+ let n := Int64.xor n1 n2 in
+ if Int64.eq n Int64.zero then t2 else Eop (Oxorlimm n) (t2:::Enil)
+ | _ => Eop (Oxorlimm n1) (e2:::Enil)
+ end.
+
+Nondetfunction xorl (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.xorl e1 e2 else
+ match e1, e2 with
+ | Eop (Olongconst n1) Enil, t2 => xorlimm n1 t2
+ | t1, Eop (Olongconst n2) Enil => xorlimm n2 t1
+ | _, _ => Eop Oxorl (e1:::e2:::Enil)
+ end.
+
+(** ** Integer logical negation *)
+
+Definition notl (e: expr) :=
+ if Archi.splitlong then SplitLong.notl e else xorlimm Int64.mone e.
+
+(** ** Integer division and modulus *)
+
+Definition divlu_base (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.divlu_base e1 e2 else Eop Odivlu (e1:::e2:::Enil).
+Definition modlu_base (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.modlu_base e1 e2 else Eop Omodlu (e1:::e2:::Enil).
+Definition divls_base (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.divls_base e1 e2 else Eop Odivl (e1:::e2:::Enil).
+Definition modls_base (e1: expr) (e2: expr) :=
+ if Archi.splitlong then SplitLong.modls_base e1 e2 else Eop Omodl (e1:::e2:::Enil).
+
+Definition shrxlimm (e: expr) (n: int) :=
+ if Archi.splitlong then SplitLong.shrxlimm e n else
+ if Int.eq n Int.zero then e else Eop (Oshrxlimm n) (e ::: Enil).
+
+(** ** Comparisons *)
+
+Definition cmplu (c: comparison) (e1 e2: expr) :=
+ if Archi.splitlong then SplitLong.cmplu c e1 e2 else
+ match is_longconst e1, is_longconst e2 with
+ | Some n1, Some n2 =>
+ Eop (Ointconst (if Int64.cmpu c n1 n2 then Int.one else Int.zero)) Enil
+ | Some n1, None => Eop (Ocmp (Ccompluimm (swap_comparison c) n1)) (e2:::Enil)
+ | None, Some n2 => Eop (Ocmp (Ccompluimm c n2)) (e1:::Enil)
+ | None, None => Eop (Ocmp (Ccomplu c)) (e1:::e2:::Enil)
+ end.
+
+Definition cmpl (c: comparison) (e1 e2: expr) :=
+ if Archi.splitlong then SplitLong.cmpl c e1 e2 else
+ match is_longconst e1, is_longconst e2 with
+ | Some n1, Some n2 =>
+ Eop (Ointconst (if Int64.cmp c n1 n2 then Int.one else Int.zero)) Enil
+ | Some n1, None => Eop (Ocmp (Ccomplimm (swap_comparison c) n1)) (e2:::Enil)
+ | None, Some n2 => Eop (Ocmp (Ccomplimm c n2)) (e1:::Enil)
+ | None, None => Eop (Ocmp (Ccompl c)) (e1:::e2:::Enil)
+ end.
+
+(** ** Floating-point conversions *)
+
+Definition longoffloat (e: expr) :=
+ if Archi.splitlong then SplitLong.longoffloat e else
+ Eop Olongoffloat (e:::Enil).
+
+Definition longuoffloat (e: expr) :=
+ if Archi.splitlong then SplitLong.longuoffloat e else
+ Eop Olonguoffloat (e:::Enil).
+
+Definition floatoflong (e: expr) :=
+ if Archi.splitlong then SplitLong.floatoflong e else
+ Eop Ofloatoflong (e:::Enil).
+
+Definition floatoflongu (e: expr) :=
+ if Archi.splitlong then SplitLong.floatoflongu e else
+ Eop Ofloatoflongu (e:::Enil).
+
+Definition longofsingle (e: expr) :=
+ if Archi.splitlong then SplitLong.longofsingle e else
+ Eop Olongofsingle (e:::Enil).
+
+Definition longuofsingle (e: expr) :=
+ if Archi.splitlong then SplitLong.longuofsingle e else
+ Eop Olonguofsingle (e:::Enil).
+
+Definition singleoflong (e: expr) :=
+ if Archi.splitlong then SplitLong.singleoflong e else
+ Eop Osingleoflong (e:::Enil).
+
+Definition singleoflongu (e: expr) :=
+ if Archi.splitlong then SplitLong.singleoflongu e else
+ Eop Osingleoflongu (e:::Enil).
+
+End SELECT.
diff --git a/riscV/SelectLongproof.v b/riscV/SelectLongproof.v
new file mode 100644
index 00000000..78a1935d
--- /dev/null
+++ b/riscV/SelectLongproof.v
@@ -0,0 +1,619 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Correctness of instruction selection for 64-bit integer operations *)
+
+Require Import String Coqlib Maps Integers Floats Errors.
+Require Archi.
+Require Import AST Values Memory Globalenvs Events.
+Require Import Cminor Op CminorSel.
+Require Import SelectOp SelectOpproof SplitLong SplitLongproof.
+Require Import SelectLong.
+
+Local Open Scope cminorsel_scope.
+Local Open Scope string_scope.
+
+(** * Correctness of the instruction selection functions for 64-bit operators *)
+
+Section CMCONSTR.
+
+Variable prog: program.
+Variable hf: helper_functions.
+Hypothesis HELPERS: helper_functions_declared prog hf.
+Let ge := Genv.globalenv prog.
+Variable sp: val.
+Variable e: env.
+Variable m: mem.
+
+Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop :=
+ forall le a x,
+ eval_expr ge sp e m le a x ->
+ exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v.
+
+Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop :=
+ forall le a x b y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v.
+
+Definition partial_unary_constructor_sound (cstr: expr -> expr) (sem: val -> option val) : Prop :=
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ sem x = Some y ->
+ exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef y v.
+
+Definition partial_binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> option val) : Prop :=
+ forall le a x b y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ sem x y = Some z ->
+ exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef z v.
+
+Theorem eval_longconst:
+ forall le n, eval_expr ge sp e m le (longconst n) (Vlong n).
+Proof.
+ unfold longconst; intros; destruct Archi.splitlong.
+ apply SplitLongproof.eval_longconst.
+ EvalOp.
+Qed.
+
+Lemma is_longconst_sound:
+ forall v a n le,
+ is_longconst a = Some n -> eval_expr ge sp e m le a v -> v = Vlong n.
+Proof with (try discriminate).
+ intros. unfold is_longconst in *. destruct Archi.splitlong.
+ eapply SplitLongproof.is_longconst_sound; eauto.
+ assert (a = Eop (Olongconst n) Enil).
+ { destruct a... destruct o... destruct e0... congruence. }
+ subst a. InvEval. auto.
+Qed.
+
+Theorem eval_intoflong: unary_constructor_sound intoflong Val.loword.
+Proof.
+ unfold intoflong; destruct Archi.splitlong. apply SplitLongproof.eval_intoflong.
+ red; intros. destruct (is_longconst a) as [n|] eqn:C.
+- TrivialExists. simpl. erewrite (is_longconst_sound x) by eauto. auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_longofintu: unary_constructor_sound longofintu Val.longofintu.
+Proof.
+ unfold longofintu; destruct Archi.splitlong. apply SplitLongproof.eval_longofintu.
+ red; intros. destruct (is_intconst a) as [n|] eqn:C.
+- econstructor; split. apply eval_longconst.
+ exploit is_intconst_sound; eauto. intros; subst x. auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_longofint: unary_constructor_sound longofint Val.longofint.
+Proof.
+ unfold longofint; destruct Archi.splitlong. apply SplitLongproof.eval_longofint.
+ red; intros. destruct (is_intconst a) as [n|] eqn:C.
+- econstructor; split. apply eval_longconst.
+ exploit is_intconst_sound; eauto. intros; subst x. auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_negl: unary_constructor_sound negl Val.negl.
+Proof.
+ unfold negl. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_negl; auto.
+ red; intros. destruct (is_longconst a) as [n|] eqn:C.
+- exploit is_longconst_sound; eauto. intros EQ; subst x.
+ econstructor; split. apply eval_longconst. auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_addlimm: forall n, unary_constructor_sound (addlimm n) (fun v => Val.addl v (Vlong n)).
+Proof.
+ unfold addlimm; intros; red; intros.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero.
+ subst. exists x; split; auto.
+ destruct x; simpl; rewrite ?Int64.add_zero, ?Ptrofs.add_zero; auto.
+ destruct Archi.ptr64; auto.
+ destruct (addlimm_match a); InvEval.
+- econstructor; split. apply eval_longconst. rewrite Int64.add_commut; auto.
+- econstructor; split. EvalOp. simpl; eauto.
+ unfold Genv.symbol_address. destruct (Genv.find_symbol ge s); simpl; auto.
+ destruct Archi.ptr64; auto. rewrite Ptrofs.add_commut; auto.
+- econstructor; split. EvalOp. simpl; eauto.
+ destruct sp; simpl; auto. destruct Archi.ptr64; auto.
+ rewrite Ptrofs.add_assoc, (Ptrofs.add_commut m0). auto.
+- subst x. rewrite Val.addl_assoc. rewrite Int64.add_commut. TrivialExists.
+- TrivialExists.
+Qed.
+
+Theorem eval_addl: binary_constructor_sound addl Val.addl.
+Proof.
+ unfold addl. destruct Archi.splitlong eqn:SL.
+ apply SplitLongproof.eval_addl. apply Archi.splitlong_ptr32; auto.
+(*
+ assert (SF: Archi.ptr64 = true).
+ { Local Transparent Archi.splitlong. unfold Archi.splitlong in SL.
+ destruct Archi.ptr64; simpl in *; congruence. }
+*)
+(*
+ assert (B: forall id ofs n,
+ Genv.symbol_address ge id (Ptrofs.add ofs (Ptrofs.repr n)) =
+ Val.addl (Genv.symbol_address ge id ofs) (Vlong (Int64.repr n))).
+ { intros. replace (Ptrofs.repr n) with (Ptrofs.of_int64 (Int64.repr n)) by auto with ptrofs.
+ apply Genv.shift_symbol_address_64; auto. }
+
+*)
+ red; intros until y.
+ case (addl_match a b); intros; InvEval.
+ - rewrite Val.addl_commut. apply eval_addlimm; auto.
+ - apply eval_addlimm; auto.
+ - subst.
+ replace (Val.addl (Val.addl v1 (Vlong n1)) (Val.addl v0 (Vlong n2)))
+ with (Val.addl (Val.addl v1 v0) (Val.addl (Vlong n1) (Vlong n2))).
+ apply eval_addlimm. EvalOp.
+ repeat rewrite Val.addl_assoc. decEq. apply Val.addl_permut.
+ - subst. econstructor; split.
+ EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto.
+ rewrite Val.addl_commut. destruct sp; simpl; auto.
+ destruct v1; simpl; auto.
+ destruct Archi.ptr64 eqn:SF; auto.
+ apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal.
+ rewrite (Ptrofs.add_commut (Ptrofs.of_int64 n1)), Ptrofs.add_assoc. f_equal. auto with ptrofs.
+ destruct Archi.ptr64 eqn:SF; auto.
+ - subst. econstructor; split.
+ EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto.
+ destruct sp; simpl; auto.
+ destruct v1; simpl; auto.
+ destruct Archi.ptr64 eqn:SF; auto.
+ apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. f_equal.
+ rewrite Ptrofs.add_commut. auto with ptrofs.
+ destruct Archi.ptr64 eqn:SF; auto.
+ - subst.
+ replace (Val.addl (Val.addl v1 (Vlong n1)) y)
+ with (Val.addl (Val.addl v1 y) (Vlong n1)).
+ apply eval_addlimm. EvalOp.
+ repeat rewrite Val.addl_assoc. decEq. apply Val.addl_commut.
+ - subst.
+ replace (Val.addl x (Val.addl v1 (Vlong n2)))
+ with (Val.addl (Val.addl x v1) (Vlong n2)).
+ apply eval_addlimm. EvalOp.
+ repeat rewrite Val.addl_assoc. reflexivity.
+ - TrivialExists.
+Qed.
+
+Theorem eval_subl: binary_constructor_sound subl Val.subl.
+Proof.
+ unfold subl. destruct Archi.splitlong eqn:SL.
+ apply SplitLongproof.eval_subl. apply Archi.splitlong_ptr32; auto.
+ red; intros; destruct (subl_match a b); InvEval.
+- rewrite Val.subl_addl_opp. apply eval_addlimm; auto.
+- subst. rewrite Val.subl_addl_l. rewrite Val.subl_addl_r.
+ rewrite Val.addl_assoc. simpl. rewrite Int64.add_commut. rewrite <- Int64.sub_add_opp.
+ apply eval_addlimm; EvalOp.
+- subst. rewrite Val.subl_addl_l. apply eval_addlimm; EvalOp.
+- subst. rewrite Val.subl_addl_r.
+ apply eval_addlimm; EvalOp.
+- TrivialExists.
+Qed.
+
+Theorem eval_shllimm: forall n, unary_constructor_sound (fun e => shllimm e n) (fun v => Val.shll v (Vint n)).
+Proof.
+ intros; unfold shllimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shllimm; auto.
+ red; intros.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ exists x; split; auto. subst n; destruct x; simpl; auto.
+ destruct (Int.ltu Int.zero Int64.iwordsize'); auto.
+ change (Int64.shl' i Int.zero) with (Int64.shl i Int64.zero). rewrite Int64.shl_zero; auto.
+ destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl.
+ assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshllimm n) (a:::Enil)) v
+ /\ Val.lessdef (Val.shll x (Vint n)) v) by TrivialExists.
+ destruct (shllimm_match a); InvEval.
+- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto.
+- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto.
+ subst. econstructor; split. EvalOp. simpl; eauto.
+ destruct v1; simpl; auto. rewrite LT'.
+ destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto.
+ simpl; rewrite LT. rewrite Int.add_commut, Int64.shl'_shl'; auto. rewrite Int.add_commut; auto.
+- apply DEFAULT.
+- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto.
+Qed.
+
+Theorem eval_shrluimm: forall n, unary_constructor_sound (fun e => shrluimm e n) (fun v => Val.shrlu v (Vint n)).
+Proof.
+ intros; unfold shrluimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrluimm; auto.
+ red; intros.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ exists x; split; auto. subst n; destruct x; simpl; auto.
+ destruct (Int.ltu Int.zero Int64.iwordsize'); auto.
+ change (Int64.shru' i Int.zero) with (Int64.shru i Int64.zero). rewrite Int64.shru_zero; auto.
+ destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl.
+ assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshrluimm n) (a:::Enil)) v
+ /\ Val.lessdef (Val.shrlu x (Vint n)) v) by TrivialExists.
+ destruct (shrluimm_match a); InvEval.
+- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto.
+- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto.
+ subst. econstructor; split. EvalOp. simpl; eauto.
+ destruct v1; simpl; auto. rewrite LT'.
+ destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto.
+ simpl; rewrite LT. rewrite Int.add_commut, Int64.shru'_shru'; auto. rewrite Int.add_commut; auto.
+- apply DEFAULT.
+- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto.
+Qed.
+
+Theorem eval_shrlimm: forall n, unary_constructor_sound (fun e => shrlimm e n) (fun v => Val.shrl v (Vint n)).
+Proof.
+ intros; unfold shrlimm. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrlimm; auto.
+ red; intros.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ exists x; split; auto. subst n; destruct x; simpl; auto.
+ destruct (Int.ltu Int.zero Int64.iwordsize'); auto.
+ change (Int64.shr' i Int.zero) with (Int64.shr i Int64.zero). rewrite Int64.shr_zero; auto.
+ destruct (Int.ltu n Int64.iwordsize') eqn:LT; simpl.
+ assert (DEFAULT: exists v, eval_expr ge sp e m le (Eop (Oshrlimm n) (a:::Enil)) v
+ /\ Val.lessdef (Val.shrl x (Vint n)) v) by TrivialExists.
+ destruct (shrlimm_match a); InvEval.
+- econstructor; split. apply eval_longconst. simpl; rewrite LT; auto.
+- destruct (Int.ltu (Int.add n n1) Int64.iwordsize') eqn:LT'; auto.
+ subst. econstructor; split. EvalOp. simpl; eauto.
+ destruct v1; simpl; auto. rewrite LT'.
+ destruct (Int.ltu n1 Int64.iwordsize') eqn:LT1; auto.
+ simpl; rewrite LT. rewrite Int.add_commut, Int64.shr'_shr'; auto. rewrite Int.add_commut; auto.
+- apply DEFAULT.
+- TrivialExists. constructor; eauto. constructor. EvalOp. simpl; eauto. constructor. auto.
+Qed.
+
+Theorem eval_shll: binary_constructor_sound shll Val.shll.
+Proof.
+ unfold shll. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shll; auto.
+ red; intros. destruct (is_intconst b) as [n2|] eqn:C.
+- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shllimm; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_shrlu: binary_constructor_sound shrlu Val.shrlu.
+Proof.
+ unfold shrlu. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrlu; auto.
+ red; intros. destruct (is_intconst b) as [n2|] eqn:C.
+- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shrluimm; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_shrl: binary_constructor_sound shrl Val.shrl.
+Proof.
+ unfold shrl. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_shrl; auto.
+ red; intros. destruct (is_intconst b) as [n2|] eqn:C.
+- exploit is_intconst_sound; eauto. intros EQ; subst y. apply eval_shrlimm; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_mullimm_base: forall n, unary_constructor_sound (mullimm_base n) (fun v => Val.mull v (Vlong n)).
+Proof.
+ intros; unfold mullimm_base. red; intros.
+ assert (DEFAULT: exists v,
+ eval_expr ge sp e m le (Eop Omull (a ::: longconst n ::: Enil)) v
+ /\ Val.lessdef (Val.mull x (Vlong n)) v).
+ { econstructor; split. EvalOp. constructor. eauto. constructor. apply eval_longconst. constructor. simpl; eauto.
+ auto. }
+ generalize (Int64.one_bits'_decomp n); intros D.
+ destruct (Int64.one_bits' n) as [ | i [ | j [ | ? ? ]]] eqn:B.
+- apply DEFAULT.
+- replace (Val.mull x (Vlong n)) with (Val.shll x (Vint i)).
+ apply eval_shllimm; auto.
+ simpl in D. rewrite D, Int64.add_zero. destruct x; simpl; auto.
+ rewrite (Int64.one_bits'_range n) by (rewrite B; auto with coqlib).
+ rewrite Int64.shl'_mul; auto.
+- set (le' := x :: le).
+ assert (A0: eval_expr ge sp e m le' (Eletvar O) x) by (constructor; reflexivity).
+ exploit (eval_shllimm i). eexact A0. intros (v1 & A1 & B1).
+ exploit (eval_shllimm j). eexact A0. intros (v2 & A2 & B2).
+ exploit (eval_addl). eexact A1. eexact A2. intros (v3 & A3 & B3).
+ exists v3; split. econstructor; eauto.
+ rewrite D. simpl. rewrite Int64.add_zero. destruct x; auto.
+ simpl in *.
+ rewrite (Int64.one_bits'_range n) in B1 by (rewrite B; auto with coqlib).
+ rewrite (Int64.one_bits'_range n) in B2 by (rewrite B; auto with coqlib).
+ inv B1; inv B2. simpl in B3; inv B3.
+ rewrite Int64.mul_add_distr_r. rewrite <- ! Int64.shl'_mul. auto.
+- apply DEFAULT.
+Qed.
+
+Theorem eval_mullimm: forall n, unary_constructor_sound (mullimm n) (fun v => Val.mull v (Vlong n)).
+Proof.
+ unfold mullimm. intros; red; intros.
+ destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_mullimm; eauto.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero.
+ exists (Vlong Int64.zero); split. apply eval_longconst.
+ destruct x; simpl; auto. subst n; rewrite Int64.mul_zero; auto.
+ predSpec Int64.eq Int64.eq_spec n Int64.one.
+ exists x; split; auto.
+ destruct x; simpl; auto. subst n; rewrite Int64.mul_one; auto.
+ destruct (mullimm_match a); InvEval.
+- econstructor; split. apply eval_longconst. rewrite Int64.mul_commut; auto.
+- exploit (eval_mullimm_base n); eauto. intros (v2 & A2 & B2).
+ exploit (eval_addlimm (Int64.mul n n2)). eexact A2. intros (v3 & A3 & B3).
+ exists v3; split; auto.
+ subst x. destruct v1; simpl; auto.
+ simpl in B2; inv B2. simpl in B3; inv B3. rewrite Int64.mul_add_distr_l.
+ rewrite (Int64.mul_commut n). auto.
+ destruct Archi.ptr64; simpl; auto.
+- apply eval_mullimm_base; auto.
+Qed.
+
+Theorem eval_mull: binary_constructor_sound mull Val.mull.
+Proof.
+ unfold mull. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mull; auto.
+ red; intros; destruct (mull_match a b); InvEval.
+- rewrite Val.mull_commut. apply eval_mullimm; auto.
+- apply eval_mullimm; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_mullhu:
+ forall n, unary_constructor_sound (fun a => mullhu a n) (fun v => Val.mullhu v (Vlong n)).
+Proof.
+ unfold mullhu; intros. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mullhu; auto.
+ red; intros. TrivialExists. constructor. eauto. constructor. apply eval_longconst. constructor. auto.
+Qed.
+
+Theorem eval_mullhs:
+ forall n, unary_constructor_sound (fun a => mullhs a n) (fun v => Val.mullhs v (Vlong n)).
+Proof.
+ unfold mullhs; intros. destruct Archi.splitlong eqn:SL. apply SplitLongproof.eval_mullhs; auto.
+ red; intros. TrivialExists. constructor. eauto. constructor. apply eval_longconst. constructor. auto.
+Qed.
+
+Theorem eval_andlimm: forall n, unary_constructor_sound (andlimm n) (fun v => Val.andl v (Vlong n)).
+Proof.
+ unfold andlimm; intros; red; intros.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero.
+ exists (Vlong Int64.zero); split. apply eval_longconst.
+ subst. destruct x; simpl; auto. rewrite Int64.and_zero; auto.
+ predSpec Int64.eq Int64.eq_spec n Int64.mone.
+ exists x; split. assumption.
+ subst. destruct x; simpl; auto. rewrite Int64.and_mone; auto.
+ destruct (andlimm_match a); InvEval; subst.
+- econstructor; split. apply eval_longconst. simpl. rewrite Int64.and_commut; auto.
+- TrivialExists. simpl. rewrite Val.andl_assoc. rewrite Int64.and_commut; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_andl: binary_constructor_sound andl Val.andl.
+Proof.
+ unfold andl; destruct Archi.splitlong. apply SplitLongproof.eval_andl.
+ red; intros. destruct (andl_match a b).
+- InvEval. rewrite Val.andl_commut. apply eval_andlimm; auto.
+- InvEval. apply eval_andlimm; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_orlimm: forall n, unary_constructor_sound (orlimm n) (fun v => Val.orl v (Vlong n)).
+Proof.
+ unfold orlimm; intros; red; intros.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero.
+ exists x; split; auto. subst. destruct x; simpl; auto. rewrite Int64.or_zero; auto.
+ predSpec Int64.eq Int64.eq_spec n Int64.mone.
+ econstructor; split. apply eval_longconst. subst. destruct x; simpl; auto. rewrite Int64.or_mone; auto.
+ destruct (orlimm_match a); InvEval; subst.
+- econstructor; split. apply eval_longconst. simpl. rewrite Int64.or_commut; auto.
+- TrivialExists. simpl. rewrite Val.orl_assoc. rewrite Int64.or_commut; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_orl: binary_constructor_sound orl Val.orl.
+Proof.
+ unfold orl; destruct Archi.splitlong. apply SplitLongproof.eval_orl.
+ red; intros.
+ destruct (orl_match a b).
+- InvEval. rewrite Val.orl_commut. apply eval_orlimm; auto.
+- InvEval. apply eval_orlimm; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_xorlimm: forall n, unary_constructor_sound (xorlimm n) (fun v => Val.xorl v (Vlong n)).
+Proof.
+ unfold xorlimm; intros; red; intros.
+ predSpec Int64.eq Int64.eq_spec n Int64.zero.
+ exists x; split; auto. subst. destruct x; simpl; auto. rewrite Int64.xor_zero; auto.
+ destruct (xorlimm_match a); InvEval; subst.
+- econstructor; split. apply eval_longconst. simpl. rewrite Int64.xor_commut; auto.
+- rewrite Val.xorl_assoc. simpl. rewrite (Int64.xor_commut n2).
+ predSpec Int64.eq Int64.eq_spec (Int64.xor n n2) Int64.zero.
++ rewrite H. exists v1; split; auto. destruct v1; simpl; auto. rewrite Int64.xor_zero; auto.
++ TrivialExists.
+- TrivialExists.
+Qed.
+
+Theorem eval_xorl: binary_constructor_sound xorl Val.xorl.
+Proof.
+ unfold xorl; destruct Archi.splitlong. apply SplitLongproof.eval_xorl.
+ red; intros. destruct (xorl_match a b).
+- InvEval. rewrite Val.xorl_commut. apply eval_xorlimm; auto.
+- InvEval. apply eval_xorlimm; auto.
+- TrivialExists.
+Qed.
+
+Theorem eval_notl: unary_constructor_sound notl Val.notl.
+Proof.
+ unfold notl; destruct Archi.splitlong. apply SplitLongproof.eval_notl.
+ red; intros. rewrite Val.notl_xorl. apply eval_xorlimm; auto.
+Qed.
+
+Theorem eval_divls_base: partial_binary_constructor_sound divls_base Val.divls.
+Proof.
+ unfold divls_base; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_divls_base; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_modls_base: partial_binary_constructor_sound modls_base Val.modls.
+Proof.
+ unfold modls_base; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_modls_base; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_divlu_base: partial_binary_constructor_sound divlu_base Val.divlu.
+Proof.
+ unfold divlu_base; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_divlu_base; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_modlu_base: partial_binary_constructor_sound modlu_base Val.modlu.
+Proof.
+ unfold modlu_base; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_modlu_base; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_shrxlimm:
+ forall le a n x z,
+ eval_expr ge sp e m le a x ->
+ Val.shrxl x (Vint n) = Some z ->
+ exists v, eval_expr ge sp e m le (shrxlimm a n) v /\ Val.lessdef z v.
+Proof.
+ unfold shrxlimm; intros. destruct Archi.splitlong eqn:SL.
++ eapply SplitLongproof.eval_shrxlimm; eauto using Archi.splitlong_ptr32.
++ predSpec Int.eq Int.eq_spec n Int.zero.
+- subst n. destruct x; simpl in H0; inv H0. econstructor; split; eauto.
+ change (Int.ltu Int.zero (Int.repr 63)) with true. simpl. rewrite Int64.shrx'_zero; auto.
+- TrivialExists.
+(*
+ intros. unfold shrxlimm. destruct Archi.splitlong eqn:SL.
++ eapply SplitLongproof.eval_shrxlimm; eauto using Archi.splitlong_ptr32.
++ destruct x; simpl in H0; try discriminate.
+ destruct (Int.ltu n (Int.repr 63)) eqn:LTU; inv H0.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ - subst n. exists (Vlong i); split; auto. rewrite Int64.shrx'_zero. auto.
+ - assert (NZ: Int.unsigned n <> 0).
+ { intro EQ; elim H0. rewrite <- (Int.repr_unsigned n). rewrite EQ; auto. }
+ assert (LT: 0 <= Int.unsigned n < 63) by (apply Int.ltu_inv in LTU; assumption).
+ assert (LTU2: Int.ltu (Int.sub Int64.iwordsize' n) Int64.iwordsize' = true).
+ { unfold Int.ltu; apply zlt_true.
+ unfold Int.sub. change (Int.unsigned Int64.iwordsize') with 64.
+ rewrite Int.unsigned_repr. omega.
+ assert (64 < Int.max_unsigned) by reflexivity. omega. }
+ assert (X: eval_expr ge sp e m le
+ (Eop (Oshrlimm (Int.repr (Int64.zwordsize - 1))) (a ::: Enil))
+ (Vlong (Int64.shr' i (Int.repr (Int64.zwordsize - 1))))).
+ { EvalOp. }
+ assert (Y: eval_expr ge sp e m le (shrxlimm_inner a n)
+ (Vlong (Int64.shru' (Int64.shr' i (Int.repr (Int64.zwordsize - 1))) (Int.sub Int64.iwordsize' n)))).
+ { EvalOp. simpl. rewrite LTU2. auto. }
+ TrivialExists.
+ constructor. EvalOp. simpl; eauto. constructor.
+ simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int64.shrx'_shr_2 by auto. reflexivity.
+ change (Int.unsigned Int64.iwordsize') with 64; omega.
+*)
+Qed.
+
+Theorem eval_cmplu:
+ forall c le a x b y v,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.cmplu (Mem.valid_pointer m) c x y = Some v ->
+ eval_expr ge sp e m le (cmplu c a b) v.
+Proof.
+ unfold cmplu; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_cmplu; eauto using Archi.splitlong_ptr32.
+ unfold Val.cmplu in H1.
+ destruct (Val.cmplu_bool (Mem.valid_pointer m) c x y) as [vb|] eqn:C; simpl in H1; inv H1.
+ destruct (is_longconst a) as [n1|] eqn:LC1; destruct (is_longconst b) as [n2|] eqn:LC2;
+ try (assert (x = Vlong n1) by (eapply is_longconst_sound; eauto));
+ try (assert (y = Vlong n2) by (eapply is_longconst_sound; eauto));
+ subst.
+- simpl in C; inv C. EvalOp. destruct (Int64.cmpu c n1 n2); reflexivity.
+- EvalOp. simpl. rewrite Val.swap_cmplu_bool. rewrite C; auto.
+- EvalOp. simpl; rewrite C; auto.
+- EvalOp. simpl; rewrite C; auto.
+Qed.
+
+Theorem eval_cmpl:
+ forall c le a x b y v,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.cmpl c x y = Some v ->
+ eval_expr ge sp e m le (cmpl c a b) v.
+Proof.
+ unfold cmpl; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_cmpl; eauto.
+ unfold Val.cmpl in H1.
+ destruct (Val.cmpl_bool c x y) as [vb|] eqn:C; simpl in H1; inv H1.
+ destruct (is_longconst a) as [n1|] eqn:LC1; destruct (is_longconst b) as [n2|] eqn:LC2;
+ try (assert (x = Vlong n1) by (eapply is_longconst_sound; eauto));
+ try (assert (y = Vlong n2) by (eapply is_longconst_sound; eauto));
+ subst.
+- simpl in C; inv C. EvalOp. destruct (Int64.cmp c n1 n2); reflexivity.
+- EvalOp. simpl. rewrite Val.swap_cmpl_bool. rewrite C; auto.
+- EvalOp. simpl; rewrite C; auto.
+- EvalOp. simpl; rewrite C; auto.
+Qed.
+
+Theorem eval_longoffloat: partial_unary_constructor_sound longoffloat Val.longoffloat.
+Proof.
+ unfold longoffloat; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_longoffloat; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_longuoffloat: partial_unary_constructor_sound longuoffloat Val.longuoffloat.
+Proof.
+ unfold longuoffloat; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_longuoffloat; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_floatoflong: partial_unary_constructor_sound floatoflong Val.floatoflong.
+Proof.
+ unfold floatoflong; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_floatoflong; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_floatoflongu: partial_unary_constructor_sound floatoflongu Val.floatoflongu.
+Proof.
+ unfold floatoflongu; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_floatoflongu; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_longofsingle: partial_unary_constructor_sound longofsingle Val.longofsingle.
+Proof.
+ unfold longofsingle; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_longofsingle; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_longuofsingle: partial_unary_constructor_sound longuofsingle Val.longuofsingle.
+Proof.
+ unfold longuofsingle; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_longuofsingle; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_singleoflong: partial_unary_constructor_sound singleoflong Val.singleoflong.
+Proof.
+ unfold singleoflong; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_singleoflong; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_singleoflongu: partial_unary_constructor_sound singleoflongu Val.singleoflongu.
+Proof.
+ unfold singleoflongu; red; intros. destruct Archi.splitlong eqn:SL.
+ eapply SplitLongproof.eval_singleoflongu; eauto.
+ TrivialExists.
+Qed.
+
+End CMCONSTR.
diff --git a/riscV/SelectOp.vp b/riscV/SelectOp.vp
new file mode 100644
index 00000000..b38b7b4d
--- /dev/null
+++ b/riscV/SelectOp.vp
@@ -0,0 +1,446 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Instruction selection for operators *)
+
+(** The instruction selection pass recognizes opportunities for using
+ combined arithmetic and logical operations and addressing modes
+ offered by the target processor. For instance, the expression [x + 1]
+ can take advantage of the "immediate add" instruction of the processor,
+ and on the PowerPC, the expression [(x >> 6) & 0xFF] can be turned
+ into a "rotate and mask" instruction.
+
+ This file defines functions for building CminorSel expressions and
+ statements, especially expressions consisting of operator
+ applications. These functions examine their arguments to choose
+ cheaper forms of operators whenever possible.
+
+ For instance, [add e1 e2] will return a CminorSel expression semantically
+ equivalent to [Eop Oadd (e1 ::: e2 ::: Enil)], but will use a
+ [Oaddimm] operator if one of the arguments is an integer constant,
+ or suppress the addition altogether if one of the arguments is the
+ null integer. In passing, we perform operator reassociation
+ ([(e + c1) * c2] becomes [(e * c2) + (c1 * c2)]) and a small amount
+ of constant propagation.
+
+ On top of the "smart constructor" functions defined below,
+ module [Selection] implements the actual instruction selection pass.
+*)
+
+Require Archi.
+Require Import Coqlib.
+Require Import Compopts.
+Require Import AST.
+Require Import Integers.
+Require Import Floats.
+Require Import Op.
+Require Import CminorSel.
+
+Local Open Scope cminorsel_scope.
+
+(** ** Constants **)
+
+Definition addrsymbol (id: ident) (ofs: ptrofs) :=
+ Eop (Oaddrsymbol id ofs) Enil.
+
+Definition addrstack (ofs: ptrofs) :=
+ Eop (Oaddrstack ofs) Enil.
+
+(** ** Integer addition and pointer addition *)
+
+Nondetfunction addimm (n: int) (e: expr) :=
+ if Int.eq n Int.zero then e else
+ match e with
+ | Eop (Ointconst m) Enil => Eop (Ointconst (Int.add n m)) Enil
+ | Eop (Oaddrsymbol s m) Enil => Eop (Oaddrsymbol s (Ptrofs.add (Ptrofs.of_int n) m)) Enil
+ | Eop (Oaddrstack m) Enil => Eop (Oaddrstack (Ptrofs.add (Ptrofs.of_int n) m)) Enil
+ | Eop (Oaddimm m) (t ::: Enil) => Eop (Oaddimm(Int.add n m)) (t ::: Enil)
+ | _ => Eop (Oaddimm n) (e ::: Enil)
+ end.
+
+Nondetfunction add (e1: expr) (e2: expr) :=
+ match e1, e2 with
+ | Eop (Ointconst n1) Enil, t2 => addimm n1 t2
+ | t1, Eop (Ointconst n2) Enil => addimm n2 t1
+ | Eop (Oaddimm n1) (t1:::Enil), Eop (Oaddimm n2) (t2:::Enil) =>
+ addimm (Int.add n1 n2) (Eop Oadd (t1:::t2:::Enil))
+ | Eop (Oaddimm n1) (t1:::Enil), Eop (Oaddrstack n2) Enil =>
+ Eop Oadd (Eop (Oaddrstack (Ptrofs.add (Ptrofs.of_int n1) n2)) Enil ::: t1 ::: Enil)
+ | Eop (Oaddrstack n1) Enil, Eop (Oaddimm n2) (t2:::Enil) =>
+ Eop Oadd (Eop (Oaddrstack (Ptrofs.add n1 (Ptrofs.of_int n2))) Enil ::: t2 ::: Enil)
+ | Eop (Oaddimm n1) (t1:::Enil), t2 =>
+ addimm n1 (Eop Oadd (t1:::t2:::Enil))
+ | t1, Eop (Oaddimm n2) (t2:::Enil) =>
+ addimm n2 (Eop Oadd (t1:::t2:::Enil))
+ | _, _ => Eop Oadd (e1:::e2:::Enil)
+ end.
+
+(** ** Integer and pointer subtraction *)
+
+Nondetfunction sub (e1: expr) (e2: expr) :=
+ match e1, e2 with
+ | t1, Eop (Ointconst n2) Enil =>
+ addimm (Int.neg n2) t1
+ | Eop (Oaddimm n1) (t1:::Enil), Eop (Oaddimm n2) (t2:::Enil) =>
+ addimm (Int.sub n1 n2) (Eop Osub (t1:::t2:::Enil))
+ | Eop (Oaddimm n1) (t1:::Enil), t2 =>
+ addimm n1 (Eop Osub (t1:::t2:::Enil))
+ | t1, Eop (Oaddimm n2) (t2:::Enil) =>
+ addimm (Int.neg n2) (Eop Osub (t1:::t2:::Enil))
+ | _, _ => Eop Osub (e1:::e2:::Enil)
+ end.
+
+Nondetfunction negint (e: expr) :=
+ match e with
+ | Eop (Ointconst n) Enil => Eop (Ointconst (Int.neg n)) Enil
+ | _ => Eop Oneg (e ::: Enil)
+ end.
+
+(** ** Immediate shifts *)
+
+Nondetfunction shlimm (e1: expr) (n: int) :=
+ if Int.eq n Int.zero then
+ e1
+ else if negb (Int.ltu n Int.iwordsize) then
+ Eop Oshl (e1 ::: Eop (Ointconst n) Enil ::: Enil)
+ else match e1 with
+ | Eop (Ointconst n1) Enil =>
+ Eop (Ointconst (Int.shl n1 n)) Enil
+ | Eop (Oshlimm n1) (t1:::Enil) =>
+ if Int.ltu (Int.add n n1) Int.iwordsize
+ then Eop (Oshlimm (Int.add n n1)) (t1:::Enil)
+ else Eop (Oshlimm n) (e1:::Enil)
+ | _ =>
+ Eop (Oshlimm n) (e1:::Enil)
+ end.
+
+Nondetfunction shruimm (e1: expr) (n: int) :=
+ if Int.eq n Int.zero then
+ e1
+ else if negb (Int.ltu n Int.iwordsize) then
+ Eop Oshru (e1 ::: Eop (Ointconst n) Enil ::: Enil)
+ else match e1 with
+ | Eop (Ointconst n1) Enil =>
+ Eop (Ointconst (Int.shru n1 n)) Enil
+ | Eop (Oshruimm n1) (t1:::Enil) =>
+ if Int.ltu (Int.add n n1) Int.iwordsize
+ then Eop (Oshruimm (Int.add n n1)) (t1:::Enil)
+ else Eop (Oshruimm n) (e1:::Enil)
+ | _ =>
+ Eop (Oshruimm n) (e1:::Enil)
+ end.
+
+Nondetfunction shrimm (e1: expr) (n: int) :=
+ if Int.eq n Int.zero then
+ e1
+ else if negb (Int.ltu n Int.iwordsize) then
+ Eop Oshr (e1 ::: Eop (Ointconst n) Enil ::: Enil)
+ else match e1 with
+ | Eop (Ointconst n1) Enil =>
+ Eop (Ointconst (Int.shr n1 n)) Enil
+ | Eop (Oshrimm n1) (t1:::Enil) =>
+ if Int.ltu (Int.add n n1) Int.iwordsize
+ then Eop (Oshrimm (Int.add n n1)) (t1:::Enil)
+ else Eop (Oshrimm n) (e1:::Enil)
+ | _ =>
+ Eop (Oshrimm n) (e1:::Enil)
+ end.
+
+(** ** Integer multiply *)
+
+Definition mulimm_base (n1: int) (e2: expr) :=
+ match Int.one_bits n1 with
+ | i :: nil =>
+ shlimm e2 i
+ | i :: j :: nil =>
+ Elet e2 (add (shlimm (Eletvar 0) i) (shlimm (Eletvar 0) j))
+ | _ =>
+ Eop Omul (Eop (Ointconst n1) Enil ::: e2 ::: Enil)
+ end.
+
+Nondetfunction mulimm (n1: int) (e2: expr) :=
+ if Int.eq n1 Int.zero then Eop (Ointconst Int.zero) Enil
+ else if Int.eq n1 Int.one then e2
+ else match e2 with
+ | Eop (Ointconst n2) Enil => Eop (Ointconst (Int.mul n1 n2)) Enil
+ | Eop (Oaddimm n2) (t2:::Enil) => addimm (Int.mul n1 n2) (mulimm_base n1 t2)
+ | _ => mulimm_base n1 e2
+ end.
+
+Nondetfunction mul (e1: expr) (e2: expr) :=
+ match e1, e2 with
+ | Eop (Ointconst n1) Enil, t2 => mulimm n1 t2
+ | t1, Eop (Ointconst n2) Enil => mulimm n2 t1
+ | _, _ => Eop Omul (e1:::e2:::Enil)
+ end.
+
+Definition mulhs (e1: expr) (e2: expr) :=
+ if Archi.ptr64 then
+ Eop Olowlong
+ (Eop (Oshrlimm (Int.repr 32))
+ (Eop Omull (Eop Ocast32signed (e1 ::: Enil) :::
+ Eop Ocast32signed (e2 ::: Enil) ::: Enil) ::: Enil)
+ ::: Enil)
+ else
+ Eop Omulhs (e1 ::: e2 ::: Enil).
+
+Definition mulhu (e1: expr) (e2: expr) :=
+ if Archi.ptr64 then
+ Eop Olowlong
+ (Eop (Oshrluimm (Int.repr 32))
+ (Eop Omull (Eop Ocast32unsigned (e1 ::: Enil) :::
+ Eop Ocast32unsigned (e2 ::: Enil) ::: Enil) ::: Enil)
+ ::: Enil)
+ else
+ Eop Omulhu (e1 ::: e2 ::: Enil).
+
+(** ** Bitwise and, or, xor *)
+
+Nondetfunction andimm (n1: int) (e2: expr) :=
+ if Int.eq n1 Int.zero then Eop (Ointconst Int.zero) Enil
+ else if Int.eq n1 Int.mone then e2
+ else match e2 with
+ | Eop (Ointconst n2) Enil => Eop (Ointconst (Int.and n1 n2)) Enil
+ | Eop (Oandimm n2) (t2:::Enil) => Eop (Oandimm (Int.and n1 n2)) (t2:::Enil)
+ | _ => Eop (Oandimm n1) (e2:::Enil)
+ end.
+
+Nondetfunction and (e1: expr) (e2: expr) :=
+ match e1, e2 with
+ | Eop (Ointconst n1) Enil, t2 => andimm n1 t2
+ | t1, Eop (Ointconst n2) Enil => andimm n2 t1
+ | _, _ => Eop Oand (e1:::e2:::Enil)
+ end.
+
+Nondetfunction orimm (n1: int) (e2: expr) :=
+ if Int.eq n1 Int.zero then e2
+ else if Int.eq n1 Int.mone then Eop (Ointconst Int.mone) Enil
+ else match e2 with
+ | Eop (Ointconst n2) Enil => Eop (Ointconst (Int.or n1 n2)) Enil
+ | Eop (Oorimm n2) (t2:::Enil) => Eop (Oorimm (Int.or n1 n2)) (t2:::Enil)
+ | _ => Eop (Oorimm n1) (e2:::Enil)
+ end.
+
+Nondetfunction or (e1: expr) (e2: expr) :=
+ match e1, e2 with
+ | Eop (Ointconst n1) Enil, t2 => orimm n1 t2
+ | t1, Eop (Ointconst n2) Enil => orimm n2 t1
+ | _, _ => Eop Oor (e1:::e2:::Enil)
+ end.
+
+Nondetfunction xorimm (n1: int) (e2: expr) :=
+ if Int.eq n1 Int.zero then e2 else
+ match e2 with
+ | Eop (Ointconst n2) Enil => Eop (Ointconst (Int.xor n1 n2)) Enil
+ | Eop (Oxorimm n2) (t2:::Enil) =>
+ let n := Int.xor n1 n2 in
+ if Int.eq n Int.zero then t2 else Eop (Oxorimm n) (t2:::Enil)
+ | _ => Eop (Oxorimm n1) (e2:::Enil)
+ end.
+
+Nondetfunction xor (e1: expr) (e2: expr) :=
+ match e1, e2 with
+ | Eop (Ointconst n1) Enil, t2 => xorimm n1 t2
+ | t1, Eop (Ointconst n2) Enil => xorimm n2 t1
+ | _, _ => Eop Oxor (e1:::e2:::Enil)
+ end.
+
+(** ** Integer logical negation *)
+
+Definition notint (e: expr) := xorimm Int.mone e.
+
+(** ** Integer division and modulus *)
+
+Definition divs_base (e1: expr) (e2: expr) := Eop Odiv (e1:::e2:::Enil).
+Definition mods_base (e1: expr) (e2: expr) := Eop Omod (e1:::e2:::Enil).
+Definition divu_base (e1: expr) (e2: expr) := Eop Odivu (e1:::e2:::Enil).
+Definition modu_base (e1: expr) (e2: expr) := Eop Omodu (e1:::e2:::Enil).
+
+Definition shrximm (e1: expr) (n2: int) :=
+ if Int.eq n2 Int.zero then e1 else Eop (Oshrximm n2) (e1:::Enil).
+
+(* Alternate definition, not convenient for strength reduction during constant propagation *)
+(*
+(* n2 will be less than 31. *)
+
+Definition shrximm_inner (e1: expr) (n2: int) :=
+ Eop (Oshruimm (Int.sub Int.iwordsize n2))
+ ((Eop (Oshrimm (Int.repr (Int.zwordsize - 1)))
+ (e1 ::: Enil))
+ ::: Enil).
+
+Definition shrximm (e1: expr) (n2: int) :=
+ if Int.eq n2 Int.zero then e1
+ else Eop (Oshrimm n2)
+ ((Eop Oadd (e1 ::: shrximm_inner e1 n2 ::: Enil))
+ ::: Enil).
+*)
+
+(** ** General shifts *)
+
+Nondetfunction shl (e1: expr) (e2: expr) :=
+ match e2 with
+ | Eop (Ointconst n2) Enil => shlimm e1 n2
+ | _ => Eop Oshl (e1:::e2:::Enil)
+ end.
+
+Nondetfunction shr (e1: expr) (e2: expr) :=
+ match e2 with
+ | Eop (Ointconst n2) Enil => shrimm e1 n2
+ | _ => Eop Oshr (e1:::e2:::Enil)
+ end.
+
+Nondetfunction shru (e1: expr) (e2: expr) :=
+ match e2 with
+ | Eop (Ointconst n2) Enil => shruimm e1 n2
+ | _ => Eop Oshru (e1:::e2:::Enil)
+ end.
+
+(** ** Floating-point arithmetic *)
+
+Definition negf (e: expr) := Eop Onegf (e ::: Enil).
+Definition absf (e: expr) := Eop Oabsf (e ::: Enil).
+Definition addf (e1 e2: expr) := Eop Oaddf (e1 ::: e2 ::: Enil).
+Definition subf (e1 e2: expr) := Eop Osubf (e1 ::: e2 ::: Enil).
+Definition mulf (e1 e2: expr) := Eop Omulf (e1 ::: e2 ::: Enil).
+
+Definition negfs (e: expr) := Eop Onegfs (e ::: Enil).
+Definition absfs (e: expr) := Eop Oabsfs (e ::: Enil).
+Definition addfs (e1 e2: expr) := Eop Oaddfs (e1 ::: e2 ::: Enil).
+Definition subfs (e1 e2: expr) := Eop Osubfs (e1 ::: e2 ::: Enil).
+Definition mulfs (e1 e2: expr) := Eop Omulfs (e1 ::: e2 ::: Enil).
+
+(** ** Comparisons *)
+
+Nondetfunction compimm (default: comparison -> int -> condition)
+ (sem: comparison -> int -> int -> bool)
+ (c: comparison) (e1: expr) (n2: int) :=
+ match c, e1 with
+ | c, Eop (Ointconst n1) Enil =>
+ Eop (Ointconst (if sem c n1 n2 then Int.one else Int.zero)) Enil
+ | Ceq, Eop (Ocmp c) el =>
+ if Int.eq_dec n2 Int.zero then
+ Eop (Ocmp (negate_condition c)) el
+ else if Int.eq_dec n2 Int.one then
+ Eop (Ocmp c) el
+ else
+ Eop (Ointconst Int.zero) Enil
+ | Cne, Eop (Ocmp c) el =>
+ if Int.eq_dec n2 Int.zero then
+ Eop (Ocmp c) el
+ else if Int.eq_dec n2 Int.one then
+ Eop (Ocmp (negate_condition c)) el
+ else
+ Eop (Ointconst Int.one) Enil
+ | _, _ =>
+ Eop (Ocmp (default c n2)) (e1 ::: Enil)
+ end.
+
+Nondetfunction comp (c: comparison) (e1: expr) (e2: expr) :=
+ match e1, e2 with
+ | Eop (Ointconst n1) Enil, t2 =>
+ compimm Ccompimm Int.cmp (swap_comparison c) t2 n1
+ | t1, Eop (Ointconst n2) Enil =>
+ compimm Ccompimm Int.cmp c t1 n2
+ | _, _ =>
+ Eop (Ocmp (Ccomp c)) (e1 ::: e2 ::: Enil)
+ end.
+
+Nondetfunction compu (c: comparison) (e1: expr) (e2: expr) :=
+ match e1, e2 with
+ | Eop (Ointconst n1) Enil, t2 =>
+ compimm Ccompuimm Int.cmpu (swap_comparison c) t2 n1
+ | t1, Eop (Ointconst n2) Enil =>
+ compimm Ccompuimm Int.cmpu c t1 n2
+ | _, _ =>
+ Eop (Ocmp (Ccompu c)) (e1 ::: e2 ::: Enil)
+ end.
+
+Definition compf (c: comparison) (e1: expr) (e2: expr) :=
+ Eop (Ocmp (Ccompf c)) (e1 ::: e2 ::: Enil).
+
+Definition compfs (c: comparison) (e1: expr) (e2: expr) :=
+ Eop (Ocmp (Ccompfs c)) (e1 ::: e2 ::: Enil).
+
+(** ** Integer conversions *)
+
+Definition cast8unsigned (e: expr) := andimm (Int.repr 255) e.
+
+Nondetfunction cast8signed (e: expr) :=
+ match e with
+ | Eop (Ointconst n) Enil => Eop (Ointconst (Int.sign_ext 8 n)) Enil
+ | _ => Eop Ocast8signed (e ::: Enil)
+ end.
+
+Definition cast16unsigned (e: expr) := andimm (Int.repr 65535) e.
+
+Nondetfunction cast16signed (e: expr) :=
+ match e with
+ | Eop (Ointconst n) Enil => Eop (Ointconst (Int.sign_ext 16 n)) Enil
+ | _ => Eop Ocast16signed (e ::: Enil)
+ end.
+
+(** ** Floating-point conversions *)
+
+Definition intoffloat (e: expr) := Eop Ointoffloat (e ::: Enil).
+Definition intuoffloat (e: expr) := Eop Ointuoffloat (e ::: Enil).
+
+Nondetfunction floatofintu (e: expr) :=
+ match e with
+ | Eop (Ointconst n) Enil => Eop (Ofloatconst (Float.of_intu n)) Enil
+ | _ => Eop Ofloatofintu (e ::: Enil)
+ end.
+
+Nondetfunction floatofint (e: expr) :=
+ match e with
+ | Eop (Ointconst n) Enil => Eop (Ofloatconst (Float.of_int n)) Enil
+ | _ => Eop Ofloatofint (e ::: Enil)
+ end.
+
+Definition intofsingle (e: expr) := Eop Ointofsingle (e ::: Enil).
+Definition singleofint (e: expr) := Eop Osingleofint (e ::: Enil).
+
+Definition intuofsingle (e: expr) := Eop Ointuofsingle (e ::: Enil).
+Definition singleofintu (e: expr) := Eop Osingleofintu (e ::: Enil).
+
+Definition singleoffloat (e: expr) := Eop Osingleoffloat (e ::: Enil).
+Definition floatofsingle (e: expr) := Eop Ofloatofsingle (e ::: Enil).
+
+(** ** Recognition of addressing modes for load and store operations *)
+
+Nondetfunction addressing (chunk: memory_chunk) (e: expr) :=
+ match e with
+ | Eop (Oaddrstack n) Enil => (Ainstack n, Enil)
+ | Eop (Oaddrsymbol id ofs) Enil => if Archi.pic_code tt then (Aindexed Ptrofs.zero, e:::Enil) else (Aglobal id ofs, Enil)
+ | Eop (Oaddimm n) (e1:::Enil) => (Aindexed (Ptrofs.of_int n), e1:::Enil)
+ | Eop (Oaddlimm n) (e1:::Enil) => (Aindexed (Ptrofs.of_int64 n), e1:::Enil)
+ | _ => (Aindexed Ptrofs.zero, e:::Enil)
+ end.
+
+(** ** Arguments of builtins *)
+
+Nondetfunction builtin_arg (e: expr) :=
+ match e with
+ | Eop (Ointconst n) Enil => BA_int n
+ | Eop (Oaddrsymbol id ofs) Enil => BA_addrglobal id ofs
+ | Eop (Oaddrstack ofs) Enil => BA_addrstack ofs
+ | Eop Omakelong (Eop (Ointconst h) Enil ::: Eop (Ointconst l) Enil ::: Enil) =>
+ BA_long (Int64.ofwords h l)
+ | Eop Omakelong (h ::: l ::: Enil) => BA_splitlong (BA h) (BA l)
+ | Eload chunk (Ainstack ofs) Enil => BA_loadstack chunk ofs
+ | _ => BA e
+ end.
diff --git a/riscV/SelectOpproof.v b/riscV/SelectOpproof.v
new file mode 100644
index 00000000..b9652b34
--- /dev/null
+++ b/riscV/SelectOpproof.v
@@ -0,0 +1,915 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Correctness of instruction selection for operators *)
+
+Require Import Coqlib.
+Require Import Maps.
+Require Import AST.
+Require Import Integers.
+Require Import Floats.
+Require Import Values.
+Require Import Memory.
+Require Import Globalenvs.
+Require Import Cminor.
+Require Import Op.
+Require Import CminorSel.
+Require Import SelectOp.
+
+Local Open Scope cminorsel_scope.
+
+(** * Useful lemmas and tactics *)
+
+(** The following are trivial lemmas and custom tactics that help
+ perform backward (inversion) and forward reasoning over the evaluation
+ of operator applications. *)
+
+Ltac EvalOp := eapply eval_Eop; eauto with evalexpr.
+
+Ltac InvEval1 :=
+ match goal with
+ | [ H: (eval_expr _ _ _ _ _ (Eop _ Enil) _) |- _ ] =>
+ inv H; InvEval1
+ | [ H: (eval_expr _ _ _ _ _ (Eop _ (_ ::: Enil)) _) |- _ ] =>
+ inv H; InvEval1
+ | [ H: (eval_expr _ _ _ _ _ (Eop _ (_ ::: _ ::: Enil)) _) |- _ ] =>
+ inv H; InvEval1
+ | [ H: (eval_exprlist _ _ _ _ _ Enil _) |- _ ] =>
+ inv H; InvEval1
+ | [ H: (eval_exprlist _ _ _ _ _ (_ ::: _) _) |- _ ] =>
+ inv H; InvEval1
+ | _ =>
+ idtac
+ end.
+
+Ltac InvEval2 :=
+ match goal with
+ | [ H: (eval_operation _ _ _ nil _ = Some _) |- _ ] =>
+ simpl in H; inv H
+ | [ H: (eval_operation _ _ _ (_ :: nil) _ = Some _) |- _ ] =>
+ simpl in H; FuncInv
+ | [ H: (eval_operation _ _ _ (_ :: _ :: nil) _ = Some _) |- _ ] =>
+ simpl in H; FuncInv
+ | [ H: (eval_operation _ _ _ (_ :: _ :: _ :: nil) _ = Some _) |- _ ] =>
+ simpl in H; FuncInv
+ | _ =>
+ idtac
+ end.
+
+Ltac InvEval := InvEval1; InvEval2; InvEval2.
+
+Ltac TrivialExists :=
+ match goal with
+ | [ |- exists v, _ /\ Val.lessdef ?a v ] => exists a; split; [EvalOp | auto]
+ end.
+
+(** * Correctness of the smart constructors *)
+
+Section CMCONSTR.
+
+Variable ge: genv.
+Variable sp: val.
+Variable e: env.
+Variable m: mem.
+
+(** We now show that the code generated by "smart constructor" functions
+ such as [Selection.notint] behaves as expected. Continuing the
+ [notint] example, we show that if the expression [e]
+ evaluates to some integer value [Vint n], then [Selection.notint e]
+ evaluates to a value [Vint (Int.not n)] which is indeed the integer
+ negation of the value of [e].
+
+ All proofs follow a common pattern:
+- Reasoning by case over the result of the classification functions
+ (such as [add_match] for integer addition), gathering additional
+ information on the shape of the argument expressions in the non-default
+ cases.
+- Inversion of the evaluations of the arguments, exploiting the additional
+ information thus gathered.
+- Equational reasoning over the arithmetic operations performed,
+ using the lemmas from the [Int] and [Float] modules.
+- Construction of an evaluation derivation for the expression returned
+ by the smart constructor.
+*)
+
+Definition unary_constructor_sound (cstr: expr -> expr) (sem: val -> val) : Prop :=
+ forall le a x,
+ eval_expr ge sp e m le a x ->
+ exists v, eval_expr ge sp e m le (cstr a) v /\ Val.lessdef (sem x) v.
+
+Definition binary_constructor_sound (cstr: expr -> expr -> expr) (sem: val -> val -> val) : Prop :=
+ forall le a x b y,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ exists v, eval_expr ge sp e m le (cstr a b) v /\ Val.lessdef (sem x y) v.
+
+Theorem eval_addrsymbol:
+ forall le id ofs,
+ exists v, eval_expr ge sp e m le (addrsymbol id ofs) v /\ Val.lessdef (Genv.symbol_address ge id ofs) v.
+Proof.
+ intros. unfold addrsymbol. econstructor; split.
+ EvalOp. simpl; eauto.
+ auto.
+Qed.
+
+Theorem eval_addrstack:
+ forall le ofs,
+ exists v, eval_expr ge sp e m le (addrstack ofs) v /\ Val.lessdef (Val.offset_ptr sp ofs) v.
+Proof.
+ intros. unfold addrstack. econstructor; split.
+ EvalOp. simpl; eauto.
+ auto.
+Qed.
+
+Theorem eval_addimm:
+ forall n, unary_constructor_sound (addimm n) (fun x => Val.add x (Vint n)).
+Proof.
+ red; unfold addimm; intros until x.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ - subst n. intros. exists x; split; auto.
+ destruct x; simpl; auto.
+ rewrite Int.add_zero; auto.
+ destruct Archi.ptr64; auto. rewrite Ptrofs.add_zero; auto.
+ - case (addimm_match a); intros; InvEval; simpl.
+ + TrivialExists; simpl. rewrite Int.add_commut. auto.
+ + econstructor; split. EvalOp. simpl; eauto.
+ unfold Genv.symbol_address. destruct (Genv.find_symbol ge s); simpl; auto.
+ destruct Archi.ptr64; auto. rewrite Ptrofs.add_commut; auto.
+ + econstructor; split. EvalOp. simpl; eauto.
+ destruct sp; simpl; auto. destruct Archi.ptr64; auto.
+ rewrite Ptrofs.add_assoc. rewrite (Ptrofs.add_commut m0). auto.
+ + TrivialExists; simpl. subst x. rewrite Val.add_assoc. rewrite Int.add_commut. auto.
+ + TrivialExists.
+Qed.
+
+Theorem eval_add: binary_constructor_sound add Val.add.
+Proof.
+ red; intros until y.
+ unfold add; case (add_match a b); intros; InvEval.
+ - rewrite Val.add_commut. apply eval_addimm; auto.
+ - apply eval_addimm; auto.
+ - subst.
+ replace (Val.add (Val.add v1 (Vint n1)) (Val.add v0 (Vint n2)))
+ with (Val.add (Val.add v1 v0) (Val.add (Vint n1) (Vint n2))).
+ apply eval_addimm. EvalOp.
+ repeat rewrite Val.add_assoc. decEq. apply Val.add_permut.
+ - subst. econstructor; split.
+ EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto.
+ rewrite Val.add_commut. destruct sp; simpl; auto.
+ destruct v1; simpl; auto.
+ destruct Archi.ptr64 eqn:SF; auto.
+ apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal.
+ rewrite (Ptrofs.add_commut (Ptrofs.of_int n1)), Ptrofs.add_assoc. f_equal. auto with ptrofs.
+ destruct Archi.ptr64 eqn:SF; auto.
+ - subst. econstructor; split.
+ EvalOp. constructor. EvalOp. simpl; eauto. constructor. eauto. constructor. simpl; eauto.
+ destruct sp; simpl; auto.
+ destruct v1; simpl; auto.
+ destruct Archi.ptr64 eqn:SF; auto.
+ apply Val.lessdef_same. f_equal. rewrite ! Ptrofs.add_assoc. f_equal. f_equal.
+ rewrite Ptrofs.add_commut. auto with ptrofs.
+ destruct Archi.ptr64 eqn:SF; auto.
+ - subst.
+ replace (Val.add (Val.add v1 (Vint n1)) y)
+ with (Val.add (Val.add v1 y) (Vint n1)).
+ apply eval_addimm. EvalOp.
+ repeat rewrite Val.add_assoc. decEq. apply Val.add_commut.
+ - subst.
+ replace (Val.add x (Val.add v1 (Vint n2)))
+ with (Val.add (Val.add x v1) (Vint n2)).
+ apply eval_addimm. EvalOp.
+ repeat rewrite Val.add_assoc. reflexivity.
+ - TrivialExists.
+Qed.
+
+Theorem eval_sub: binary_constructor_sound sub Val.sub.
+Proof.
+ red; intros until y.
+ unfold sub; case (sub_match a b); intros; InvEval.
+ - rewrite Val.sub_add_opp. apply eval_addimm; auto.
+ - subst. rewrite Val.sub_add_l. rewrite Val.sub_add_r.
+ rewrite Val.add_assoc. simpl. rewrite Int.add_commut. rewrite <- Int.sub_add_opp.
+ apply eval_addimm; EvalOp.
+ - subst. rewrite Val.sub_add_l. apply eval_addimm; EvalOp.
+ - subst. rewrite Val.sub_add_r. apply eval_addimm; EvalOp.
+ - TrivialExists.
+Qed.
+
+Theorem eval_negint: unary_constructor_sound negint (fun v => Val.sub Vzero v).
+Proof.
+ red; intros until x. unfold negint. case (negint_match a); intros; InvEval.
+ TrivialExists.
+ TrivialExists.
+Qed.
+
+Theorem eval_shlimm:
+ forall n, unary_constructor_sound (fun a => shlimm a n)
+ (fun x => Val.shl x (Vint n)).
+Proof.
+ red; intros until x. unfold shlimm.
+
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shl_zero; auto.
+
+ destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl.
+ destruct (shlimm_match a); intros; InvEval.
+ - exists (Vint (Int.shl n1 n)); split. EvalOp.
+ simpl. rewrite LT. auto.
+ - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?.
+ + exists (Val.shl v1 (Vint (Int.add n n1))); split. EvalOp.
+ subst. destruct v1; simpl; auto.
+ rewrite Heqb.
+ destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto.
+ destruct (Int.ltu n Int.iwordsize) eqn:?; simpl; auto.
+ rewrite Int.add_commut. rewrite Int.shl_shl; auto. rewrite Int.add_commut; auto.
+ + subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor.
+ simpl. auto.
+ - TrivialExists.
+ - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
+ auto.
+Qed.
+
+Theorem eval_shruimm:
+ forall n, unary_constructor_sound (fun a => shruimm a n)
+ (fun x => Val.shru x (Vint n)).
+Proof.
+ red; intros until x. unfold shruimm.
+
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shru_zero; auto.
+
+ destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl.
+ destruct (shruimm_match a); intros; InvEval.
+ - exists (Vint (Int.shru n1 n)); split. EvalOp.
+ simpl. rewrite LT; auto.
+ - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?.
+ exists (Val.shru v1 (Vint (Int.add n n1))); split. EvalOp.
+ subst. destruct v1; simpl; auto.
+ rewrite Heqb.
+ destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto.
+ rewrite LT. rewrite Int.add_commut. rewrite Int.shru_shru; auto. rewrite Int.add_commut; auto.
+ subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor.
+ simpl. auto.
+ - TrivialExists.
+ - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
+ auto.
+Qed.
+
+Theorem eval_shrimm:
+ forall n, unary_constructor_sound (fun a => shrimm a n)
+ (fun x => Val.shr x (Vint n)).
+Proof.
+ red; intros until x. unfold shrimm.
+
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ intros; subst. exists x; split; auto. destruct x; simpl; auto. rewrite Int.shr_zero; auto.
+
+ destruct (Int.ltu n Int.iwordsize) eqn:LT; simpl.
+ destruct (shrimm_match a); intros; InvEval.
+ - exists (Vint (Int.shr n1 n)); split. EvalOp.
+ simpl. rewrite LT; auto.
+ - destruct (Int.ltu (Int.add n n1) Int.iwordsize) eqn:?.
+ exists (Val.shr v1 (Vint (Int.add n n1))); split. EvalOp.
+ subst. destruct v1; simpl; auto.
+ rewrite Heqb.
+ destruct (Int.ltu n1 Int.iwordsize) eqn:?; simpl; auto.
+ rewrite LT.
+ rewrite Int.add_commut. rewrite Int.shr_shr; auto. rewrite Int.add_commut; auto.
+ subst. TrivialExists. econstructor. EvalOp. simpl; eauto. constructor.
+ simpl. auto.
+ - TrivialExists.
+ - intros; TrivialExists. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
+ auto.
+Qed.
+
+Lemma eval_mulimm_base:
+ forall n, unary_constructor_sound (mulimm_base n) (fun x => Val.mul x (Vint n)).
+Proof.
+ intros; red; intros; unfold mulimm_base.
+
+ assert (DFL: exists v, eval_expr ge sp e m le (Eop Omul (Eop (Ointconst n) Enil ::: a ::: Enil)) v /\ Val.lessdef (Val.mul x (Vint n)) v).
+ TrivialExists. econstructor. EvalOp. simpl; eauto. econstructor. eauto. constructor.
+ rewrite Val.mul_commut. auto.
+
+ generalize (Int.one_bits_decomp n).
+ generalize (Int.one_bits_range n).
+ destruct (Int.one_bits n).
+ - intros. auto.
+ - destruct l.
+ + intros. rewrite H1. simpl.
+ rewrite Int.add_zero.
+ replace (Vint (Int.shl Int.one i)) with (Val.shl Vone (Vint i)). rewrite Val.shl_mul.
+ apply eval_shlimm. auto. simpl. rewrite H0; auto with coqlib.
+ + destruct l.
+ intros. rewrite H1. simpl.
+ exploit (eval_shlimm i (x :: le) (Eletvar 0) x). constructor; auto. intros [v1 [A1 B1]].
+ exploit (eval_shlimm i0 (x :: le) (Eletvar 0) x). constructor; auto. intros [v2 [A2 B2]].
+ exploit (eval_add (x :: le)). eexact A1. eexact A2. intros [v [A B]].
+ exists v; split. econstructor; eauto.
+ rewrite Int.add_zero.
+ replace (Vint (Int.add (Int.shl Int.one i) (Int.shl Int.one i0)))
+ with (Val.add (Val.shl Vone (Vint i)) (Val.shl Vone (Vint i0))).
+ rewrite Val.mul_add_distr_r.
+ repeat rewrite Val.shl_mul. eapply Val.lessdef_trans. 2: eauto. apply Val.add_lessdef; auto.
+ simpl. repeat rewrite H0; auto with coqlib.
+ intros. auto.
+Qed.
+
+Theorem eval_mulimm:
+ forall n, unary_constructor_sound (mulimm n) (fun x => Val.mul x (Vint n)).
+Proof.
+ intros; red; intros until x; unfold mulimm.
+
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ intros. exists (Vint Int.zero); split. EvalOp.
+ destruct x; simpl; auto. subst n. rewrite Int.mul_zero. auto.
+
+ predSpec Int.eq Int.eq_spec n Int.one.
+ intros. exists x; split; auto.
+ destruct x; simpl; auto. subst n. rewrite Int.mul_one. auto.
+
+ case (mulimm_match a); intros; InvEval.
+ - TrivialExists. simpl. rewrite Int.mul_commut; auto.
+ - subst. rewrite Val.mul_add_distr_l.
+ exploit eval_mulimm_base; eauto. instantiate (1 := n). intros [v' [A1 B1]].
+ exploit (eval_addimm (Int.mul n n2) le (mulimm_base n t2) v'). auto. intros [v'' [A2 B2]].
+ exists v''; split; auto. eapply Val.lessdef_trans. eapply Val.add_lessdef; eauto.
+ rewrite Val.mul_commut; auto.
+ - apply eval_mulimm_base; auto.
+Qed.
+
+Theorem eval_mul: binary_constructor_sound mul Val.mul.
+Proof.
+ red; intros until y.
+ unfold mul; case (mul_match a b); intros; InvEval.
+ rewrite Val.mul_commut. apply eval_mulimm. auto.
+ apply eval_mulimm. auto.
+ TrivialExists.
+Qed.
+
+Theorem eval_mulhs: binary_constructor_sound mulhs Val.mulhs.
+Proof.
+ red; intros. unfold mulhs; destruct Archi.ptr64 eqn:SF.
+- econstructor; split.
+ EvalOp. constructor. EvalOp. constructor. EvalOp. constructor. EvalOp. simpl; eauto.
+ constructor. EvalOp. simpl; eauto. constructor.
+ simpl; eauto. constructor. simpl; eauto. constructor. simpl; eauto.
+ destruct x; simpl; auto. destruct y; simpl; auto.
+ change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl.
+ apply Val.lessdef_same. f_equal.
+ transitivity (Int.repr (Z.shiftr (Int.signed i * Int.signed i0) 32)).
+ unfold Int.mulhs; f_equal. rewrite Int.Zshiftr_div_two_p by omega. reflexivity.
+ apply Int.same_bits_eq; intros n N.
+ change Int.zwordsize with 32 in *.
+ assert (N1: 0 <= n < 64) by omega.
+ rewrite Int64.bits_loword by auto.
+ rewrite Int64.bits_shr' by auto.
+ change (Int.unsigned (Int.repr 32)) with 32. change Int64.zwordsize with 64.
+ rewrite zlt_true by omega.
+ rewrite Int.testbit_repr by auto.
+ unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; omega).
+ transitivity (Z.testbit (Int.signed i * Int.signed i0) (n + 32)).
+ rewrite Z.shiftr_spec by omega. auto.
+ apply Int64.same_bits_eqm. apply Int64.eqm_mult; apply Int64.eqm_unsigned_repr.
+ change Int64.zwordsize with 64; omega.
+- TrivialExists.
+Qed.
+
+Theorem eval_mulhu: binary_constructor_sound mulhu Val.mulhu.
+Proof.
+ red; intros. unfold mulhu; destruct Archi.ptr64 eqn:SF.
+- econstructor; split.
+ EvalOp. constructor. EvalOp. constructor. EvalOp. constructor. EvalOp. simpl; eauto.
+ constructor. EvalOp. simpl; eauto. constructor.
+ simpl; eauto. constructor. simpl; eauto. constructor. simpl; eauto.
+ destruct x; simpl; auto. destruct y; simpl; auto.
+ change (Int.ltu (Int.repr 32) Int64.iwordsize') with true; simpl.
+ apply Val.lessdef_same. f_equal.
+ transitivity (Int.repr (Z.shiftr (Int.unsigned i * Int.unsigned i0) 32)).
+ unfold Int.mulhu; f_equal. rewrite Int.Zshiftr_div_two_p by omega. reflexivity.
+ apply Int.same_bits_eq; intros n N.
+ change Int.zwordsize with 32 in *.
+ assert (N1: 0 <= n < 64) by omega.
+ rewrite Int64.bits_loword by auto.
+ rewrite Int64.bits_shru' by auto.
+ change (Int.unsigned (Int.repr 32)) with 32. change Int64.zwordsize with 64.
+ rewrite zlt_true by omega.
+ rewrite Int.testbit_repr by auto.
+ unfold Int64.mul. rewrite Int64.testbit_repr by (change Int64.zwordsize with 64; omega).
+ transitivity (Z.testbit (Int.unsigned i * Int.unsigned i0) (n + 32)).
+ rewrite Z.shiftr_spec by omega. auto.
+ apply Int64.same_bits_eqm. apply Int64.eqm_mult; apply Int64.eqm_unsigned_repr.
+ change Int64.zwordsize with 64; omega.
+- TrivialExists.
+Qed.
+
+Theorem eval_andimm:
+ forall n, unary_constructor_sound (andimm n) (fun x => Val.and x (Vint n)).
+Proof.
+ intros; red; intros until x. unfold andimm.
+
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ intros. exists (Vint Int.zero); split. EvalOp.
+ destruct x; simpl; auto. subst n. rewrite Int.and_zero. auto.
+
+ predSpec Int.eq Int.eq_spec n Int.mone.
+ intros. exists x; split; auto.
+ subst. destruct x; simpl; auto. rewrite Int.and_mone; auto.
+
+ case (andimm_match a); intros.
+ - InvEval. TrivialExists. simpl. rewrite Int.and_commut; auto.
+ - InvEval. subst. rewrite Val.and_assoc. simpl. rewrite Int.and_commut. TrivialExists.
+ - TrivialExists.
+Qed.
+
+Theorem eval_and: binary_constructor_sound and Val.and.
+Proof.
+ red; intros until y; unfold and; case (and_match a b); intros; InvEval.
+ - rewrite Val.and_commut. apply eval_andimm; auto.
+ - apply eval_andimm; auto.
+ - TrivialExists.
+Qed.
+
+Theorem eval_orimm:
+ forall n, unary_constructor_sound (orimm n) (fun x => Val.or x (Vint n)).
+Proof.
+ intros; red; intros until x. unfold orimm.
+
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ intros. subst. exists x; split; auto.
+ destruct x; simpl; auto. rewrite Int.or_zero; auto.
+
+ predSpec Int.eq Int.eq_spec n Int.mone.
+ intros. exists (Vint Int.mone); split. EvalOp.
+ destruct x; simpl; auto. subst n. rewrite Int.or_mone. auto.
+
+ destruct (orimm_match a); intros; InvEval.
+ - TrivialExists. simpl. rewrite Int.or_commut; auto.
+ - subst. rewrite Val.or_assoc. simpl. rewrite Int.or_commut. TrivialExists.
+ - TrivialExists.
+Qed.
+
+Theorem eval_or: binary_constructor_sound or Val.or.
+Proof.
+ red; intros until y; unfold or; case (or_match a b); intros; InvEval.
+ - rewrite Val.or_commut. apply eval_orimm; auto.
+ - apply eval_orimm; auto.
+ - TrivialExists.
+Qed.
+
+Theorem eval_xorimm:
+ forall n, unary_constructor_sound (xorimm n) (fun x => Val.xor x (Vint n)).
+Proof.
+ intros; red; intros until x. unfold xorimm.
+
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ intros. exists x; split. auto.
+ destruct x; simpl; auto. subst n. rewrite Int.xor_zero. auto.
+
+ intros. destruct (xorimm_match a); intros; InvEval.
+ - TrivialExists. simpl. rewrite Int.xor_commut; auto.
+ - subst. rewrite Val.xor_assoc. simpl. rewrite Int.xor_commut.
+ predSpec Int.eq Int.eq_spec (Int.xor n2 n) Int.zero.
+ + exists v1; split; auto. destruct v1; simpl; auto. rewrite H0, Int.xor_zero; auto.
+ + TrivialExists.
+ - TrivialExists.
+Qed.
+
+Theorem eval_xor: binary_constructor_sound xor Val.xor.
+Proof.
+ red; intros until y; unfold xor; case (xor_match a b); intros; InvEval.
+ - rewrite Val.xor_commut. apply eval_xorimm; auto.
+ - apply eval_xorimm; auto.
+ - TrivialExists.
+Qed.
+
+Theorem eval_notint: unary_constructor_sound notint Val.notint.
+Proof.
+ unfold notint; red; intros. rewrite Val.not_xor. apply eval_xorimm; auto.
+Qed.
+
+Theorem eval_divs_base:
+ forall le a b x y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.divs x y = Some z ->
+ exists v, eval_expr ge sp e m le (divs_base a b) v /\ Val.lessdef z v.
+Proof.
+ intros. unfold divs_base. exists z; split. EvalOp. auto.
+Qed.
+
+Theorem eval_mods_base:
+ forall le a b x y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.mods x y = Some z ->
+ exists v, eval_expr ge sp e m le (mods_base a b) v /\ Val.lessdef z v.
+Proof.
+ intros. unfold mods_base. exists z; split. EvalOp. auto.
+Qed.
+
+Theorem eval_divu_base:
+ forall le a b x y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.divu x y = Some z ->
+ exists v, eval_expr ge sp e m le (divu_base a b) v /\ Val.lessdef z v.
+Proof.
+ intros. unfold divu_base. exists z; split. EvalOp. auto.
+Qed.
+
+Theorem eval_modu_base:
+ forall le a b x y z,
+ eval_expr ge sp e m le a x ->
+ eval_expr ge sp e m le b y ->
+ Val.modu x y = Some z ->
+ exists v, eval_expr ge sp e m le (modu_base a b) v /\ Val.lessdef z v.
+Proof.
+ intros. unfold modu_base. exists z; split. EvalOp. auto.
+Qed.
+
+Theorem eval_shrximm:
+ forall le a n x z,
+ eval_expr ge sp e m le a x ->
+ Val.shrx x (Vint n) = Some z ->
+ exists v, eval_expr ge sp e m le (shrximm a n) v /\ Val.lessdef z v.
+Proof.
+ intros. unfold shrximm.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ subst n. exists x; split; auto.
+ destruct x; simpl in H0; try discriminate.
+ destruct (Int.ltu Int.zero (Int.repr 31)); inv H0.
+ replace (Int.shrx i Int.zero) with i. auto.
+ unfold Int.shrx, Int.divs. rewrite Int.shl_zero.
+ change (Int.signed Int.one) with 1. rewrite Z.quot_1_r. rewrite Int.repr_signed; auto.
+ econstructor; split. EvalOp. auto.
+(*
+ intros. destruct x; simpl in H0; try discriminate.
+ destruct (Int.ltu n (Int.repr 31)) eqn:LTU; inv H0.
+ unfold shrximm.
+ predSpec Int.eq Int.eq_spec n Int.zero.
+ - subst n. exists (Vint i); split; auto.
+ unfold Int.shrx, Int.divs. rewrite Z.quot_1_r. rewrite Int.repr_signed. auto.
+ - assert (NZ: Int.unsigned n <> 0).
+ { intro EQ; elim H0. rewrite <- (Int.repr_unsigned n). rewrite EQ; auto. }
+ assert (LT: 0 <= Int.unsigned n < 31) by (apply Int.ltu_inv in LTU; assumption).
+ assert (LTU2: Int.ltu (Int.sub Int.iwordsize n) Int.iwordsize = true).
+ { unfold Int.ltu; apply zlt_true.
+ unfold Int.sub. change (Int.unsigned Int.iwordsize) with 32.
+ rewrite Int.unsigned_repr. omega.
+ assert (32 < Int.max_unsigned) by reflexivity. omega. }
+ assert (X: eval_expr ge sp e m le
+ (Eop (Oshrimm (Int.repr (Int.zwordsize - 1))) (a ::: Enil))
+ (Vint (Int.shr i (Int.repr (Int.zwordsize - 1))))).
+ { EvalOp. }
+ assert (Y: eval_expr ge sp e m le (shrximm_inner a n)
+ (Vint (Int.shru (Int.shr i (Int.repr (Int.zwordsize - 1))) (Int.sub Int.iwordsize n)))).
+ { EvalOp. simpl. rewrite LTU2. auto. }
+ TrivialExists.
+ constructor. EvalOp. simpl; eauto. constructor.
+ simpl. unfold Int.ltu; rewrite zlt_true. rewrite Int.shrx_shr_2 by auto. reflexivity.
+ change (Int.unsigned Int.iwordsize) with 32; omega.
+*)
+Qed.
+
+Theorem eval_shl: binary_constructor_sound shl Val.shl.
+Proof.
+ red; intros until y; unfold shl; case (shl_match b); intros.
+ InvEval. apply eval_shlimm; auto.
+ TrivialExists.
+Qed.
+
+Theorem eval_shr: binary_constructor_sound shr Val.shr.
+Proof.
+ red; intros until y; unfold shr; case (shr_match b); intros.
+ InvEval. apply eval_shrimm; auto.
+ TrivialExists.
+Qed.
+
+Theorem eval_shru: binary_constructor_sound shru Val.shru.
+Proof.
+ red; intros until y; unfold shru; case (shru_match b); intros.
+ InvEval. apply eval_shruimm; auto.
+ TrivialExists.
+Qed.
+
+Theorem eval_negf: unary_constructor_sound negf Val.negf.
+Proof.
+ red; intros. TrivialExists.
+Qed.
+
+Theorem eval_absf: unary_constructor_sound absf Val.absf.
+Proof.
+ red; intros. TrivialExists.
+Qed.
+
+Theorem eval_addf: binary_constructor_sound addf Val.addf.
+Proof.
+ red; intros; TrivialExists.
+Qed.
+
+Theorem eval_subf: binary_constructor_sound subf Val.subf.
+Proof.
+ red; intros; TrivialExists.
+Qed.
+
+Theorem eval_mulf: binary_constructor_sound mulf Val.mulf.
+Proof.
+ red; intros; TrivialExists.
+Qed.
+
+Theorem eval_negfs: unary_constructor_sound negfs Val.negfs.
+Proof.
+ red; intros. TrivialExists.
+Qed.
+
+Theorem eval_absfs: unary_constructor_sound absfs Val.absfs.
+Proof.
+ red; intros. TrivialExists.
+Qed.
+
+Theorem eval_addfs: binary_constructor_sound addfs Val.addfs.
+Proof.
+ red; intros; TrivialExists.
+Qed.
+
+Theorem eval_subfs: binary_constructor_sound subfs Val.subfs.
+Proof.
+ red; intros; TrivialExists.
+Qed.
+
+Theorem eval_mulfs: binary_constructor_sound mulfs Val.mulfs.
+Proof.
+ red; intros; TrivialExists.
+Qed.
+
+Section COMP_IMM.
+
+Variable default: comparison -> int -> condition.
+Variable intsem: comparison -> int -> int -> bool.
+Variable sem: comparison -> val -> val -> val.
+
+Hypothesis sem_int: forall c x y, sem c (Vint x) (Vint y) = Val.of_bool (intsem c x y).
+Hypothesis sem_undef: forall c v, sem c Vundef v = Vundef.
+Hypothesis sem_eq: forall x y, sem Ceq (Vint x) (Vint y) = Val.of_bool (Int.eq x y).
+Hypothesis sem_ne: forall x y, sem Cne (Vint x) (Vint y) = Val.of_bool (negb (Int.eq x y)).
+Hypothesis sem_default: forall c v n, sem c v (Vint n) = Val.of_optbool (eval_condition (default c n) (v :: nil) m).
+
+Lemma eval_compimm:
+ forall le c a n2 x,
+ eval_expr ge sp e m le a x ->
+ exists v, eval_expr ge sp e m le (compimm default intsem c a n2) v
+ /\ Val.lessdef (sem c x (Vint n2)) v.
+Proof.
+ intros until x.
+ unfold compimm; case (compimm_match c a); intros.
+(* constant *)
+ - InvEval. rewrite sem_int. TrivialExists. simpl. destruct (intsem c0 n1 n2); auto.
+(* eq cmp *)
+ - InvEval. inv H. simpl in H5. inv H5.
+ destruct (Int.eq_dec n2 Int.zero).
+ + subst n2. TrivialExists.
+ simpl. rewrite eval_negate_condition.
+ destruct (eval_condition c0 vl m); simpl.
+ unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_eq; auto.
+ rewrite sem_undef; auto.
+ + destruct (Int.eq_dec n2 Int.one). subst n2. TrivialExists.
+ simpl. destruct (eval_condition c0 vl m); simpl.
+ unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_eq; auto.
+ rewrite sem_undef; auto.
+ exists (Vint Int.zero); split. EvalOp.
+ destruct (eval_condition c0 vl m); simpl.
+ unfold Vtrue, Vfalse. destruct b; rewrite sem_eq; rewrite Int.eq_false; auto.
+ rewrite sem_undef; auto.
+(* ne cmp *)
+ - InvEval. inv H. simpl in H5. inv H5.
+ destruct (Int.eq_dec n2 Int.zero).
+ + subst n2. TrivialExists.
+ simpl. destruct (eval_condition c0 vl m); simpl.
+ unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_ne; auto.
+ rewrite sem_undef; auto.
+ + destruct (Int.eq_dec n2 Int.one). subst n2. TrivialExists.
+ simpl. rewrite eval_negate_condition. destruct (eval_condition c0 vl m); simpl.
+ unfold Vtrue, Vfalse. destruct b; simpl; rewrite sem_ne; auto.
+ rewrite sem_undef; auto.
+ exists (Vint Int.one); split. EvalOp.
+ destruct (eval_condition c0 vl m); simpl.
+ unfold Vtrue, Vfalse. destruct b; rewrite sem_ne; rewrite Int.eq_false; auto.
+ rewrite sem_undef; auto.
+(* default *)
+ - TrivialExists. simpl. rewrite sem_default. auto.
+Qed.
+
+Hypothesis sem_swap:
+ forall c x y, sem (swap_comparison c) x y = sem c y x.
+
+Lemma eval_compimm_swap:
+ forall le c a n2 x,
+ eval_expr ge sp e m le a x ->
+ exists v, eval_expr ge sp e m le (compimm default intsem (swap_comparison c) a n2) v
+ /\ Val.lessdef (sem c (Vint n2) x) v.
+Proof.
+ intros. rewrite <- sem_swap. eapply eval_compimm; eauto.
+Qed.
+
+End COMP_IMM.
+
+Theorem eval_comp:
+ forall c, binary_constructor_sound (comp c) (Val.cmp c).
+Proof.
+ intros; red; intros until y. unfold comp; case (comp_match a b); intros; InvEval.
+ eapply eval_compimm_swap; eauto.
+ intros. unfold Val.cmp. rewrite Val.swap_cmp_bool; auto.
+ eapply eval_compimm; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_compu:
+ forall c, binary_constructor_sound (compu c) (Val.cmpu (Mem.valid_pointer m) c).
+Proof.
+ intros; red; intros until y. unfold compu; case (compu_match a b); intros; InvEval.
+ eapply eval_compimm_swap; eauto.
+ intros. unfold Val.cmpu. rewrite Val.swap_cmpu_bool; auto.
+ eapply eval_compimm; eauto.
+ TrivialExists.
+Qed.
+
+Theorem eval_compf:
+ forall c, binary_constructor_sound (compf c) (Val.cmpf c).
+Proof.
+ intros; red; intros. unfold compf. TrivialExists.
+Qed.
+
+Theorem eval_compfs:
+ forall c, binary_constructor_sound (compfs c) (Val.cmpfs c).
+Proof.
+ intros; red; intros. unfold compfs. TrivialExists.
+Qed.
+
+Theorem eval_cast8signed: unary_constructor_sound cast8signed (Val.sign_ext 8).
+Proof.
+ red; intros until x. unfold cast8signed. case (cast8signed_match a); intros; InvEval.
+ TrivialExists.
+ TrivialExists.
+Qed.
+
+Theorem eval_cast8unsigned: unary_constructor_sound cast8unsigned (Val.zero_ext 8).
+Proof.
+ red; intros until x. unfold cast8unsigned.
+ rewrite Val.zero_ext_and. apply eval_andimm. compute; auto.
+Qed.
+
+Theorem eval_cast16signed: unary_constructor_sound cast16signed (Val.sign_ext 16).
+Proof.
+ red; intros until x. unfold cast16signed. case (cast16signed_match a); intros; InvEval.
+ TrivialExists.
+ TrivialExists.
+Qed.
+
+Theorem eval_cast16unsigned: unary_constructor_sound cast16unsigned (Val.zero_ext 16).
+Proof.
+ red; intros until x. unfold cast8unsigned.
+ rewrite Val.zero_ext_and. apply eval_andimm. compute; auto.
+Qed.
+
+Theorem eval_intoffloat:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.intoffloat x = Some y ->
+ exists v, eval_expr ge sp e m le (intoffloat a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold intoffloat. TrivialExists.
+Qed.
+
+Theorem eval_intuoffloat:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.intuoffloat x = Some y ->
+ exists v, eval_expr ge sp e m le (intuoffloat a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold intuoffloat. TrivialExists.
+Qed.
+
+Theorem eval_floatofintu:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.floatofintu x = Some y ->
+ exists v, eval_expr ge sp e m le (floatofintu a) v /\ Val.lessdef y v.
+Proof.
+ intros until y; unfold floatofintu. case (floatofintu_match a); intros.
+ InvEval. simpl in H0. TrivialExists.
+ TrivialExists.
+Qed.
+
+Theorem eval_floatofint:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.floatofint x = Some y ->
+ exists v, eval_expr ge sp e m le (floatofint a) v /\ Val.lessdef y v.
+Proof.
+ intros until y; unfold floatofint. case (floatofint_match a); intros.
+ InvEval. simpl in H0. TrivialExists.
+ TrivialExists.
+Qed.
+
+Theorem eval_intofsingle:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.intofsingle x = Some y ->
+ exists v, eval_expr ge sp e m le (intofsingle a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold intofsingle. TrivialExists.
+Qed.
+
+Theorem eval_singleofint:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.singleofint x = Some y ->
+ exists v, eval_expr ge sp e m le (singleofint a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold singleofint; TrivialExists.
+Qed.
+
+Theorem eval_intuofsingle:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.intuofsingle x = Some y ->
+ exists v, eval_expr ge sp e m le (intuofsingle a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold intuofsingle. TrivialExists.
+Qed.
+
+Theorem eval_singleofintu:
+ forall le a x y,
+ eval_expr ge sp e m le a x ->
+ Val.singleofintu x = Some y ->
+ exists v, eval_expr ge sp e m le (singleofintu a) v /\ Val.lessdef y v.
+Proof.
+ intros; unfold intuofsingle. TrivialExists.
+Qed.
+
+Theorem eval_singleoffloat: unary_constructor_sound singleoffloat Val.singleoffloat.
+Proof.
+ red; intros. unfold singleoffloat. TrivialExists.
+Qed.
+
+Theorem eval_floatofsingle: unary_constructor_sound floatofsingle Val.floatofsingle.
+Proof.
+ red; intros. unfold floatofsingle. TrivialExists.
+Qed.
+
+Theorem eval_addressing:
+ forall le chunk a v b ofs,
+ eval_expr ge sp e m le a v ->
+ v = Vptr b ofs ->
+ match addressing chunk a with (mode, args) =>
+ exists vl,
+ eval_exprlist ge sp e m le args vl /\
+ eval_addressing ge sp mode vl = Some v
+ end.
+Proof.
+ intros until v. unfold addressing; case (addressing_match a); intros; InvEval.
+ - exists (@nil val); split. eauto with evalexpr. simpl. auto.
+ - destruct (Archi.pic_code tt).
+ + exists (Vptr b ofs0 :: nil); split.
+ constructor. EvalOp. simpl. congruence. constructor. simpl. rewrite Ptrofs.add_zero. congruence.
+ + exists (@nil val); split. constructor. simpl; auto.
+ - exists (v1 :: nil); split. eauto with evalexpr. simpl.
+ destruct v1; simpl in H; try discriminate. destruct Archi.ptr64 eqn:SF; inv H.
+ simpl. auto.
+ - exists (v1 :: nil); split. eauto with evalexpr. simpl.
+ destruct v1; simpl in H; try discriminate. destruct Archi.ptr64 eqn:SF; inv H.
+ simpl. auto.
+ - exists (v :: nil); split. eauto with evalexpr. subst. simpl. rewrite Ptrofs.add_zero; auto.
+Qed.
+
+Theorem eval_builtin_arg:
+ forall a v,
+ eval_expr ge sp e m nil a v ->
+ CminorSel.eval_builtin_arg ge sp e m (builtin_arg a) v.
+Proof.
+ intros until v. unfold builtin_arg; case (builtin_arg_match a); intros; InvEval.
+- constructor.
+- constructor.
+- constructor.
+- simpl in H5. inv H5. constructor.
+- subst v. constructor; auto.
+- inv H. InvEval. simpl in H6; inv H6. constructor; auto.
+- constructor; auto.
+Qed.
+
+End CMCONSTR.
diff --git a/riscV/Stacklayout.v b/riscV/Stacklayout.v
new file mode 100644
index 00000000..ba116380
--- /dev/null
+++ b/riscV/Stacklayout.v
@@ -0,0 +1,147 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris *)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Machine- and ABI-dependent layout information for activation records. *)
+
+Require Import Coqlib.
+Require Import AST Memory Separation.
+Require Import Bounds.
+
+Local Open Scope sep_scope.
+
+(** The general shape of activation records is as follows,
+ from bottom (lowest offsets) to top:
+- Space for outgoing arguments to function calls.
+- Back link to parent frame
+- Return address
+- Saved values of callee-save registers used by the function.
+- Local stack slots.
+- Space for the stack-allocated data declared in Cminor.
+
+The stack pointer is kept 16-aligned.
+*)
+
+Definition fe_ofs_arg := 0.
+
+Definition make_env (b: bounds) : frame_env :=
+ let w := if Archi.ptr64 then 8 else 4 in
+ let olink := align (4 * b.(bound_outgoing)) w in (* back link *)
+ let oretaddr := olink + w in (* return address *)
+ let ocs := oretaddr + w in (* callee-saves *)
+ let ol := align (size_callee_save_area b ocs) 8 in (* locals *)
+ let ostkdata := align (ol + 4 * b.(bound_local)) 8 in (* stack data *)
+ let sz := align (ostkdata + b.(bound_stack_data)) 16 in
+ {| fe_size := sz;
+ fe_ofs_link := olink;
+ fe_ofs_retaddr := oretaddr;
+ fe_ofs_local := ol;
+ fe_ofs_callee_save := ocs;
+ fe_stack_data := ostkdata;
+ fe_used_callee_save := b.(used_callee_save) |}.
+
+Lemma frame_env_separated:
+ forall b sp m P,
+ let fe := make_env b in
+ m |= range sp 0 (fe_stack_data fe) ** range sp (fe_stack_data fe + bound_stack_data b) (fe_size fe) ** P ->
+ m |= range sp (fe_ofs_local fe) (fe_ofs_local fe + 4 * bound_local b)
+ ** range sp fe_ofs_arg (fe_ofs_arg + 4 * bound_outgoing b)
+ ** range sp (fe_ofs_link fe) (fe_ofs_link fe + size_chunk Mptr)
+ ** range sp (fe_ofs_retaddr fe) (fe_ofs_retaddr fe + size_chunk Mptr)
+ ** range sp (fe_ofs_callee_save fe) (size_callee_save_area b (fe_ofs_callee_save fe))
+ ** P.
+Proof.
+Local Opaque Z.add Z.mul sepconj range.
+ intros; simpl.
+ set (w := if Archi.ptr64 then 8 else 4).
+ set (olink := align (4 * b.(bound_outgoing)) w).
+ set (oretaddr := olink + w).
+ set (ocs := oretaddr + w).
+ set (ol := align (size_callee_save_area b ocs) 8).
+ set (ostkdata := align (ol + 4 * b.(bound_local)) 8).
+ replace (size_chunk Mptr) with w by (rewrite size_chunk_Mptr; auto).
+ assert (0 < w) by (unfold w; destruct Archi.ptr64; omega).
+ generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros.
+ assert (0 <= 4 * b.(bound_outgoing)) by omega.
+ assert (4 * b.(bound_outgoing) <= olink) by (apply align_le; omega).
+ assert (olink + w <= oretaddr) by (unfold oretaddr; omega).
+ assert (oretaddr + w <= ocs) by (unfold ocs; omega).
+ assert (ocs <= size_callee_save_area b ocs) by (apply size_callee_save_area_incr).
+ assert (size_callee_save_area b ocs <= ol) by (apply align_le; omega).
+ assert (ol + 4 * b.(bound_local) <= ostkdata) by (apply align_le; omega).
+(* Reorder as:
+ outgoing
+ back link
+ retaddr
+ callee-save
+ local *)
+ rewrite sep_swap12.
+ rewrite sep_swap23.
+ rewrite sep_swap34.
+ rewrite sep_swap45.
+(* Apply range_split and range_split2 repeatedly *)
+ unfold fe_ofs_arg.
+ apply range_split_2. fold olink; omega. omega.
+ apply range_split. omega.
+ apply range_split. omega.
+ apply range_split_2. fold ol. omega. omega.
+ apply range_drop_right with ostkdata. omega.
+ eapply sep_drop2. eexact H.
+Qed.
+
+Lemma frame_env_range:
+ forall b,
+ let fe := make_env b in
+ 0 <= fe_stack_data fe /\ fe_stack_data fe + bound_stack_data b <= fe_size fe.
+Proof.
+ intros; simpl.
+ set (w := if Archi.ptr64 then 8 else 4).
+ set (olink := align (4 * b.(bound_outgoing)) w).
+ set (oretaddr := olink + w).
+ set (ocs := oretaddr + w).
+ set (ol := align (size_callee_save_area b ocs) 8).
+ set (ostkdata := align (ol + 4 * b.(bound_local)) 8).
+ assert (0 < w) by (unfold w; destruct Archi.ptr64; omega).
+ generalize b.(bound_local_pos) b.(bound_outgoing_pos) b.(bound_stack_data_pos); intros.
+ assert (0 <= 4 * b.(bound_outgoing)) by omega.
+ assert (4 * b.(bound_outgoing) <= olink) by (apply align_le; omega).
+ assert (olink + w <= oretaddr) by (unfold oretaddr; omega).
+ assert (oretaddr + w <= ocs) by (unfold ocs; omega).
+ assert (ocs <= size_callee_save_area b ocs) by (apply size_callee_save_area_incr).
+ assert (size_callee_save_area b ocs <= ol) by (apply align_le; omega).
+ assert (ol + 4 * b.(bound_local) <= ostkdata) by (apply align_le; omega).
+ split. omega. apply align_le. omega.
+Qed.
+
+Lemma frame_env_aligned:
+ forall b,
+ let fe := make_env b in
+ (8 | fe_ofs_arg)
+ /\ (8 | fe_ofs_local fe)
+ /\ (8 | fe_stack_data fe)
+ /\ (align_chunk Mptr | fe_ofs_link fe)
+ /\ (align_chunk Mptr | fe_ofs_retaddr fe).
+Proof.
+ intros; simpl.
+ set (w := if Archi.ptr64 then 8 else 4).
+ set (olink := align (4 * b.(bound_outgoing)) w).
+ set (oretaddr := olink + w).
+ set (ocs := oretaddr + w).
+ set (ol := align (size_callee_save_area b ocs) 8).
+ set (ostkdata := align (ol + 4 * b.(bound_local)) 8).
+ assert (0 < w) by (unfold w; destruct Archi.ptr64; omega).
+ replace (align_chunk Mptr) with w by (rewrite align_chunk_Mptr; auto).
+ split. apply Zdivide_0.
+ split. apply align_divides; omega.
+ split. apply align_divides; omega.
+ split. apply align_divides; omega.
+ apply Z.divide_add_r. apply align_divides; omega. apply Z.divide_refl.
+Qed.
diff --git a/riscV/TargetPrinter.ml b/riscV/TargetPrinter.ml
new file mode 100644
index 00000000..7a369832
--- /dev/null
+++ b/riscV/TargetPrinter.ml
@@ -0,0 +1,695 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(* Printing RISC-V assembly code in asm syntax *)
+
+open Printf
+open Camlcoq
+open Sections
+open AST
+open Asm
+open PrintAsmaux
+open Fileinfo
+
+(* Module containing the printing functions *)
+
+module Target : TARGET =
+ struct
+
+(* Basic printing functions *)
+
+ let comment = "#"
+
+ let symbol = elf_symbol
+ let symbol_offset = elf_symbol_offset
+ let label = elf_label
+
+ let print_label oc lbl = label oc (transl_label lbl)
+
+ let use_abi_name = false
+
+ let int_reg_num_name = function
+ | X1 -> "x1" | X2 -> "x2" | X3 -> "x3"
+ | X4 -> "x4" | X5 -> "x5" | X6 -> "x6" | X7 -> "x7"
+ | X8 -> "x8" | X9 -> "x9" | X10 -> "x10" | X11 -> "x11"
+ | X12 -> "x12" | X13 -> "x13" | X14 -> "x14" | X15 -> "x15"
+ | X16 -> "x16" | X17 -> "x17" | X18 -> "x18" | X19 -> "x19"
+ | X20 -> "x20" | X21 -> "x21" | X22 -> "x22" | X23 -> "x23"
+ | X24 -> "x24" | X25 -> "x25" | X26 -> "x26" | X27 -> "x27"
+ | X28 -> "x28" | X29 -> "x29" | X30 -> "x30" | X31 -> "x31"
+
+ let int_reg_abi_name = function
+ | X1 -> "ra" | X2 -> "sp" | X3 -> "gp"
+ | X4 -> "tp" | X5 -> "t0" | X6 -> "t1" | X7 -> "t2"
+ | X8 -> "s0" | X9 -> "s1" | X10 -> "a0" | X11 -> "a1"
+ | X12 -> "a2" | X13 -> "a3" | X14 -> "a4" | X15 -> "a5"
+ | X16 -> "a6" | X17 -> "a7" | X18 -> "s2" | X19 -> "s3"
+ | X20 -> "s4" | X21 -> "s5" | X22 -> "s6" | X23 -> "s7"
+ | X24 -> "s8" | X25 -> "s9" | X26 -> "s10" | X27 -> "s11"
+ | X28 -> "t3" | X29 -> "t4" | X30 -> "t5" | X31 -> "t6"
+
+ let float_reg_num_name = function
+ | F0 -> "f0" | F1 -> "f1" | F2 -> "f2" | F3 -> "f3"
+ | F4 -> "f4" | F5 -> "f5" | F6 -> "f6" | F7 -> "f7"
+ | F8 -> "f8" | F9 -> "f9" | F10 -> "f10" | F11 -> "f11"
+ | F12 -> "f12" | F13 -> "f13" | F14 -> "f14" | F15 -> "f15"
+ | F16 -> "f16" | F17 -> "f17" | F18 -> "f18" | F19 -> "f19"
+ | F20 -> "f20" | F21 -> "f21" | F22 -> "f22" | F23 -> "f23"
+ | F24 -> "f24" | F25 -> "f25" | F26 -> "f26" | F27 -> "f27"
+ | F28 -> "f28" | F29 -> "f29" | F30 -> "f30" | F31 -> "f31"
+
+ let float_reg_abi_name = function
+ | F0 -> "ft0" | F1 -> "ft1" | F2 -> "ft2" | F3 -> "ft3"
+ | F4 -> "ft4" | F5 -> "ft5" | F6 -> "ft6" | F7 -> "ft7"
+ | F8 -> "fs0" | F9 -> "fs1" | F10 -> "fa0" | F11 -> "fa1"
+ | F12 -> "fa2" | F13 -> "fa3" | F14 -> "fa4" | F15 -> "fa5"
+ | F16 -> "fa6" | F17 -> "fa7" | F18 -> "fs2" | F19 -> "fs3"
+ | F20 -> "fs4" | F21 -> "fs5" | F22 -> "fs6" | F23 -> "fs7"
+ | F24 -> "fs8" | F25 -> "fs9" | F26 ->"fs10" | F27 -> "fs11"
+ | F28 -> "ft3" | F29 -> "ft4" | F30 -> "ft5" | F31 -> "ft6"
+
+ let int_reg_name = if use_abi_name then int_reg_abi_name else int_reg_num_name
+ let float_reg_name = if use_abi_name then float_reg_abi_name else float_reg_num_name
+
+ let ireg oc r = output_string oc (int_reg_name r)
+ let freg oc r = output_string oc (float_reg_name r)
+
+ let ireg0 oc = function
+ | X0 -> output_string oc "x0"
+ | X r -> ireg oc r
+
+ let preg oc = function
+ | IR r -> ireg oc r
+ | FR r -> freg oc r
+ | _ -> assert false
+
+(* Names of sections *)
+
+ let name_of_section = function
+ | Section_text -> ".text"
+ | Section_data i | Section_small_data i ->
+ if i then ".data" else "COMM"
+ | Section_const i | Section_small_const i ->
+ if i then ".section .rodata" else "COMM"
+ | Section_string -> ".section .rodata"
+ | Section_literal -> ".section .rodata"
+ | Section_jumptable -> ".section .rodata"
+ | Section_debug_info _ -> ".section .debug_info,\"\",%progbits"
+ | Section_debug_loc -> ".section .debug_loc,\"\",%progbits"
+ | Section_debug_abbrev -> ".section .debug_abbrev,\"\",%progbits"
+ | Section_debug_line _ -> ".section .debug_line,\"\",%progbits"
+ | Section_debug_ranges -> ".section .debug_ranges,\"\",%progbits"
+ | Section_debug_str -> ".section .debug_str,\"MS\",%progbits,1"
+ | Section_user(s, wr, ex) ->
+ sprintf ".section \"%s\",\"a%s%s\",%%progbits"
+ s (if wr then "w" else "") (if ex then "x" else "")
+
+ let section oc sec =
+ fprintf oc " %s\n" (name_of_section sec)
+
+(* Associate labels to floating-point constants and to symbols. *)
+
+ let label_constant (h: ('a, int) Hashtbl.t) (cst: 'a) =
+ try
+ Hashtbl.find h cst
+ with Not_found ->
+ let lbl = new_label() in
+ Hashtbl.add h cst lbl;
+ lbl
+
+ let float_labels = (Hashtbl.create 39 : (int64, int) Hashtbl.t)
+ let float32_labels = (Hashtbl.create 39 : (int32, int) Hashtbl.t)
+ let int64_labels = (Hashtbl.create 39 : (int64, int) Hashtbl.t)
+
+ let label_float bf = label_constant float_labels bf
+ let label_float32 bf = label_constant float32_labels bf
+ let label_int64 n = label_constant int64_labels n
+
+ let reset_constants () =
+ Hashtbl.clear float_labels;
+ Hashtbl.clear float32_labels;
+ Hashtbl.clear int64_labels
+
+ let emit_constants oc =
+ if Hashtbl.length int64_labels > 0 then
+ begin
+ fprintf oc " .align 3\n";
+ Hashtbl.iter
+ (fun bf lbl -> fprintf oc "%a: .quad 0x%Lx\n" label lbl bf)
+ int64_labels
+ end;
+ if Hashtbl.length float_labels > 0 then
+ begin
+ fprintf oc " .align 3\n";
+ Hashtbl.iter
+ (fun bf lbl -> fprintf oc "%a: .quad 0x%Lx\n" label lbl bf)
+ float_labels
+ end;
+ if Hashtbl.length float32_labels > 0 then
+ begin
+ fprintf oc " .align 2\n";
+ Hashtbl.iter
+ (fun bf lbl ->
+ fprintf oc "%a: .long 0x%lx\n" label lbl bf)
+ float32_labels
+ end;
+ reset_constants ()
+
+(* Generate code to load the address of id + ofs in register r *)
+
+ let loadsymbol oc r id ofs =
+ if Archi.pic_code () then begin
+ assert (ofs = Integers.Ptrofs.zero);
+ fprintf oc " la %a, %s\n" ireg r (extern_atom id)
+ end else begin
+ fprintf oc " lui %a, %%hi(%a)\n"
+ ireg r symbol_offset (id, ofs);
+ fprintf oc " addi %a, %a, %%lo(%a)\n"
+ ireg r ireg r symbol_offset (id, ofs)
+ end
+
+(* Emit .file / .loc debugging directives *)
+
+ let print_file_line oc file line =
+ print_file_line oc comment file line
+
+(*
+ let print_location oc loc =
+ if loc <> Cutil.no_loc then print_file_line oc (fst loc) (snd loc)
+*)
+
+(* Add "w" suffix to 32-bit instructions if we are in 64-bit mode *)
+
+ let w oc =
+ if Archi.ptr64 then output_string oc "w"
+
+(* Offset part of a load or store *)
+
+ let offset oc = function
+ | Ofsimm n -> ptrofs oc n
+ | Ofslow(id, ofs) -> fprintf oc "%%lo(%a)" symbol_offset (id, ofs)
+
+(* Printing of instructions *)
+ let print_instruction oc = function
+ | Pmv(rd, rs) ->
+ fprintf oc " mv %a, %a\n" ireg rd ireg rs
+
+ (* 32-bit integer register-immediate instructions *)
+ | Paddiw (rd, rs, imm) ->
+ fprintf oc " addi%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm
+ | Psltiw (rd, rs, imm) ->
+ fprintf oc " slti %a, %a, %a\n" ireg rd ireg0 rs coqint imm
+ | Psltiuw (rd, rs, imm) ->
+ fprintf oc " sltiu %a, %a, %a\n" ireg rd ireg0 rs coqint imm
+ | Pandiw (rd, rs, imm) ->
+ fprintf oc " andi %a, %a, %a\n" ireg rd ireg0 rs coqint imm
+ | Poriw (rd, rs, imm) ->
+ fprintf oc " ori %a, %a, %a\n" ireg rd ireg0 rs coqint imm
+ | Pxoriw (rd, rs, imm) ->
+ fprintf oc " xori %a, %a, %a\n" ireg rd ireg0 rs coqint imm
+ | Pslliw (rd, rs, imm) ->
+ fprintf oc " slli%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm
+ | Psrliw (rd, rs, imm) ->
+ fprintf oc " srli%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm
+ | Psraiw (rd, rs, imm) ->
+ fprintf oc " srai%t %a, %a, %a\n" w ireg rd ireg0 rs coqint imm
+ | Pluiw (rd, imm) ->
+ fprintf oc " lui %a, %a\n" ireg rd coqint imm
+
+ (* 32-bit integer register-register instructions *)
+ | Paddw(rd, rs1, rs2) ->
+ fprintf oc " add%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+ | Psubw(rd, rs1, rs2) ->
+ fprintf oc " sub%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+
+ | Pmulw(rd, rs1, rs2) ->
+ fprintf oc " mul%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+ | Pmulhw(rd, rs1, rs2) -> assert (not Archi.ptr64);
+ fprintf oc " mulh %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Pmulhuw(rd, rs1, rs2) -> assert (not Archi.ptr64);
+ fprintf oc " mulhu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+
+ | Pdivw(rd, rs1, rs2) ->
+ fprintf oc " div%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+ | Pdivuw(rd, rs1, rs2) ->
+ fprintf oc " divu%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+ | Premw(rd, rs1, rs2) ->
+ fprintf oc " rem%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+ | Premuw(rd, rs1, rs2) ->
+ fprintf oc " remu%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+
+ | Psltw(rd, rs1, rs2) ->
+ fprintf oc " slt %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Psltuw(rd, rs1, rs2) ->
+ fprintf oc " sltu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+
+ | Pandw(rd, rs1, rs2) ->
+ fprintf oc " and %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Porw(rd, rs1, rs2) ->
+ fprintf oc " or %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Pxorw(rd, rs1, rs2) ->
+ fprintf oc " xor %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Psllw(rd, rs1, rs2) ->
+ fprintf oc " sll%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+ | Psrlw(rd, rs1, rs2) ->
+ fprintf oc " srl%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+ | Psraw(rd, rs1, rs2) ->
+ fprintf oc " sra%t %a, %a, %a\n" w ireg rd ireg0 rs1 ireg0 rs2
+
+ (* 64-bit integer register-immediate instructions *)
+ | Paddil (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " addi %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Psltil (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " slti %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Psltiul (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " sltiu %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Pandil (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " andi %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Poril (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " ori %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Pxoril (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " xori %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Psllil (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " slli %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Psrlil (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " srli %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Psrail (rd, rs, imm) -> assert Archi.ptr64;
+ fprintf oc " srai %a, %a, %a\n" ireg rd ireg0 rs coqint64 imm
+ | Pluil (rd, imm) -> assert Archi.ptr64;
+ fprintf oc " lui %a, %a\n" ireg rd coqint64 imm
+
+ (* 64-bit integer register-register instructions *)
+ | Paddl(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " add %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Psubl(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " sub %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+
+ | Pmull(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " mul %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Pmulhl(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " mulh %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Pmulhul(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " mulhu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+
+ | Pdivl(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " div %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Pdivul(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " divu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Preml(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " rem %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Premul(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " remu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+
+ | Psltl(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " slt %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Psltul(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " sltu %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+
+ | Pandl(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " and %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Porl(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " or %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Pxorl(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " xor %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Pslll(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " sll %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Psrll(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " srl %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+ | Psral(rd, rs1, rs2) -> assert Archi.ptr64;
+ fprintf oc " sra %a, %a, %a\n" ireg rd ireg0 rs1 ireg0 rs2
+
+ (* Unconditional jumps. Links are always to X1/RA. *)
+ (* TODO: fix up arguments for calls to variadics, to move *)
+ (* floating point arguments to integer registers. How? *)
+ | Pj_l(l) ->
+ fprintf oc " j %a\n" print_label l
+ | Pj_s(s, sg) ->
+ fprintf oc " j %a\n" symbol s
+ | Pj_r(r, sg) ->
+ fprintf oc " jr %a\n" ireg r
+ | Pjal_s(s, sg) ->
+ fprintf oc " call %a\n" symbol s
+ | Pjal_r(r, sg) ->
+ fprintf oc " jalr %a\n" ireg r
+
+ (* Conditional branches, 32-bit comparisons *)
+ | Pbeqw(rs1, rs2, l) ->
+ fprintf oc " beq %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbnew(rs1, rs2, l) ->
+ fprintf oc " bne %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbltw(rs1, rs2, l) ->
+ fprintf oc " blt %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbltuw(rs1, rs2, l) ->
+ fprintf oc " bltu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbgew(rs1, rs2, l) ->
+ fprintf oc " bge %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbgeuw(rs1, rs2, l) ->
+ fprintf oc " bgeu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+
+ (* Conditional branches, 64-bit comparisons *)
+ | Pbeql(rs1, rs2, l) -> assert Archi.ptr64;
+ fprintf oc " beq %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbnel(rs1, rs2, l) -> assert Archi.ptr64;
+ fprintf oc " bne %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbltl(rs1, rs2, l) -> assert Archi.ptr64;
+ fprintf oc " blt %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbltul(rs1, rs2, l) -> assert Archi.ptr64;
+ fprintf oc " bltu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbgel(rs1, rs2, l) -> assert Archi.ptr64;
+ fprintf oc " bge %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+ | Pbgeul(rs1, rs2, l) -> assert Archi.ptr64;
+ fprintf oc " bgeu %a, %a, %a\n" ireg0 rs1 ireg0 rs2 print_label l
+
+ (* Loads and stores *)
+ | Plb(rd, ra, ofs) ->
+ fprintf oc " lb %a, %a(%a)\n" ireg rd offset ofs ireg ra
+ | Plbu(rd, ra, ofs) ->
+ fprintf oc " lbu %a, %a(%a)\n" ireg rd offset ofs ireg ra
+ | Plh(rd, ra, ofs) ->
+ fprintf oc " lh %a, %a(%a)\n" ireg rd offset ofs ireg ra
+ | Plhu(rd, ra, ofs) ->
+ fprintf oc " lhu %a, %a(%a)\n" ireg rd offset ofs ireg ra
+ | Plw(rd, ra, ofs) | Plw_a(rd, ra, ofs) ->
+ fprintf oc " lw %a, %a(%a)\n" ireg rd offset ofs ireg ra
+ | Pld(rd, ra, ofs) | Pld_a(rd, ra, ofs) -> assert Archi.ptr64;
+ fprintf oc " ld %a, %a(%a)\n" ireg rd offset ofs ireg ra
+
+ | Psb(rd, ra, ofs) ->
+ fprintf oc " sb %a, %a(%a)\n" ireg rd offset ofs ireg ra
+ | Psh(rd, ra, ofs) ->
+ fprintf oc " sh %a, %a(%a)\n" ireg rd offset ofs ireg ra
+ | Psw(rd, ra, ofs) | Psw_a(rd, ra, ofs) ->
+ fprintf oc " sw %a, %a(%a)\n" ireg rd offset ofs ireg ra
+ | Psd(rd, ra, ofs) | Psd_a(rd, ra, ofs) -> assert Archi.ptr64;
+ fprintf oc " sd %a, %a(%a)\n" ireg rd offset ofs ireg ra
+
+
+ (* Synchronization *)
+ | Pfence ->
+ fprintf oc " fence\n"
+
+ (* floating point register move.
+ fmv.d preserves single-precision register contents, and hence
+ is applicable to both single- and double-precision moves.
+ *)
+ | Pfmv (fd,fs) ->
+ fprintf oc " fmv.d %a, %a\n" freg fd freg fs
+ | Pfmvxs (rd,fs) ->
+ fprintf oc " fmv.x.s %a, %a\n" ireg rd freg fs
+ | Pfmvxd (rd,fs) ->
+ fprintf oc " fmv.x.d %a, %a\n" ireg rd freg fs
+
+ (* 32-bit (single-precision) floating point *)
+ | Pfls (fd, ra, ofs) ->
+ fprintf oc " flw %a, %a(%a)\n" freg fd offset ofs ireg ra
+ | Pfss (fs, ra, ofs) ->
+ fprintf oc " fsw %a, %a(%a)\n" freg fs offset ofs ireg ra
+
+ | Pfnegs (fd, fs) ->
+ fprintf oc " fneg.s %a, %a\n" freg fd freg fs
+ | Pfabss (fd, fs) ->
+ fprintf oc " fabs.s %a, %a\n" freg fd freg fs
+
+ | Pfadds (fd, fs1, fs2) ->
+ fprintf oc " fadd.s %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfsubs (fd, fs1, fs2) ->
+ fprintf oc " fsub.s %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfmuls (fd, fs1, fs2) ->
+ fprintf oc " fmul.s %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfdivs (fd, fs1, fs2) ->
+ fprintf oc " fdiv.s %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfmins (fd, fs1, fs2) ->
+ fprintf oc " fmin.s %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfmaxs (fd, fs1, fs2) ->
+ fprintf oc " fmax.s %a, %a, %a\n" freg fd freg fs1 freg fs2
+
+ | Pfeqs (rd, fs1, fs2) ->
+ fprintf oc " feq.s %a, %a, %a\n" ireg rd freg fs1 freg fs2
+ | Pflts (rd, fs1, fs2) ->
+ fprintf oc " flt.s %a, %a, %a\n" ireg rd freg fs1 freg fs2
+ | Pfles (rd, fs1, fs2) ->
+ fprintf oc " fle.s %a, %a, %a\n" ireg rd freg fs1 freg fs2
+
+ | Pfsqrts (fd, fs) ->
+ fprintf oc " fsqrt.s %a, %a\n" freg fd freg fs
+
+ | Pfmadds (fd, fs1, fs2, fs3) ->
+ fprintf oc " fmadd.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3
+ | Pfmsubs (fd, fs1, fs2, fs3) ->
+ fprintf oc " fmsub.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3
+ | Pfnmadds (fd, fs1, fs2, fs3) ->
+ fprintf oc " fnmadd.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3
+ | Pfnmsubs (fd, fs1, fs2, fs3) ->
+ fprintf oc " fnmsub.s %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3
+
+ | Pfcvtws (rd, fs) ->
+ fprintf oc " fcvt.w.s %a, %a, rtz\n" ireg rd freg fs
+ | Pfcvtwus (rd, fs) ->
+ fprintf oc " fcvt.wu.s %a, %a, rtz\n" ireg rd freg fs
+ | Pfcvtsw (fd, rs) ->
+ fprintf oc " fcvt.s.w %a, %a\n" freg fd ireg0 rs
+ | Pfcvtswu (fd, rs) ->
+ fprintf oc " fcvt.s.wu %a, %a\n" freg fd ireg0 rs
+
+ | Pfcvtls (rd, fs) -> assert Archi.ptr64;
+ fprintf oc " fcvt.l.s %a, %a, rtz\n" ireg rd freg fs
+ | Pfcvtlus (rd, fs) -> assert Archi.ptr64;
+ fprintf oc " fcvt.lu.s %a, %a, rtz\n" ireg rd freg fs
+ | Pfcvtsl (fd, rs) -> assert Archi.ptr64;
+ fprintf oc " fcvt.s.l %a, %a\n" freg fd ireg0 rs
+ | Pfcvtslu (fd, rs) -> assert Archi.ptr64;
+ fprintf oc " fcvt.s.lu %a, %a\n" freg fd ireg0 rs
+
+ (* 64-bit (double-precision) floating point *)
+ | Pfld (fd, ra, ofs) | Pfld_a (fd, ra, ofs) ->
+ fprintf oc " fld %a, %a(%a)\n" freg fd offset ofs ireg ra
+ | Pfsd (fs, ra, ofs) | Pfsd_a (fs, ra, ofs) ->
+ fprintf oc " fsd %a, %a(%a)\n" freg fs offset ofs ireg ra
+
+ | Pfnegd (fd, fs) ->
+ fprintf oc " fneg.d %a, %a\n" freg fd freg fs
+ | Pfabsd (fd, fs) ->
+ fprintf oc " fabs.d %a, %a\n" freg fd freg fs
+
+ | Pfaddd (fd, fs1, fs2) ->
+ fprintf oc " fadd.d %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfsubd (fd, fs1, fs2) ->
+ fprintf oc " fsub.d %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfmuld (fd, fs1, fs2) ->
+ fprintf oc " fmul.d %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfdivd (fd, fs1, fs2) ->
+ fprintf oc " fdiv.d %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfmind (fd, fs1, fs2) ->
+ fprintf oc " fmin.d %a, %a, %a\n" freg fd freg fs1 freg fs2
+ | Pfmaxd (fd, fs1, fs2) ->
+ fprintf oc " fmax.d %a, %a, %a\n" freg fd freg fs1 freg fs2
+
+ | Pfeqd (rd, fs1, fs2) ->
+ fprintf oc " feq.d %a, %a, %a\n" ireg rd freg fs1 freg fs2
+ | Pfltd (rd, fs1, fs2) ->
+ fprintf oc " flt.d %a, %a, %a\n" ireg rd freg fs1 freg fs2
+ | Pfled (rd, fs1, fs2) ->
+ fprintf oc " fle.d %a, %a, %a\n" ireg rd freg fs1 freg fs2
+
+ | Pfsqrtd (fd, fs) ->
+ fprintf oc " fsqrt.d %a, %a\n" freg fd freg fs
+
+ | Pfmaddd (fd, fs1, fs2, fs3) ->
+ fprintf oc " fmadd.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3
+ | Pfmsubd (fd, fs1, fs2, fs3) ->
+ fprintf oc " fmsub.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3
+ | Pfnmaddd (fd, fs1, fs2, fs3) ->
+ fprintf oc " fnmadd.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3
+ | Pfnmsubd (fd, fs1, fs2, fs3) ->
+ fprintf oc " fnmsub.d %a, %a, %a, %a\n" freg fd freg fs1 freg fs2 freg fs3
+
+ | Pfcvtwd (rd, fs) ->
+ fprintf oc " fcvt.w.d %a, %a, rtz\n" ireg rd freg fs
+ | Pfcvtwud (rd, fs) ->
+ fprintf oc " fcvt.wu.d %a, %a, rtz\n" ireg rd freg fs
+ | Pfcvtdw (fd, rs) ->
+ fprintf oc " fcvt.d.w %a, %a\n" freg fd ireg0 rs
+ | Pfcvtdwu (fd, rs) ->
+ fprintf oc " fcvt.d.wu %a, %a\n" freg fd ireg0 rs
+
+ | Pfcvtld (rd, fs) -> assert Archi.ptr64;
+ fprintf oc " fcvt.l.d %a, %a, rtz\n" ireg rd freg fs
+ | Pfcvtlud (rd, fs) -> assert Archi.ptr64;
+ fprintf oc " fcvt.lu.d %a, %a, rtz\n" ireg rd freg fs
+ | Pfcvtdl (fd, rs) -> assert Archi.ptr64;
+ fprintf oc " fcvt.d.l %a, %a\n" freg fd ireg0 rs
+ | Pfcvtdlu (fd, rs) -> assert Archi.ptr64;
+ fprintf oc " fcvt.d.lu %a, %a\n" freg fd ireg0 rs
+
+ | Pfcvtds (fd, fs) ->
+ fprintf oc " fcvt.d.s %a, %a\n" freg fd freg fs
+ | Pfcvtsd (fd, fs) ->
+ fprintf oc " fcvt.s.d %a, %a\n" freg fd freg fs
+
+ (* Pseudo-instructions expanded in Asmexpand *)
+ | Pallocframe(sz, ofs) ->
+ assert false
+ | Pfreeframe(sz, ofs) ->
+ assert false
+ | Pseqw _ | Psnew _ | Pseql _ | Psnel _ | Pcvtl2w _ | Pcvtw2l _ ->
+ assert false
+
+ (* Pseudo-instructions that remain *)
+ | Plabel lbl ->
+ fprintf oc "%a:\n" print_label lbl
+ | Ploadsymbol(rd, id, ofs) ->
+ loadsymbol oc rd id ofs
+ | Ploadsymbol_high(rd, id, ofs) ->
+ fprintf oc " lui %a, %%hi(%a)\n" ireg rd symbol_offset (id, ofs)
+ | Ploadli(rd, n) ->
+ let d = camlint64_of_coqint n in
+ let lbl = label_int64 d in
+ fprintf oc " ld %a, %a %s %Lx\n" ireg rd label lbl comment d
+ | Ploadfi(rd, f) ->
+ let d = camlint64_of_coqint(Floats.Float.to_bits f) in
+ let lbl = label_float d in
+ fprintf oc " fld %a, %a, x31 %s %.18g\n"
+ freg rd label lbl comment (camlfloat_of_coqfloat f)
+ | Ploadsi(rd, f) ->
+ let s = camlint_of_coqint(Floats.Float32.to_bits f) in
+ let lbl = label_float32 s in
+ fprintf oc " flw %a, %a, x31 %s %.18g\n"
+ freg rd label lbl comment (camlfloat_of_coqfloat32 f)
+ | Pbtbl(r, tbl) ->
+ let lbl = new_label() in
+ fprintf oc "%s jumptable [ " comment;
+ List.iter (fun l -> fprintf oc "%a " print_label l) tbl;
+ fprintf oc "]\n";
+ fprintf oc " sll x5, %a, 2\n" ireg r;
+ fprintf oc " la x31, %a\n" label lbl;
+ fprintf oc " add x5, x31, x5\n";
+ fprintf oc " lw x5, 0(x5)\n";
+ fprintf oc " add x5, x31, x5\n";
+ fprintf oc " jr x5\n";
+ jumptables := (lbl, tbl) :: !jumptables;
+ fprintf oc "%s end pseudoinstr btbl\n" comment
+ | Pbuiltin(ef, args, res) ->
+ begin match ef with
+ | EF_annot(txt, targs) ->
+ fprintf oc "%s annotation: " comment;
+ print_annot_text preg "sp" oc (camlstring_of_coqstring txt) args
+ | EF_debug(kind, txt, targs) ->
+ print_debug_info comment print_file_line preg "sp" oc
+ (P.to_int kind) (extern_atom txt) args
+ | EF_inline_asm(txt, sg, clob) ->
+ fprintf oc "%s begin inline assembly\n\t" comment;
+ print_inline_asm preg oc (camlstring_of_coqstring txt) sg args res;
+ fprintf oc "%s end inline assembly\n" comment
+ | _ ->
+ assert false
+ end
+
+ let get_section_names name =
+ let (text, lit) =
+ match C2C.atom_sections name with
+ | t :: l :: _ -> (t, l)
+ | _ -> (Section_text, Section_literal) in
+ text,lit,Section_jumptable
+
+ let print_align oc alignment =
+ fprintf oc " .balign %d\n" alignment
+
+ let print_jumptable oc jmptbl =
+ let print_tbl oc (lbl, tbl) =
+ fprintf oc "%a:\n" label lbl;
+ List.iter
+ (fun l -> fprintf oc " .long %a - %a\n"
+ print_label l label lbl)
+ tbl in
+ if !jumptables <> [] then
+ begin
+ section oc jmptbl;
+ fprintf oc " .balign 4\n";
+ List.iter (print_tbl oc) !jumptables;
+ jumptables := []
+ end
+
+ let print_fun_info = elf_print_fun_info
+
+ let print_optional_fun_info _ = ()
+
+ let print_var_info = elf_print_var_info
+
+ let print_comm_symb oc sz name align =
+ if C2C.atom_is_static name then
+ fprintf oc " .local %a\n" symbol name;
+ fprintf oc " .comm %a, %s, %d\n"
+ symbol name
+ (Z.to_string sz)
+ align
+
+ let print_instructions oc fn =
+ current_function_sig := fn.fn_sig;
+ List.iter (print_instruction oc) fn.fn_code
+
+ let emit_constants oc lit =
+ section oc lit;
+ emit_constants oc
+
+(* Data *)
+
+ let address = if Archi.ptr64 then ".quad" else ".long"
+
+ let print_init oc = function
+ | Init_int8 n ->
+ fprintf oc " .byte %ld\n" (camlint_of_coqint n)
+ | Init_int16 n ->
+ fprintf oc " .short %ld\n" (camlint_of_coqint n)
+ | Init_int32 n ->
+ fprintf oc " .long %ld\n" (camlint_of_coqint n)
+ | Init_int64 n ->
+ fprintf oc " .quad %Ld\n" (camlint64_of_coqint n)
+ | Init_float32 n ->
+ fprintf oc " .long 0x%lx %s %.15g \n" (camlint_of_coqint (Floats.Float32.to_bits n))
+ comment (camlfloat_of_coqfloat n)
+ | Init_float64 n ->
+ fprintf oc " .quad %Ld %s %.18g\n" (camlint64_of_coqint (Floats.Float.to_bits n))
+ comment (camlfloat_of_coqfloat n)
+ | Init_space n ->
+ if Z.gt n Z.zero then
+ fprintf oc " .space %s\n" (Z.to_string n)
+ | Init_addrof(symb, ofs) ->
+ fprintf oc " %s %a\n" address symbol_offset (symb, ofs)
+
+ let print_prologue oc =
+ fprintf oc " .option %s\n" (if Archi.pic_code() then "pic" else "nopic");
+ if !Clflags.option_g then begin
+ section oc Section_text;
+ end
+
+ let print_epilogue oc =
+ if !Clflags.option_g then begin
+ Debug.compute_gnu_file_enum (fun f -> ignore (print_file oc f));
+ section oc Section_text;
+ end
+
+ let default_falignment = 2
+
+ let new_label = new_label
+
+ let cfi_startproc oc = ()
+ let cfi_endproc oc = ()
+
+ end
+
+let sel_target () =
+ (module Target:TARGET)
diff --git a/riscV/ValueAOp.v b/riscV/ValueAOp.v
new file mode 100644
index 00000000..5670b5fe
--- /dev/null
+++ b/riscV/ValueAOp.v
@@ -0,0 +1,218 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, INRIA Paris *)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+Require Import Coqlib Compopts.
+Require Import AST Integers Floats Values Memory Globalenvs.
+Require Import Op RTL ValueDomain.
+
+(** Value analysis for RISC V operators *)
+
+Definition eval_static_condition (cond: condition) (vl: list aval): abool :=
+ match cond, vl with
+ | Ccomp c, v1 :: v2 :: nil => cmp_bool c v1 v2
+ | Ccompu c, v1 :: v2 :: nil => cmpu_bool c v1 v2
+ | Ccompimm c n, v1 :: nil => cmp_bool c v1 (I n)
+ | Ccompuimm c n, v1 :: nil => cmpu_bool c v1 (I n)
+ | Ccompl c, v1 :: v2 :: nil => cmpl_bool c v1 v2
+ | Ccomplu c, v1 :: v2 :: nil => cmplu_bool c v1 v2
+ | Ccomplimm c n, v1 :: nil => cmpl_bool c v1 (L n)
+ | Ccompluimm c n, v1 :: nil => cmplu_bool c v1 (L n)
+ | Ccompf c, v1 :: v2 :: nil => cmpf_bool c v1 v2
+ | Cnotcompf c, v1 :: v2 :: nil => cnot (cmpf_bool c v1 v2)
+ | Ccompfs c, v1 :: v2 :: nil => cmpfs_bool c v1 v2
+ | Cnotcompfs c, v1 :: v2 :: nil => cnot (cmpfs_bool c v1 v2)
+ | _, _ => Bnone
+ end.
+
+Definition eval_static_addressing (addr: addressing) (vl: list aval): aval :=
+ match addr, vl with
+ | Aindexed n, v1::nil => offset_ptr v1 n
+ | Aglobal s ofs, nil => Ptr (Gl s ofs)
+ | Ainstack ofs, nil => Ptr (Stk ofs)
+ | _, _ => Vbot
+ end.
+
+Definition eval_static_operation (op: operation) (vl: list aval): aval :=
+ match op, vl with
+ | Omove, v1::nil => v1
+ | Ointconst n, nil => I n
+ | Olongconst n, nil => L n
+ | Ofloatconst n, nil => if propagate_float_constants tt then F n else ntop
+ | Osingleconst n, nil => if propagate_float_constants tt then FS n else ntop
+ | Oaddrsymbol id ofs, nil => Ptr (Gl id ofs)
+ | Oaddrstack ofs, nil => Ptr (Stk ofs)
+ | Ocast8signed, v1 :: nil => sign_ext 8 v1
+ | Ocast16signed, v1 :: nil => sign_ext 16 v1
+ | Oadd, v1::v2::nil => add v1 v2
+ | Oaddimm n, v1::nil => add v1 (I n)
+ | Oneg, v1::nil => neg v1
+ | Osub, v1::v2::nil => sub v1 v2
+ | Omul, v1::v2::nil => mul v1 v2
+ | Omulhs, v1::v2::nil => mulhs v1 v2
+ | Omulhu, v1::v2::nil => mulhu v1 v2
+ | Odiv, v1::v2::nil => divs v1 v2
+ | Odivu, v1::v2::nil => divu v1 v2
+ | Omod, v1::v2::nil => mods v1 v2
+ | Omodu, v1::v2::nil => modu v1 v2
+ | Oand, v1::v2::nil => and v1 v2
+ | Oandimm n, v1::nil => and v1 (I n)
+ | Oor, v1::v2::nil => or v1 v2
+ | Oorimm n, v1::nil => or v1 (I n)
+ | Oxor, v1::v2::nil => xor v1 v2
+ | Oxorimm n, v1::nil => xor v1 (I n)
+ | Oshl, v1::v2::nil => shl v1 v2
+ | Oshlimm n, v1::nil => shl v1 (I n)
+ | Oshr, v1::v2::nil => shr v1 v2
+ | Oshrimm n, v1::nil => shr v1 (I n)
+ | Oshru, v1::v2::nil => shru v1 v2
+ | Oshruimm n, v1::nil => shru v1 (I n)
+ | Oshrximm n, v1::nil => shrx v1 (I n)
+ | Omakelong, v1::v2::nil => longofwords v1 v2
+ | Olowlong, v1::nil => loword v1
+ | Ohighlong, v1::nil => hiword v1
+ | Ocast32signed, v1::nil => longofint v1
+ | Ocast32unsigned, v1::nil => longofintu v1
+ | Oaddl, v1::v2::nil => addl v1 v2
+ | Oaddlimm n, v1::nil => addl v1 (L n)
+ | Onegl, v1::nil => negl v1
+ | Osubl, v1::v2::nil => subl v1 v2
+ | Omull, v1::v2::nil => mull v1 v2
+ | Omullhs, v1::v2::nil => mullhs v1 v2
+ | Omullhu, v1::v2::nil => mullhu v1 v2
+ | Odivl, v1::v2::nil => divls v1 v2
+ | Odivlu, v1::v2::nil => divlu v1 v2
+ | Omodl, v1::v2::nil => modls v1 v2
+ | Omodlu, v1::v2::nil => modlu v1 v2
+ | Oandl, v1::v2::nil => andl v1 v2
+ | Oandlimm n, v1::nil => andl v1 (L n)
+ | Oorl, v1::v2::nil => orl v1 v2
+ | Oorlimm n, v1::nil => orl v1 (L n)
+ | Oxorl, v1::v2::nil => xorl v1 v2
+ | Oxorlimm n, v1::nil => xorl v1 (L n)
+ | Oshll, v1::v2::nil => shll v1 v2
+ | Oshllimm n, v1::nil => shll v1 (I n)
+ | Oshrl, v1::v2::nil => shrl v1 v2
+ | Oshrlimm n, v1::nil => shrl v1 (I n)
+ | Oshrlu, v1::v2::nil => shrlu v1 v2
+ | Oshrluimm n, v1::nil => shrlu v1 (I n)
+ | Oshrxlimm n, v1::nil => shrxl v1 (I n)
+ | Onegf, v1::nil => negf v1
+ | Oabsf, v1::nil => absf v1
+ | Oaddf, v1::v2::nil => addf v1 v2
+ | Osubf, v1::v2::nil => subf v1 v2
+ | Omulf, v1::v2::nil => mulf v1 v2
+ | Odivf, v1::v2::nil => divf v1 v2
+ | Onegfs, v1::nil => negfs v1
+ | Oabsfs, v1::nil => absfs v1
+ | Oaddfs, v1::v2::nil => addfs v1 v2
+ | Osubfs, v1::v2::nil => subfs v1 v2
+ | Omulfs, v1::v2::nil => mulfs v1 v2
+ | Odivfs, v1::v2::nil => divfs v1 v2
+ | Osingleoffloat, v1::nil => singleoffloat v1
+ | Ofloatofsingle, v1::nil => floatofsingle v1
+ | Ointoffloat, v1::nil => intoffloat v1
+ | Ointuoffloat, v1::nil => intuoffloat v1
+ | Ofloatofint, v1::nil => floatofint v1
+ | Ofloatofintu, v1::nil => floatofintu v1
+ | Ointofsingle, v1::nil => intofsingle v1
+ | Ointuofsingle, v1::nil => intuofsingle v1
+ | Osingleofint, v1::nil => singleofint v1
+ | Osingleofintu, v1::nil => singleofintu v1
+ | Olongoffloat, v1::nil => longoffloat v1
+ | Olonguoffloat, v1::nil => longuoffloat v1
+ | Ofloatoflong, v1::nil => floatoflong v1
+ | Ofloatoflongu, v1::nil => floatoflongu v1
+ | Olongofsingle, v1::nil => longofsingle v1
+ | Olonguofsingle, v1::nil => longuofsingle v1
+ | Osingleoflong, v1::nil => singleoflong v1
+ | Osingleoflongu, v1::nil => singleoflongu v1
+ | Ocmp c, _ => of_optbool (eval_static_condition c vl)
+ | _, _ => Vbot
+ end.
+
+Section SOUNDNESS.
+
+Variable bc: block_classification.
+Variable ge: genv.
+Hypothesis GENV: genv_match bc ge.
+Variable sp: block.
+Hypothesis STACK: bc sp = BCstack.
+
+Theorem eval_static_condition_sound:
+ forall cond vargs m aargs,
+ list_forall2 (vmatch bc) vargs aargs ->
+ cmatch (eval_condition cond vargs m) (eval_static_condition cond aargs).
+Proof.
+ intros until aargs; intros VM. inv VM.
+ destruct cond; auto with va.
+ inv H0.
+ destruct cond; simpl; eauto with va.
+ inv H2.
+ destruct cond; simpl; eauto with va.
+ destruct cond; auto with va.
+Qed.
+
+Lemma symbol_address_sound:
+ forall id ofs,
+ vmatch bc (Genv.symbol_address ge id ofs) (Ptr (Gl id ofs)).
+Proof.
+ intros; apply symbol_address_sound; apply GENV.
+Qed.
+
+Lemma symbol_address_sound_2:
+ forall id ofs,
+ vmatch bc (Genv.symbol_address ge id ofs) (Ifptr (Gl id ofs)).
+Proof.
+ intros. unfold Genv.symbol_address. destruct (Genv.find_symbol ge id) as [b|] eqn:F.
+ constructor. constructor. apply GENV; auto.
+ constructor.
+Qed.
+
+Hint Resolve symbol_address_sound symbol_address_sound_2: va.
+
+Ltac InvHyps :=
+ match goal with
+ | [H: None = Some _ |- _ ] => discriminate
+ | [H: Some _ = Some _ |- _] => inv H
+ | [H1: match ?vl with nil => _ | _ :: _ => _ end = Some _ ,
+ H2: list_forall2 _ ?vl _ |- _ ] => inv H2; InvHyps
+ | [H: (if Archi.ptr64 then _ else _) = Some _ |- _] => destruct Archi.ptr64 eqn:?; InvHyps
+ | _ => idtac
+ end.
+
+Theorem eval_static_addressing_sound:
+ forall addr vargs vres aargs,
+ eval_addressing ge (Vptr sp Ptrofs.zero) addr vargs = Some vres ->
+ list_forall2 (vmatch bc) vargs aargs ->
+ vmatch bc vres (eval_static_addressing addr aargs).
+Proof.
+ unfold eval_addressing, eval_static_addressing; intros;
+ destruct addr; InvHyps; eauto with va.
+ rewrite Ptrofs.add_zero_l; eauto with va.
+Qed.
+
+Theorem eval_static_operation_sound:
+ forall op vargs m vres aargs,
+ eval_operation ge (Vptr sp Ptrofs.zero) op vargs m = Some vres ->
+ list_forall2 (vmatch bc) vargs aargs ->
+ vmatch bc vres (eval_static_operation op aargs).
+Proof.
+ unfold eval_operation, eval_static_operation; intros;
+ destruct op; InvHyps; eauto with va.
+ destruct (propagate_float_constants tt); constructor.
+ destruct (propagate_float_constants tt); constructor.
+ rewrite Ptrofs.add_zero_l; eauto with va.
+ apply of_optbool_sound. eapply eval_static_condition_sound; eauto.
+Qed.
+
+End SOUNDNESS.
+
diff --git a/riscV/extractionMachdep.v b/riscV/extractionMachdep.v
new file mode 100644
index 00000000..81cfc88c
--- /dev/null
+++ b/riscV/extractionMachdep.v
@@ -0,0 +1,24 @@
+(* *********************************************************************)
+(* *)
+(* 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. *)
+(* *)
+(* *********************************************************************)
+
+(* Additional extraction directives specific to the RISC-V port *)
+
+Require Archi Asm.
+
+(* Archi *)
+
+Extract Constant Archi.ptr64 => " Configuration.model = ""64"" ".
+Extract Constant Archi.pic_code => "fun () -> false". (* for the time being *)
+
+(* Asm *)
+Extract Constant Asm.low_half => "fun _ _ _ -> assert false".
+Extract Constant Asm.high_half => "fun _ _ _ -> assert false".
diff --git a/runtime/Makefile b/runtime/Makefile
index 641c9fdc..b819991d 100644
--- a/runtime/Makefile
+++ b/runtime/Makefile
@@ -52,8 +52,19 @@ $(LIB): $(OBJS)
%.o: %.S
$(CASMRUNTIME) -DMODEL_$(MODEL) -DABI_$(ABI) -DENDIANNESS_$(ENDIANNESS) -DSYS_$(SYSTEM) -o $@ $^
+# If no asm implementation available, compile the portable C implementation
+# with CompCert. Since CompCert rejects the "__i64_" identifiers, the C
+# implementation uses "i64_" identifiers, and we rename them in the
+# generated assembly
+
+%.o: c/%.c c/i64.h ../ccomp
+ ../ccomp -O2 -S -o $*.s -I./c c/$*.c
+ sed -i -e 's/i64_/__i64_/g' $*.s
+ $(CASMRUNTIME) -o $*.o $*.s
+ @rm -f $*.s
+
clean::
- rm -f *.o $(LIB)
+ rm -f *.o $(LIB) *.tmp?.s
ifeq ($(strip $(HAS_RUNTIME_LIB)),true)
install::
diff --git a/runtime/c/i64.h b/runtime/c/i64.h
index a75214fe..3f59164b 100644
--- a/runtime/c/i64.h
+++ b/runtime/c/i64.h
@@ -34,12 +34,12 @@
/* Helper functions for 64-bit integer arithmetic. Reference C implementation */
-extern unsigned long long __i64_shl(unsigned long long x, int amount);
-extern unsigned long long __i64_shr(unsigned long long x, int amount);
-extern signed long long __i64_sar(signed long long x, int amount);
+extern unsigned long long i64_shl(unsigned long long x, int amount);
+extern unsigned long long i64_shr(unsigned long long x, int amount);
+extern signed long long i64_sar(signed long long x, int amount);
-extern unsigned long long __i64_udivmod(unsigned long long n,
+extern unsigned long long i64_udivmod(unsigned long long n,
unsigned long long d,
unsigned long long * rp);
-extern unsigned long long __i64_umulh(unsigned long long u,
+extern unsigned long long i64_umulh(unsigned long long u,
unsigned long long v);
diff --git a/runtime/c/i64_dtos.c b/runtime/c/i64_dtos.c
index d428e744..5d12bf4b 100644
--- a/runtime/c/i64_dtos.c
+++ b/runtime/c/i64_dtos.c
@@ -38,7 +38,7 @@
/* Conversion float64 -> signed int64 */
-long long __i64_dtos(double d)
+long long i64_dtos(double d)
{
/* Extract bits of d's representation */
union { double d; unsigned long long i; } buf;
@@ -63,9 +63,9 @@ long long __i64_dtos(double d)
(buf.i & ~0xFFF0000000000000LL) | 0x0010000000000000LL;
/* Shift it appropriately */
if (e >= 0)
- m = __i64_shl(m, e);
+ m = i64_shl(m, e);
else
- m = __i64_shr(m, -e);
+ m = i64_shr(m, -e);
/* Apply sign to result */
if ((int) h < 0)
return -m;
diff --git a/runtime/c/i64_dtou.c b/runtime/c/i64_dtou.c
index c5f9df4d..5f4901b2 100644
--- a/runtime/c/i64_dtou.c
+++ b/runtime/c/i64_dtou.c
@@ -38,7 +38,7 @@
/* Conversion float64 -> unsigned int64 */
-unsigned long long __i64_dtou(double d)
+unsigned long long i64_dtou(double d)
{
/* Extract bits of d's representation */
union { double d; unsigned long long i; } buf;
@@ -62,8 +62,8 @@ unsigned long long __i64_dtou(double d)
(buf.i & ~0xFFF0000000000000LL) | 0x0010000000000000LL;
/* Shift it appropriately */
if (e >= 0)
- return __i64_shl(m, e);
+ return i64_shl(m, e);
else
- return __i64_shr(m, -e);
+ return i64_shr(m, -e);
}
diff --git a/runtime/c/i64_sar.c b/runtime/c/i64_sar.c
index a5f3364a..611aa047 100644
--- a/runtime/c/i64_sar.c
+++ b/runtime/c/i64_sar.c
@@ -38,7 +38,7 @@
/* Shift right signed */
-signed long long __i64_sar(signed long long x, int amount)
+signed long long i64_sar(signed long long x, int amount)
{
unsigned xl = x;
int xh = x >> 32;
diff --git a/runtime/c/i64_sdiv.c b/runtime/c/i64_sdiv.c
index 3f21d9b7..20faf121 100644
--- a/runtime/c/i64_sdiv.c
+++ b/runtime/c/i64_sdiv.c
@@ -38,14 +38,14 @@
/* Signed division */
-signed long long __i64_sdiv(signed long long n, signed long long d)
+signed long long i64_sdiv(signed long long n, signed long long d)
{
unsigned long long un, ud, uq, ur;
int nh = n >> 32, dh = d >> 32;
/* Take absolute values of n and d */
un = nh < 0 ? -n : n;
ud = dh < 0 ? -d : d;
- uq = __i64_udivmod(un, ud, &ur);
+ uq = i64_udivmod(un, ud, &ur);
/* Apply sign to quotient */
return (nh ^ dh) < 0 ? -uq : uq;
}
diff --git a/runtime/c/i64_shl.c b/runtime/c/i64_shl.c
index 9b9aae57..186d2d32 100644
--- a/runtime/c/i64_shl.c
+++ b/runtime/c/i64_shl.c
@@ -38,7 +38,7 @@
/* Shift left */
-unsigned long long __i64_shl(unsigned long long x, int amount)
+unsigned long long i64_shl(unsigned long long x, int amount)
{
unsigned xl = x, xh = x >> 32;
amount = amount & 63;
diff --git a/runtime/c/i64_shr.c b/runtime/c/i64_shr.c
index c1db2a5f..bb2facce 100644
--- a/runtime/c/i64_shr.c
+++ b/runtime/c/i64_shr.c
@@ -38,7 +38,7 @@
/* Shift right unsigned */
-unsigned long long __i64_shr(unsigned long long x, int amount)
+unsigned long long i64_shr(unsigned long long x, int amount)
{
unsigned xl = x, xh = x >> 32;
amount = amount & 63;
diff --git a/runtime/c/i64_smod.c b/runtime/c/i64_smod.c
index ab15b6e6..b0f2c3cd 100644
--- a/runtime/c/i64_smod.c
+++ b/runtime/c/i64_smod.c
@@ -38,14 +38,14 @@
/* Signed remainder */
-signed long long __i64_smod(signed long long n, signed long long d)
+signed long long i64_smod(signed long long n, signed long long d)
{
unsigned long long un, ud, ur;
int nh = n >> 32, dh = d >> 32;
/* Take absolute values of n and d */
un = nh < 0 ? -n : n;
ud = dh < 0 ? -d : d;
- (void) __i64_udivmod(un, ud, &ur);
+ (void) i64_udivmod(un, ud, &ur);
/* Apply sign to remainder */
return nh < 0 ? -ur : ur;
}
diff --git a/runtime/c/i64_smulh.c b/runtime/c/i64_smulh.c
index b7a42474..3b33d375 100644
--- a/runtime/c/i64_smulh.c
+++ b/runtime/c/i64_smulh.c
@@ -47,9 +47,9 @@ typedef unsigned long long u64;
* - subtract Y if X < 0
*/
-s64 __i64_smulh(s64 x, s64 y)
+s64 i64_smulh(s64 x, s64 y)
{
- s64 t = (s64) __i64_umulh(x, y);
+ s64 t = (s64) i64_umulh(x, y);
if (y < 0) t = t - x;
if (x < 0) t = t - y;
return t;
diff --git a/runtime/c/i64_stod.c b/runtime/c/i64_stod.c
index 158b6892..b5b6a7bc 100644
--- a/runtime/c/i64_stod.c
+++ b/runtime/c/i64_stod.c
@@ -38,7 +38,7 @@
/* Conversion from signed int64 to float64 */
-double __i64_stod(signed long long x)
+double i64_stod(signed long long x)
{
unsigned xl = x;
signed xh = x >> 32;
diff --git a/runtime/c/i64_stof.c b/runtime/c/i64_stof.c
index 8410ba13..12406421 100644
--- a/runtime/c/i64_stof.c
+++ b/runtime/c/i64_stof.c
@@ -38,7 +38,7 @@
/* Conversion from signed int64 to float32 */
-float __i64_stof(signed long long x)
+float i64_stof(signed long long x)
{
if (x < -(1LL << 53) || x >= (1LL << 53)) {
/* x is large enough that double rounding can occur.
diff --git a/runtime/c/i64_udiv.c b/runtime/c/i64_udiv.c
index 91fbf6e4..19d9264d 100644
--- a/runtime/c/i64_udiv.c
+++ b/runtime/c/i64_udiv.c
@@ -38,8 +38,8 @@
/* Unsigned division */
-unsigned long long __i64_udiv(unsigned long long n, unsigned long long d)
+unsigned long long i64_udiv(unsigned long long n, unsigned long long d)
{
unsigned long long r;
- return __i64_udivmod(n, d, &r);
+ return i64_udivmod(n, d, &r);
}
diff --git a/runtime/c/i64_udivmod.c b/runtime/c/i64_udivmod.c
index d8f5073a..03efbc9d 100644
--- a/runtime/c/i64_udivmod.c
+++ b/runtime/c/i64_udivmod.c
@@ -37,13 +37,13 @@
#include <stddef.h>
#include "i64.h"
-static unsigned __i64_udiv6432(unsigned u1, unsigned u0,
+static unsigned i64_udiv6432(unsigned u1, unsigned u0,
unsigned v, unsigned *r);
-static int __i64_nlz(unsigned x);
+static int i64_nlz(unsigned x);
/* Unsigned division and remainder */
-unsigned long long __i64_udivmod(unsigned long long n,
+unsigned long long i64_udivmod(unsigned long long n,
unsigned long long d,
unsigned long long * rp)
{
@@ -62,7 +62,7 @@ unsigned long long __i64_udivmod(unsigned long long n,
unsigned dl = d;
unsigned qh = nh / dl;
unsigned rl;
- unsigned ql = __i64_udiv6432(nh % dl, nl, dl, &rl);
+ unsigned ql = i64_udiv6432(nh % dl, nl, dl, &rl);
*rp = (unsigned long long) rl; /* high word of remainder is 0 */
return ((unsigned long long) qh) << 32 | ql;
}
@@ -70,11 +70,11 @@ unsigned long long __i64_udivmod(unsigned long long n,
/* General case 64 / 64 */
unsigned dl = d;
/* Scale N and D down, giving N' and D' such that 2^31 <= D' < 2^32 */
- int s = 32 - __i64_nlz(dh); /* shift amount, between 1 and 32 */
- unsigned long long np = __i64_shr(n, s);
- unsigned dp = (unsigned) __i64_shr(d, s);
+ int s = 32 - i64_nlz(dh); /* shift amount, between 1 and 32 */
+ unsigned long long np = i64_shr(n, s);
+ unsigned dp = (unsigned) i64_shr(d, s);
/* Divide N' by D' to get an approximate quotient Q */
- unsigned q = __i64_udiv6432(np >> 32, np, dp, NULL);
+ unsigned q = i64_udiv6432(np >> 32, np, dp, NULL);
again: ;
/* Tentative quotient Q is either correct or one too high */
/* Compute Q * D, checking for overflow */
@@ -99,7 +99,7 @@ unsigned long long __i64_udivmod(unsigned long long n,
/* Unsigned division and remainder for 64 bits divided by 32 bits. */
/* This is algorithm "divlu" from _Hacker's Delight_, fig 9.3 */
-static unsigned __i64_udiv6432(unsigned u1, unsigned u0,
+static unsigned i64_udiv6432(unsigned u1, unsigned u0,
unsigned v, unsigned *r)
{
const unsigned b = 65536; // Number base (16 bits).
@@ -114,7 +114,7 @@ static unsigned __i64_udiv6432(unsigned u1, unsigned u0,
if (r != NULL) *r = 0xFFFFFFFFU; // set rem to an impossible value,
return 0xFFFFFFFFU; // and return largest possible quotient.
}
- s = __i64_nlz(v); // 0 <= s <= 31.
+ s = i64_nlz(v); // 0 <= s <= 31.
v = v << s; // Normalize divisor.
vn1 = v >> 16; // Break divisor up into
vn0 = v & 0xFFFF; // two 16-bit digits.
@@ -145,7 +145,7 @@ again2:
/* Number of leading zeroes */
-static int __i64_nlz(unsigned x)
+static int i64_nlz(unsigned x)
{
if (x == 0) return 32;
int n = 0;
diff --git a/runtime/c/i64_umod.c b/runtime/c/i64_umod.c
index d30e29c4..37640d5f 100644
--- a/runtime/c/i64_umod.c
+++ b/runtime/c/i64_umod.c
@@ -38,9 +38,9 @@
/* Unsigned remainder */
-unsigned long long __i64_umod(unsigned long long n, unsigned long long d)
+unsigned long long i64_umod(unsigned long long n, unsigned long long d)
{
unsigned long long r;
- (void) __i64_udivmod(n, d, &r);
+ (void) i64_udivmod(n, d, &r);
return r;
}
diff --git a/runtime/c/i64_umulh.c b/runtime/c/i64_umulh.c
index d2394d09..a2e33f5d 100644
--- a/runtime/c/i64_umulh.c
+++ b/runtime/c/i64_umulh.c
@@ -43,7 +43,7 @@ typedef unsigned int u32;
/* Hacker's Delight, algorithm 8.1, specialized to two 32-bit words */
-u64 __i64_umulh(u64 u, u64 v)
+u64 i64_umulh(u64 u, u64 v)
{
u32 u0 = u, u1 = u >> 32;
u32 v0 = v, v1 = v >> 32;
diff --git a/runtime/c/i64_utod.c b/runtime/c/i64_utod.c
index e70820a9..7210c73c 100644
--- a/runtime/c/i64_utod.c
+++ b/runtime/c/i64_utod.c
@@ -38,7 +38,7 @@
/* Conversion from unsigned int64 to float64 */
-double __i64_utod(unsigned long long x)
+double i64_utod(unsigned long long x)
{
unsigned xl = x, xh = x >> 32;
return (double) xl + 0x1.0p32 * (double) xh;
diff --git a/runtime/c/i64_utof.c b/runtime/c/i64_utof.c
index 87b85bfc..77a23765 100644
--- a/runtime/c/i64_utof.c
+++ b/runtime/c/i64_utof.c
@@ -38,7 +38,7 @@
/* Conversion from unsigned int64 to float32 */
-float __i64_utof(unsigned long long x)
+float i64_utof(unsigned long long x)
{
if (x >= 1ULL << 53) {
/* x is large enough that double rounding can occur.
diff --git a/runtime/riscV/sysdeps.h b/runtime/riscV/sysdeps.h
new file mode 100644
index 00000000..b95ca1b1
--- /dev/null
+++ b/runtime/riscV/sysdeps.h
@@ -0,0 +1,63 @@
+// *****************************************************************
+//
+// The Compcert verified compiler
+//
+// Xavier Leroy, INRIA Paris-Rocquencourt
+//
+// Copyright (c) 2013 Institut National de Recherche en Informatique et
+// en Automatique.
+//
+// Redistribution and use in source and binary forms, with or without
+// modification, are permitted provided that the following conditions are met:
+// * Redistributions of source code must retain the above copyright
+// notice, this list of conditions and the following disclaimer.
+// * Redistributions in binary form must reproduce the above copyright
+// notice, this list of conditions and the following disclaimer in the
+// documentation and/or other materials provided with the distribution.
+// * Neither the name of the <organization> nor the
+// names of its contributors may be used to endorse or promote products
+// derived from this software without specific prior written permission.
+//
+// THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
+// "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
+// LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
+// A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
+// HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
+// EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
+// PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
+// PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
+// LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
+// NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
+// SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+//
+// *********************************************************************
+
+// System dependencies
+
+#define FUNCTION(f) \
+ .text; \
+ .balign 16; \
+ .globl f; \
+f:
+
+#define ENDFUNCTION(f) \
+ .type f, @function; .size f, . - f
+
+#if defined(MODEL_64)
+
+#define WORDSIZE 8
+#define lptr ld
+#define sptr sd
+
+#elif defined(MODEL_32)
+
+#define WORDSIZE 4
+#define lptr lw
+#define sptr sw
+
+#else
+
+#error "Wrong MODEL"
+
+#endif
+
diff --git a/runtime/riscV/vararg.S b/runtime/riscV/vararg.S
new file mode 100644
index 00000000..a9481077
--- /dev/null
+++ b/runtime/riscV/vararg.S
@@ -0,0 +1,90 @@
+// *****************************************************************
+//
+// The Compcert verified compiler
+//
+// Xavier Leroy, INRIA Paris-Rocquencourt
+// Prashanth Mundkur, SRI International
+//
+// Copyright (c) 2013 Institut National de Recherche en Informatique et
+// en Automatique.
+//
+// The contributions by Prashanth Mundkur are reused and adapted
+// under the terms of a Contributor License Agreement between
+// SRI International and INRIA.
+//
+// Redistribution and use in source and binary forms, with or without
+// modification, are permitted provided that the following conditions are met:
+// * Redistributions of source code must retain the above copyright
+// notice, this list of conditions and the following disclaimer.
+// * Redistributions in binary form must reproduce the above copyright
+// notice, this list of conditions and the following disclaimer in the
+// documentation and/or other materials provided with the distribution.
+// * Neither the name of the <organization> nor the
+// names of its contributors may be used to endorse or promote products
+// derived from this software without specific prior written permission.
+//
+// THIS SOFTWARE IS PROVIDED BY THE COPYRIGHT HOLDERS AND CONTRIBUTORS
+// "AS IS" AND ANY EXPRESS OR IMPLIED WARRANTIES, INCLUDING, BUT NOT
+// LIMITED TO, THE IMPLIED WARRANTIES OF MERCHANTABILITY AND FITNESS FOR
+// A PARTICULAR PURPOSE ARE DISCLAIMED. IN NO EVENT SHALL <COPYRIGHT
+// HOLDER> BE LIABLE FOR ANY DIRECT, INDIRECT, INCIDENTAL, SPECIAL,
+// EXEMPLARY, OR CONSEQUENTIAL DAMAGES (INCLUDING, BUT NOT LIMITED TO,
+// PROCUREMENT OF SUBSTITUTE GOODS OR SERVICES; LOSS OF USE, DATA, OR
+// PROFITS; OR BUSINESS INTERRUPTION) HOWEVER CAUSED AND ON ANY THEORY OF
+// LIABILITY, WHETHER IN CONTRACT, STRICT LIABILITY, OR TORT (INCLUDING
+// NEGLIGENCE OR OTHERWISE) ARISING IN ANY WAY OUT OF THE USE OF THIS
+// SOFTWARE, EVEN IF ADVISED OF THE POSSIBILITY OF SUCH DAMAGE.
+//
+// *********************************************************************
+
+// Helper functions for variadic functions <stdarg.h>. RISC-V version.
+
+#include "sysdeps.h"
+
+// typedef void * va_list;
+// unsigned int __compcert_va_int32(va_list * ap);
+// unsigned long long __compcert_va_int64(va_list * ap);
+// double __compcert_va_float64(va_list * ap);
+
+FUNCTION(__compcert_va_int32)
+ # a0 = ap parameter
+ lptr t5, 0(a0) # t5 = pointer to next argument
+ addi t5, t5, WORDSIZE # advance ap
+ sptr t5, 0(a0) # update ap
+ lw a0, -WORDSIZE(t5) # load it and return it in a0
+ jr ra
+ENDFUNCTION(__compcert_va_int32)
+
+FUNCTION(__compcert_va_int64)
+ # a0 = ap parameter
+ lptr t5, 0(a0) # t5 = pointer to next argument
+ addi t5, t5, 15 # 8-align and advance by 8
+ and t5, t5, -8
+ sptr t5, 0(a0) # update ap
+#ifdef MODEL_64
+ ld a0, -8(t5) # return it in a0
+#else
+ lw a0, -8(t5) # return it in [a0,a1]
+ lw a1, -4(t5)
+#endif
+ jr ra
+ENDFUNCTION(__compcert_va_int64)
+
+FUNCTION(__compcert_va_float64)
+ # a0 = ap parameter
+ lptr t5, 0(a0) # t5 = pointer to next argument
+ addi t5, t5, 15 # 8-align and advance by 8
+ and t5, t5, -8
+ sw t5, 0(a0) # update ap
+ fld fa0, -8(t5) # return it in fa0
+ jr ra
+ENDFUNCTION(__compcert_va_float64)
+
+// Right now we pass structs by reference. This is not ABI conformant.
+FUNCTION(__compcert_va_composite)
+#ifdef MODEL_64
+ j __compcert_va_int64
+#else
+ j __compcert_va_int32
+#endif
+ENDFUNCTION(__compcert_va_composite)
diff --git a/test/Makefile b/test/Makefile
index b469eec2..5e9e0555 100644
--- a/test/Makefile
+++ b/test/Makefile
@@ -4,7 +4,7 @@ all:
for i in $(DIRS); do $(MAKE) -C $$i all; done
test:
- set -e; for i in $(DIRS); do $(MAKE) -C $$i test; done
+ set -e; for i in $(DIRS); do $(MAKE) SIMU='$(SIMU)' -C $$i test; done
bench:
for i in $(DIRS); do $(MAKE) -C $$i bench; done
diff --git a/test/c/Makefile b/test/c/Makefile
index 5979dfd4..94feb993 100644
--- a/test/c/Makefile
+++ b/test/c/Makefile
@@ -32,7 +32,7 @@ all_gcc: $(PROGS:%=%.gcc)
test:
@for i in $(PROGS); do \
- if ./$$i.compcert | cmp -s - Results/$$i; \
+ if $(SIMU) ./$$i.compcert | cmp -s - Results/$$i; \
then echo "$$i: passed"; \
else echo "$$i: FAILED"; exit 2; \
fi; \
diff --git a/test/c/aes.c b/test/c/aes.c
index 053324f3..5bd57cbc 100644
--- a/test/c/aes.c
+++ b/test/c/aes.c
@@ -38,7 +38,8 @@ typedef unsigned int u32;
#if defined(__ppc__) || defined(__PPC__) || defined(__ARMEB__)
#define ARCH_BIG_ENDIAN
-#elif defined(__i386__) || defined(__x86_64__) || defined(__ARMEL__)
+#elif defined(__i386__) || defined(__x86_64__) || defined(__ARMEL__) \
+ || defined(__riscv)
#undef ARCH_BIG_ENDIAN
#else
#error "unknown endianness"
diff --git a/test/compression/Makefile b/test/compression/Makefile
index 784f7e73..fa73f0e6 100644
--- a/test/compression/Makefile
+++ b/test/compression/Makefile
@@ -35,9 +35,9 @@ test:
@echo "Test data: $(TESTFILE)"
@for i in $(EXE); do \
echo "$$i: compression..."; \
- ./$$i -c -i $(TESTFILE) -o $(TESTCOMPR); \
+ $(SIMU) ./$$i -c -i $(TESTFILE) -o $(TESTCOMPR); \
echo "$$i: decompression..."; \
- ./$$i -d -i $(TESTCOMPR) -o $(TESTEXPND); \
+ $(SIMU) ./$$i -d -i $(TESTCOMPR) -o $(TESTEXPND); \
if cmp $(TESTFILE) $(TESTEXPND); \
then echo "$$i: passed"; \
else echo "$$i: FAILED"; exit 2; \
diff --git a/test/raytracer/Makefile b/test/raytracer/Makefile
index dc007fd7..c481ff86 100644
--- a/test/raytracer/Makefile
+++ b/test/raytracer/Makefile
@@ -23,7 +23,7 @@ depend:
gcc -MM *.c > .depend
test:
- ./render < kal.gml
+ $(SIMU) ./render < kal.gml
@if cmp kal.ppm Results/kal.ppm; \
then echo "raytracer: passed"; \
else echo "raytracer: FAILED"; exit 2; \
diff --git a/test/regression/Makefile b/test/regression/Makefile
index 5def966b..54745863 100644
--- a/test/regression/Makefile
+++ b/test/regression/Makefile
@@ -71,15 +71,15 @@ clean:
test:
@echo "----------- Compiled tests -------------"
@for i in $(TESTS) $(TESTS_COMP); do \
- ./Runtest $$i ./$$i.compcert; \
+ SIMU='$(SIMU)' ./Runtest $$i ./$$i.compcert; \
done
@echo "----------- Interpreted tests -------------"
@for i in $(TESTS); do \
- ./Runtest $$i $(CCOMP) -fall -interp -quiet $$i.c; \
+ SIMU='' ./Runtest $$i $(CCOMP) -fall -interp -quiet $$i.c; \
done
@for i in $(TESTS_DIFF); do \
if $(CCOMP) -fall -interp -quiet $$i.c > _cinterp.log; then \
- if ./$$i.compcert | cmp -s _cinterp.log -; \
+ if $(SIMU) ./$$i.compcert | cmp -s _cinterp.log -; \
then echo "$$i: compiler and interpreter agree"; \
else echo "$$i: compiler and interpreter DISAGREE"; \
fi; \
diff --git a/test/regression/Results/builtins-riscV b/test/regression/Results/builtins-riscV
new file mode 100644
index 00000000..1576b252
--- /dev/null
+++ b/test/regression/Results/builtins-riscV
@@ -0,0 +1,12 @@
+bswap16(1234) = 3412
+bswap32(12345678) = 78563412
+bswap64(123456789abcdef0) = f0debc9a78563412
+fmadd(3.141590, 2.718000, 1.414000) = 9.952842
+fmsub(3.141590, 2.718000, 1.414000) = 7.124842
+fnmadd(3.141590, 2.718000, 1.414000) = -9.952842
+fnmsub(3.141590, 2.718000, 1.414000) = -7.124842
+fabs(3.141590) = 3.141590
+fabs(-3.141590) = 3.141590
+fsqrt(3.141590) = 1.772453
+fmax(3.141590, 2.718000) = 3.141590
+fmin(3.141590, 2.718000) = 2.718000
diff --git a/test/regression/Runtest b/test/regression/Runtest
index 9051b5b7..ad2a58f1 100755
--- a/test/regression/Runtest
+++ b/test/regression/Runtest
@@ -27,7 +27,7 @@ else
fi
# Administer the test
-if $* > $out
+if $SIMU $* > $out
then
if test -n "$ref"; then
if cmp -s "$out" "$ref"; then
diff --git a/test/regression/alignas.c b/test/regression/alignas.c
index a6a2e690..b3754039 100644
--- a/test/regression/alignas.c
+++ b/test/regression/alignas.c
@@ -7,6 +7,13 @@
#define _Alignof(x) __alignof__(x)
#endif
+#ifdef _Alignas
+#undef _Alignas
+#endif
+#ifdef _Alignof
+#undef _Alignof
+#endif
+
/* Base type */
int _Alignas(16) a;
char filler1;
diff --git a/test/regression/builtins-riscV.c b/test/regression/builtins-riscV.c
new file mode 100644
index 00000000..a302a6c4
--- /dev/null
+++ b/test/regression/builtins-riscV.c
@@ -0,0 +1,30 @@
+/* Fun with builtins */
+
+#include <stdio.h>
+
+int main(int argc, char ** argv)
+{
+ unsigned int x = 0x12345678;
+ unsigned short s = 0x1234;
+ unsigned long long zz = 0x123456789ABCDEF0ULL;
+ double a = 3.14159;
+ double b = 2.718;
+ double c = 1.414;
+
+ printf("bswap16(%x) = %x\n", s, __builtin_bswap16(s));
+ printf("bswap32(%x) = %x\n", x, __builtin_bswap32(x));
+ printf("bswap64(%llx) = %llx\n", zz, __builtin_bswap64(zz));
+ 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("fnmadd(%f, %f, %f) = %f\n", a, b, c, __builtin_fnmadd(a, b, c));
+ printf("fnmsub(%f, %f, %f) = %f\n", a, b, c, __builtin_fnmsub(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("fmax(%f, %f) = %f\n", a, b, __builtin_fmax(a, b));
+ printf("fmin(%f, %f) = %f\n", a, b, __builtin_fmin(a, b));
+ /* Make sure that ignoring the result of a builtin
+ doesn't cause an internal error */
+ (void) __builtin_fsqrt(a);
+ return 0;
+}
diff --git a/test/regression/extasm.c b/test/regression/extasm.c
index c0534047..a41c4202 100644
--- a/test/regression/extasm.c
+++ b/test/regression/extasm.c
@@ -18,7 +18,7 @@ int clobbers(int x, int z)
return y + z;
}
-#if defined(__x86_64__)
+#if defined(__x86_64__) || __riscv_xlen == 64
#define SIXTYFOUR
#else
#undef SIXTYFOUR
diff --git a/test/regression/floats-basics.c b/test/regression/floats-basics.c
index 5aa91d14..a7ba3623 100644
--- a/test/regression/floats-basics.c
+++ b/test/regression/floats-basics.c
@@ -6,7 +6,8 @@
#if defined(__ppc__) || defined(__PPC__) || defined(__ARMEB__)
#define ARCH_BIG_ENDIAN
-#elif defined(__i386__) || defined(__x86_64__) || defined(__ARMEL__)
+#elif defined(__i386__) || defined(__x86_64__) || defined(__ARMEL__) \
+ || defined(__riscv)
#undef ARCH_BIG_ENDIAN
#else
#error "unknown endianness"
diff --git a/test/regression/floats.c b/test/regression/floats.c
index 68d60f65..84c4e062 100644
--- a/test/regression/floats.c
+++ b/test/regression/floats.c
@@ -5,7 +5,8 @@
#if defined(__ppc__) || defined(__PPC__) || defined(__ARMEB__)
#define ARCH_BIG_ENDIAN
-#elif defined(__i386__) || defined(__x86_64__) || defined(__ARMEL__)
+#elif defined(__i386__) || defined(__x86_64__) || defined(__ARMEL__) \
+ || defined(__riscv)
#undef ARCH_BIG_ENDIAN
#else
#error "unknown endianness"
diff --git a/test/spass/Makefile b/test/spass/Makefile
index f6acc551..110359ad 100644
--- a/test/spass/Makefile
+++ b/test/spass/Makefile
@@ -22,7 +22,7 @@ clean:
rm -f *.o *.s *.parsed.c *.light.c *.sdump
test:
- ./spass small_problem.dfg | grep 'Proof found'
+ $(SIMU) ./spass small_problem.dfg | grep 'Proof found'
TIME=xtime -o /dev/null # Xavier's hack
#TIME=time >/dev/null # Otherwise
diff --git a/x86/SelectOp.vp b/x86/SelectOp.vp
index f8010f0a..2037760f 100644
--- a/x86/SelectOp.vp
+++ b/x86/SelectOp.vp
@@ -226,6 +226,9 @@ Nondetfunction mul (e1: expr) (e2: expr) :=
| _, _ => Eop Omul (e1:::e2:::Enil)
end.
+Definition mulhs (e1: expr) (e2: expr) := Eop Omulhs (e1 ::: e2 ::: Enil).
+Definition mulhu (e1: expr) (e2: expr) := Eop Omulhu (e1 ::: e2 ::: Enil).
+
(** ** Bitwise and, or, xor *)
Nondetfunction andimm (n1: int) (e2: expr) :=
diff --git a/x86/SelectOpproof.v b/x86/SelectOpproof.v
index cdb79c6f..1728c39d 100644
--- a/x86/SelectOpproof.v
+++ b/x86/SelectOpproof.v
@@ -363,6 +363,16 @@ Proof.
- TrivialExists.
Qed.
+Theorem eval_mulhs: binary_constructor_sound mulhs Val.mulhs.
+Proof.
+ unfold mulhs; red; intros; TrivialExists.
+Qed.
+
+Theorem eval_mulhu: binary_constructor_sound mulhu Val.mulhu.
+Proof.
+ unfold mulhu; red; intros; TrivialExists.
+Qed.
+
Theorem eval_andimm:
forall n, unary_constructor_sound (andimm n) (fun x => Val.and x (Vint n)).
Proof.