From 82f9d1f96b30106a338e77ec83b7321c2c65f929 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 17 May 2016 15:37:56 +0200 Subject: 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. --- backend/Linear.v | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) (limited to 'backend/Linear.v') diff --git a/backend/Linear.v b/backend/Linear.v index 8c91a809..da1b4c04 100644 --- a/backend/Linear.v +++ b/backend/Linear.v @@ -246,9 +246,9 @@ Inductive step: state -> trace -> state -> Prop := E0 (State s f (Vptr stk Int.zero) f.(fn_code) rs' m') | exec_function_external: forall s ef args res rs1 rs2 m t m', - args = map rs1 (loc_arguments (ef_sig ef)) -> - external_call' ef ge args m t res m' -> - rs2 = Locmap.setlist (map R (loc_result (ef_sig ef))) res rs1 -> + args = map (fun p => Locmap.getpair p rs1) (loc_arguments (ef_sig ef)) -> + external_call ef ge args m t res m' -> + rs2 = Locmap.setpair (loc_result (ef_sig ef)) res rs1 -> step (Callstate s (External ef) rs1 m) t (Returnstate s rs2 m') | exec_return: @@ -268,9 +268,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) := -- cgit