From a74f6b45d72834b5b8417297017bd81424123d98 Mon Sep 17 00:00:00 2001 From: xleroy Date: Sun, 7 Mar 2010 15:52:58 +0000 Subject: Merge of the newmem and newextcalls branches: - Revised memory model with concrete representation of ints & floats, and per-byte access permissions - Revised Globalenvs implementation - Matching changes in all languages and proofs. git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@1282 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e --- backend/Selectionproof.v | 17 +++++++++-------- 1 file changed, 9 insertions(+), 8 deletions(-) (limited to 'backend/Selectionproof.v') diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v index 70cbeb41..1da7884e 100644 --- a/backend/Selectionproof.v +++ b/backend/Selectionproof.v @@ -18,7 +18,7 @@ Require Import AST. Require Import Integers. Require Import Floats. Require Import Values. -Require Import Mem. +Require Import Memory. Require Import Events. Require Import Globalenvs. Require Import Smallstep. @@ -300,7 +300,7 @@ Lemma functions_translated: Genv.find_funct tge v = Some (sel_fundef f). Proof. intros. - exact (Genv.find_funct_transf sel_fundef H). + exact (Genv.find_funct_transf sel_fundef _ _ H). Qed. Lemma function_ptr_translated: @@ -309,7 +309,7 @@ Lemma function_ptr_translated: Genv.find_funct_ptr tge b = Some (sel_fundef f). Proof. intros. - exact (Genv.find_funct_ptr_transf sel_fundef H). + exact (Genv.find_funct_ptr_transf sel_fundef _ _ H). Qed. Lemma sig_function_translated: @@ -428,6 +428,7 @@ Proof. econstructor; split. econstructor. destruct k; simpl in H; simpl; auto. rewrite <- H0; reflexivity. + simpl. eauto. constructor; auto. (* (* assign *) @@ -457,11 +458,11 @@ Proof. constructor; auto. destruct b; auto. (* Sreturn None *) econstructor; split. - econstructor. + econstructor. simpl; eauto. constructor; auto. apply call_cont_commut. (* Sreturn Some *) econstructor; split. - econstructor. simpl. eauto with evalexpr. + econstructor. simpl. eauto with evalexpr. simpl; eauto. constructor; auto. apply call_cont_commut. (* Sgoto *) econstructor; split. @@ -477,10 +478,10 @@ Proof. induction 1. econstructor; split. econstructor. - simpl. fold tge. rewrite symbols_preserved. eexact H. + apply Genv.init_mem_transf; eauto. + simpl. fold tge. rewrite symbols_preserved. eexact H0. apply function_ptr_translated. eauto. - rewrite <- H1. apply sig_function_translated; auto. - unfold tprog, sel_program. rewrite Genv.init_mem_transf. + rewrite <- H2. apply sig_function_translated; auto. constructor; auto. Qed. -- cgit