aboutsummaryrefslogtreecommitdiffstats
path: root/backend/CSE2proof.v
diff options
context:
space:
mode:
authorDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-04-16 15:29:24 +0200
committerDavid Monniaux <david.monniaux@univ-grenoble-alpes.fr>2020-04-16 15:29:24 +0200
commitb3431b1d9ee5121883d307cff0b62b7e53369891 (patch)
treed0951b0952d7f5b12b2c03901723ceb46f7d0909 /backend/CSE2proof.v
parentba32e5daa1ff343a1a0b89e65c2ba5764c9cef04 (diff)
downloadcompcert-kvx-b3431b1d9ee5121883d307cff0b62b7e53369891.tar.gz
compcert-kvx-b3431b1d9ee5121883d307cff0b62b7e53369891.zip
refine the rules for builtins
Diffstat (limited to 'backend/CSE2proof.v')
-rw-r--r--backend/CSE2proof.v23
1 files changed, 19 insertions, 4 deletions
diff --git a/backend/CSE2proof.v b/backend/CSE2proof.v
index e61cde3d..9e0ad909 100644
--- a/backend/CSE2proof.v
+++ b/backend/CSE2proof.v
@@ -1125,6 +1125,22 @@ Definition is_killed_in_fmap fmap pc res :=
| Some map => is_killed_in_map map pc res
end.
+Lemma external_call_sound:
+ forall ef (rel : RELATION.t) sp (m m' : mem) (rs : regset) vargs t vres
+ (REL : sem_rel fundef unit ge sp m rel rs)
+ (CALL : external_call ef ge vargs m t vres m'),
+ sem_rel fundef unit ge sp m' (apply_external_call ef rel) rs.
+Proof.
+ destruct ef; intros; simpl in *.
+ all: eauto using kill_mem_sound.
+ all: unfold builtin_or_external_sem in *.
+ 1, 2: destruct (Builtins.lookup_builtin_function name sg);
+ eauto using kill_mem_sound;
+ inv CALL; eauto using kill_mem_sound.
+ all: inv CALL.
+ all: eauto using kill_mem_sound.
+Qed.
+
Definition sem_rel_b' := sem_rel_b fundef unit ge.
Definition fmap_sem' := fmap_sem fundef unit ge.
Definition subst_arg_ok' := subst_arg_ok fundef unit ge.
@@ -1587,9 +1603,9 @@ Proof.
destruct (forward_map _) as [map |] eqn:MAP in *; trivial.
destruct (map # pc) as [mpc |] eqn:MPC in *; try contradiction.
- apply sem_rel_b_ge with (rb2 := Some (kill_builtin_res res (kill_mem mpc))).
+ apply sem_rel_b_ge with (rb2 := Some (kill_builtin_res res (apply_external_call ef mpc))).
{
- replace (Some (kill_builtin_res res (kill_mem mpc))) with (apply_instr' (fn_code f) pc (map # pc)).
+ replace (Some (kill_builtin_res res (apply_external_call ef mpc))) with (apply_instr' (fn_code f) pc (map # pc)).
{
eapply DS.fixpoint_solution with (code := fn_code f) (successors := successors_instr); try eassumption.
2: apply apply_instr'_bot.
@@ -1601,8 +1617,7 @@ Proof.
reflexivity.
}
apply kill_builtin_res_sound.
- apply kill_mem_sound with (m := m).
- assumption.
+ eapply external_call_sound with (m := m); eassumption.
(* cond *)
- econstructor; split.