aboutsummaryrefslogtreecommitdiffstats
path: root/backend
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 /backend
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.
Diffstat (limited to 'backend')
-rw-r--r--backend/PrintAsmaux.ml7
-rw-r--r--backend/SelectDiv.vp58
-rw-r--r--backend/SelectDivproof.v78
-rw-r--r--backend/ValueDomain.v17
4 files changed, 117 insertions, 43 deletions
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