aboutsummaryrefslogtreecommitdiffstats
path: root/backend/Mach.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2015-08-21 11:05:36 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2015-08-21 11:05:36 +0200
commit54f97d1988f623ba7422e13a504caeb5701ba93c (patch)
treed8dea46837352979490f4ed4516f9649fef41b98 /backend/Mach.v
parentb4846ffadfa3fbb73ffa7d9c43e5218adeece8da (diff)
downloadcompcert-kvx-54f97d1988f623ba7422e13a504caeb5701ba93c.tar.gz
compcert-kvx-54f97d1988f623ba7422e13a504caeb5701ba93c.zip
Refactoring of builtins and annotations in the back-end.
Before, the back-end languages had distinct instructions - Iannot for annotations, taking structured expressions (annot_arg) as arguments, and producing no results' - Ibuiltin for other builtins, using simple pseudoregs/locations/registers as arguments and results. This branch enriches Ibuiltin instructions so that they take structured expressions (builtin_arg and builtin_res) as arguments and results. This way, - Annotations fit the general pattern of builtin functions, so Iannot instructions are removed. - EF_vload_global and EF_vstore_global become useless, as the same optimization can be achieved by EF_vload/vstore taking a structured argument of the "address of global" kind. - Better code can be generated for builtin_memcpy between stack locations, or volatile accesses to stack locations. Finally, this commit also introduces a new kind of external function, EF_debug, which is like EF_annot but produces no observable events. It will be used later to transport debug info through the back-end, without preventing optimizations.
Diffstat (limited to 'backend/Mach.v')
-rw-r--r--backend/Mach.v23
1 files changed, 12 insertions, 11 deletions
diff --git a/backend/Mach.v b/backend/Mach.v
index fe00d42d..08fe7c0a 100644
--- a/backend/Mach.v
+++ b/backend/Mach.v
@@ -60,8 +60,7 @@ Inductive instruction: Type :=
| Mstore: memory_chunk -> addressing -> list mreg -> mreg -> instruction
| Mcall: signature -> mreg + ident -> instruction
| Mtailcall: signature -> mreg + ident -> instruction
- | Mbuiltin: external_function -> list mreg -> list mreg -> instruction
- | Mannot: external_function -> list (annot_arg mreg) -> instruction
+ | Mbuiltin: external_function -> list (builtin_arg mreg) -> builtin_res mreg -> instruction
| Mlabel: label -> instruction
| Mgoto: label -> instruction
| Mcond: condition -> list mreg -> label -> instruction
@@ -163,6 +162,13 @@ Fixpoint set_regs (rl: list mreg) (vl: list val) (rs: regset) : regset :=
| _, _ => rs
end.
+Fixpoint set_res (res: builtin_res mreg) (v: val) (rs: regset) : regset :=
+ match res with
+ | BR r => Regmap.set r v rs
+ | BR_none => rs
+ | BR_longofwords hi lo => set_res lo (Val.loword v) (set_res hi (Val.hiword v) rs)
+ end.
+
Definition is_label (lbl: label) (instr: instruction) : bool :=
match instr with
| Mlabel lbl' => if peq lbl lbl' then true else false
@@ -328,17 +334,12 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s fb (Vptr stk soff) (Mtailcall sig ros :: c) rs m)
E0 (Callstate s f' rs m')
| exec_Mbuiltin:
- forall s f sp rs m ef args res b t vl rs' m',
- external_call' ef ge rs##args m t vl m' ->
- rs' = set_regs res vl (undef_regs (destroyed_by_builtin ef) rs) ->
+ forall s f sp rs m ef args res b vargs t vres rs' m',
+ eval_builtin_args ge rs sp m args vargs ->
+ external_call ef ge vargs m t vres m' ->
+ rs' = set_res res vres (undef_regs (destroyed_by_builtin ef) rs) ->
step (State s f sp (Mbuiltin ef args res :: b) rs m)
t (State s f sp b rs' m')
- | exec_Mannot:
- forall s f sp rs m ef args b vargs t v m',
- eval_annot_args ge rs sp m args vargs ->
- external_call ef ge vargs m t v m' ->
- step (State s f sp (Mannot ef args :: b) rs m)
- t (State s f sp b rs m')
| exec_Mgoto:
forall s fb f sp lbl c rs m c',
Genv.find_funct_ptr ge fb = Some (Internal f) ->