aboutsummaryrefslogtreecommitdiffstats
path: root/backend
diff options
context:
space:
mode:
Diffstat (limited to 'backend')
-rw-r--r--backend/Asmaux.v5
-rw-r--r--backend/Asmexpandaux.ml4
-rw-r--r--backend/IRC.ml5
-rw-r--r--backend/IRC.mli1
-rw-r--r--backend/Lineartyping.v6
-rw-r--r--backend/OpHelpers.v42
-rw-r--r--backend/OpHelpersproof.v78
-rw-r--r--backend/Regalloc.ml2
-rw-r--r--backend/SelectDiv.vp39
-rw-r--r--backend/SelectDivproof.v31
-rw-r--r--backend/Selection.v12
-rw-r--r--backend/Selectionaux.ml13
-rw-r--r--backend/Selectionproof.v23
-rw-r--r--backend/SplitLong.vp31
-rw-r--r--backend/SplitLongproof.v11
-rw-r--r--backend/Stackingproof.v13
16 files changed, 217 insertions, 99 deletions
diff --git a/backend/Asmaux.v b/backend/Asmaux.v
new file mode 100644
index 00000000..51e94f6b
--- /dev/null
+++ b/backend/Asmaux.v
@@ -0,0 +1,5 @@
+Require Import Asm.
+Require Import AST.
+
+(* Constant only needed by Asmexpandaux.ml *)
+Definition dummy_function := {| fn_code := nil; fn_sig := signature_main |}. \ No newline at end of file
diff --git a/backend/Asmexpandaux.ml b/backend/Asmexpandaux.ml
index f5c76925..b1d822db 100644
--- a/backend/Asmexpandaux.ml
+++ b/backend/Asmexpandaux.ml
@@ -15,6 +15,7 @@
pseudo-instructions *)
open Asm
+open Asmaux
open AST
open Camlcoq
@@ -26,7 +27,10 @@ let emit i = current_code := i :: !current_code
(* Generation of fresh labels *)
+(* now imported from Asmaux.ml
let dummy_function = { fn_code = []; fn_sig = signature_main }
+*)
+
let current_function = ref dummy_function
let next_label = ref (None: label option)
diff --git a/backend/IRC.ml b/backend/IRC.ml
index 43955897..67da47da 100644
--- a/backend/IRC.ml
+++ b/backend/IRC.ml
@@ -15,6 +15,7 @@ open Camlcoq
open AST
open Registers
open Machregs
+open Machregsaux
open Locations
open Conventions1
open XTL
@@ -237,10 +238,6 @@ type graph = {
according to their types. A variable can be forced into class 2
by giving it a negative spill cost. *)
-let class_of_type = function
- | Tint | Tlong -> 0
- | Tfloat | Tsingle -> 1
- | Tany32 | Tany64 -> assert false
let class_of_reg r =
if Conventions1.is_float_reg r then 1 else 0
diff --git a/backend/IRC.mli b/backend/IRC.mli
index 30b6d5c1..f7bbf9c5 100644
--- a/backend/IRC.mli
+++ b/backend/IRC.mli
@@ -43,5 +43,4 @@ val coloring: graph -> (var -> loc)
val reserved_registers: mreg list ref
(* Auxiliaries to deal with register classes *)
-val class_of_type: AST.typ -> int
val class_of_loc: loc -> int
diff --git a/backend/Lineartyping.v b/backend/Lineartyping.v
index 55fa7a67..1fe23a9d 100644
--- a/backend/Lineartyping.v
+++ b/backend/Lineartyping.v
@@ -321,11 +321,11 @@ Local Opaque mreg_type.
+ (* other ops *)
destruct (type_of_operation op) as [ty_args ty_res] eqn:TYOP. InvBooleans.
econstructor; eauto.
- apply wt_setreg; auto. eapply Val.has_subtype; eauto.
+ apply wt_setreg; auto; try (apply wt_undef_regs; auto).
+ eapply Val.has_subtype; eauto.
change ty_res with (snd (ty_args, ty_res)). rewrite <- TYOP. eapply type_of_operation_sound; eauto.
red; intros; subst op. simpl in ISMOVE.
- destruct args; try discriminate. destruct args; discriminate.
- apply wt_undef_regs; auto.
+ destruct args; try discriminate. destruct args; discriminate.
- (* load *)
simpl in *; InvBooleans.
econstructor; eauto.
diff --git a/backend/OpHelpers.v b/backend/OpHelpers.v
new file mode 100644
index 00000000..53414dab
--- /dev/null
+++ b/backend/OpHelpers.v
@@ -0,0 +1,42 @@
+Require Import Coqlib.
+Require Import AST Integers Floats.
+Require Import Op CminorSel.
+
+(** Some arithmetic operations are transformed into calls to
+ runtime library functions. The following type class collects
+ the names of these functions. *)
+
+Definition sig_l_l := mksignature (Tlong :: nil) (Some Tlong) cc_default.
+Definition sig_l_f := mksignature (Tlong :: nil) (Some Tfloat) cc_default.
+Definition sig_l_s := mksignature (Tlong :: nil) (Some Tsingle) cc_default.
+Definition sig_f_l := mksignature (Tfloat :: nil) (Some Tlong) cc_default.
+Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_default.
+Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default.
+Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default.
+Definition sig_ii_i := mksignature (Tint :: Tint :: nil) (Some Tint) cc_default.
+Definition sig_ff_f := mksignature (Tfloat :: Tfloat :: nil) (Some Tfloat) cc_default.
+Definition sig_ss_s := mksignature (Tsingle :: Tsingle :: nil) (Some Tsingle) cc_default.
+
+Class helper_functions := mk_helper_functions {
+ i64_dtos: ident; (**r float64 -> signed long *)
+ i64_dtou: ident; (**r float64 -> unsigned long *)
+ i64_stod: ident; (**r signed long -> float64 *)
+ i64_utod: ident; (**r unsigned long -> float64 *)
+ i64_stof: ident; (**r signed long -> float32 *)
+ i64_utof: ident; (**r unsigned long -> float32 *)
+ i64_sdiv: ident; (**r signed division *)
+ i64_udiv: ident; (**r unsigned division *)
+ i64_smod: ident; (**r signed remainder *)
+ i64_umod: ident; (**r unsigned remainder *)
+ i64_shl: ident; (**r shift left *)
+ i64_shr: ident; (**r shift right unsigned *)
+ i64_sar: ident; (**r shift right signed *)
+ i64_umulh: ident; (**r unsigned multiply high *)
+ i64_smulh: ident; (**r signed multiply high *)
+ i32_sdiv: ident; (**r signed division *)
+ i32_udiv: ident; (**r unsigned division *)
+ i32_smod: ident; (**r signed remainder *)
+ i32_umod: ident; (**r unsigned remainder *)
+ f64_div: ident; (**float division*)
+ f32_div: ident; (**float division*)
+}.
diff --git a/backend/OpHelpersproof.v b/backend/OpHelpersproof.v
new file mode 100644
index 00000000..63040c5f
--- /dev/null
+++ b/backend/OpHelpersproof.v
@@ -0,0 +1,78 @@
+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 Events.
+Require Import OpHelpers.
+
+(** * Axiomatization of the helper functions *)
+
+Definition external_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop :=
+ forall F V (ge: Genv.t F V) m,
+ external_call (EF_runtime name sg) ge vargs m E0 vres m.
+
+Definition builtin_implements (name: string) (sg: signature) (vargs: list val) (vres: val) : Prop :=
+ forall F V (ge: Genv.t F V) m,
+ external_call (EF_builtin name sg) ge vargs m E0 vres m.
+
+Axiom arith_helpers_correct :
+ (forall x z, Val.longoffloat x = Some z -> external_implements "__compcert_i64_dtos" sig_f_l (x::nil) z)
+ /\ (forall x z, Val.longuoffloat x = Some z -> external_implements "__compcert_i64_dtou" sig_f_l (x::nil) z)
+ /\ (forall x z, Val.floatoflong x = Some z -> external_implements "__compcert_i64_stod" sig_l_f (x::nil) z)
+ /\ (forall x z, Val.floatoflongu x = Some z -> external_implements "__compcert_i64_utod" sig_l_f (x::nil) z)
+ /\ (forall x z, Val.singleoflong x = Some z -> external_implements "__compcert_i64_stof" sig_l_s (x::nil) z)
+ /\ (forall x z, Val.singleoflongu x = Some z -> external_implements "__compcert_i64_utof" sig_l_s (x::nil) z)
+ /\ (forall x, builtin_implements "__builtin_negl" sig_l_l (x::nil) (Val.negl x))
+ /\ (forall x y, builtin_implements "__builtin_addl" sig_ll_l (x::y::nil) (Val.addl x y))
+ /\ (forall x y, builtin_implements "__builtin_subl" sig_ll_l (x::y::nil) (Val.subl x y))
+ /\ (forall x y, builtin_implements "__builtin_mull" sig_ii_l (x::y::nil) (Val.mull' x y))
+ /\ (forall x y z, Val.divls x y = Some z -> external_implements "__compcert_i64_sdiv" sig_ll_l (x::y::nil) z)
+ /\ (forall x y z, Val.divlu x y = Some z -> external_implements "__compcert_i64_udiv" sig_ll_l (x::y::nil) z)
+ /\ (forall x y z, Val.modls x y = Some z -> external_implements "__compcert_i64_smod" sig_ll_l (x::y::nil) z)
+ /\ (forall x y z, Val.modlu x y = Some z -> external_implements "__compcert_i64_umod" sig_ll_l (x::y::nil) z)
+ /\ (forall x y, external_implements "__compcert_i64_shl" sig_li_l (x::y::nil) (Val.shll x y))
+ /\ (forall x y, external_implements "__compcert_i64_shr" sig_li_l (x::y::nil) (Val.shrlu x y))
+ /\ (forall x y, external_implements "__compcert_i64_sar" sig_li_l (x::y::nil) (Val.shrl x y))
+ /\ (forall x y, external_implements "__compcert_i64_umulh" sig_ll_l (x::y::nil) (Val.mullhu x y))
+ /\ (forall x y, external_implements "__compcert_i64_smulh" sig_ll_l (x::y::nil) (Val.mullhs x y))
+ /\ (forall x y z, Val.divs x y = Some z -> external_implements "__compcert_i32_sdiv" sig_ii_i (x::y::nil) z)
+ /\ (forall x y z, Val.divu x y = Some z -> external_implements "__compcert_i32_udiv" sig_ii_i (x::y::nil) z)
+ /\ (forall x y z, Val.mods x y = Some z -> external_implements "__compcert_i32_smod" sig_ii_i (x::y::nil) z)
+ /\ (forall x y z, Val.modu x y = Some z -> external_implements "__compcert_i32_umod" sig_ii_i (x::y::nil) z)
+ /\ (forall x y z, Val.divfs x y = z -> external_implements "__compcert_f32_div" sig_ss_s (x::y::nil) z)
+ /\ (forall x y z, Val.divf x y = z -> external_implements "__compcert_f64_div" sig_ff_f (x::y::nil) z)
+.
+
+Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop :=
+ (prog_defmap p)!id = Some (Gfun (External (EF_runtime name sg))).
+
+Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F) V) (hf: helper_functions) : Prop :=
+ helper_declared p i64_dtos "__compcert_i64_dtos" sig_f_l
+ /\ helper_declared p i64_dtou "__compcert_i64_dtou" sig_f_l
+ /\ helper_declared p i64_stod "__compcert_i64_stod" sig_l_f
+ /\ helper_declared p i64_utod "__compcert_i64_utod" sig_l_f
+ /\ helper_declared p i64_stof "__compcert_i64_stof" sig_l_s
+ /\ helper_declared p i64_utof "__compcert_i64_utof" sig_l_s
+ /\ helper_declared p i64_sdiv "__compcert_i64_sdiv" sig_ll_l
+ /\ helper_declared p i64_udiv "__compcert_i64_udiv" sig_ll_l
+ /\ helper_declared p i64_smod "__compcert_i64_smod" sig_ll_l
+ /\ helper_declared p i64_umod "__compcert_i64_umod" sig_ll_l
+ /\ helper_declared p i64_shl "__compcert_i64_shl" sig_li_l
+ /\ helper_declared p i64_shr "__compcert_i64_shr" sig_li_l
+ /\ helper_declared p i64_sar "__compcert_i64_sar" sig_li_l
+ /\ helper_declared p i64_umulh "__compcert_i64_umulh" sig_ll_l
+ /\ helper_declared p i64_smulh "__compcert_i64_smulh" sig_ll_l
+ /\ helper_declared p i32_sdiv "__compcert_i32_sdiv" sig_ii_i
+ /\ helper_declared p i32_udiv "__compcert_i32_udiv" sig_ii_i
+ /\ helper_declared p i32_smod "__compcert_i32_smod" sig_ii_i
+ /\ helper_declared p i32_umod "__compcert_i32_umod" sig_ii_i
+ /\ helper_declared p f32_div "__compcert_f32_div" sig_ss_s
+ /\ helper_declared p f64_div "__compcert_f64_div" sig_ff_f
+.
diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml
index 19aba4f6..7db8a866 100644
--- a/backend/Regalloc.ml
+++ b/backend/Regalloc.ml
@@ -1067,7 +1067,7 @@ let make_parmove srcs dsts itmp ftmp k =
| Locations.S(sl, ofs, ty), R rd ->
code := LTL.Lgetstack(sl, ofs, ty, rd) :: !code
| Locations.S(sls, ofss, tys), Locations.S(sld, ofsd, tyd) ->
- let tmp = temp_for (class_of_type tys) in
+ let tmp = temp_for (Machregsaux.class_of_type tys) in
(* code will be reversed at the end *)
code := LTL.Lsetstack(tmp, sld, ofsd, tyd) ::
LTL.Lgetstack(sls, ofss, tys, tmp) :: !code
diff --git a/backend/SelectDiv.vp b/backend/SelectDiv.vp
index d91797c5..9f852616 100644
--- a/backend/SelectDiv.vp
+++ b/backend/SelectDiv.vp
@@ -15,10 +15,13 @@
Require Import Coqlib.
Require Import Compopts.
Require Import AST Integers Floats.
-Require Import Op CminorSel SelectOp SplitLong SelectLong.
+Require Import Op CminorSel OpHelpers SelectOp SplitLong SelectLong.
Local Open Scope cminorsel_scope.
+Section SELECT.
+Context {hf: helper_functions}.
+
Definition is_intconst (e: expr) : option int :=
match e with
| Eop (Ointconst n) _ => Some n
@@ -221,10 +224,6 @@ Definition mods (e1: expr) (e2: expr) :=
(** 64-bit integer divisions *)
-Section SELECT.
-
-Context {hf: helper_functions}.
-
Definition modl_from_divl (equo: expr) (n: int64) :=
subl (Eletvar O) (mullimm n equo).
@@ -241,8 +240,8 @@ Definition divlu (e1 e2: expr) :=
divlu_base e1 e2
else
match divlu_mul_params (Int64.unsigned n2) with
- | None => divlu_base e1 e2
- | Some(p, m) => Elet e1 (divlu_mull p m)
+ | _ => divlu_base e1 e2
+ (* | Some(p, m) => Elet e1 (divlu_mull p m) *) (* FIXME - hack K1 *)
end
end
| _, _ => divlu_base e1 e2
@@ -258,8 +257,8 @@ Definition modlu (e1 e2: expr) :=
modlu_base e1 e2
else
match divlu_mul_params (Int64.unsigned n2) with
- | None => modlu_base e1 e2
- | Some(p, m) => Elet e1 (modl_from_divl (divlu_mull p m) n2)
+ | _ => modlu_base e1 e2
+ (* | Some(p, m) => Elet e1 (modl_from_divl (divlu_mull p m) n2) *) (* FIXME - hack K1 *)
end
end
| _, _ => modlu_base e1 e2
@@ -285,8 +284,8 @@ Definition divls (e1 e2: expr) :=
divls_base e1 e2
else
match divls_mul_params (Int64.signed n2) with
- | None => divls_base e1 e2
- | Some(p, m) => Elet e1 (divls_mull p m)
+ | _ => divls_base e1 e2
+ (* | Some(p, m) => Elet e1 (divls_mull p m) *) (* FIXME - hack K1 *)
end
end
| _, _ => divls_base e1 e2
@@ -304,40 +303,38 @@ Definition modls (e1 e2: expr) :=
modls_base e1 e2
else
match divls_mul_params (Int64.signed n2) with
- | None => modls_base e1 e2
- | Some(p, m) => Elet e1 (modl_from_divl (divls_mull p m) n2)
+ | _ => modls_base e1 e2
+ (* | Some(p, m) => Elet e1 (modl_from_divl (divls_mull p m) n2) *) (* FIXME - hack K1 *)
end
end
| _, _ => modls_base e1 e2
end.
-
-End SELECT.
-
+
(** Floating-point division by a constant can also be turned into a FP
multiplication by the inverse constant, but only for powers of 2. *)
Definition divfimm (e: expr) (n: float) :=
match Float.exact_inverse n with
| Some n' => Eop Omulf (e ::: Eop (Ofloatconst n') Enil ::: Enil)
- | None => Eop Odivf (e ::: Eop (Ofloatconst n) Enil ::: Enil)
+ | None => divf_base e (Eop (Ofloatconst n) Enil)
end.
Nondetfunction divf (e1: expr) (e2: expr) :=
match e2 with
| Eop (Ofloatconst n2) Enil => divfimm e1 n2
- | _ => Eop Odivf (e1 ::: e2 ::: Enil)
+ | _ => divf_base e1 e2
end.
Definition divfsimm (e: expr) (n: float32) :=
match Float32.exact_inverse n with
| Some n' => Eop Omulfs (e ::: Eop (Osingleconst n') Enil ::: Enil)
- | None => Eop Odivfs (e ::: Eop (Osingleconst n) Enil ::: Enil)
+ | None => divfs_base e (Eop (Osingleconst n) Enil)
end.
Nondetfunction divfs (e1: expr) (e2: expr) :=
match e2 with
| Eop (Osingleconst n2) Enil => divfsimm e1 n2
- | _ => Eop Odivfs (e1 ::: e2 ::: Enil)
+ | _ => divfs_base e1 e2
end.
-
+End SELECT. \ No newline at end of file
diff --git a/backend/SelectDivproof.v b/backend/SelectDivproof.v
index f4ff2c86..a8ee8453 100644
--- a/backend/SelectDivproof.v
+++ b/backend/SelectDivproof.v
@@ -15,6 +15,7 @@
Require Import Zquot Coqlib Zbits.
Require Import AST Integers Floats Values Memory Globalenvs Events.
Require Import Cminor Op CminorSel.
+Require Import OpHelpers OpHelpersproof.
Require Import SelectOp SelectOpproof SplitLong SplitLongproof SelectLong SelectLongproof SelectDiv.
Local Open Scope cminorsel_scope.
@@ -587,12 +588,12 @@ Proof.
- destruct (Compopts.optim_for_size tt).
+ eapply eval_modu_base; eauto. EvalOp.
+ destruct (divu_mul_params (Int.unsigned n2)) as [[p M] | ] eqn:PARAMS.
- * econstructor; split.
+ * econstructor; split.
econstructor; eauto. eapply eval_mod_from_div.
eapply eval_divu_mul; eauto. simpl; eauto. simpl; eauto.
rewrite Int.modu_divu. auto.
red; intros; subst n2; discriminate.
- * eapply eval_modu_base; eauto. EvalOp.
+ * eapply eval_modu_base; eauto. EvalOp.
Qed.
Theorem eval_modu:
@@ -704,7 +705,7 @@ Proof.
|| Int.eq i (Int.repr Int.min_signed) && Int.eq n2 Int.mone) eqn:Z2; inv DIV.
destruct (Int.is_power2 n2) as [l | ] eqn:P2.
- destruct (Int.ltu l (Int.repr 31)) eqn:LT31.
- + exploit (eval_shrximm ge sp e m (Vint i :: le) (Eletvar O)).
+ + exploit (eval_shrximm prog sp e m (Vint i :: le) (Eletvar O)).
constructor. simpl; eauto. eapply Val.divs_pow2; eauto.
intros [v1 [X LD]]. inv LD.
econstructor; split. econstructor. eauto.
@@ -788,10 +789,10 @@ Proof.
+ destruct (Int64.is_power2' n2) as [l|] eqn:POW.
* exploit Val.divlu_pow2; eauto. intros EQ; subst z. apply eval_shrluimm; auto.
* destruct (Compopts.optim_for_size tt). eapply eval_divlu_base; eauto.
- destruct (divlu_mul_params (Int64.unsigned n2)) as [[p M]|] eqn:PARAMS.
+ (* destruct (divlu_mul_params (Int64.unsigned n2)) as [[p M]|] eqn:PARAMS.
** destruct x; simpl in H1; try discriminate.
destruct (Int64.eq n2 Int64.zero); inv H1.
- econstructor; split; eauto. econstructor. eauto. eapply eval_divlu_mull; eauto.
+ econstructor; split; eauto. econstructor. eauto. eapply eval_divlu_mull; eauto. *) (* FIXME - K1 hack *)
** eapply eval_divlu_base; eauto.
- eapply eval_divlu_base; eauto.
Qed.
@@ -813,14 +814,14 @@ Proof.
+ destruct (Int64.is_power2 n2) as [l|] eqn:POW.
* exploit Val.modlu_pow2; eauto. intros EQ; subst z. eapply eval_andl; eauto. apply eval_longconst.
* destruct (Compopts.optim_for_size tt). eapply eval_modlu_base; eauto.
- destruct (divlu_mul_params (Int64.unsigned n2)) as [[p M]|] eqn:PARAMS.
+ (* destruct (divlu_mul_params (Int64.unsigned n2)) as [[p M]|] eqn:PARAMS.
** destruct x; simpl in H1; try discriminate.
destruct (Int64.eq n2 Int64.zero) eqn:Z; inv H1.
rewrite Int64.modu_divu.
econstructor; split; eauto. econstructor. eauto.
eapply eval_modl_from_divl; eauto.
eapply eval_divlu_mull; eauto.
- red; intros; subst n2; discriminate Z.
+ red; intros; subst n2; discriminate Z. *)
** eapply eval_modlu_base; eauto.
- eapply eval_modlu_base; eauto.
Qed.
@@ -883,12 +884,12 @@ Proof.
** exploit Val.divls_pow2; eauto. intros EQ. eapply eval_shrxlimm; eauto.
** eapply eval_divls_base; eauto.
* destruct (Compopts.optim_for_size tt). eapply eval_divls_base; eauto.
- destruct (divls_mul_params (Int64.signed n2)) as [[p M]|] eqn:PARAMS.
+ (* destruct (divls_mul_params (Int64.signed n2)) as [[p M]|] eqn:PARAMS.
** destruct x; simpl in H1; try discriminate.
destruct (Int64.eq n2 Int64.zero
|| Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq n2 Int64.mone); inv H1.
econstructor; split; eauto. econstructor. eauto.
- eapply eval_divls_mull; eauto.
+ eapply eval_divls_mull; eauto. *)
** eapply eval_divls_base; eauto.
- eapply eval_divls_base; eauto.
Qed.
@@ -925,14 +926,14 @@ Proof.
rewrite Int64.mods_divs. auto.
**eapply eval_modls_base; eauto.
* destruct (Compopts.optim_for_size tt). eapply eval_modls_base; eauto.
- destruct (divls_mul_params (Int64.signed n2)) as [[p M]|] eqn:PARAMS.
+ (* destruct (divls_mul_params (Int64.signed n2)) as [[p M]|] eqn:PARAMS.
** destruct x; simpl in H1; try discriminate.
destruct (Int64.eq n2 Int64.zero
|| Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq n2 Int64.mone); inv H1.
econstructor; split; eauto. econstructor. eauto.
rewrite Int64.mods_divs.
eapply eval_modl_from_divl; auto.
- eapply eval_divls_mull; eauto.
+ eapply eval_divls_mull; eauto. *)
** eapply eval_modls_base; eauto.
- eapply eval_modls_base; eauto.
Qed.
@@ -951,8 +952,8 @@ Proof.
EvalOp. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
simpl; eauto.
destruct x; simpl; auto. erewrite Float.div_mul_inverse; eauto.
- + TrivialExists.
-- TrivialExists.
+ + apply eval_divf_base; trivial.
+- apply eval_divf_base; trivial.
Qed.
Theorem eval_divfs:
@@ -967,8 +968,8 @@ Proof.
EvalOp. constructor. eauto. constructor. EvalOp. simpl; eauto. constructor.
simpl; eauto.
destruct x; simpl; auto. erewrite Float32.div_mul_inverse; eauto.
- + TrivialExists.
-- TrivialExists.
+ + apply eval_divfs_base; trivial.
+- apply eval_divfs_base; trivial.
Qed.
End CMCONSTRS.
diff --git a/backend/Selection.v b/backend/Selection.v
index c9771d99..4ab3331e 100644
--- a/backend/Selection.v
+++ b/backend/Selection.v
@@ -26,7 +26,7 @@ Require String.
Require Import Coqlib Maps.
Require Import AST Errors Integers Globalenvs Builtins Switch.
Require Cminor.
-Require Import Op CminorSel Cminortyping.
+Require Import Op CminorSel OpHelpers Cminortyping.
Require Import SelectOp SplitLong SelectLong SelectDiv.
Require Machregs.
@@ -507,11 +507,19 @@ Definition get_helpers (defmap: PTree.t globdef) : res helper_functions :=
do i64_sar <- lookup_helper globs "__compcert_i64_sar" sig_li_l ;
do i64_umulh <- lookup_helper globs "__compcert_i64_umulh" sig_ll_l ;
do i64_smulh <- lookup_helper globs "__compcert_i64_smulh" sig_ll_l ;
+ do i32_sdiv <- lookup_helper globs "__compcert_i32_sdiv" sig_ii_i ;
+ do i32_udiv <- lookup_helper globs "__compcert_i32_udiv" sig_ii_i ;
+ do i32_smod <- lookup_helper globs "__compcert_i32_smod" sig_ii_i ;
+ do i32_umod <- lookup_helper globs "__compcert_i32_umod" sig_ii_i ;
+ do f64_div <- lookup_helper globs "__compcert_f64_div" sig_ff_f ;
+ do f32_div <- lookup_helper globs "__compcert_f32_div" sig_ss_s ;
OK (mk_helper_functions
i64_dtos i64_dtou i64_stod i64_utod i64_stof i64_utof
i64_sdiv i64_udiv i64_smod i64_umod
i64_shl i64_shr i64_sar
- i64_umulh i64_smulh).
+ i64_umulh i64_smulh
+ i32_sdiv i32_udiv i32_smod i32_umod
+ f64_div f32_div).
(** Conversion of programs. *)
diff --git a/backend/Selectionaux.ml b/backend/Selectionaux.ml
index 4ca7dd21..60a1eccd 100644
--- a/backend/Selectionaux.ml
+++ b/backend/Selectionaux.ml
@@ -71,13 +71,13 @@ let fast_cmove ty =
| "arm", _ ->
(match ty with Tint | Tfloat | Tsingle -> true | _ -> false)
| "powerpc", "e5500" ->
- (match ty with Tint -> true | Tlong -> true | _ -> false)
+ (match ty with Tint | Tlong -> true | _ -> false)
| "powerpc", _ -> false
| "riscV", _ -> false
| "x86", _ ->
- (match ty with Tint -> true | Tlong -> Archi.ptr64 | _ -> false)
- | _, _ ->
- assert false
+ (match ty with Tint -> true | Tlong -> Archi.ptr64 | _ -> false)
+ | "mppa_k1c", _ -> true
+ | a, m -> failwith (Printf.sprintf "fast_cmove: unknown arch %s %s" a m)
(* The if-conversion heuristic depend on the
-fif-conversion and -Obranchless flags.
@@ -96,12 +96,17 @@ If-conversion seems beneficial if:
Intuition: on a modern processor, the "then" and the "else" branches
can generally be computed in parallel, there is enough ILP for that.
So, the bad case is if the most taken branch is much cheaper than the
+<<<<<<< HEAD
+other branch. Since our cost estimates are very imprecise, the
+bound on the total cost acts as a safety guard,
+=======
other branch. Another bad case is if both branches are big: since the
code for one branch precedes entirely the code for the other branch,
if the first branch contains a lot of instructions,
dynamic reordering of instructions will not look ahead far enough
to execute instructions from the other branch in parallel with
instructions from the first branch.
+>>>>>>> ddb2c968e6c57d2117434f169471d87f643d831a
*)
let if_conversion_heuristic cond ifso ifnot ty =
diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v
index ee3ed358..622992e5 100644
--- a/backend/Selectionproof.v
+++ b/backend/Selectionproof.v
@@ -17,6 +17,7 @@ Require Import Coqlib Maps.
Require Import AST Linking Errors Integers.
Require Import Values Memory Builtins Events Globalenvs Smallstep.
Require Import Switch Cminor Op CminorSel Cminortyping.
+Require Import OpHelpers OpHelpersproof.
Require Import SelectOp SelectDiv SplitLong SelectLong Selection.
Require Import SelectOpproof SelectDivproof SplitLongproof SelectLongproof.
@@ -87,7 +88,7 @@ Lemma get_helpers_correct:
forall p hf,
get_helpers (prog_defmap p) = OK hf -> helper_functions_declared p hf.
Proof.
- intros. monadInv H. red; simpl. auto 20 using lookup_helper_correct.
+ intros. monadInv H. red; simpl. auto 22 using lookup_helper_correct.
Qed.
Theorem transf_program_match:
@@ -107,7 +108,7 @@ Proof.
{ unfold helper_declared; intros.
destruct (prog_defmap_linkorder _ _ _ _ H0 H1) as (gd & P & Q).
inv Q. inv H3. auto. }
- red in H. decompose [Logic.and] H; clear H. red; auto 20.
+ red in H. decompose [Logic.and] H; clear H. red; auto 22.
Qed.
(** * Correctness of the instruction selection functions for expressions *)
@@ -175,7 +176,7 @@ Proof.
generalize (match_program_defmap _ _ _ _ _ TRANSF id).
unfold Cminor.fundef; rewrite H; intros R; inv R. inv H2.
destruct H4 as (cu & A & B). monadInv B. auto. }
- unfold helper_functions_declared; intros. decompose [Logic.and] H; clear H. auto 20.
+ unfold helper_functions_declared; intros. decompose [Logic.and] H; clear H. auto 22.
Qed.
Section CMCONSTR.
@@ -1186,7 +1187,7 @@ Remark find_label_commut:
Proof.
induction s; intros until k'; simpl; intros MC SE; try (monadInv SE); simpl; auto.
(* store *)
- unfold store. destruct (addressing m (sel_expr e)); simpl; auto.
+- unfold store. destruct (addressing m (sel_expr e)); simpl; auto.
(* call *)
destruct (classify_call (prog_defmap cunit) e); simpl; auto.
rewrite sel_builtin_nolabel; auto.
@@ -1195,12 +1196,12 @@ Proof.
(* builtin *)
rewrite sel_builtin_nolabel; auto.
(* seq *)
- exploit (IHs1 (Cminor.Kseq s2 k)). constructor; eauto. eauto.
+- exploit (IHs1 (Cminor.Kseq s2 k)). constructor; eauto. eauto.
destruct (Cminor.find_label lbl s1 (Cminor.Kseq s2 k)) as [[sx kx] | ];
destruct (find_label lbl x (Kseq x0 k')) as [[sy ky] | ];
intuition. apply IHs2; auto.
(* ifthenelse *)
- destruct (if_conversion ki env e s1 s2) as [s|] eqn:IFC.
+- destruct (if_conversion ki env e s1 s2) as [s|] eqn:IFC.
inv SE. exploit if_conversion_nolabel; eauto. intros (A & B & C).
rewrite A, B, C. auto.
monadInv SE; simpl.
@@ -1209,19 +1210,19 @@ Proof.
destruct (find_label lbl x k') as [[sy ky] | ];
intuition. apply IHs2; auto.
(* loop *)
- apply IHs. constructor; auto. simpl; rewrite EQ; auto. auto.
+- apply IHs. constructor; auto. simpl; rewrite EQ; auto. auto.
(* block *)
- apply IHs. constructor; auto. auto.
+- apply IHs. constructor; auto. auto.
(* switch *)
- destruct b.
+- destruct b.
destruct (validate_switch Int64.modulus n l (compile_switch Int64.modulus n l)); inv SE.
simpl; auto.
destruct (validate_switch Int.modulus n l (compile_switch Int.modulus n l)); inv SE.
simpl; auto.
(* return *)
- destruct o; inv SE; simpl; auto.
+- destruct o; inv SE; simpl; auto.
(* label *)
- destruct (ident_eq lbl l). auto. apply IHs; auto.
+- destruct (ident_eq lbl l). auto. apply IHs; auto.
Qed.
Definition measure (s: Cminor.state) : nat :=
diff --git a/backend/SplitLong.vp b/backend/SplitLong.vp
index de954482..dfe42df0 100644
--- a/backend/SplitLong.vp
+++ b/backend/SplitLong.vp
@@ -16,41 +16,12 @@ Require String.
Require Import Coqlib.
Require Import AST Integers Floats.
Require Import Op CminorSel.
+Require Import OpHelpers.
Require Import SelectOp.
Local Open Scope cminorsel_scope.
Local Open Scope string_scope.
-(** Some operations on 64-bit integers are transformed into calls to
- runtime library functions. The following type class collects
- the names of these functions. *)
-
-Class helper_functions := mk_helper_functions {
- i64_dtos: ident; (**r float64 -> signed long *)
- i64_dtou: ident; (**r float64 -> unsigned long *)
- i64_stod: ident; (**r signed long -> float64 *)
- i64_utod: ident; (**r unsigned long -> float64 *)
- i64_stof: ident; (**r signed long -> float32 *)
- i64_utof: ident; (**r unsigned long -> float32 *)
- i64_sdiv: ident; (**r signed division *)
- i64_udiv: ident; (**r unsigned division *)
- i64_smod: ident; (**r signed remainder *)
- i64_umod: ident; (**r unsigned remainder *)
- i64_shl: ident; (**r shift left *)
- i64_shr: ident; (**r shift right unsigned *)
- i64_sar: ident; (**r shift right signed *)
- i64_umulh: ident; (**r unsigned multiply high *)
- i64_smulh: ident; (**r signed multiply high *)
-}.
-
-Definition sig_l_l := mksignature (Tlong :: nil) (Some Tlong) cc_default.
-Definition sig_l_f := mksignature (Tlong :: nil) (Some Tfloat) cc_default.
-Definition sig_l_s := mksignature (Tlong :: nil) (Some Tsingle) cc_default.
-Definition sig_f_l := mksignature (Tfloat :: nil) (Some Tlong) cc_default.
-Definition sig_ll_l := mksignature (Tlong :: Tlong :: nil) (Some Tlong) cc_default.
-Definition sig_li_l := mksignature (Tlong :: Tint :: nil) (Some Tlong) cc_default.
-Definition sig_ii_l := mksignature (Tint :: Tint :: nil) (Some Tlong) cc_default.
-
Section SELECT.
Context {hf: helper_functions}.
diff --git a/backend/SplitLongproof.v b/backend/SplitLongproof.v
index 18c1f18d..df77b322 100644
--- a/backend/SplitLongproof.v
+++ b/backend/SplitLongproof.v
@@ -16,11 +16,15 @@ Require Import String.
Require Import Coqlib Maps.
Require Import AST Errors Integers Floats.
Require Import Values Memory Globalenvs Builtins Events Cminor Op CminorSel.
+Require Import OpHelpers OpHelpersproof.
+Require Import Values Memory Globalenvs Builtins Events Cminor Op CminorSel.
Require Import SelectOp SelectOpproof SplitLong.
Local Open Scope cminorsel_scope.
Local Open Scope string_scope.
+<<<<<<< HEAD
+=======
(** * Properties of the helper functions *)
Definition helper_declared {F V: Type} (p: AST.program (AST.fundef F) V) (id: ident) (name: string) (sg: signature) : Prop :=
@@ -43,6 +47,7 @@ Definition helper_functions_declared {F V: Type} (p: AST.program (AST.fundef F)
/\ helper_declared p i64_umulh "__compcert_i64_umulh" sig_ll_l
/\ helper_declared p i64_smulh "__compcert_i64_smulh" sig_ll_l.
+>>>>>>> 91381b65f5aa76e5195caae9ef331b3f5f95afaf
(** * Correctness of the instruction selection functions for 64-bit operators *)
Section CMCONSTR.
@@ -55,6 +60,10 @@ Variable sp: val.
Variable e: env.
Variable m: mem.
+<<<<<<< HEAD
+Ltac UseHelper := decompose [Logic.and] arith_helpers_correct; eauto.
+=======
+>>>>>>> 91381b65f5aa76e5195caae9ef331b3f5f95afaf
Ltac DeclHelper := red in HELPERS; decompose [Logic.and] HELPERS; eauto.
Lemma eval_helper:
@@ -342,7 +351,7 @@ Theorem eval_longofint: unary_constructor_sound longofint Val.longofint.
Proof.
red; intros. unfold longofint. destruct (longofint_match a).
- InvEval. econstructor; split. apply eval_longconst. auto.
-- exploit (eval_shrimm ge sp e m (Int.repr 31) (x :: le) (Eletvar 0)). EvalOp.
+- exploit (eval_shrimm prog sp e m (Int.repr 31) (x :: le) (Eletvar 0)). EvalOp.
intros [v1 [A B]].
econstructor; split. EvalOp.
destruct x; simpl; auto.
diff --git a/backend/Stackingproof.v b/backend/Stackingproof.v
index ffd9b227..326fab61 100644
--- a/backend/Stackingproof.v
+++ b/backend/Stackingproof.v
@@ -1891,12 +1891,13 @@ Proof.
apply plus_one. econstructor.
instantiate (1 := v'). rewrite <- A. apply eval_operation_preserved.
exact symbols_preserved. eauto.
- econstructor; eauto with coqlib.
- apply agree_regs_set_reg; auto.
- rewrite transl_destroyed_by_op. apply agree_regs_undef_regs; auto.
- apply agree_locs_set_reg; auto. apply agree_locs_undef_locs. auto. apply destroyed_by_op_caller_save.
- apply frame_set_reg. apply frame_undef_regs. exact SEP.
-
+ econstructor; eauto with coqlib;
+ try (apply agree_regs_set_reg; auto);
+ (* generic proof *)
+ solve [
+ (rewrite transl_destroyed_by_op; apply agree_regs_undef_regs; auto) |
+ (apply agree_locs_set_reg; auto; apply agree_locs_undef_locs; auto; apply destroyed_by_op_caller_save) |
+ (apply frame_set_reg; apply frame_undef_regs; exact SEP) ].
- (* Lload *)
assert (exists a',
eval_addressing ge (Vptr sp' Ptrofs.zero) (transl_addr (make_env (function_bounds f)) addr) rs0##args = Some a'