aboutsummaryrefslogtreecommitdiffstats
path: root/riscV/Builtins1.v
diff options
context:
space:
mode:
authorCyril SIX <cyril.six@kalray.eu>2019-09-03 10:52:23 +0200
committerCyril SIX <cyril.six@kalray.eu>2019-09-03 10:52:23 +0200
commit0cb5a0b65b4fbeb5bc1c14f75951798f20500177 (patch)
tree767381d2490c86dcee95da2631ac5c94e14de8f5 /riscV/Builtins1.v
parent1fbd5d18a9f4398d7ecb9b9ab148a96f575fd1e0 (diff)
parent2f7f68f69b6408e4de6210c827b108eff011af51 (diff)
downloadcompcert-kvx-0cb5a0b65b4fbeb5bc1c14f75951798f20500177.tar.gz
compcert-kvx-0cb5a0b65b4fbeb5bc1c14f75951798f20500177.zip
Merge remote-tracking branch 'origin/mppa-work-upstream-merge' into mppa-work
Conflicts: configure mppa_k1c/Archi.v mppa_k1c/Asmexpand.ml
Diffstat (limited to 'riscV/Builtins1.v')
-rw-r--r--riscV/Builtins1.v33
1 files changed, 33 insertions, 0 deletions
diff --git a/riscV/Builtins1.v b/riscV/Builtins1.v
new file mode 100644
index 00000000..f6e643d2
--- /dev/null
+++ b/riscV/Builtins1.v
@@ -0,0 +1,33 @@
+(* *********************************************************************)
+(* *)
+(* The Compcert verified compiler *)
+(* *)
+(* Xavier Leroy, Collège de France and 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. *)
+(* *)
+(* *********************************************************************)
+
+(** Platform-specific built-in functions *)
+
+Require Import String Coqlib.
+Require Import AST Integers Floats Values.
+Require Import Builtins0.
+
+Inductive platform_builtin : Type := .
+
+Local Open Scope string_scope.
+
+Definition platform_builtin_table : list (string * platform_builtin) :=
+ nil.
+
+Definition platform_builtin_sig (b: platform_builtin) : signature :=
+ match b with end.
+
+Definition platform_builtin_sem (b: platform_builtin) : builtin_sem (proj_sig_res (platform_builtin_sig b)) :=
+ match b with end.