aboutsummaryrefslogtreecommitdiffstats
path: root/backend/OpHelpers.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-03-22 08:12:41 +0100
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2019-03-22 08:12:41 +0100
commit5a425ba34f3f2520de752ea95a4636a9437c312a (patch)
tree91a1ef113e01f6e7a1006e8ef4e84e9d1f6bab04 /backend/OpHelpers.v
parentba2b825a42807102a93cd6df1ce90b41d415569c (diff)
downloadcompcert-kvx-5a425ba34f3f2520de752ea95a4636a9437c312a.tar.gz
compcert-kvx-5a425ba34f3f2520de752ea95a4636a9437c312a.zip
ça recompile sur x86
Diffstat (limited to 'backend/OpHelpers.v')
-rw-r--r--backend/OpHelpers.v42
1 files changed, 42 insertions, 0 deletions
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*)
+}.