summaryrefslogtreecommitdiff
path: root/powerpc
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2010-05-23 15:26:33 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2010-05-23 15:26:33 +0000
commit3a050b22f37f3c79a10a8ebae3d292fa77e91b76 (patch)
tree16a0d9366f6805cfc9726c0b80dd8da84cb63a3d /powerpc
parent7999c9ee1f09f7d555e3efc39f030564f76a3354 (diff)
More faithful semantics for volatile reads and writes.
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1346 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'powerpc')
-rw-r--r--powerpc/Asm.v2
-rw-r--r--powerpc/Asmgenproof.v13
2 files changed, 12 insertions, 3 deletions
diff --git a/powerpc/Asm.v b/powerpc/Asm.v
index f4a8a1f..21c237e 100644
--- a/powerpc/Asm.v
+++ b/powerpc/Asm.v
@@ -881,7 +881,7 @@ Inductive step: state -> trace -> state -> Prop :=
forall b ef args res rs m t rs' m',
rs PC = Vptr b Int.zero ->
Genv.find_funct_ptr ge b = Some (External ef) ->
- external_call ef (Genv.find_symbol ge) args m t res m' ->
+ external_call ef ge args m t res m' ->
extcall_arguments rs m ef.(ef_sig) args ->
rs' = (rs#(loc_external_result ef.(ef_sig)) <- res
#PC <- (rs LR)) ->
diff --git a/powerpc/Asmgenproof.v b/powerpc/Asmgenproof.v
index 6a48451..fcbbbd7 100644
--- a/powerpc/Asmgenproof.v
+++ b/powerpc/Asmgenproof.v
@@ -50,6 +50,14 @@ Proof.
exact TRANSF.
Qed.
+Lemma varinfo_preserved:
+ forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
+Proof.
+ intros. unfold ge, tge.
+ apply Genv.find_var_info_transf_partial with transf_fundef.
+ exact TRANSF.
+Qed.
+
Lemma functions_translated:
forall b f,
Genv.find_funct_ptr ge b = Some f ->
@@ -1335,7 +1343,7 @@ Lemma exec_function_external_prop:
(m : mem) (t0 : trace) (ms' : RegEq.t -> val)
(ef : external_function) (args : list val) (res : val) (m': mem),
Genv.find_funct_ptr ge fb = Some (External ef) ->
- external_call ef (Genv.find_symbol ge) args m t0 res m' ->
+ external_call ef ge args m t0 res m' ->
Machconcr.extcall_arguments ms m (parent_sp s) (ef_sig ef) args ->
ms' = Regmap.set (Conventions.loc_result (ef_sig ef)) res ms ->
exec_instr_prop (Machconcr.Callstate s fb ms m)
@@ -1347,7 +1355,8 @@ Proof.
left; exists (State (rs#(loc_external_result (ef_sig ef)) <- res #PC <- (rs LR))
m'); split.
apply plus_one. eapply exec_step_external; eauto.
- eapply external_call_symbols_preserved; eauto. exact symbols_preserved.
+ eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
eapply extcall_arguments_match; eauto.
econstructor; eauto.
unfold loc_external_result. auto with ppcgen.