aboutsummaryrefslogtreecommitdiffstats
path: root/backend/LTL.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2016-05-17 15:37:56 +0200
committerXavier Leroy <xavier.leroy@inria.fr>2016-05-17 15:37:56 +0200
commit82f9d1f96b30106a338e77ec83b7321c2c65f929 (patch)
tree6b8bb30473b1385f8b84fe1600f592c2bd4abed7 /backend/LTL.v
parent672393ef623acb3e230a8019d51c87e051a7567a (diff)
downloadcompcert-kvx-82f9d1f96b30106a338e77ec83b7321c2c65f929.tar.gz
compcert-kvx-82f9d1f96b30106a338e77ec83b7321c2c65f929.zip
Introduce register pairs to describe calling conventions more precisely
This commit changes the loc_arguments and loc_result functions that describe calling conventions so that each argument/result can be mapped either to a single location or (in the case of a 64-bit integer) to a pair of two 32-bit locations. In the current CompCert, all arguments/results of type Tlong are systematically split in two 32-bit halves. We will need to change this in the future to support 64-bit processors. The alternative approach implemented by this commit enables the loc_arguments and loc_result functions to describe precisely which arguments need splitting. Eventually, the remainder of CompCert should not assume anything about splitting 64-bit types in two halves. Summary of changes: - AST: introduce the type "rpair A" of register pairs - Conventions1, Conventions: use it when describing calling conventions - LTL, Linear, Mach, Asm: honor the new calling conventions when observing external calls - Events: suppress external_call', no longer useful - All passes from Allocation to Asmgen: adapt accordingly.
Diffstat (limited to 'backend/LTL.v')
-rw-r--r--backend/LTL.v11
1 files changed, 5 insertions, 6 deletions
diff --git a/backend/LTL.v b/backend/LTL.v
index bb596fa2..5f7116ae 100644
--- a/backend/LTL.v
+++ b/backend/LTL.v
@@ -266,9 +266,9 @@ Inductive step: state -> trace -> state -> Prop :=
step (Callstate s (Internal f) rs m)
E0 (State s f (Vptr sp Int.zero) f.(fn_entrypoint) rs' m')
| exec_function_external: forall s ef t args res rs m rs' m',
- args = map rs (loc_arguments (ef_sig ef)) ->
- external_call' ef ge args m t res m' ->
- rs' = Locmap.setlist (map R (loc_result (ef_sig ef))) res rs ->
+ args = map (fun p => Locmap.getpair p rs) (loc_arguments (ef_sig ef)) ->
+ external_call ef ge args m t res m' ->
+ rs' = Locmap.setpair (loc_result (ef_sig ef)) res rs ->
step (Callstate s (External ef) rs m)
t (Returnstate s rs' m')
| exec_return: forall f sp rs1 bb s rs m,
@@ -292,9 +292,8 @@ Inductive initial_state (p: program): state -> Prop :=
initial_state p (Callstate nil f (Locmap.init Vundef) m0).
Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall rs m r retcode,
- loc_result signature_main = r :: nil ->
- rs (R r) = Vint retcode ->
+ | final_state_intro: forall rs m retcode,
+ Locmap.getpair (map_rpair R (loc_result signature_main)) rs = Vint retcode ->
final_state (Returnstate nil rs m) retcode.
Definition semantics (p: program) :=