aboutsummaryrefslogtreecommitdiffstats
path: root/cfrontend/Cexec.v
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@inria.fr>2016-03-06 10:23:18 +0100
committerXavier Leroy <xavier.leroy@inria.fr>2016-03-06 10:23:18 +0100
commit2b5f5cb0bb6d8dbf302ab6d84c27eda30252912d (patch)
treeaea549d4f8a35fe5ca70a4eb7c21c2bafdc9781d /cfrontend/Cexec.v
parentbe2315351761e4fc0430b91ac791d66eec0e0cd7 (diff)
downloadcompcert-kvx-2b5f5cb0bb6d8dbf302ab6d84c27eda30252912d.tar.gz
compcert-kvx-2b5f5cb0bb6d8dbf302ab6d84c27eda30252912d.zip
Add support for EF_runtime externals
Also: in Events, use Senv.equiv to state invariance wrt changes of global envs.
Diffstat (limited to 'cfrontend/Cexec.v')
-rw-r--r--cfrontend/Cexec.v5
1 files changed, 5 insertions, 0 deletions
diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v
index 7e966ffe..36f7bf7d 100644
--- a/cfrontend/Cexec.v
+++ b/cfrontend/Cexec.v
@@ -536,6 +536,7 @@ Definition do_external (ef: external_function):
match ef with
| EF_external name sg => do_external_function name sg ge
| EF_builtin name sg => do_external_function name sg ge
+ | EF_runtime name sg => do_external_function name sg ge
| EF_vload chunk => do_ef_volatile_load chunk
| EF_vstore chunk => do_ef_volatile_store chunk
| EF_malloc => do_ef_malloc
@@ -558,6 +559,8 @@ Proof with try congruence.
eapply do_external_function_sound; eauto.
(* EF_builtin *)
eapply do_external_function_sound; eauto.
+(* EF_runtime *)
+ eapply do_external_function_sound; eauto.
(* EF_vload *)
unfold do_ef_volatile_load. destruct vargs... destruct v... destruct vargs...
mydestr. destruct p as [[w'' t''] v]; mydestr.
@@ -604,6 +607,8 @@ Proof.
eapply do_external_function_complete; eauto.
(* EF_builtin *)
eapply do_external_function_complete; eauto.
+(* EF_runtime *)
+ eapply do_external_function_complete; eauto.
(* EF_vload *)
inv H; unfold do_ef_volatile_load.
exploit do_volatile_load_complete; eauto. intros EQ; rewrite EQ; auto.