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 --- common/Determinism.v | 142 +-- common/Events.v | 755 ++++++++++--- common/Globalenvs.v | 1733 ++++++++++++++++-------------- common/Mem.v | 2887 -------------------------------------------------- common/Memdata.v | 1058 ++++++++++++++++++ common/Memdataaux.ml | 68 ++ common/Memory.v | 2844 +++++++++++++++++++++++++++++++++++++++++++++++++ common/Memtype.v | 989 +++++++++++++++++ common/Values.v | 81 ++ 9 files changed, 6688 insertions(+), 3869 deletions(-) delete mode 100644 common/Mem.v create mode 100644 common/Memdata.v create mode 100644 common/Memdataaux.ml create mode 100644 common/Memory.v create mode 100644 common/Memtype.v (limited to 'common') diff --git a/common/Determinism.v b/common/Determinism.v index 430ee93..862d5a5 100644 --- a/common/Determinism.v +++ b/common/Determinism.v @@ -32,8 +32,8 @@ Axiom traceinf_extensionality: (** * Deterministic worlds *) (** One source of possible nondeterminism is that our semantics leave - unspecified the results of calls to external - functions. Different results to e.g. a "read" operation can of + unspecified the results of system calls. + Different results to e.g. a "read" operation can of course lead to different behaviors for the program. We address this issue by modeling a notion of deterministic external world that uniquely determines the results of external calls. *) @@ -61,13 +61,21 @@ Definition nextworld (w: world) (evname: ident) (evargs: list eventval) : world and [T] the infinite trace of interest. *) +Inductive possible_event: world -> event -> world -> Prop := + | possible_event_syscall: forall w1 evname evargs evres w2, + nextworld w1 evname evargs = Some (evres, w2) -> + possible_event w1 (Event_syscall evname evargs evres) w2 + | possible_event_load: forall w label, + possible_event w (Event_load label) w + | possible_event_store: forall w label, + possible_event w (Event_store label) w. + Inductive possible_trace: world -> trace -> world -> Prop := | possible_trace_nil: forall w, possible_trace w E0 w - | possible_trace_cons: forall w0 evname evargs evres w1 t w2, - nextworld w0 evname evargs = Some (evres, w1) -> - possible_trace w1 t w2 -> - possible_trace w0 (mkevent evname evargs evres :: t) w2. + | possible_trace_cons: forall w1 ev w2 t w3, + possible_event w1 ev w2 -> possible_trace w2 t w3 -> + possible_trace w1 (ev :: t) w3. Lemma possible_trace_app: forall t2 w2 w0 t1 w1, @@ -90,11 +98,28 @@ Proof. exists w1; split. econstructor; eauto. auto. Qed. +Lemma possible_event_final_world: + forall w ev w1 w2, + possible_event w ev w1 -> possible_event w ev w2 -> w1 = w2. +Proof. + intros. inv H; inv H0; congruence. +Qed. + +Lemma possible_trace_final_world: + forall w0 t w1, possible_trace w0 t w1 -> + forall w2, possible_trace w0 t w2 -> w1 = w2. +Proof. + induction 1; intros. + inv H. auto. + inv H1. assert (w2 = w5) by (eapply possible_event_final_world; eauto). + subst; eauto. +Qed. + CoInductive possible_traceinf: world -> traceinf -> Prop := - | possible_traceinf_cons: forall w0 evname evargs evres w1 T, - nextworld w0 evname evargs = Some (evres, w1) -> - possible_traceinf w1 T -> - possible_traceinf w0 (Econsinf (mkevent evname evargs evres) T). + | possible_traceinf_cons: forall w1 ev w2 T, + possible_event w1 ev w2 -> + possible_traceinf w2 T -> + possible_traceinf w1 (Econsinf ev T). Lemma possible_traceinf_app: forall t2 w0 t1 w1, @@ -149,34 +174,13 @@ Definition possible_behavior (w: world) (b: program_behavior) : Prop := | Goes_wrong t => exists w', possible_trace w t w' end. -(** Determinism properties of [event_match]. *) - -Remark eventval_match_deterministic: - forall ev1 ev2 ty v1 v2, - eventval_match ev1 ty v1 -> eventval_match ev2 ty v2 -> - (ev1 = ev2 <-> v1 = v2). -Proof. - intros. inv H; inv H0; intuition congruence. -Qed. - -Remark eventval_list_match_deterministic: - forall ev1 ty v, eventval_list_match ev1 ty v -> - forall ev2, eventval_list_match ev2 ty v -> ev1 = ev2. -Proof. - induction 1; intros. - inv H. auto. - inv H1. decEq. - rewrite (eventval_match_deterministic _ _ _ _ _ H H6). auto. - eauto. -Qed. - (** * Deterministic semantics *) Section DETERM_SEM. (** We assume given a transition semantics that is internally deterministic: the only source of non-determinism is the return - value of external calls. *) + value of system calls. *) Variable genv: Type. Variable state: Type. @@ -184,17 +188,9 @@ Variable step: genv -> state -> trace -> state -> Prop. Variable initial_state: state -> Prop. Variable final_state: state -> int -> Prop. -Inductive internal_determinism: trace -> state -> trace -> state -> Prop := - | int_determ_0: forall s, - internal_determinism E0 s E0 s - | int_determ_1: forall s s' id arg res res', - (res = res' -> s = s') -> - internal_determinism (mkevent id arg res :: nil) s - (mkevent id arg res' :: nil) s'. - Hypothesis step_internal_deterministic: forall ge s t1 s1 t2 s2, - step ge s t1 s1 -> step ge s t2 s2 -> internal_determinism t1 s1 t2 s2. + step ge s t1 s1 -> step ge s t2 s2 -> matching_traces t1 t2 -> s1 = s2 /\ t1 = t2. Hypothesis initial_state_determ: forall s1 s2, initial_state s1 -> initial_state s2 -> s1 = s2. @@ -208,18 +204,29 @@ Hypothesis final_state_nostep: (** Consequently, the [step] relation is deterministic if restricted to traces that are possible in a deterministic world. *) +Remark matching_possible_traces: + forall w0 t1 w1, possible_trace w0 t1 w1 -> + forall t2 w2, possible_trace w0 t2 w2 -> + matching_traces t1 t2. +Proof. + induction 1; intros. + destruct t2; simpl; auto. + destruct t2; simpl. destruct ev; auto. inv H1. + inv H; inv H5; auto; intros. + subst. rewrite H in H1; inv H1. split; eauto. + eauto. + eauto. +Qed. + Lemma step_deterministic: forall ge s0 t1 s1 t2 s2 w0 w1 w2, step ge s0 t1 s1 -> step ge s0 t2 s2 -> possible_trace w0 t1 w1 -> possible_trace w0 t2 w2 -> s1 = s2 /\ t1 = t2 /\ w1 = w2. Proof. - intros. exploit step_internal_deterministic. eexact H. eexact H0. intro ID. - inv ID. - inv H1. inv H2. auto. - inv H2. inv H11. inv H1. inv H11. - rewrite H10 in H9. inv H9. - intuition. + intros. exploit step_internal_deterministic. eexact H. eexact H0. + eapply matching_possible_traces; eauto. intuition. + subst. eapply possible_trace_final_world; eauto. Qed. Ltac use_step_deterministic := @@ -378,44 +385,55 @@ Lemma forever_reactive_inv2: t1 <> E0 -> t2 <> E0 -> forever_reactive step ge s1 T1 -> possible_traceinf w1 T1 -> forever_reactive step ge s2 T2 -> possible_traceinf w2 T2 -> - exists s', exists e, exists T1', exists T2', exists w', + exists s', exists t, exists T1', exists T2', exists w', + t <> E0 /\ forever_reactive step ge s' T1' /\ possible_traceinf w' T1' /\ forever_reactive step ge s' T2' /\ possible_traceinf w' T2' /\ - t1 *** T1 = Econsinf e T1' /\ - t2 *** T2 = Econsinf e T2'. + t1 *** T1 = t *** T1' /\ + t2 *** T2 = t *** T2'. Proof. induction 1; intros. congruence. inv H3. congruence. possibleTraceInv. - assert (ID: internal_determinism t3 s5 t1 s2). eauto. - inv ID. - possibleTraceInv. eauto. - inv P. inv P1. inv H17. inv H19. rewrite H18 in H16; inv H16. - assert (s5 = s2) by auto. subst s5. - exists s2; exists (mkevent id arg res'); - exists (t2 *** T1); exists (t4 *** T2); exists w0. + use_step_deterministic. + destruct t3. + (* inductive case *) + simpl in *. inv P1; inv P. eapply IHstar; eauto. + (* base case *) + exists s5; exists (e :: t3); + exists (t2 *** T1); exists (t4 *** T2); exists w3. + split. unfold E0; congruence. split. eapply star_forever_reactive; eauto. split. eapply possible_traceinf_app; eauto. split. eapply star_forever_reactive; eauto. split. eapply possible_traceinf_app; eauto. - auto. + split; traceEq. Qed. -Lemma forever_reactive_determ: +Lemma forever_reactive_determ': forall ge s T1 T2 w, forever_reactive step ge s T1 -> possible_traceinf w T1 -> forever_reactive step ge s T2 -> possible_traceinf w T2 -> - traceinf_sim T1 T2. + traceinf_sim' T1 T2. Proof. cofix COINDHYP; intros. inv H. inv H1. possibleTraceInv. destruct (forever_reactive_inv2 _ _ _ _ H _ _ _ _ _ _ _ P H3 P1 H6 H4 H7 P0 H5 P2) - as [s' [e [T1' [T2' [w' [A [B [C [D [E G]]]]]]]]]]. - rewrite E; rewrite G. constructor. + as [s' [t' [T1' [T2' [w' [A [B [C [D [E [G K]]]]]]]]]]]. + rewrite G; rewrite K. constructor. auto. eapply COINDHYP; eauto. Qed. +Lemma forever_reactive_determ: + forall ge s T1 T2 w, + forever_reactive step ge s T1 -> possible_traceinf w T1 -> + forever_reactive step ge s T2 -> possible_traceinf w T2 -> + traceinf_sim T1 T2. +Proof. + intros. apply traceinf_sim'_sim. eapply forever_reactive_determ'; eauto. +Qed. + Lemma star_forever_reactive_inv: forall ge s t s', star step ge s t s' -> forall w w' T, possible_trace w t w' -> forever_reactive step ge s T -> diff --git a/common/Events.v b/common/Events.v index 855c013..ad1fc51 100644 --- a/common/Events.v +++ b/common/Events.v @@ -13,34 +13,44 @@ (* *) (* *********************************************************************) -(** Representation of observable events and execution traces. *) +(** Observable events, execution traces, and semantics of external calls. *) Require Import Coqlib. +Require Intv. Require Import AST. Require Import Integers. Require Import Floats. Require Import Values. +Require Import Memory. + +(** * Events and traces *) (** The observable behaviour of programs is stated in terms of - input/output events, which can also be thought of as system calls - to the operating system. An event is generated each time an - external function (see module AST) is invoked. The event records - the name of the external function, the arguments to the function - invocation provided by the program, and the return value provided by - the outside world (e.g. the operating system). Arguments and values - are either integers or floating-point numbers. We currently do not - allow pointers to be exchanged between the program and the outside - world. *) + input/output events, which represent the actions of the program + that the external world can observe. CompCert leaves much flexibility as to + the exact content of events: the only requirement is that they + do not expose pointer values nor memory states, because these + are not preserved literally during compilation. For concreteness, + we use the following type for events. Each event represents either: + +- A system call (e.g. an input/output operation), recording the + name of the system call, its int-or-float parameters, + and its int-or-float result. + +- A volatile load from a memory location, recording a label + associated with the read (e.g. a global variable name or a source code position). + +- A volatile store to a memory location, also recording a label. +*) Inductive eventval: Type := | EVint: int -> eventval | EVfloat: float -> eventval. -Record event : Type := mkevent { - ev_name: ident; - ev_args: list eventval; - ev_res: eventval -}. +Inductive event: Type := + | Event_syscall: ident -> list eventval -> eventval -> event + | Event_load: ident -> event + | Event_store: ident -> event. (** The dynamic semantics for programs collect traces of events. Traces are of two kinds: finite (type [trace]) or infinite (type [traceinf]). *) @@ -49,10 +59,6 @@ Definition trace := list event. Definition E0 : trace := nil. -Definition Eextcall - (name: ident) (args: list eventval) (res: eventval) : trace := - mkevent name args res :: nil. - Definition Eapp (t1 t2: trace) : trace := t1 ++ t2. CoInductive traceinf : Type := @@ -93,7 +99,7 @@ Qed. Hint Rewrite E0_left E0_right Eapp_assoc E0_left_inf Eappinf_assoc: trace_rewrite. -Opaque trace E0 Eextcall Eapp Eappinf. +Opaque trace E0 Eapp Eappinf. (** The following [traceEq] tactic proves equalities between traces or infinite traces. *) @@ -115,115 +121,6 @@ Ltac decomposeTraceEq := Ltac traceEq := repeat substTraceHyp; autorewrite with trace_rewrite; decomposeTraceEq. -(** The predicate [event_match ef vargs t vres] expresses that - the event [t] is generated when invoking external function [ef] - with arguments [vargs], and obtaining [vres] as a return value - from the operating system. *) - -Inductive eventval_match: eventval -> typ -> val -> Prop := - | ev_match_int: - forall i, eventval_match (EVint i) Tint (Vint i) - | ev_match_float: - forall f, eventval_match (EVfloat f) Tfloat (Vfloat f). - -Inductive eventval_list_match: list eventval -> list typ -> list val -> Prop := - | evl_match_nil: - eventval_list_match nil nil nil - | evl_match_cons: - forall ev1 evl ty1 tyl v1 vl, - eventval_match ev1 ty1 v1 -> - eventval_list_match evl tyl vl -> - eventval_list_match (ev1::evl) (ty1::tyl) (v1::vl). - -Inductive event_match: - external_function -> list val -> trace -> val -> Prop := - event_match_intro: - forall ef vargs vres eargs eres, - eventval_list_match eargs (sig_args ef.(ef_sig)) vargs -> - eventval_match eres (proj_sig_res ef.(ef_sig)) vres -> - event_match ef vargs (Eextcall ef.(ef_id) eargs eres) vres. - -(** The following section shows that [event_match] is stable under - relocation of pointer values, as performed by memory injections - (see module [Mem]). *) - -Require Import Mem. - -Section EVENT_MATCH_INJECT. - -Variable f: meminj. - -Remark eventval_match_inject: - forall ev ty v1, eventval_match ev ty v1 -> - forall v2, val_inject f v1 v2 -> - eventval_match ev ty v2. -Proof. - induction 1; intros; inversion H; constructor. -Qed. - -Remark eventval_list_match_inject: - forall evl tyl vl1, eventval_list_match evl tyl vl1 -> - forall vl2, val_list_inject f vl1 vl2 -> - eventval_list_match evl tyl vl2. -Proof. - induction 1; intros. - inversion H; constructor. - inversion H1; constructor. - eapply eventval_match_inject; eauto. - eauto. -Qed. - -Lemma event_match_inject: - forall ef args1 t res args2, - event_match ef args1 t res -> - val_list_inject f args1 args2 -> - event_match ef args2 t res /\ val_inject f res res. -Proof. - intros. inversion H; subst. - split. constructor. eapply eventval_list_match_inject; eauto. auto. - inversion H2; constructor. -Qed. - -End EVENT_MATCH_INJECT. - -(** The following section shows that [event_match] is stable under - replacement of [Vundef] values by more defined values. *) - -Section EVENT_MATCH_LESSDEF. - -Remark eventval_match_lessdef: - forall ev ty v1, eventval_match ev ty v1 -> - forall v2, Val.lessdef v1 v2 -> - eventval_match ev ty v2. -Proof. - induction 1; intros; inv H; constructor. -Qed. - -Remark eventval_list_match_moredef: - forall evl tyl vl1, eventval_list_match evl tyl vl1 -> - forall vl2, Val.lessdef_list vl1 vl2 -> - eventval_list_match evl tyl vl2. -Proof. - induction 1; intros. - inversion H; constructor. - inversion H1; constructor. - eapply eventval_match_lessdef; eauto. - eauto. -Qed. - -Lemma event_match_lessdef: - forall ef args1 t res1 args2, - event_match ef args1 t res1 -> - Val.lessdef_list args1 args2 -> - exists res2, event_match ef args2 t res2 /\ Val.lessdef res1 res2. -Proof. - intros. inversion H; subst. exists res1; split. - constructor. eapply eventval_list_match_moredef; eauto. auto. - auto. -Qed. - -End EVENT_MATCH_LESSDEF. - (** Bisimilarity between infinite traces. *) CoInductive traceinf_sim: traceinf -> traceinf -> Prop := @@ -251,6 +148,23 @@ Proof. cofix COINDHYP;intros. inv H; inv H0; constructor; eauto. Qed. +CoInductive traceinf_sim': traceinf -> traceinf -> Prop := + | traceinf_sim'_cons: forall t T1 T2, + t <> E0 -> traceinf_sim' T1 T2 -> traceinf_sim' (t *** T1) (t *** T2). + +Lemma traceinf_sim'_sim: + forall T1 T2, traceinf_sim' T1 T2 -> traceinf_sim T1 T2. +Proof. + cofix COINDHYP; intros. inv H. + destruct t. elim H0; auto. +Transparent Eappinf. +Transparent E0. + simpl. + destruct t. simpl. constructor. apply COINDHYP; auto. + constructor. apply COINDHYP. + constructor. unfold E0; congruence. auto. +Qed. + (** The "is prefix of" relation between a finite and an infinite trace. *) Inductive traceinf_prefix: trace -> traceinf -> Prop := @@ -321,3 +235,586 @@ Proof. Transparent Eappinf. simpl. f_equal. apply IHt. Qed. + +(** * Semantics of external functions *) + +(** Each external function is of one of the following kinds: *) + +Inductive extfun_kind: signature -> Type := + | EF_syscall (sg: signature) (name: ident): extfun_kind sg + (** A system call. Takes integer-or-float arguments, produces a + result that is an integer or a float, does not modify + the memory, and produces an [Event_syscall] event in the trace. *) + | EF_load (label: ident) (chunk: memory_chunk): extfun_kind (mksignature (Tint :: nil) (Some (type_of_chunk chunk))) + (** A volatile read operation. Reads and returns the given memory + chunk from the address given as first argument. + Produces an [Event_load] event containing the given label. *) + | EF_store (label: ident) (chunk: memory_chunk): extfun_kind (mksignature (Tint :: type_of_chunk chunk :: nil) None) + (** A volatile store operation. Store the value given as second + argument at the address given as first argument, using the + given memory chunk. + Produces an [Event_store] event containing the given label. *) + | EF_malloc: extfun_kind (mksignature (Tint :: nil) (Some Tint)) + (** Dynamic memory allocation. Takes the requested size in bytes + as argument; returns a pointer to a fresh block of the given size. + Produces no observable event. *) + | EF_free: extfun_kind (mksignature (Tint :: nil) None). + (** Dynamic memory deallocation. Takes a pointer to a block + allocated by an [EF_malloc] external call and frees the + corresponding block. + Produces no observable event. *) + +Parameter classify_external_function: + forall (ef: external_function), extfun_kind (ef.(ef_sig)). + +(** For each external function, its behavior is defined by a predicate relating: +- the values of the arguments passed to this function +- the memory state before the call +- the result value of the call +- the memory state after the call +- the trace generated by the call (can be empty). + +We now specify the expected properties of this predicate. +*) + +Definition mem_unchanged_on (P: block -> Z -> Prop) (m_before m_after: mem): Prop := + (forall b ofs p, + P b ofs -> Mem.perm m_before b ofs p -> Mem.perm m_after b ofs p) +/\(forall chunk b ofs v, + (forall i, ofs <= i < ofs + size_chunk chunk -> P b i) -> + Mem.load chunk m_before b ofs = Some v -> + Mem.load chunk m_after b ofs = Some v). + +Definition loc_out_of_bounds (m: mem) (b: block) (ofs: Z) : Prop := + ofs < Mem.low_bound m b \/ ofs > Mem.high_bound m b. + +Definition loc_unmapped (f: meminj) (b: block) (ofs: Z): Prop := + f b = None. + +Definition loc_out_of_reach (f: meminj) (m: mem) (b: block) (ofs: Z): Prop := + forall b0 delta, + f b0 = Some(b, delta) -> + ofs < Mem.low_bound m b0 + delta \/ ofs >= Mem.high_bound m b0 + delta. + +Definition inject_separated (f f': meminj) (m1 m2: mem): Prop := + forall b1 b2 delta, + f b1 = None -> f' b1 = Some(b2, delta) -> + ~Mem.valid_block m1 b1 /\ ~Mem.valid_block m2 b2. + +Fixpoint matching_traces (t1 t2: trace) {struct t1} : Prop := + match t1, t2 with + | Event_syscall name1 args1 res1 :: t1', Event_syscall name2 args2 res2 :: t2' => + name1 = name2 -> args1 = args2 -> res1 = res2 /\ matching_traces t1' t2' + | Event_load name1 :: t1', Event_load name2 :: t2' => + name1 = name2 -> matching_traces t1' t2' + | Event_store name1 :: t1', Event_store name2 :: t2' => + name1 = name2 -> matching_traces t1' t2' + | _, _ => + True + end. + +Record extcall_properties (sem: list val -> mem -> trace -> val -> mem -> Prop) + (sg: signature) : Prop := mk_extcall_properties { + +(** The return value of an external call must agree with its signature. *) + ec_well_typed: + forall vargs m1 t vres m2, + sem vargs m1 t vres m2 -> + Val.has_type vres (proj_sig_res sg); + +(** The number of arguments of an external call must agree with its signature. *) + ec_arity: + forall vargs m1 t vres m2, + sem vargs m1 t vres m2 -> + List.length vargs = List.length sg.(sig_args); + +(** External calls cannot invalidate memory blocks. (Remember that + freeing a block does not invalidate its block identifier.) *) + ec_valid_block: + forall vargs m1 t vres m2 b, + sem vargs m1 t vres m2 -> + Mem.valid_block m1 b -> Mem.valid_block m2 b; + +(** External calls preserve the bounds of valid blocks. *) + ec_bounds: + forall vargs m1 t vres m2 b, + sem vargs m1 t vres m2 -> + Mem.valid_block m1 b -> Mem.bounds m2 b = Mem.bounds m1 b; + +(** External calls must commute with memory extensions, in the + following sense. *) + ec_mem_extends: + forall vargs m1 t vres m2 m1' vargs', + sem vargs m1 t vres m2 -> + Mem.extends m1 m1' -> + Val.lessdef_list vargs vargs' -> + exists vres', exists m2', + sem vargs' m1' t vres' m2' + /\ Val.lessdef vres vres' + /\ Mem.extends m2 m2' + /\ mem_unchanged_on (loc_out_of_bounds m1) m1' m2'; + +(** External calls must commute with memory injections, + in the following sense. *) + ec_mem_inject: + forall vargs m1 t vres m2 f m1' vargs', + sem vargs m1 t vres m2 -> + Mem.inject f m1 m1' -> + val_list_inject f vargs vargs' -> + exists f', exists vres', exists m2', + sem vargs' m1' t vres' m2' + /\ val_inject f' vres vres' + /\ Mem.inject f' m2 m2' + /\ mem_unchanged_on (loc_unmapped f) m1 m2 + /\ mem_unchanged_on (loc_out_of_reach f m1) m1' m2' + /\ inject_incr f f' + /\ inject_separated f f' m1 m1'; + +(** External calls must be internally deterministic: + if the observable traces match, the return states must be + identical. *) + ec_determ: + forall vargs m t1 vres1 m1 t2 vres2 m2, + sem vargs m t1 vres1 m1 -> sem vargs m t2 vres2 m2 -> + matching_traces t1 t2 -> t1 = t2 /\ vres1 = vres2 /\ m1 = m2 +}. + +(** ** Semantics of volatile loads *) + +Inductive extcall_load_sem (label: ident) (chunk: memory_chunk): + list val -> mem -> trace -> val -> mem -> Prop := + | extcall_load_sem_intro: forall b ofs m vres, + Mem.load chunk m b (Int.signed ofs) = Some vres -> + extcall_load_sem label chunk (Vptr b ofs :: nil) m (Event_load label :: nil) vres m. + +Lemma extcall_load_ok: + forall label chunk, + extcall_properties (extcall_load_sem label chunk) + (mksignature (Tint :: nil) (Some (type_of_chunk chunk))). +Proof. + intros; constructor; intros. + + inv H. unfold proj_sig_res. simpl. eapply Mem.load_type; eauto. + + inv H. simpl. auto. + + inv H. auto. + + inv H. auto. + + inv H. inv H1. inv H6. inv H4. + exploit Mem.load_extends; eauto. intros [v2 [A B]]. + exists v2; exists m1'; intuition. + constructor; auto. + red. auto. + + inv H. inv H1. inv H6. + assert (Mem.loadv chunk m2 (Vptr b ofs) = Some vres). auto. + exploit Mem.loadv_inject; eauto. intros [v2 [A B]]. + inv H4. + exists f; exists v2; exists m1'; intuition. + constructor. auto. + red; auto. + red; auto. + red; intros. congruence. + + inv H; inv H0. intuition congruence. +Qed. + +(** ** Semantics of volatile stores *) + +Inductive extcall_store_sem (label: ident) (chunk: memory_chunk): + list val -> mem -> trace -> val -> mem -> Prop := + | extcall_store_sem_intro: forall b ofs v m m', + Mem.store chunk m b (Int.signed ofs) v = Some m' -> + extcall_store_sem label chunk (Vptr b ofs :: v :: nil) m (Event_store label :: nil) Vundef m'. + +Lemma extcall_store_ok: + forall label chunk, + extcall_properties (extcall_store_sem label chunk) + (mksignature (Tint :: type_of_chunk chunk :: nil) None). +Proof. + intros; constructor; intros. + + inv H. unfold proj_sig_res. simpl. auto. + + inv H. simpl. auto. + + inv H. eauto with mem. + + inv H. eapply Mem.bounds_store; eauto. + + inv H. inv H1. inv H6. inv H7. inv H4. + exploit Mem.store_within_extends; eauto. intros [m' [A B]]. + exists Vundef; exists m'; intuition. + constructor; auto. + red; split; intros. + eapply Mem.perm_store_1; eauto. + rewrite <- H1. eapply Mem.load_store_other; eauto. + destruct (eq_block b0 b); auto. subst b0; right. + exploit Mem.valid_access_in_bounds. + eapply Mem.store_valid_access_3. eexact H2. + intros [C D]. + generalize (size_chunk_pos chunk0). intro E. + generalize (size_chunk_pos chunk). intro F. + apply (Intv.range_disjoint' (ofs0, ofs0 + size_chunk chunk0) + (Int.signed ofs, Int.signed ofs + size_chunk chunk)). + red; intros. generalize (H x H4). unfold loc_out_of_bounds, Intv.In; simpl. omega. + simpl; omega. simpl; omega. + + inv H. inv H1. inv H6. inv H7. + assert (Mem.storev chunk m1 (Vptr b ofs) v = Some m2). simpl; auto. + exploit Mem.storev_mapped_inject; eauto. intros [m2' [A B]]. + inv H4. + exists f; exists Vundef; exists m2'; intuition. + constructor; auto. + split; intros. eapply Mem.perm_store_1; eauto. + rewrite <- H4. eapply Mem.load_store_other; eauto. + left. exploit (H1 ofs0). generalize (size_chunk_pos chunk0). omega. + unfold loc_unmapped. congruence. + split; intros. eapply Mem.perm_store_1; eauto. + rewrite <- H4. eapply Mem.load_store_other; eauto. + destruct (eq_block b0 b2); auto. subst b0; right. + assert (EQ: Int.signed (Int.add ofs (Int.repr delta)) = Int.signed ofs + delta). + eapply Mem.address_inject; eauto with mem. + simpl in A. rewrite EQ in A. rewrite EQ. + exploit Mem.valid_access_in_bounds. + eapply Mem.store_valid_access_3. eexact H2. + intros [C D]. + generalize (size_chunk_pos chunk0). intro E. + generalize (size_chunk_pos chunk). intro F. + apply (Intv.range_disjoint' (ofs0, ofs0 + size_chunk chunk0) + (Int.signed ofs + delta, Int.signed ofs + delta + size_chunk chunk)). + red; intros. exploit (H1 x H5). eauto. unfold Intv.In; simpl. omega. + simpl; omega. simpl; omega. + + red; intros. congruence. + + inv H; inv H0. intuition congruence. +Qed. + +(** ** Semantics of dynamic memory allocation (malloc) *) + +Inductive extcall_malloc_sem: + list val -> mem -> trace -> val -> mem -> Prop := + | extcall_malloc_sem_intro: forall n m m' b m'', + Mem.alloc m (-4) (Int.signed n) = (m', b) -> + Mem.store Mint32 m' b (-4) (Vint n) = Some m'' -> + extcall_malloc_sem (Vint n :: nil) m E0 (Vptr b Int.zero) m''. + +Lemma extcall_malloc_ok: + extcall_properties extcall_malloc_sem + (mksignature (Tint :: nil) (Some Tint)). +Proof. + assert (UNCHANGED: + forall (P: block -> Z -> Prop) m n m' b m'', + Mem.alloc m (-4) (Int.signed n) = (m', b) -> + Mem.store Mint32 m' b (-4) (Vint n) = Some m'' -> + mem_unchanged_on P m m''). + intros; split; intros. + eauto with mem. + transitivity (Mem.load chunk m' b0 ofs). + eapply Mem.load_store_other; eauto. left. + apply Mem.valid_not_valid_diff with m; eauto with mem. + eapply Mem.load_alloc_other; eauto. + + constructor; intros. + + inv H. unfold proj_sig_res; simpl. auto. + + inv H. auto. + + inv H. eauto with mem. + + inv H. transitivity (Mem.bounds m' b). + eapply Mem.bounds_store; eauto. + eapply Mem.bounds_alloc_other; eauto. + apply Mem.valid_not_valid_diff with m1; eauto with mem. + + inv H. inv H1. inv H5. inv H7. + exploit Mem.alloc_extends; eauto. apply Zle_refl. apply Zle_refl. + intros [m3' [A B]]. + exploit Mem.store_within_extends. eexact B. eauto. + instantiate (1 := Vint n). auto. + intros [m2' [C D]]. + exists (Vptr b Int.zero); exists m2'; intuition. + econstructor; eauto. + eapply UNCHANGED; eauto. + + inv H. inv H1. inv H5. inv H7. + exploit Mem.alloc_parallel_inject; eauto. apply Zle_refl. apply Zle_refl. + intros [f' [m3' [b' [ALLOC [A [B [C D]]]]]]]. + exploit Mem.store_mapped_inject. eexact A. eauto. eauto. + instantiate (1 := Vint n). auto. + intros [m2' [E F]]. + exists f'; exists (Vptr b' Int.zero); exists m2'; intuition. + econstructor; eauto. + econstructor. eauto. auto. + eapply UNCHANGED; eauto. + eapply UNCHANGED; eauto. + red; intros. destruct (eq_block b1 b). + subst b1. rewrite C in H1. inv H1. eauto with mem. + rewrite D in H1. congruence. auto. + + inv H; inv H0. intuition congruence. +Qed. + +(** ** Semantics of dynamic memory deallocation (free) *) + +Inductive extcall_free_sem: + list val -> mem -> trace -> val -> mem -> Prop := + | extcall_free_sem_intro: forall b lo sz m m', + Mem.load Mint32 m b (Int.signed lo - 4) = Some (Vint sz) -> + Int.signed sz > 0 -> + Mem.free m b (Int.signed lo - 4) (Int.signed lo + Int.signed sz) = Some m' -> + extcall_free_sem (Vptr b lo :: nil) m E0 Vundef m'. + +Lemma extcall_free_ok: + extcall_properties extcall_free_sem + (mksignature (Tint :: nil) None). +Proof. + assert (UNCHANGED: + forall (P: block -> Z -> Prop) m b lo hi m', + Mem.free m b lo hi = Some m' -> + lo < hi -> + (forall b' ofs, P b' ofs -> b' <> b \/ ofs < lo \/ hi <= ofs) -> + mem_unchanged_on P m m'). + intros; split; intros. + eapply Mem.perm_free_1; eauto. + rewrite <- H3. eapply Mem.load_free; eauto. + destruct (eq_block b0 b); auto. right. right. + apply (Intv.range_disjoint' (ofs, ofs + size_chunk chunk) (lo, hi)). + red; intros. apply Intv.notin_range. simpl. exploit H1; eauto. intuition. + simpl; generalize (size_chunk_pos chunk); omega. + simpl; omega. + + constructor; intros. + + inv H. unfold proj_sig_res. simpl. auto. + + inv H. auto. + + inv H. eauto with mem. + + inv H. eapply Mem.bounds_free; eauto. + + inv H. inv H1. inv H8. inv H6. + exploit Mem.load_extends; eauto. intros [vsz [A B]]. inv B. + exploit Mem.free_parallel_extends; eauto. intros [m2' [C D]]. + exists Vundef; exists m2'; intuition. + econstructor; eauto. + eapply UNCHANGED; eauto. omega. + intros. destruct (eq_block b' b); auto. subst b; right. + red in H. + exploit Mem.range_perm_in_bounds. + eapply Mem.free_range_perm. eexact H4. omega. omega. + + inv H. inv H1. inv H8. inv H6. + exploit Mem.load_inject; eauto. intros [vsz [A B]]. inv B. + assert (Mem.range_perm m1 b (Int.signed lo - 4) (Int.signed lo + Int.signed sz) Freeable). + eapply Mem.free_range_perm; eauto. + exploit Mem.address_inject; eauto. + apply Mem.perm_implies with Freeable; auto with mem. + apply H. instantiate (1 := lo). omega. + intro EQ. + assert (Mem.range_perm m1' b2 (Int.signed lo + delta - 4) (Int.signed lo + delta + Int.signed sz) Freeable). + red; intros. + replace ofs with ((ofs - delta) + delta) by omega. + eapply Mem.perm_inject; eauto. apply H. omega. + destruct (Mem.range_perm_free _ _ _ _ H1) as [m2' FREE]. + exists f; exists Vundef; exists m2'; intuition. + + econstructor. + rewrite EQ. replace (Int.signed lo + delta - 4) with (Int.signed lo - 4 + delta) by omega. + eauto. auto. + rewrite EQ. auto. + + assert (Mem.free_list m1 ((b, Int.signed lo - 4, Int.signed lo + Int.signed sz) :: nil) = Some m2). + simpl. rewrite H4. auto. + eapply Mem.free_inject; eauto. + intros. destruct (eq_block b b1). + subst b. assert (delta0 = delta) by congruence. subst delta0. + exists (Int.signed lo - 4); exists (Int.signed lo + Int.signed sz); split. + simpl; auto. omega. + elimtype False. + exploit Mem.inject_no_overlap. eauto. eauto. eauto. eauto. + instantiate (1 := ofs + delta0 - delta). + apply Mem.perm_implies with Freeable; auto with mem. + apply H. omega. eauto with mem. + unfold block; omega. + + eapply UNCHANGED; eauto. omega. intros. + red in H6. left. congruence. + + eapply UNCHANGED; eauto. omega. intros. + destruct (eq_block b' b2); auto. subst b'. right. + red in H6. generalize (H6 _ _ H5). intros. + exploit Mem.range_perm_in_bounds. eexact H. omega. intros. omega. + + red; intros. congruence. + + inv H; inv H0. intuition congruence. +Qed. + +(** ** Semantics of system calls. *) + +Inductive eventval_match: eventval -> typ -> val -> Prop := + | ev_match_int: + forall i, eventval_match (EVint i) Tint (Vint i) + | ev_match_float: + forall f, eventval_match (EVfloat f) Tfloat (Vfloat f). + +Inductive eventval_list_match: list eventval -> list typ -> list val -> Prop := + | evl_match_nil: + eventval_list_match nil nil nil + | evl_match_cons: + forall ev1 evl ty1 tyl v1 vl, + eventval_match ev1 ty1 v1 -> + eventval_list_match evl tyl vl -> + eventval_list_match (ev1::evl) (ty1::tyl) (v1::vl). + +Inductive extcall_io_sem (name: ident) (sg: signature): + list val -> mem -> trace -> val -> mem -> Prop := + | extcall_io_sem_intro: forall vargs m args res vres, + eventval_list_match args (sig_args sg) vargs -> + eventval_match res (proj_sig_res sg) vres -> + extcall_io_sem name sg vargs m (Event_syscall name args res :: E0) vres m. + +Remark eventval_match_lessdef: + forall ev ty v1 v2, + eventval_match ev ty v1 -> Val.lessdef v1 v2 -> eventval_match ev ty v2. +Proof. + intros. inv H; inv H0; constructor. +Qed. + +Remark eventval_list_match_lessdef: + forall evl tyl vl1, eventval_list_match evl tyl vl1 -> + forall vl2, Val.lessdef_list vl1 vl2 -> eventval_list_match evl tyl vl2. +Proof. + induction 1; intros. inv H; constructor. + inv H1. constructor. eapply eventval_match_lessdef; eauto. eauto. +Qed. + +Remark eventval_match_inject: + forall f ev ty v1 v2, + eventval_match ev ty v1 -> val_inject f v1 v2 -> eventval_match ev ty v2. +Proof. + intros. inv H; inv H0; constructor. +Qed. + +Remark eventval_match_inject_2: + forall f ev ty v, + eventval_match ev ty v -> val_inject f v v. +Proof. + induction 1; constructor. +Qed. + +Remark eventval_list_match_inject: + forall f evl tyl vl1, eventval_list_match evl tyl vl1 -> + forall vl2, val_list_inject f vl1 vl2 -> eventval_list_match evl tyl vl2. +Proof. + induction 1; intros. inv H; constructor. + inv H1. constructor. eapply eventval_match_inject; eauto. eauto. +Qed. + +Remark eventval_list_match_length: + forall evl tyl vl, eventval_list_match evl tyl vl -> List.length vl = List.length tyl. +Proof. + induction 1; simpl; eauto. +Qed. + +Remark eventval_match_determ_1: + forall ev ty v1 v2, eventval_match ev ty v1 -> eventval_match ev ty v2 -> v1 = v2. +Proof. + intros. inv H; inv H0; auto. +Qed. + +Remark eventval_match_determ_2: + forall ev1 ev2 ty v, eventval_match ev1 ty v -> eventval_match ev2 ty v -> ev1 = ev2. +Proof. + intros. inv H; inv H0; auto. +Qed. + +Remark eventval_list_match_determ_2: + forall evl1 tyl vl, eventval_list_match evl1 tyl vl -> + forall evl2, eventval_list_match evl2 tyl vl -> evl1 = evl2. +Proof. + induction 1; intros. inv H. auto. inv H1. f_equal; eauto. + eapply eventval_match_determ_2; eauto. +Qed. + +Lemma extcall_io_ok: + forall name sg, + extcall_properties (extcall_io_sem name sg) sg. +Proof. + intros; constructor; intros. + + inv H. inv H1; constructor. + + inv H. eapply eventval_list_match_length; eauto. + + inv H; auto. + + inv H; auto. + + inv H. + exists vres; exists m1'; intuition. + econstructor; eauto. eapply eventval_list_match_lessdef; eauto. + red; auto. + + inv H. + exists f; exists vres; exists m1'; intuition. + econstructor; eauto. eapply eventval_list_match_inject; eauto. + eapply eventval_match_inject_2; eauto. + red; auto. + red; auto. + red; intros; congruence. + + inv H; inv H0. simpl in H1. + assert (args = args0) by (eapply eventval_list_match_determ_2; eauto). + destruct H1; auto. subst. + intuition. eapply eventval_match_determ_1; eauto. +Qed. + +(** ** Combined semantics of external calls *) + +(** Combining the semantics given above for the various kinds of external calls, + we define the predicate [external_call] that relates: +- the external function being invoked +- the values of the arguments passed to this function +- the memory state before the call +- the result value of the call +- the memory state after the call +- the trace generated by the call (can be empty). + +This predicate is used in the semantics of all CompCert languages. *) + +Definition external_call (ef: external_function): + list val -> mem -> trace -> val -> mem -> Prop := + match classify_external_function ef with + | EF_syscall sg name => extcall_io_sem name sg + | EF_load label chunk => extcall_load_sem label chunk + | EF_store label chunk => extcall_store_sem label chunk + | EF_malloc => extcall_malloc_sem + | EF_free => extcall_free_sem + end. + +Theorem external_call_spec: + forall ef, + extcall_properties (external_call ef) (ef.(ef_sig)). +Proof. + intros. unfold external_call. destruct (classify_external_function ef). + apply extcall_io_ok. + apply extcall_load_ok. + apply extcall_store_ok. + apply extcall_malloc_ok. + apply extcall_free_ok. +Qed. + +Definition external_call_well_typed ef := ec_well_typed _ _ (external_call_spec ef). +Definition external_call_arity ef := ec_arity _ _ (external_call_spec ef). +Definition external_call_valid_block ef := ec_valid_block _ _ (external_call_spec ef). +Definition external_call_bounds ef := ec_bounds _ _ (external_call_spec ef). +Definition external_call_mem_extends ef := ec_mem_extends _ _ (external_call_spec ef). +Definition external_call_mem_inject ef := ec_mem_inject _ _ (external_call_spec ef). +Definition external_call_determ ef := ec_determ _ _ (external_call_spec ef). diff --git a/common/Globalenvs.v b/common/Globalenvs.v index 1ce7bf5..9dbf902 100644 --- a/common/Globalenvs.v +++ b/common/Globalenvs.v @@ -38,530 +38,259 @@ Require Import Errors. Require Import Maps. Require Import AST. Require Import Integers. +Require Import Floats. Require Import Values. -Require Import Mem. +Require Import Memory. -Set Implicit Arguments. +Notation "s #1" := (fst s) (at level 9, format "s '#1'") : pair_scope. +Notation "s #2" := (snd s) (at level 9, format "s '#2'") : pair_scope. -Module Type GENV. - -(** ** Types and operations *) - - Variable t: Type -> Type. - (** The type of global environments. The parameter [F] is the type - of function descriptions. *) - - Variable globalenv: forall (F V: Type), program F V -> t F. - (** Return the global environment for the given program. *) - - Variable init_mem: forall (F V: Type), program F V -> mem. - (** Return the initial memory state for the given program. *) - - Variable find_funct_ptr: forall (F: Type), t F -> block -> option F. - (** Return the function description associated with the given address, - if any. *) - - Variable find_funct: forall (F: Type), t F -> val -> option F. - (** Same as [find_funct_ptr] but the function address is given as - a value, which must be a pointer with offset 0. *) - - Variable find_symbol: forall (F: Type), t F -> ident -> option block. - (** Return the address of the given global symbol, if any. *) - -(** ** Properties of the operations. *) - - Hypothesis find_funct_inv: - forall (F: Type) (ge: t F) (v: val) (f: F), - find_funct ge v = Some f -> exists b, v = Vptr b Int.zero. - Hypothesis find_funct_find_funct_ptr: - forall (F: Type) (ge: t F) (b: block), - find_funct ge (Vptr b Int.zero) = find_funct_ptr ge b. - - Hypothesis find_symbol_exists: - forall (F V: Type) (p: program F V) - (id: ident) (init: list init_data) (v: V), - In (id, init, v) (prog_vars p) -> - exists b, find_symbol (globalenv p) id = Some b. - Hypothesis find_funct_ptr_exists: - forall (F V: Type) (p: program F V) (id: ident) (f: F), - list_norepet (prog_funct_names p) -> - list_disjoint (prog_funct_names p) (prog_var_names p) -> - In (id, f) (prog_funct p) -> - exists b, find_symbol (globalenv p) id = Some b - /\ find_funct_ptr (globalenv p) b = Some f. - - Hypothesis find_funct_ptr_inversion: - forall (F V: Type) (P: F -> Prop) (p: program F V) (b: block) (f: F), - find_funct_ptr (globalenv p) b = Some f -> - exists id, In (id, f) (prog_funct p). - Hypothesis find_funct_inversion: - forall (F V: Type) (P: F -> Prop) (p: program F V) (v: val) (f: F), - find_funct (globalenv p) v = Some f -> - exists id, In (id, f) (prog_funct p). - Hypothesis find_funct_ptr_symbol_inversion: - forall (F V: Type) (p: program F V) (id: ident) (b: block) (f: F), - find_symbol (globalenv p) id = Some b -> - find_funct_ptr (globalenv p) b = Some f -> - In (id, f) p.(prog_funct). - - Hypothesis find_funct_ptr_prop: - forall (F V: Type) (P: F -> Prop) (p: program F V) (b: block) (f: F), - (forall id f, In (id, f) (prog_funct p) -> P f) -> - find_funct_ptr (globalenv p) b = Some f -> - P f. - Hypothesis find_funct_prop: - forall (F V: Type) (P: F -> Prop) (p: program F V) (v: val) (f: F), - (forall id f, In (id, f) (prog_funct p) -> P f) -> - find_funct (globalenv p) v = Some f -> - P f. - - Hypothesis initmem_nullptr: - forall (F V: Type) (p: program F V), - let m := init_mem p in - valid_block m nullptr /\ - m.(blocks) nullptr = empty_block 0 0. - Hypothesis initmem_inject_neutral: - forall (F V: Type) (p: program F V), - mem_inject_neutral (init_mem p). - Hypothesis find_funct_ptr_negative: - forall (F V: Type) (p: program F V) (b: block) (f: F), - find_funct_ptr (globalenv p) b = Some f -> b < 0. - Hypothesis find_symbol_not_fresh: - forall (F V: Type) (p: program F V) (id: ident) (b: block), - find_symbol (globalenv p) id = Some b -> b < nextblock (init_mem p). - Hypothesis find_symbol_not_nullptr: - forall (F V: Type) (p: program F V) (id: ident) (b: block), - find_symbol (globalenv p) id = Some b -> b <> nullptr. - Hypothesis global_addresses_distinct: - forall (F V: Type) (p: program F V) id1 id2 b1 b2, - id1<>id2 -> - find_symbol (globalenv p) id1 = Some b1 -> - find_symbol (globalenv p) id2 = Some b2 -> - b1<>b2. - -(** Commutation properties between program transformations - and operations over global environments. *) - - Hypothesis find_funct_ptr_transf: - forall (A B V: Type) (transf: A -> B) (p: program A V), - forall (b: block) (f: A), - find_funct_ptr (globalenv p) b = Some f -> - find_funct_ptr (globalenv (transform_program transf p)) b = Some (transf f). - Hypothesis find_funct_transf: - forall (A B V: Type) (transf: A -> B) (p: program A V), - forall (v: val) (f: A), - find_funct (globalenv p) v = Some f -> - find_funct (globalenv (transform_program transf p)) v = Some (transf f). - Hypothesis find_symbol_transf: - forall (A B V: Type) (transf: A -> B) (p: program A V), - forall (s: ident), - find_symbol (globalenv (transform_program transf p)) s = - find_symbol (globalenv p) s. - Hypothesis init_mem_transf: - forall (A B V: Type) (transf: A -> B) (p: program A V), - init_mem (transform_program transf p) = init_mem p. - Hypothesis find_funct_ptr_rev_transf: - forall (A B V: Type) (transf: A -> B) (p: program A V), - forall (b : block) (tf : B), - find_funct_ptr (globalenv (transform_program transf p)) b = Some tf -> - exists f : A, find_funct_ptr (globalenv p) b = Some f /\ transf f = tf. - Hypothesis find_funct_rev_transf: - forall (A B V: Type) (transf: A -> B) (p: program A V), - forall (v : val) (tf : B), - find_funct (globalenv (transform_program transf p)) v = Some tf -> - exists f : A, find_funct (globalenv p) v = Some f /\ transf f = tf. - -(** Commutation properties between partial program transformations - and operations over global environments. *) - - Hypothesis find_funct_ptr_transf_partial: - forall (A B V: Type) (transf: A -> res B) (p: program A V) (p': program B V), - transform_partial_program transf p = OK p' -> - forall (b: block) (f: A), - find_funct_ptr (globalenv p) b = Some f -> - exists f', - find_funct_ptr (globalenv p') b = Some f' /\ transf f = OK f'. - Hypothesis find_funct_transf_partial: - forall (A B V: Type) (transf: A -> res B) (p: program A V) (p': program B V), - transform_partial_program transf p = OK p' -> - forall (v: val) (f: A), - find_funct (globalenv p) v = Some f -> - exists f', - find_funct (globalenv p') v = Some f' /\ transf f = OK f'. - Hypothesis find_symbol_transf_partial: - forall (A B V: Type) (transf: A -> res B) (p: program A V) (p': program B V), - transform_partial_program transf p = OK p' -> - forall (s: ident), - find_symbol (globalenv p') s = find_symbol (globalenv p) s. - Hypothesis init_mem_transf_partial: - forall (A B V: Type) (transf: A -> res B) (p: program A V) (p': program B V), - transform_partial_program transf p = OK p' -> - init_mem p' = init_mem p. - Hypothesis find_funct_ptr_rev_transf_partial: - forall (A B V: Type) (transf: A -> res B) (p: program A V) (p': program B V), - transform_partial_program transf p = OK p' -> - forall (b : block) (tf : B), - find_funct_ptr (globalenv p') b = Some tf -> - exists f : A, - find_funct_ptr (globalenv p) b = Some f /\ transf f = OK tf. - Hypothesis find_funct_rev_transf_partial: - forall (A B V: Type) (transf: A -> res B) (p: program A V) (p': program B V), - transform_partial_program transf p = OK p' -> - forall (v : val) (tf : B), - find_funct (globalenv p') v = Some tf -> - exists f : A, - find_funct (globalenv p) v = Some f /\ transf f = OK tf. - - Hypothesis find_funct_ptr_transf_partial2: - forall (A B V W: Type) (transf_fun: A -> res B) (transf_var: V -> res W) - (p: program A V) (p': program B W), - transform_partial_program2 transf_fun transf_var p = OK p' -> - forall (b: block) (f: A), - find_funct_ptr (globalenv p) b = Some f -> - exists f', - find_funct_ptr (globalenv p') b = Some f' /\ transf_fun f = OK f'. - Hypothesis find_funct_transf_partial2: - forall (A B V W: Type) (transf_fun: A -> res B) (transf_var: V -> res W) - (p: program A V) (p': program B W), - transform_partial_program2 transf_fun transf_var p = OK p' -> - forall (v: val) (f: A), - find_funct (globalenv p) v = Some f -> - exists f', - find_funct (globalenv p') v = Some f' /\ transf_fun f = OK f'. - Hypothesis find_symbol_transf_partial2: - forall (A B V W: Type) (transf_fun: A -> res B) (transf_var: V -> res W) - (p: program A V) (p': program B W), - transform_partial_program2 transf_fun transf_var p = OK p' -> - forall (s: ident), - find_symbol (globalenv p') s = find_symbol (globalenv p) s. - Hypothesis init_mem_transf_partial2: - forall (A B V W: Type) (transf_fun: A -> res B) (transf_var: V -> res W) - (p: program A V) (p': program B W), - transform_partial_program2 transf_fun transf_var p = OK p' -> - init_mem p' = init_mem p. - Hypothesis find_funct_ptr_rev_transf_partial2: - forall (A B V W: Type) (transf_fun: A -> res B) (transf_var: V -> res W) - (p: program A V) (p': program B W), - transform_partial_program2 transf_fun transf_var p = OK p' -> - forall (b : block) (tf : B), - find_funct_ptr (globalenv p') b = Some tf -> - exists f : A, - find_funct_ptr (globalenv p) b = Some f /\ transf_fun f = OK tf. - Hypothesis find_funct_rev_transf_partial2: - forall (A B V W: Type) (transf_fun: A -> res B) (transf_var: V -> res W) - (p: program A V) (p': program B W), - transform_partial_program2 transf_fun transf_var p = OK p' -> - forall (v : val) (tf : B), - find_funct (globalenv p') v = Some tf -> - exists f : A, - find_funct (globalenv p) v = Some f /\ transf_fun f = OK tf. - -(** Commutation properties between matching between programs - and operations over global environments. *) - - Hypothesis find_funct_ptr_match: - forall (A B V W: Type) (match_fun: A -> B -> Prop) - (match_var: V -> W -> Prop) (p: program A V) (p': program B W), - match_program match_fun match_var p p' -> - forall (b : block) (f : A), - find_funct_ptr (globalenv p) b = Some f -> - exists tf : B, - find_funct_ptr (globalenv p') b = Some tf /\ match_fun f tf. - Hypothesis find_funct_ptr_rev_match: - forall (A B V W: Type) (match_fun: A -> B -> Prop) - (match_var: V -> W -> Prop) (p: program A V) (p': program B W), - match_program match_fun match_var p p' -> - forall (b : block) (tf : B), - find_funct_ptr (globalenv p') b = Some tf -> - exists f : A, - find_funct_ptr (globalenv p) b = Some f /\ match_fun f tf. - Hypothesis find_funct_match: - forall (A B V W: Type) (match_fun: A -> B -> Prop) - (match_var: V -> W -> Prop) (p: program A V) (p': program B W), - match_program match_fun match_var p p' -> - forall (v : val) (f : A), - find_funct (globalenv p) v = Some f -> - exists tf : B, find_funct (globalenv p') v = Some tf /\ match_fun f tf. - Hypothesis find_funct_rev_match: - forall (A B V W: Type) (match_fun: A -> B -> Prop) - (match_var: V -> W -> Prop) (p: program A V) (p': program B W), - match_program match_fun match_var p p' -> - forall (v : val) (tf : B), - find_funct (globalenv p') v = Some tf -> - exists f : A, find_funct (globalenv p) v = Some f /\ match_fun f tf. - Hypothesis find_symbol_match: - forall (A B V W: Type) (match_fun: A -> B -> Prop) - (match_var: V -> W -> Prop) (p: program A V) (p': program B W), - match_program match_fun match_var p p' -> - forall (s : ident), - find_symbol (globalenv p') s = find_symbol (globalenv p) s. - Hypothesis init_mem_match: - forall (A B V W: Type) (match_fun: A -> B -> Prop) - (match_var: V -> W -> Prop) (p: program A V) (p': program B W), - match_program match_fun match_var p p' -> - init_mem p' = init_mem p. +Local Open Scope pair_scope. +Local Open Scope error_monad_scope. -End GENV. +Set Implicit Arguments. -(** The rest of this library is a straightforward implementation of - the module signature above. *) +Module Genv. -Module Genv: GENV. +(** * Global environments *) Section GENV. -Variable F: Type. (* The type of functions *) -Variable V: Type. (* The type of information over variables *) - -Record genv : Type := mkgenv { - functions: ZMap.t (option F); (* mapping function ptr -> function *) - nextfunction: Z; - symbols: PTree.t block (* mapping symbol -> block *) +Variable F: Type. (**r The type of function descriptions *) +Variable V: Type. (**r The type of information attached to variables *) + +(** The type of global environments. *) + +Record t: Type := mkgenv { + genv_symb: PTree.t block; (**r mapping symbol -> block *) + genv_funs: ZMap.t (option F); (**r mapping function pointer -> definition *) + genv_vars: ZMap.t (option V); (**r mapping variable pointer -> info *) + genv_nextfun: block; (**r next function pointer *) + genv_nextvar: block; (**r next variable pointer *) + genv_nextfun_neg: genv_nextfun < 0; + genv_nextvar_pos: genv_nextvar > 0; + genv_symb_range: forall id b, PTree.get id genv_symb = Some b -> b <> 0 /\ genv_nextfun < b /\ b < genv_nextvar; + genv_funs_range: forall b f, ZMap.get b genv_funs = Some f -> genv_nextfun < b < 0; + genv_vars_range: forall b v, ZMap.get b genv_vars = Some v -> 0 < b < genv_nextvar }. -Definition t := genv. +(** ** Lookup functions *) -Definition add_funct (name_fun: (ident * F)) (g: genv) : genv := - let b := g.(nextfunction) in - mkgenv (ZMap.set b (Some (snd name_fun)) g.(functions)) - (Zpred b) - (PTree.set (fst name_fun) b g.(symbols)). +(** [find_symbol ge id] returns the block associated with the given name, if any *) -Definition add_symbol (name: ident) (b: block) (g: genv) : genv := - mkgenv g.(functions) - g.(nextfunction) - (PTree.set name b g.(symbols)). +Definition find_symbol (ge: t) (id: ident) : option block := + PTree.get id ge.(genv_symb). -Definition find_funct_ptr (g: genv) (b: block) : option F := - ZMap.get b g.(functions). +(** [find_funct_ptr ge b] returns the function description associated with + the given address. *) -Definition find_funct (g: genv) (v: val) : option F := +Definition find_funct_ptr (ge: t) (b: block) : option F := + ZMap.get b ge.(genv_funs). + +(** [find_funct] is similar to [find_funct_ptr], but the function address + is given as a value, which must be a pointer with offset 0. *) + +Definition find_funct (ge: t) (v: val) : option F := match v with - | Vptr b ofs => - if Int.eq ofs Int.zero then find_funct_ptr g b else None - | _ => - None + | Vptr b ofs => if Int.eq_dec ofs Int.zero then find_funct_ptr ge b else None + | _ => None end. -Definition find_symbol (g: genv) (symb: ident) : option block := - PTree.get symb g.(symbols). +(** [find_var_info ge b] returns the information attached to the variable + at address [b]. *) + +Definition find_var_info (ge: t) (b: block) : option V := + ZMap.get b ge.(genv_vars). + +(** ** Constructing the global environment *) + +Program Definition add_function (ge: t) (idf: ident * F) : t := + @mkgenv + (PTree.set idf#1 ge.(genv_nextfun) ge.(genv_symb)) + (ZMap.set ge.(genv_nextfun) (Some idf#2) ge.(genv_funs)) + ge.(genv_vars) + (ge.(genv_nextfun) - 1) + ge.(genv_nextvar) + _ _ _ _ _. +Next Obligation. + destruct ge; simpl; omega. +Qed. +Next Obligation. + destruct ge; auto. +Qed. +Next Obligation. + destruct ge; simpl in *. + rewrite PTree.gsspec in H. destruct (peq id i). inv H. unfold block; omega. + exploit genv_symb_range0; eauto. unfold block; omega. +Qed. +Next Obligation. + destruct ge; simpl in *. rewrite ZMap.gsspec in H. + destruct (ZIndexed.eq b genv_nextfun0). subst; omega. + exploit genv_funs_range0; eauto. omega. +Qed. +Next Obligation. + destruct ge; eauto. +Qed. + +Program Definition add_variable (ge: t) (idv: ident * list init_data * V) : t := + @mkgenv + (PTree.set idv#1#1 ge.(genv_nextvar) ge.(genv_symb)) + ge.(genv_funs) + (ZMap.set ge.(genv_nextvar) (Some idv#2) ge.(genv_vars)) + ge.(genv_nextfun) + (ge.(genv_nextvar) + 1) + _ _ _ _ _. +Next Obligation. + destruct ge; auto. +Qed. +Next Obligation. + destruct ge; simpl; omega. +Qed. +Next Obligation. + destruct ge; simpl in *. + rewrite PTree.gsspec in H. destruct (peq id i). inv H. unfold block; omega. + exploit genv_symb_range0; eauto. unfold block; omega. +Qed. +Next Obligation. + destruct ge; eauto. +Qed. +Next Obligation. + destruct ge; simpl in *. rewrite ZMap.gsspec in H. + destruct (ZIndexed.eq b genv_nextvar0). subst; omega. + exploit genv_vars_range0; eauto. omega. +Qed. + +Program Definition empty_genv : t := + @mkgenv (PTree.empty block) (ZMap.init None) (ZMap.init None) (-1) 1 _ _ _ _ _. +Next Obligation. + omega. +Qed. +Next Obligation. + omega. +Qed. +Next Obligation. + rewrite PTree.gempty in H. discriminate. +Qed. +Next Obligation. + rewrite ZMap.gi in H. discriminate. +Qed. +Next Obligation. + rewrite ZMap.gi in H. discriminate. +Qed. + +Definition add_functions (ge: t) (fl: list (ident * F)) : t := + List.fold_left add_function fl ge. + +Definition add_variables (ge: t) (vl: list (ident * list init_data * V)) : t := + List.fold_left add_variable vl ge. + +Definition globalenv (p: program F V) := + add_variables (add_functions empty_genv p.(prog_funct)) p.(prog_vars). -Lemma find_funct_inv: - forall (ge: t) (v: val) (f: F), +(** ** Properties of the operations over global environments *) + +Theorem find_funct_inv: + forall ge v f, find_funct ge v = Some f -> exists b, v = Vptr b Int.zero. Proof. - intros until f. unfold find_funct. destruct v; try (intros; discriminate). - generalize (Int.eq_spec i Int.zero). case (Int.eq i Int.zero); intros. - exists b. congruence. - discriminate. -Qed. - -Lemma find_funct_find_funct_ptr: - forall (ge: t) (b: block), - find_funct ge (Vptr b Int.zero) = find_funct_ptr ge b. -Proof. - intros. simpl. - generalize (Int.eq_spec Int.zero Int.zero). - case (Int.eq Int.zero Int.zero); intros. - auto. tauto. -Qed. - -(* Construct environment and initial memory store *) - -Definition empty : genv := - mkgenv (ZMap.init None) (-1) (PTree.empty block). - -Definition add_functs (init: genv) (fns: list (ident * F)) : genv := - List.fold_right add_funct init fns. - -Definition add_globals - (init: genv * mem) (vars: list (ident * list init_data * V)) - : genv * mem := - List.fold_right - (fun (id_init: ident * list init_data * V) (g_st: genv * mem) => - match id_init, g_st with - | ((id, init), info), (g, st) => - let (st', b) := Mem.alloc_init_data st init in - (add_symbol id b g, st') - end) - init vars. - -Definition globalenv_initmem (p: program F V) : (genv * mem) := - add_globals - (add_functs empty p.(prog_funct), Mem.empty) - p.(prog_vars). - -Definition globalenv (p: program F V) : genv := - fst (globalenv_initmem p). -Definition init_mem (p: program F V) : mem := - snd (globalenv_initmem p). - -Lemma functions_globalenv: - forall (p: program F V), - functions (globalenv p) = functions (add_functs empty p.(prog_funct)). -Proof. - assert (forall init vars, - functions (fst (add_globals init vars)) = functions (fst init)). - induction vars; simpl. - reflexivity. - destruct a as [[id1 init1] info1]. destruct (add_globals init vars). - simpl. exact IHvars. - - unfold add_globals; simpl. - intros. unfold globalenv; unfold globalenv_initmem. - rewrite H. reflexivity. -Qed. - -Lemma initmem_nullptr: - forall (p: program F V), - let m := init_mem p in - valid_block m nullptr /\ - m.(blocks) nullptr = mkblock 0 0 (fun y => Undef). -Proof. - pose (P := fun m => valid_block m nullptr /\ - m.(blocks) nullptr = mkblock 0 0 (fun y => Undef)). - assert (forall init, P (snd init) -> forall vars, P (snd (add_globals init vars))). - induction vars; simpl; intros. - auto. - destruct a as [[id1 in1] inf1]. - destruct (add_globals init vars) as [g st]. - simpl in *. destruct IHvars. split. - red; simpl. red in H0. omega. - simpl. rewrite update_o. auto. unfold block. red in H0. omega. - - intro. unfold init_mem, globalenv_initmem. apply H. - red; simpl. split. compute. auto. reflexivity. -Qed. + intros until f; unfold find_funct. + destruct v; try congruence. + destruct (Int.eq_dec i Int.zero); try congruence. + intros. exists b; congruence. +Qed. -Lemma initmem_inject_neutral: - forall (p: program F V), - mem_inject_neutral (init_mem p). -Proof. - assert (forall g0 vars g1 m, - add_globals (g0, Mem.empty) vars = (g1, m) -> - mem_inject_neutral m). - Opaque alloc_init_data. - induction vars; simpl. - intros. inv H. red; intros. destruct (load_inv _ _ _ _ _ H). - simpl in H1. rewrite Mem.getN_init in H1. - replace v with Vundef. auto. destruct chunk; simpl in H1; auto. - destruct a as [[id1 init1] info1]. - caseEq (add_globals (g0, Mem.empty) vars). intros g1 m1 EQ. - caseEq (alloc_init_data m1 init1). intros m2 b ALLOC. - intros. inv H. - eapply Mem.alloc_init_data_neutral; eauto. - intros. caseEq (globalenv_initmem p). intros g m EQ. - unfold init_mem; rewrite EQ; simpl. - unfold globalenv_initmem in EQ. eauto. -Qed. - -Remark nextfunction_add_functs_neg: - forall fns, nextfunction (add_functs empty fns) < 0. -Proof. - induction fns; simpl; intros. omega. unfold Zpred. omega. +Theorem find_funct_find_funct_ptr: + forall ge b, + find_funct ge (Vptr b Int.zero) = find_funct_ptr ge b. +Proof. + intros; simpl. apply dec_eq_true. Qed. -Theorem find_funct_ptr_negative: - forall (p: program F V) (b: block) (f: F), - find_funct_ptr (globalenv p) b = Some f -> b < 0. +Theorem find_symbol_exists: + forall p id init v, + In (id, init, v) (prog_vars p) -> + exists b, find_symbol (globalenv p) id = Some b. Proof. - intros until f. - assert (forall fns, ZMap.get b (functions (add_functs empty fns)) = Some f -> b < 0). - induction fns; simpl. - rewrite ZMap.gi. congruence. - rewrite ZMap.gsspec. case (ZIndexed.eq b (nextfunction (add_functs empty fns))); intro. - intro. rewrite e. apply nextfunction_add_functs_neg. - auto. - unfold find_funct_ptr. rewrite functions_globalenv. - intros. eauto. -Qed. - -Remark find_symbol_add_functs_negative: - forall (fns: list (ident * F)) s b, - (symbols (add_functs empty fns)) ! s = Some b -> b < 0. -Proof. - induction fns; simpl; intros until b. - rewrite PTree.gempty. congruence. - rewrite PTree.gsspec. destruct a; simpl. case (peq s i); intro. - intro EQ; inversion EQ. apply nextfunction_add_functs_neg. + intros until v. + assert (forall vl ge, + (exists b, find_symbol ge id = Some b) -> + exists b, find_symbol (add_variables ge vl) id = Some b). + unfold find_symbol; induction vl; simpl; intros. auto. apply IHvl. + simpl. rewrite PTree.gsspec. fold ident. destruct (peq id a#1#1). + exists (genv_nextvar ge); auto. auto. + + assert (forall vl ge, In (id, init, v) vl -> + exists b, find_symbol (add_variables ge vl) id = Some b). + unfold find_symbol; induction vl; simpl; intros. contradiction. + destruct H0. apply H. subst; unfold find_symbol; simpl. + rewrite PTree.gss. exists (genv_nextvar ge); auto. eauto. + + intros. unfold globalenv; eauto. Qed. -Remark find_symbol_add_symbols_not_fresh: - forall fns vars g m s b, - add_globals (add_functs empty fns, Mem.empty) vars = (g, m) -> - (symbols g)!s = Some b -> - b < nextblock m. +Remark add_functions_same_symb: + forall id fl ge, + ~In id (map (@fst ident F) fl) -> + find_symbol (add_functions ge fl) id = find_symbol ge id. Proof. - induction vars; simpl; intros until b. - intros. inversion H. subst g m. simpl. - generalize (find_symbol_add_functs_negative fns s H0). omega. - Transparent alloc_init_data. - destruct a as [[id1 init1] info1]. - caseEq (add_globals (add_functs empty fns, Mem.empty) vars). - intros g1 m1 ADG EQ. inversion EQ; subst g m; clear EQ. - unfold add_symbol; simpl. rewrite PTree.gsspec. case (peq s id1); intro. - intro EQ; inversion EQ. omega. - intro. generalize (IHvars _ _ _ _ ADG H). omega. + induction fl; simpl; intros. auto. + rewrite IHfl. unfold find_symbol; simpl. apply PTree.gso. intuition. intuition. Qed. -Theorem find_symbol_not_fresh: - forall (p: program F V) (id: ident) (b: block), - find_symbol (globalenv p) id = Some b -> b < nextblock (init_mem p). +Remark add_functions_same_address: + forall b fl ge, + b > ge.(genv_nextfun) -> + find_funct_ptr (add_functions ge fl) b = find_funct_ptr ge b. Proof. - intros until b. unfold find_symbol, globalenv, init_mem, globalenv_initmem; simpl. - caseEq (add_globals (add_functs empty (prog_funct p), Mem.empty) - (prog_vars p)); intros g m EQ. - simpl; intros. eapply find_symbol_add_symbols_not_fresh; eauto. + induction fl; simpl; intros. auto. + rewrite IHfl. unfold find_funct_ptr; simpl. apply ZMap.gso. + red; intros; subst b; omegaContradiction. + simpl. omega. Qed. -Lemma find_symbol_exists: - forall (p: program F V) - (id: ident) (init: list init_data) (v: V), - In (id, init, v) (prog_vars p) -> - exists b, find_symbol (globalenv p) id = Some b. +Remark add_variables_same_symb: + forall id vl ge, + ~In id (map (fun idv => idv#1#1) vl) -> + find_symbol (add_variables ge vl) id = find_symbol ge id. Proof. - intros until v. - assert (forall initm vl, In (id, init, v) vl -> - exists b, PTree.get id (symbols (fst (add_globals initm vl))) = Some b). - induction vl; simpl; intros. - elim H. - destruct a as [[id0 init0] v0]. - caseEq (add_globals initm vl). intros g1 m1 EQ. simpl. - rewrite PTree.gsspec. destruct (peq id id0). econstructor; eauto. - elim H; intro. congruence. generalize (IHvl H0). rewrite EQ. auto. - intros. unfold globalenv, find_symbol, globalenv_initmem. auto. -Qed. - -Remark find_symbol_above_nextfunction: - forall (id: ident) (b: block) (fns: list (ident * F)), - let g := add_functs empty fns in - PTree.get id g.(symbols) = Some b -> - b > g.(nextfunction). -Proof. - induction fns; simpl. - rewrite PTree.gempty. congruence. - rewrite PTree.gsspec. case (peq id (fst a)); intro. - intro EQ. inversion EQ. unfold Zpred. omega. - intros. generalize (IHfns H). unfold Zpred; omega. -Qed. - -Remark find_symbol_add_globals: - forall (id: ident) (ge_m: t * mem) (vars: list (ident * list init_data * V)), - ~In id (map (fun x: ident * list init_data * V => fst(fst x)) vars) -> - find_symbol (fst (add_globals ge_m vars)) id = - find_symbol (fst ge_m) id. -Proof. - unfold find_symbol; induction vars; simpl; intros. - auto. - destruct a as [[id0 init0] var0]. simpl in *. - caseEq (add_globals ge_m vars); intros ge' m' EQ. - simpl. rewrite PTree.gso. rewrite EQ in IHvars. simpl in IHvars. - apply IHvars. tauto. intuition congruence. + induction vl; simpl; intros. auto. + rewrite IHvl. unfold find_symbol; simpl. apply PTree.gso. intuition. intuition. +Qed. + +Remark add_variables_same_address: + forall b vl ge, + b < ge.(genv_nextvar) -> + find_var_info (add_variables ge vl) b = find_var_info ge b. +Proof. + induction vl; simpl; intros. auto. + rewrite IHvl. unfold find_var_info; simpl. apply ZMap.gso. + red; intros; subst b; omegaContradiction. + simpl. omega. +Qed. + +Remark add_variables_same_funs: + forall b vl ge, find_funct_ptr (add_variables ge vl) b = find_funct_ptr ge b. +Proof. + induction vl; simpl; intros. auto. rewrite IHvl. auto. +Qed. + +Remark add_functions_nextvar: + forall fl ge, genv_nextvar (add_functions ge fl) = genv_nextvar ge. +Proof. + induction fl; simpl; intros. auto. rewrite IHfl. auto. +Qed. + +Remark add_variables_nextvar: + forall vl ge, genv_nextvar (add_variables ge vl) = genv_nextvar ge + Z_of_nat(List.length vl). +Proof. + induction vl; intros. + simpl. unfold block; omega. + simpl length; rewrite inj_S; simpl. rewrite IHvl. simpl. unfold block; omega. Qed. -Lemma find_funct_ptr_exists: - forall (p: program F V) (id: ident) (f: F), +Theorem find_funct_ptr_exists: + forall p id f, list_norepet (prog_funct_names p) -> list_disjoint (prog_funct_names p) (prog_var_names p) -> In (id, f) (prog_funct p) -> @@ -569,384 +298,784 @@ Lemma find_funct_ptr_exists: /\ find_funct_ptr (globalenv p) b = Some f. Proof. intros until f. - assert (forall (fns: list (ident * F)), - list_norepet (map (@fst ident F) fns) -> - In (id, f) fns -> - exists b, find_symbol (add_functs empty fns) id = Some b - /\ find_funct_ptr (add_functs empty fns) b = Some f). - unfold find_symbol, find_funct_ptr. induction fns; intros. - elim H0. - destruct a as [id0 f0]; simpl in *. inv H. - unfold add_funct; simpl. - rewrite PTree.gsspec. destruct (peq id id0). - subst id0. econstructor; split. eauto. - replace f0 with f. apply ZMap.gss. - elim H0; intro. congruence. elim H3. - change id with (@fst ident F (id, f)). apply List.in_map. auto. - exploit IHfns; eauto. elim H0; intro. congruence. auto. - intros [b [X Y]]. exists b; split. auto. rewrite ZMap.gso. auto. - generalize (find_symbol_above_nextfunction _ _ X). - unfold block; unfold ZIndexed.t; intro; omega. - - intros. exploit H; eauto. intros [b [X Y]]. - exists b; split. - unfold globalenv, globalenv_initmem. rewrite find_symbol_add_globals. - assumption. apply list_disjoint_notin with (prog_funct_names p). assumption. - unfold prog_funct_names. change id with (fst (id, f)). - apply List.in_map; auto. - unfold find_funct_ptr. rewrite functions_globalenv. - assumption. -Qed. - -Lemma find_funct_ptr_inversion: - forall (P: F -> Prop) (p: program F V) (b: block) (f: F), - find_funct_ptr (globalenv p) b = Some f -> - exists id, In (id, f) (prog_funct p). -Proof. - intros until f. - assert (forall fns: list (ident * F), - find_funct_ptr (add_functs empty fns) b = Some f -> - exists id, In (id, f) fns). - unfold find_funct_ptr. induction fns; simpl. - rewrite ZMap.gi. congruence. - destruct a as [id0 f0]; simpl. - rewrite ZMap.gsspec. destruct (ZIndexed.eq b (nextfunction (add_functs empty fns))). - intro. inv H. exists id0; auto. - intro. exploit IHfns; eauto. intros [id A]. exists id; auto. - unfold find_funct_ptr; rewrite functions_globalenv. intros; apply H; auto. -Qed. - -Lemma find_funct_ptr_prop: - forall (P: F -> Prop) (p: program F V) (b: block) (f: F), + + assert (forall fl ge, In (id, f) fl -> list_norepet (map (@fst ident F) fl) -> + exists b, find_symbol (add_functions ge fl) id = Some b + /\ find_funct_ptr (add_functions ge fl) b = Some f). + induction fl; simpl; intros. contradiction. inv H0. + destruct H. subst a. exists (genv_nextfun ge); split. + rewrite add_functions_same_symb; auto. unfold find_symbol; simpl. apply PTree.gss. + rewrite add_functions_same_address. unfold find_funct_ptr; simpl. apply ZMap.gss. + simpl; omega. + apply IHfl; auto. + + intros. exploit (H p.(prog_funct) empty_genv); eauto. intros [b [A B]]. + unfold globalenv; exists b; split. + rewrite add_variables_same_symb. auto. eapply list_disjoint_notin; eauto. + unfold prog_funct_names. change id with (fst (id, f)). apply in_map; auto. + rewrite add_variables_same_funs. auto. +Qed. + +Theorem find_funct_ptr_prop: + forall (P: F -> Prop) p b f, (forall id f, In (id, f) (prog_funct p) -> P f) -> find_funct_ptr (globalenv p) b = Some f -> P f. Proof. - intros. exploit find_funct_ptr_inversion; eauto. intros [id A]. eauto. + intros until f. intros PROP. + assert (forall fl ge, + List.incl fl (prog_funct p) -> + match find_funct_ptr ge b with None => True | Some f => P f end -> + match find_funct_ptr (add_functions ge fl) b with None => True | Some f => P f end). + induction fl; simpl; intros. auto. + apply IHfl. eauto with coqlib. unfold find_funct_ptr; simpl. + destruct a as [id' f']; simpl. + rewrite ZMap.gsspec. destruct (ZIndexed.eq b (genv_nextfun ge)). + apply PROP with id'. apply H. auto with coqlib. + assumption. + + unfold globalenv. rewrite add_variables_same_funs. intro. + exploit (H p.(prog_funct) empty_genv). auto with coqlib. + unfold find_funct_ptr; simpl. rewrite ZMap.gi. auto. + rewrite H0. auto. Qed. -Lemma find_funct_inversion: - forall (P: F -> Prop) (p: program F V) (v: val) (f: F), +Theorem find_funct_prop: + forall (P: F -> Prop) p v f, + (forall id f, In (id, f) (prog_funct p) -> P f) -> find_funct (globalenv p) v = Some f -> + P f. +Proof. + intros. exploit find_funct_inv; eauto. intros [b EQ]. subst v. + rewrite find_funct_find_funct_ptr in H0. + eapply find_funct_ptr_prop; eauto. +Qed. + +Theorem find_funct_ptr_inversion: + forall p b f, + find_funct_ptr (globalenv p) b = Some f -> exists id, In (id, f) (prog_funct p). Proof. - intros. exploit find_funct_inv; eauto. intros [b EQ]. rewrite EQ in H. - rewrite find_funct_find_funct_ptr in H. - eapply find_funct_ptr_inversion; eauto. + intros. pattern f. apply find_funct_ptr_prop with p b; auto. + intros. exists id; auto. Qed. -Lemma find_funct_prop: - forall (P: F -> Prop) (p: program F V) (v: val) (f: F), - (forall id f, In (id, f) (prog_funct p) -> P f) -> +Theorem find_funct_inversion: + forall p v f, find_funct (globalenv p) v = Some f -> - P f. + exists id, In (id, f) (prog_funct p). +Proof. + intros. pattern f. apply find_funct_prop with p v; auto. + intros. exists id; auto. +Qed. + +Theorem find_funct_ptr_negative: + forall p b f, + find_funct_ptr (globalenv p) b = Some f -> b < 0. Proof. - intros. exploit find_funct_inversion; eauto. intros [id A]. eauto. + unfold find_funct_ptr. intros. destruct (globalenv p). simpl in H. + exploit genv_funs_range0; eauto. omega. Qed. -Lemma find_funct_ptr_symbol_inversion: - forall (p: program F V) (id: ident) (b: block) (f: F), +Theorem find_var_info_positive: + forall p b v, + find_var_info (globalenv p) b = Some v -> b > 0. +Proof. + unfold find_var_info. intros. destruct (globalenv p). simpl in H. + exploit genv_vars_range0; eauto. omega. +Qed. + +Remark add_variables_symb_neg: + forall id b vl ge, + find_symbol (add_variables ge vl) id = Some b -> b < 0 -> + find_symbol ge id = Some b. +Proof. + induction vl; simpl; intros. auto. + exploit IHvl; eauto. unfold find_symbol; simpl. rewrite PTree.gsspec. + fold ident. destruct (peq id (a#1#1)); auto. intros. inv H1. + generalize (genv_nextvar_pos ge). intros. omegaContradiction. +Qed. + +Theorem find_funct_ptr_symbol_inversion: + forall p id b f, find_symbol (globalenv p) id = Some b -> find_funct_ptr (globalenv p) b = Some f -> In (id, f) p.(prog_funct). Proof. intros until f. - assert (forall fns, - let g := add_functs empty fns in - PTree.get id g.(symbols) = Some b -> - ZMap.get b g.(functions) = Some f -> - In (id, f) fns). - induction fns; simpl. - rewrite ZMap.gi. congruence. - set (g := add_functs empty fns). - rewrite PTree.gsspec. rewrite ZMap.gsspec. - case (peq id (fst a)); intro. - intro EQ. inversion EQ. unfold ZIndexed.eq. rewrite zeq_true. - intro EQ2. left. destruct a. simpl in *. congruence. - intro. unfold ZIndexed.eq. rewrite zeq_false. intro. eauto. - generalize (find_symbol_above_nextfunction _ _ H). fold g. unfold block. omega. - assert (forall g0 m0, b < 0 -> - forall vars g m, - add_globals (g0, m0) vars = (g, m) -> - PTree.get id g.(symbols) = Some b -> - PTree.get id g0.(symbols) = Some b). - induction vars; simpl. - intros. inv H1. auto. - destruct a as [[id1 init1] info1]. caseEq (add_globals (g0, m0) vars). - intros g1 m1 EQ g m EQ1. injection EQ1; simpl; clear EQ1. - unfold add_symbol; intros A B. rewrite <- B. simpl. - rewrite PTree.gsspec. case (peq id id1); intros. - assert (b > 0). inv H1. apply nextblock_pos. - omegaContradiction. - eauto. - intros. - generalize (find_funct_ptr_negative _ _ H2). intro. - pose (g := add_functs empty (prog_funct p)). - apply H. - apply H0 with Mem.empty (prog_vars p) (globalenv p) (init_mem p). - auto. unfold globalenv, init_mem. rewrite <- surjective_pairing. - reflexivity. assumption. rewrite <- functions_globalenv. assumption. + + assert (forall fl ge, + find_symbol (add_functions ge fl) id = Some b -> + find_funct_ptr (add_functions ge fl) b = Some f -> + In (id, f) fl \/ (find_symbol ge id = Some b /\ find_funct_ptr ge b = Some f)). + induction fl; simpl; intros. + auto. + exploit IHfl; eauto. intros [A | [A B]]. auto. + destruct a as [id' f']. + unfold find_symbol in A; simpl in A. + unfold find_funct_ptr in B; simpl in B. + rewrite PTree.gsspec in A. destruct (peq id id'). inv A. + rewrite ZMap.gss in B. inv B. auto. + rewrite ZMap.gso in B. right; auto. + exploit genv_symb_range; eauto. unfold block, ZIndexed.t; omega. + + intros. assert (b < 0) by (eapply find_funct_ptr_negative; eauto). + unfold globalenv in *. rewrite add_variables_same_funs in H1. + exploit (H (prog_funct p) empty_genv). + eapply add_variables_symb_neg; eauto. auto. + intuition. unfold find_symbol in H3; simpl in H3. rewrite PTree.gempty in H3. discriminate. Qed. Theorem find_symbol_not_nullptr: - forall (p: program F V) (id: ident) (b: block), - find_symbol (globalenv p) id = Some b -> b <> nullptr. -Proof. - intros until b. - assert (forall fns, - find_symbol (add_functs empty fns) id = Some b -> - b <> nullptr). - unfold find_symbol; induction fns; simpl. - rewrite PTree.gempty. congruence. - destruct a as [id1 f1]. simpl. rewrite PTree.gsspec. - destruct (peq id id1); intros. - inversion H. generalize (nextfunction_add_functs_neg fns). - unfold block, nullptr; omega. - auto. - set (g0 := add_functs empty p.(prog_funct)). - assert (forall vars g m, - add_globals (g0, Mem.empty) vars = (g, m) -> - find_symbol g id = Some b -> - b <> nullptr). - induction vars; simpl; intros until m. - intros. inversion H0. subst g. apply H with (prog_funct p). auto. - destruct a as [[id1 init1] info1]. - caseEq (add_globals (g0, Mem.empty) vars); intros g1 m1 EQ1 EQ2. - inv EQ2. unfold find_symbol, add_symbol; simpl. rewrite PTree.gsspec. - destruct (peq id id1); intros. - inv H0. generalize (nextblock_pos m1). unfold nullptr, block; omega. - eauto. - intros. eapply H0 with (vars := prog_vars p). apply surjective_pairing. auto. -Qed. + forall p id b, + find_symbol (globalenv p) id = Some b -> b <> Mem.nullptr. +Proof. + intros until b. unfold find_symbol. destruct (globalenv p); simpl. + intros. exploit genv_symb_range0; eauto. intuition. +Qed. Theorem global_addresses_distinct: - forall (p: program F V) id1 id2 b1 b2, + forall p id1 id2 b1 b2, id1<>id2 -> find_symbol (globalenv p) id1 = Some b1 -> find_symbol (globalenv p) id2 = Some b2 -> b1<>b2. +Proof. + intros until b2; intro DIFF. + + set (P := fun ge => find_symbol ge id1 = Some b1 -> find_symbol ge id2 = Some b2 -> b1 <> b2). + + assert (forall fl ge, P ge -> P (add_functions ge fl)). + induction fl; intros; simpl. auto. + apply IHfl. red. unfold find_symbol; simpl. + repeat rewrite PTree.gsspec. + fold ident. destruct (peq id1 a#1); destruct (peq id2 a#1). + congruence. + intros. inversion H0. exploit genv_symb_range; eauto. unfold block; omega. + intros. inversion H1. exploit genv_symb_range; eauto. unfold block; omega. + auto. + + assert (forall vl ge, P ge -> P (add_variables ge vl)). + induction vl; intros; simpl. auto. + apply IHvl. red. unfold find_symbol; simpl. + repeat rewrite PTree.gsspec. + fold ident. destruct (peq id1 a#1#1); destruct (peq id2 a#1#1). + congruence. + intros. inversion H1. exploit genv_symb_range; eauto. unfold block; omega. + intros. inversion H2. exploit genv_symb_range; eauto. unfold block; omega. + auto. + + change (P (globalenv p)). unfold globalenv. apply H0. apply H. + red; unfold find_symbol; simpl; intros. rewrite PTree.gempty in H1. congruence. +Qed. + +(** * Construction of the initial memory state *) + +Section INITMEM. + +Variable ge: t. + +Definition init_data_size (i: init_data) : Z := + match i with + | Init_int8 _ => 1 + | Init_int16 _ => 2 + | Init_int32 _ => 4 + | Init_float32 _ => 4 + | Init_float64 _ => 8 + | Init_addrof _ _ => 4 + | Init_space n => Zmax n 0 + end. + +Lemma init_data_size_pos: + forall i, init_data_size i >= 0. +Proof. + destruct i; simpl; try omega. generalize (Zle_max_r z 0). omega. +Qed. + +Definition store_init_data (m: mem) (b: block) (p: Z) (id: init_data) : option mem := + match id with + | Init_int8 n => Mem.store Mint8unsigned m b p (Vint n) + | Init_int16 n => Mem.store Mint16unsigned m b p (Vint n) + | Init_int32 n => Mem.store Mint32 m b p (Vint n) + | Init_float32 n => Mem.store Mfloat32 m b p (Vfloat n) + | Init_float64 n => Mem.store Mfloat64 m b p (Vfloat n) + | Init_addrof symb ofs => + match find_symbol ge symb with + | None => None + | Some b' => Mem.store Mint32 m b p (Vptr b' ofs) + end + | Init_space n => Some m + end. + +Fixpoint store_init_data_list (m: mem) (b: block) (p: Z) (idl: list init_data) + {struct idl}: option mem := + match idl with + | nil => Some m + | id :: idl' => + match store_init_data m b p id with + | None => None + | Some m' => store_init_data_list m' b (p + init_data_size id) idl' + end + end. + +Fixpoint init_data_list_size (il: list init_data) {struct il} : Z := + match il with + | nil => 0 + | i :: il' => init_data_size i + init_data_list_size il' + end. + +Definition alloc_variable (m: mem) (idv: ident * list init_data * V) : option mem := + let (m', b) := Mem.alloc m 0 (init_data_list_size idv#1#2) in + store_init_data_list m' b 0 idv#1#2. + +Fixpoint alloc_variables (m: mem) (vl: list (ident * list init_data * V)) + {struct vl} : option mem := + match vl with + | nil => Some m + | v :: vl' => + match alloc_variable m v with + | None => None + | Some m' => alloc_variables m' vl' + end + end. + +Remark store_init_data_list_nextblock: + forall idl b m p m', + store_init_data_list m b p idl = Some m' -> + Mem.nextblock m' = Mem.nextblock m. +Proof. + induction idl; simpl; intros until m'. + intros. congruence. + caseEq (store_init_data m b p a); try congruence. intros. + transitivity (Mem.nextblock m0). eauto. + destruct a; simpl in H; try (eapply Mem.nextblock_store; eauto; fail). + congruence. + destruct (find_symbol ge i); try congruence. eapply Mem.nextblock_store; eauto. +Qed. + +Remark alloc_variables_nextblock: + forall vl m m', + alloc_variables m vl = Some m' -> + Mem.nextblock m' = Mem.nextblock m + Z_of_nat(List.length vl). +Proof. + induction vl. + simpl; intros. inv H; unfold block; omega. + simpl length; rewrite inj_S; simpl. intros m m'. + unfold alloc_variable. + caseEq (Mem.alloc m 0 (init_data_list_size (a#1)#2)). intros m1 b ALLOC. + caseEq (store_init_data_list m1 b 0 a#1#2); try congruence. intros m2 STORE REC. + rewrite (IHvl _ _ REC). + rewrite (store_init_data_list_nextblock _ _ _ _ STORE). + rewrite (Mem.nextblock_alloc _ _ _ _ _ ALLOC). + unfold block in *; omega. +Qed. + +Remark store_init_data_list_perm: + forall prm b' q idl b m p m', + store_init_data_list m b p idl = Some m' -> + Mem.perm m b' q prm -> Mem.perm m' b' q prm. +Proof. + induction idl; simpl; intros until m'. + intros. congruence. + caseEq (store_init_data m b p a); try congruence. intros. + eapply IHidl; eauto. + destruct a; simpl in H; eauto with mem. + congruence. + destruct (find_symbol ge i); try congruence. eauto with mem. +Qed. + +Remark alloc_variables_perm: + forall prm b' q vl m m', + alloc_variables m vl = Some m' -> + Mem.perm m b' q prm -> Mem.perm m' b' q prm. +Proof. + induction vl. + simpl; intros. congruence. + intros until m'. simpl. unfold alloc_variable. + caseEq (Mem.alloc m 0 (init_data_list_size (a#1)#2)). intros m1 b ALLOC. + caseEq (store_init_data_list m1 b 0 a#1#2); try congruence. intros m2 STORE REC PERM. + eapply IHvl; eauto. + eapply store_init_data_list_perm; eauto. + eauto with mem. +Qed. + +Remark store_init_data_list_outside: + forall b il m p m', + store_init_data_list m b p il = Some m' -> + forall chunk b' q, + b' <> b \/ q + size_chunk chunk <= p -> + Mem.load chunk m' b' q = Mem.load chunk m b' q. +Proof. + induction il; simpl. + intros; congruence. + intros until m'. caseEq (store_init_data m b p a); try congruence. + intros m1 A B chunk b' q C. transitivity (Mem.load chunk m1 b' q). + eapply IHil; eauto. generalize (init_data_size_pos a). intuition omega. + destruct a; simpl in A; + try (eapply Mem.load_store_other; eauto; intuition; fail). + congruence. + destruct (find_symbol ge i); try congruence. + eapply Mem.load_store_other; eauto; intuition. +Qed. + +(* +Remark alloc_variables_nextblock: + forall vl g m m', + alloc_variables m vl = Some m' -> + Mem.nextblock m = genv_nextvar g -> + Mem.nextblock m' = genv_nextvar (add_variables g vl). +Proof. + induction vl; simpl; intros until m'. + intros. congruence. + unfold alloc_variable. + caseEq (Mem.alloc m 0 (init_data_list_size (a#1)#2)). intros m1 b ALLOC. + caseEq (store_init_data_list m1 b 0 a#1#2); try congruence. intros m2 STORE REC EQ. + eapply IHvl; eauto. + rewrite (store_init_data_list_nextblock _ _ _ _ STORE). + rewrite (Mem.nextblock_alloc _ _ _ _ _ ALLOC). + simpl. unfold block in *; omega. +Qed. +*) +Fixpoint load_store_init_data (m: mem) (b: block) (p: Z) (il: list init_data) {struct il} : Prop := + match il with + | nil => True + | Init_int8 n :: il' => + Mem.load Mint8unsigned m b p = Some(Vint(Int.zero_ext 8 n)) + /\ load_store_init_data m b (p + 1) il' + | Init_int16 n :: il' => + Mem.load Mint16unsigned m b p = Some(Vint(Int.zero_ext 16 n)) + /\ load_store_init_data m b (p + 2) il' + | Init_int32 n :: il' => + Mem.load Mint32 m b p = Some(Vint n) + /\ load_store_init_data m b (p + 4) il' + | Init_float32 n :: il' => + Mem.load Mfloat32 m b p = Some(Vfloat(Float.singleoffloat n)) + /\ load_store_init_data m b (p + 4) il' + | Init_float64 n :: il' => + Mem.load Mfloat64 m b p = Some(Vfloat n) + /\ load_store_init_data m b (p + 8) il' + | Init_addrof symb ofs :: il' => + (exists b', find_symbol ge symb = Some b' /\ Mem.load Mint32 m b p = Some(Vptr b' ofs)) + /\ load_store_init_data m b (p + 4) il' + | Init_space n :: il' => + load_store_init_data m b (p + Zmax n 0) il' + end. + +Lemma store_init_data_list_charact: + forall b il m p m', + store_init_data_list m b p il = Some m' -> + load_store_init_data m' b p il. +Proof. + assert (A: forall chunk v m b p m1 il m', + Mem.store chunk m b p v = Some m1 -> + store_init_data_list m1 b (p + size_chunk chunk) il = Some m' -> + Val.has_type v (type_of_chunk chunk) -> + Mem.load chunk m' b p = Some(Val.load_result chunk v)). + intros. transitivity (Mem.load chunk m1 b p). + eapply store_init_data_list_outside; eauto. right. omega. + eapply Mem.load_store_same; eauto. + + induction il; simpl. + auto. + intros until m'. caseEq (store_init_data m b p a); try congruence. + intros m1 B C. + exploit IHil; eauto. intro D. + destruct a; simpl in B; intuition. + eapply (A Mint8unsigned (Vint i)); eauto. simpl; auto. + eapply (A Mint16unsigned (Vint i)); eauto. simpl; auto. + eapply (A Mint32 (Vint i)); eauto. simpl; auto. + eapply (A Mfloat32 (Vfloat f)); eauto. simpl; auto. + eapply (A Mfloat64 (Vfloat f)); eauto. simpl; auto. + destruct (find_symbol ge i); try congruence. exists b0; split; auto. + eapply (A Mint32 (Vptr b0 i0)); eauto. simpl; auto. +Qed. + +Remark load_alloc_variables: + forall chunk b p vl m m', + alloc_variables m vl = Some m' -> + Mem.valid_block m b -> + Mem.load chunk m' b p = Mem.load chunk m b p. +Proof. + induction vl; simpl; intros until m'. + congruence. + unfold alloc_variable. + caseEq (Mem.alloc m 0 (init_data_list_size a#1#2)); intros m1 b1 ALLOC. + caseEq (store_init_data_list m1 b1 0 a#1#2); try congruence. intros m2 STO REC VAL. + transitivity (Mem.load chunk m2 b p). + apply IHvl; auto. red. rewrite (store_init_data_list_nextblock _ _ _ _ STO). + change (Mem.valid_block m1 b). eauto with mem. + transitivity (Mem.load chunk m1 b p). + eapply store_init_data_list_outside; eauto. left. + apply Mem.valid_not_valid_diff with m; eauto with mem. + eapply Mem.load_alloc_unchanged; eauto. +Qed. + +Remark load_store_init_data_invariant: + forall m m' b, + (forall chunk ofs, Mem.load chunk m' b ofs = Mem.load chunk m b ofs) -> + forall il p, + load_store_init_data m b p il -> load_store_init_data m' b p il. +Proof. + induction il; intro p; simpl. + auto. + repeat rewrite H. destruct a; intuition. +Qed. + +Lemma alloc_variables_charact: + forall id init v vl g m m', + genv_nextvar g = Mem.nextblock m -> + alloc_variables m vl = Some m' -> + list_norepet (map (fun v => v#1#1) vl) -> + In (id, init, v) vl -> + exists b, find_symbol (add_variables g vl) id = Some b + /\ find_var_info (add_variables g vl) b = Some v + /\ Mem.range_perm m' b 0 (init_data_list_size init) Writable + /\ load_store_init_data m' b 0 init. +Proof. + induction vl; simpl. + contradiction. + intros until m'; intro NEXT. + unfold alloc_variable. destruct a as [[id' init'] v']. simpl. + caseEq (Mem.alloc m 0 (init_data_list_size init')); try congruence. + intros m1 b ALLOC. + caseEq (store_init_data_list m1 b 0 init'); try congruence. + intros m2 STORE REC NOREPET IN. inv NOREPET. + exploit Mem.alloc_result; eauto. intro BEQ. + destruct IN. inv H. + exists (Mem.nextblock m); split. + rewrite add_variables_same_symb; auto. unfold find_symbol; simpl. + rewrite PTree.gss. congruence. + split. rewrite add_variables_same_address. unfold find_var_info; simpl. + rewrite NEXT. apply ZMap.gss. + simpl. rewrite <- NEXT; omega. + split. red; intros. eapply alloc_variables_perm; eauto. + eapply store_init_data_list_perm; eauto. + apply Mem.perm_implies with Freeable; eauto with mem. + apply load_store_init_data_invariant with m2. + intros. eapply load_alloc_variables; eauto. + red. rewrite (store_init_data_list_nextblock _ _ _ _ STORE). + change (Mem.valid_block m1 (Mem.nextblock m)). eauto with mem. + eapply store_init_data_list_charact; eauto. + + apply IHvl with m2; auto. + simpl. rewrite (store_init_data_list_nextblock _ _ _ _ STORE). + rewrite (Mem.nextblock_alloc _ _ _ _ _ ALLOC). unfold block in *; omega. +Qed. + +End INITMEM. + +Definition init_mem (p: program F V) := + alloc_variables (globalenv p) Mem.empty p.(prog_vars). + +Theorem find_symbol_not_fresh: + forall p id b m, + init_mem p = Some m -> + find_symbol (globalenv p) id = Some b -> Mem.valid_block m b. +Proof. + unfold init_mem; intros. + exploit alloc_variables_nextblock; eauto. rewrite Mem.nextblock_empty. intro. + exploit genv_symb_range; eauto. intros. + generalize (add_variables_nextvar (prog_vars p) (add_functions empty_genv (prog_funct p))). + rewrite add_functions_nextvar. simpl genv_nextvar. intro. + red. rewrite H1. rewrite <- H3. intuition. +Qed. + +Theorem find_var_exists: + forall p id init v m, + list_norepet (prog_var_names p) -> + In (id, init, v) (prog_vars p) -> + init_mem p = Some m -> + exists b, find_symbol (globalenv p) id = Some b + /\ find_var_info (globalenv p) b = Some v + /\ Mem.range_perm m b 0 (init_data_list_size init) Writable + /\ load_store_init_data (globalenv p) m b 0 init. +Proof. + intros. exploit alloc_variables_charact; eauto. + instantiate (1 := Mem.empty). rewrite add_functions_nextvar. rewrite Mem.nextblock_empty; auto. + assumption. +Qed. + +(** ** Compatibility with memory injections *) + +Section INITMEM_INJ. + +Variable ge: t. +Variable thr: block. +Hypothesis symb_inject: forall id b, find_symbol ge id = Some b -> b < thr. + +Lemma store_init_data_neutral: + forall m b p id m', + Mem.inject_neutral thr m -> + b < thr -> + store_init_data ge m b p id = Some m' -> + Mem.inject_neutral thr m'. Proof. intros. - assert (forall fns, - find_symbol (add_functs empty fns) id1 = Some b1 -> - find_symbol (add_functs empty fns) id2 = Some b2 -> - b1 <> b2). - unfold find_symbol. induction fns; simpl; intros. - rewrite PTree.gempty in H2. discriminate. - destruct a as [id f]; simpl in *. - rewrite PTree.gsspec in H2. - destruct (peq id1 id). subst id. inv H2. - rewrite PTree.gso in H3; auto. - generalize (find_symbol_above_nextfunction _ _ H3). unfold block. omega. - rewrite PTree.gsspec in H3. - destruct (peq id2 id). subst id. inv H3. - generalize (find_symbol_above_nextfunction _ _ H2). unfold block. omega. - eauto. - set (ge0 := add_functs empty p.(prog_funct)). - assert (forall (vars: list (ident * list init_data * V)) ge m, - add_globals (ge0, Mem.empty) vars = (ge, m) -> - find_symbol ge id1 = Some b1 -> - find_symbol ge id2 = Some b2 -> - b1 <> b2). - unfold find_symbol. induction vars; simpl. - intros. inv H3. subst ge. apply H2 with (p.(prog_funct)); auto. - intros ge m. destruct a as [[id init] info]. - caseEq (add_globals (ge0, Mem.empty) vars). intros ge1 m1 A B. inv B. - unfold add_symbol. simpl. intros. - rewrite PTree.gsspec in H3; destruct (peq id1 id). subst id. inv H3. - rewrite PTree.gso in H4; auto. - generalize (find_symbol_add_symbols_not_fresh _ _ _ A H4). unfold block; omega. - rewrite PTree.gsspec in H4; destruct (peq id2 id). subst id. inv H4. - generalize (find_symbol_add_symbols_not_fresh _ _ _ A H3). unfold block; omega. + destruct id; simpl in H1; try (eapply Mem.store_inject_neutral; eauto; fail). + inv H1; auto. + revert H1. caseEq (find_symbol ge i); try congruence. intros b' FS ST. + eapply Mem.store_inject_neutral; eauto. + econstructor. unfold Mem.flat_inj. apply zlt_true; eauto. + rewrite Int.add_zero. auto. +Qed. + +Lemma store_init_data_list_neutral: + forall b idl m p m', + Mem.inject_neutral thr m -> + b < thr -> + store_init_data_list ge m b p idl = Some m' -> + Mem.inject_neutral thr m'. +Proof. + induction idl; simpl. + intros; congruence. + intros until m'; intros INJ FB. + caseEq (store_init_data ge m b p a); try congruence. intros. + eapply IHidl. eapply store_init_data_neutral; eauto. auto. eauto. +Qed. + +Lemma alloc_variable_neutral: + forall id m m', + alloc_variable ge m id = Some m' -> + Mem.inject_neutral thr m -> + Mem.nextblock m < thr -> + Mem.inject_neutral thr m'. +Proof. + intros until m'. unfold alloc_variable. + caseEq (Mem.alloc m 0 (init_data_list_size (id#1)#2)); intros m1 b; intros. + eapply store_init_data_list_neutral with (b := b). + eapply Mem.alloc_inject_neutral; eauto. + rewrite (Mem.alloc_result _ _ _ _ _ H). auto. eauto. - set (ge_m := add_globals (ge0, Mem.empty) p.(prog_vars)). - apply H3 with (p.(prog_vars)) (fst ge_m) (snd ge_m). - fold ge_m. apply surjective_pairing. auto. auto. +Qed. + +Lemma alloc_variables_neutral: + forall idl m m', + alloc_variables ge m idl = Some m' -> + Mem.inject_neutral thr m -> + Mem.nextblock m' <= thr -> + Mem.inject_neutral thr m'. +Proof. + induction idl; simpl. + intros. congruence. + intros until m'. caseEq (alloc_variable ge m a); try congruence. intros. + assert (Mem.nextblock m' = Mem.nextblock m + Z_of_nat(length (a :: idl))). + eapply alloc_variables_nextblock with ge. simpl. rewrite H. auto. + simpl length in H3. rewrite inj_S in H3. + exploit alloc_variable_neutral; eauto. unfold block in *; omega. +Qed. + +End INITMEM_INJ. + +Theorem initmem_inject: + forall p m, + init_mem p = Some m -> + Mem.inject (Mem.flat_inj (Mem.nextblock m)) m m. +Proof. + unfold init_mem; intros. + apply Mem.neutral_inject. + eapply alloc_variables_neutral; eauto. + intros. exploit find_symbol_not_fresh; eauto. + apply Mem.empty_inject_neutral. + omega. Qed. End GENV. -(* Global environments and program transformations. *) +(** * Commutation with program transformations *) -Section MATCH_PROGRAM. +(** ** Commutation with matching between programs. *) -Variable A B V W: Type. +Section MATCH_PROGRAMS. + +Variables A B V W: Type. Variable match_fun: A -> B -> Prop. Variable match_var: V -> W -> Prop. -Variable p: program A V. -Variable p': program B W. -Hypothesis match_prog: - match_program match_fun match_var p p'. - -Lemma add_functs_match: - forall (fns: list (ident * A)) (tfns: list (ident * B)), - list_forall2 (match_funct_entry match_fun) fns tfns -> - let r := add_functs (empty A) fns in - let tr := add_functs (empty B) tfns in - nextfunction tr = nextfunction r /\ - symbols tr = symbols r /\ - forall (b: block) (f: A), - ZMap.get b (functions r) = Some f -> - exists tf, ZMap.get b (functions tr) = Some tf /\ match_fun f tf. -Proof. - induction 1; simpl. - - split. reflexivity. split. reflexivity. - intros b f; repeat (rewrite ZMap.gi). intros; discriminate. - - destruct a1 as [id1 fn1]. destruct b1 as [id2 fn2]. - simpl. red in H. destruct H. - destruct IHlist_forall2 as [X [Y Z]]. - rewrite X. rewrite Y. - split. auto. - split. congruence. - intros b f. - repeat (rewrite ZMap.gsspec). - destruct (ZIndexed.eq b (nextfunction (add_functs (empty A) al))). - intro EQ; inv EQ. exists fn2; auto. + +Record match_genvs (ge1: t A V) (ge2: t B W): Prop := { + mge_nextfun: genv_nextfun ge1 = genv_nextfun ge2; + mge_nextvar: genv_nextvar ge1 = genv_nextvar ge2; + mge_symb: genv_symb ge1 = genv_symb ge2; + mge_funs: + forall b f, ZMap.get b (genv_funs ge1) = Some f -> + exists tf, ZMap.get b (genv_funs ge2) = Some tf /\ match_fun f tf; + mge_rev_funs: + forall b tf, ZMap.get b (genv_funs ge2) = Some tf -> + exists f, ZMap.get b (genv_funs ge1) = Some f /\ match_fun f tf; + mge_vars: + forall b v, ZMap.get b (genv_vars ge1) = Some v -> + exists tv, ZMap.get b (genv_vars ge2) = Some tv /\ match_var v tv; + mge_rev_vars: + forall b tv, ZMap.get b (genv_vars ge2) = Some tv -> + exists v, ZMap.get b (genv_vars ge1) = Some v /\ match_var v tv +}. + +Lemma add_function_match: + forall ge1 ge2 id f1 f2, + match_genvs ge1 ge2 -> + match_fun f1 f2 -> + match_genvs (add_function ge1 (id, f1)) (add_function ge2 (id, f2)). +Proof. + intros. destruct H. constructor; simpl. + congruence. congruence. congruence. + rewrite mge_nextfun0. intros. rewrite ZMap.gsspec in H. rewrite ZMap.gsspec. + destruct (ZIndexed.eq b (genv_nextfun ge2)). + exists f2; split; congruence. + eauto. + rewrite mge_nextfun0. intros. rewrite ZMap.gsspec in H. rewrite ZMap.gsspec. + destruct (ZIndexed.eq b (genv_nextfun ge2)). + exists f1; split; congruence. + eauto. + auto. auto. Qed. -Lemma add_functs_rev_match: - forall (fns: list (ident * A)) (tfns: list (ident * B)), - list_forall2 (match_funct_entry match_fun) fns tfns -> - let r := add_functs (empty A) fns in - let tr := add_functs (empty B) tfns in - nextfunction tr = nextfunction r /\ - symbols tr = symbols r /\ - forall (b: block) (tf: B), - ZMap.get b (functions tr) = Some tf -> - exists f, ZMap.get b (functions r) = Some f /\ match_fun f tf. -Proof. - induction 1; simpl. - - split. reflexivity. split. reflexivity. - intros b f; repeat (rewrite ZMap.gi). intros; discriminate. - - destruct a1 as [id1 fn1]. destruct b1 as [id2 fn2]. - simpl. red in H. destruct H. - destruct IHlist_forall2 as [X [Y Z]]. - rewrite X. rewrite Y. - split. auto. - split. congruence. - intros b f. - repeat (rewrite ZMap.gsspec). - destruct (ZIndexed.eq b (nextfunction (add_functs (empty A) al))). - intro EQ; inv EQ. exists fn1; auto. +Lemma add_functions_match: + forall fl1 fl2, list_forall2 (match_funct_entry match_fun) fl1 fl2 -> + forall ge1 ge2, match_genvs ge1 ge2 -> + match_genvs (add_functions ge1 fl1) (add_functions ge2 fl2). +Proof. + induction 1; intros; simpl. auto. + destruct a1 as [id1 f1]; destruct b1 as [id2 f2]. + destruct H. subst. apply IHlist_forall2. apply add_function_match; auto. Qed. -Lemma mem_add_globals_match: - forall (g1: genv A) (g2: genv B) (m: mem) - (vars: list (ident * list init_data * V)) - (tvars: list (ident * list init_data * W)), - list_forall2 (match_var_entry match_var) vars tvars -> - snd (add_globals (g1, m) vars) = snd (add_globals (g2, m) tvars). +Lemma add_variable_match: + forall ge1 ge2 id idl v1 v2, + match_genvs ge1 ge2 -> + match_var v1 v2 -> + match_genvs (add_variable ge1 (id, idl, v1)) (add_variable ge2 (id, idl, v2)). Proof. - induction 1; simpl. + intros. destruct H. constructor; simpl. + congruence. congruence. congruence. auto. - destruct a1 as [[id1 init1] info1]. - destruct b1 as [[id2 init2] info2]. - red in H. destruct H as [X [Y Z]]. subst id2 init2. - generalize IHlist_forall2. - destruct (add_globals (g1, m) al). - destruct (add_globals (g2, m) bl). - simpl. intro. subst m1. auto. -Qed. - -Lemma symbols_add_globals_match: - forall (g1: genv A) (g2: genv B) (m: mem), - symbols g1 = symbols g2 -> - forall (vars: list (ident * list init_data * V)) - (tvars: list (ident * list init_data * W)), - list_forall2 (match_var_entry match_var) vars tvars -> - symbols (fst (add_globals (g1, m) vars)) = - symbols (fst (add_globals (g2, m) tvars)). -Proof. - induction 2; simpl. auto. - destruct a1 as [[id1 init1] info1]. - destruct b1 as [[id2 init2] info2]. - red in H0. destruct H0 as [X [Y Z]]. subst id2 init2. - generalize IHlist_forall2. - generalize (mem_add_globals_match g1 g2 m H1). - destruct (add_globals (g1, m) al). - destruct (add_globals (g2, m) bl). - simpl. intros. congruence. + rewrite mge_nextvar0. intros. rewrite ZMap.gsspec in H. rewrite ZMap.gsspec. + destruct (ZIndexed.eq b (genv_nextvar ge2)). + exists v2; split; congruence. + eauto. + rewrite mge_nextvar0. intros. rewrite ZMap.gsspec in H. rewrite ZMap.gsspec. + destruct (ZIndexed.eq b (genv_nextvar ge2)). + exists v1; split; congruence. + eauto. Qed. -Theorem find_funct_ptr_match: - forall (b: block) (f: A), - find_funct_ptr (globalenv p) b = Some f -> - exists tf, find_funct_ptr (globalenv p') b = Some tf /\ match_fun f tf. +Lemma add_variables_match: + forall vl1 vl2, list_forall2 (match_var_entry match_var) vl1 vl2 -> + forall ge1 ge2, match_genvs ge1 ge2 -> + match_genvs (add_variables ge1 vl1) (add_variables ge2 vl2). Proof. - intros until f. destruct match_prog as [X [Y Z]]. - destruct (add_functs_match X) as [P [Q R]]. - unfold find_funct_ptr. repeat rewrite functions_globalenv. + induction 1; intros; simpl. auto. + destruct a1 as [[id1 init1] f1]; destruct b1 as [[id2 init2] f2]. + destruct H. destruct H2. subst. apply IHlist_forall2. apply add_variable_match; auto. Qed. -Theorem find_funct_ptr_rev_match: - forall (b: block) (tf: B), - find_funct_ptr (globalenv p') b = Some tf -> - exists f, find_funct_ptr (globalenv p) b = Some f /\ match_fun f tf. +Variable p: program A V. +Variable p': program B W. +Hypothesis progmatch: match_program match_fun match_var p p'. + +Lemma globalenvs_match: + match_genvs (globalenv p) (globalenv p'). Proof. - intros until tf. destruct match_prog as [X [Y Z]]. - destruct (add_functs_rev_match X) as [P [Q R]]. - unfold find_funct_ptr. repeat rewrite functions_globalenv. - auto. + unfold globalenv. destruct progmatch. destruct H0. + apply add_variables_match; auto. apply add_functions_match; auto. + constructor; simpl; auto; intros; rewrite ZMap.gi in H2; congruence. Qed. +Theorem find_funct_ptr_match: + forall (b : block) (f : A), + find_funct_ptr (globalenv p) b = Some f -> + exists tf : B, + find_funct_ptr (globalenv p') b = Some tf /\ match_fun f tf. +Proof (mge_funs globalenvs_match). + +Theorem find_funct_ptr_rev_match: + forall (b : block) (tf : B), + find_funct_ptr (globalenv p') b = Some tf -> + exists f : A, + find_funct_ptr (globalenv p) b = Some f /\ match_fun f tf. +Proof (mge_rev_funs globalenvs_match). + Theorem find_funct_match: - forall (v: val) (f: A), + forall (v : val) (f : A), find_funct (globalenv p) v = Some f -> - exists tf, find_funct (globalenv p') v = Some tf /\ match_fun f tf. + exists tf : B, find_funct (globalenv p') v = Some tf /\ match_fun f tf. Proof. - intros until f. unfold find_funct. - case v; try (intros; discriminate). - intros b ofs. - case (Int.eq ofs Int.zero); try (intros; discriminate). - apply find_funct_ptr_match. + intros. exploit find_funct_inv; eauto. intros [b EQ]. subst v. + rewrite find_funct_find_funct_ptr in H. + rewrite find_funct_find_funct_ptr. + apply find_funct_ptr_match. auto. Qed. Theorem find_funct_rev_match: - forall (v: val) (tf: B), + forall (v : val) (tf : B), find_funct (globalenv p') v = Some tf -> - exists f, find_funct (globalenv p) v = Some f /\ match_fun f tf. + exists f : A, find_funct (globalenv p) v = Some f /\ match_fun f tf. Proof. - intros until tf. unfold find_funct. - case v; try (intros; discriminate). - intros b ofs. - case (Int.eq ofs Int.zero); try (intros; discriminate). - apply find_funct_ptr_rev_match. + intros. exploit find_funct_inv; eauto. intros [b EQ]. subst v. + rewrite find_funct_find_funct_ptr in H. + rewrite find_funct_find_funct_ptr. + apply find_funct_ptr_rev_match. auto. Qed. -Lemma symbols_init_match: - symbols (globalenv p') = symbols (globalenv p). -Proof. - unfold globalenv. unfold globalenv_initmem. - destruct match_prog as [X [Y Z]]. - destruct (add_functs_match X) as [P [Q R]]. - simpl. symmetry. apply symbols_add_globals_match. auto. auto. -Qed. +Theorem find_var_info_match: + forall (b : block) (v : V), + find_var_info (globalenv p) b = Some v -> + exists tv, + find_var_info (globalenv p') b = Some tv /\ match_var v tv. +Proof (mge_vars globalenvs_match). + +Theorem find_var_info_rev_match: + forall (b : block) (tv : W), + find_var_info (globalenv p') b = Some tv -> + exists v, + find_var_info (globalenv p) b = Some v /\ match_var v tv. +Proof (mge_rev_vars globalenvs_match). Theorem find_symbol_match: - forall (s: ident), + forall (s : ident), find_symbol (globalenv p') s = find_symbol (globalenv p) s. Proof. - intros. unfold find_symbol. - rewrite symbols_init_match. auto. + intros. destruct globalenvs_match. unfold find_symbol. congruence. +Qed. + +Lemma store_init_data_list_match: + forall idl m b ofs, + store_init_data_list (globalenv p') m b ofs idl = + store_init_data_list (globalenv p) m b ofs idl. +Proof. + induction idl; simpl; intros. + auto. + assert (store_init_data (globalenv p') m b ofs a = + store_init_data (globalenv p) m b ofs a). + destruct a; simpl; auto. rewrite find_symbol_match. auto. + rewrite H. destruct (store_init_data (globalenv p) m b ofs a); auto. +Qed. + +Lemma alloc_variables_match: + forall vl1 vl2, list_forall2 (match_var_entry match_var) vl1 vl2 -> + forall m, + alloc_variables (globalenv p') m vl2 = alloc_variables (globalenv p) m vl1. +Proof. + induction 1; intros; simpl. + auto. + destruct a1 as [[id1 init1] v1]; destruct b1 as [[id2 init2] v2]. + destruct H. destruct H1. subst. + unfold alloc_variable; simpl. + destruct (Mem.alloc m 0 (init_data_list_size init2)). + rewrite store_init_data_list_match. + destruct (store_init_data_list (globalenv p) m0 b 0 init2); auto. Qed. Theorem init_mem_match: - init_mem p' = init_mem p. + forall m, init_mem p = Some m -> init_mem p' = Some m. Proof. - unfold init_mem. unfold globalenv_initmem. - destruct match_prog as [X [Y Z]]. - symmetry. apply mem_add_globals_match. auto. + intros. rewrite <- H. unfold init_mem. destruct progmatch. destruct H1. + apply alloc_variables_match; auto. Qed. -End MATCH_PROGRAM. +End MATCH_PROGRAMS. Section TRANSF_PROGRAM_PARTIAL2. @@ -1007,6 +1136,28 @@ Proof. exploit find_funct_rev_match. eexact prog_match. eauto. auto. Qed. +Theorem find_var_info_transf_partial2: + forall (b: block) (v: V), + find_var_info (globalenv p) b = Some v -> + exists v', + find_var_info (globalenv p') b = Some v' /\ transf_var v = OK v'. +Proof. + intros. + exploit find_var_info_match. eexact prog_match. eauto. + intros [tv [X Y]]. exists tv; auto. +Qed. + +Theorem find_var_info_rev_transf_partial2: + forall (b: block) (v': W), + find_var_info (globalenv p') b = Some v' -> + exists v, + find_var_info (globalenv p) b = Some v /\ transf_var v = OK v'. +Proof. + intros. + exploit find_var_info_rev_match. eexact prog_match. eauto. + intros [v [X Y]]. exists v; auto. +Qed. + Theorem find_symbol_transf_partial2: forall (s: ident), find_symbol (globalenv p') s = find_symbol (globalenv p) s. @@ -1015,9 +1166,9 @@ Proof. Qed. Theorem init_mem_transf_partial2: - init_mem p' = init_mem p. + forall m, init_mem p = Some m -> init_mem p' = Some m. Proof. - intros. eapply init_mem_match. eexact prog_match. + intros. eapply init_mem_match. eexact prog_match. auto. Qed. End TRANSF_PROGRAM_PARTIAL2. @@ -1080,7 +1231,7 @@ Proof. Qed. Theorem init_mem_transf_partial: - init_mem p' = init_mem p. + forall m, init_mem p = Some m -> init_mem p' = Some m. Proof. exact (@init_mem_transf_partial2 _ _ _ _ _ _ _ _ transf2_OK). Qed. @@ -1147,7 +1298,7 @@ Proof. Qed. Theorem init_mem_transf: - init_mem tp = init_mem p. + forall m, init_mem p = Some m -> init_mem tp = Some m. Proof. exact (@init_mem_transf_partial _ _ _ _ _ _ transf_OK). Qed. diff --git a/common/Mem.v b/common/Mem.v deleted file mode 100644 index 252ee29..0000000 --- a/common/Mem.v +++ /dev/null @@ -1,2887 +0,0 @@ -(* *********************************************************************) -(* *) -(* The Compcert verified compiler *) -(* *) -(* Xavier Leroy, INRIA Paris-Rocquencourt *) -(* Sandrine Blazy, ENSIIE and INRIA Paris-Rocquencourt *) -(* *) -(* Copyright Institut National de Recherche en Informatique et en *) -(* Automatique. All rights reserved. This file is distributed *) -(* under the terms of the GNU General Public License as published by *) -(* the Free Software Foundation, either version 2 of the License, or *) -(* (at your option) any later version. This file is also distributed *) -(* under the terms of the INRIA Non-Commercial License Agreement. *) -(* *) -(* *********************************************************************) - -(** This file develops the memory model that is used in the dynamic - semantics of all the languages used in the compiler. - It defines a type [mem] of memory states, the following 4 basic - operations over memory states, and their properties: -- [load]: read a memory chunk at a given address; -- [store]: store a memory chunk at a given address; -- [alloc]: allocate a fresh memory block; -- [free]: invalidate a memory block. -*) - -Require Import Coqlib. -Require Import Maps. -Require Import AST. -Require Import Integers. -Require Import Floats. -Require Import Values. - -Definition update (A: Type) (x: Z) (v: A) (f: Z -> A) : Z -> A := - fun y => if zeq y x then v else f y. - -Implicit Arguments update [A]. - -Lemma update_s: - forall (A: Type) (x: Z) (v: A) (f: Z -> A), - update x v f x = v. -Proof. - intros; unfold update. apply zeq_true. -Qed. - -Lemma update_o: - forall (A: Type) (x: Z) (v: A) (f: Z -> A) (y: Z), - x <> y -> update x v f y = f y. -Proof. - intros; unfold update. apply zeq_false; auto. -Qed. - -(** * Structure of memory states *) - -(** A memory state is organized in several disjoint blocks. Each block - has a low and a high bound that defines its size. Each block map - byte offsets to the contents of this byte. *) - -(** The possible contents of a byte-sized memory cell. To give intuitions, - a 4-byte value [v] stored at offset [d] will be represented by - the content [Datum(4, v)] at offset [d] and [Cont] at offsets [d+1], - [d+2] and [d+3]. The [Cont] contents enable detecting future writes - that would partially overlap the 4-byte value. *) - -Inductive content : Type := - | Undef: content (**r undefined contents *) - | Datum: nat -> val -> content (**r first byte of a value *) - | Cont: content. (**r continuation bytes for a multi-byte value *) - -Definition contentmap : Type := Z -> content. - -(** A memory block comprises the dimensions of the block (low and high bounds) - plus a mapping from byte offsets to contents. *) - -Record block_contents : Type := mkblock { - low: Z; - high: Z; - contents: contentmap -}. - -(** A memory state is a mapping from block addresses (represented by [Z] - integers) to blocks. We also maintain the address of the next - unallocated block, and a proof that this address is positive. *) - -Record mem : Type := mkmem { - blocks: Z -> block_contents; - nextblock: block; - nextblock_pos: nextblock > 0 -}. - -(** * Operations on memory stores *) - -(** Memory reads and writes are performed by quantities called memory chunks, - encoding the type, size and signedness of the chunk being addressed. - The following functions extract the size information from a chunk. *) - -Definition size_chunk (chunk: memory_chunk) : Z := - match chunk with - | Mint8signed => 1 - | Mint8unsigned => 1 - | Mint16signed => 2 - | Mint16unsigned => 2 - | Mint32 => 4 - | Mfloat32 => 4 - | Mfloat64 => 8 - end. - -Definition pred_size_chunk (chunk: memory_chunk) : nat := - match chunk with - | Mint8signed => 0%nat - | Mint8unsigned => 0%nat - | Mint16signed => 1%nat - | Mint16unsigned => 1%nat - | Mint32 => 3%nat - | Mfloat32 => 3%nat - | Mfloat64 => 7%nat - end. - -Lemma size_chunk_pred: - forall chunk, size_chunk chunk = 1 + Z_of_nat (pred_size_chunk chunk). -Proof. - destruct chunk; auto. -Qed. - -Lemma size_chunk_pos: - forall chunk, size_chunk chunk > 0. -Proof. - intros. rewrite size_chunk_pred. omega. -Qed. - -(** Memory reads and writes must respect alignment constraints: - the byte offset of the location being addressed should be an exact - multiple of the natural alignment for the chunk being addressed. - This natural alignment is defined by the following - [align_chunk] function. Some target architectures - (e.g. the PowerPC) have no alignment constraints, which we could - reflect by taking [align_chunk chunk = 1]. However, other architectures - have stronger alignment requirements. The following definition is - appropriate for PowerPC and ARM. *) - -Definition align_chunk (chunk: memory_chunk) : Z := - match chunk with - | Mint8signed => 1 - | Mint8unsigned => 1 - | Mint16signed => 2 - | Mint16unsigned => 2 - | _ => 4 - end. - -Lemma align_chunk_pos: - forall chunk, align_chunk chunk > 0. -Proof. - intro. destruct chunk; simpl; omega. -Qed. - -Lemma align_size_chunk_divides: - forall chunk, (align_chunk chunk | size_chunk chunk). -Proof. - intros. destruct chunk; simpl; try apply Zdivide_refl. exists 2; auto. -Qed. - -Lemma align_chunk_compat: - forall chunk1 chunk2, - size_chunk chunk1 = size_chunk chunk2 -> align_chunk chunk1 = align_chunk chunk2. -Proof. - intros chunk1 chunk2. - destruct chunk1; destruct chunk2; simpl; congruence. -Qed. - -(** The initial store. *) - -Remark one_pos: 1 > 0. -Proof. omega. Qed. - -Definition empty_block (lo hi: Z) : block_contents := - mkblock lo hi (fun y => Undef). - -Definition empty: mem := - mkmem (fun x => empty_block 0 0) 1 one_pos. - -Definition nullptr: block := 0. - -(** Allocation of a fresh block with the given bounds. Return an updated - memory state and the address of the fresh block, which initially contains - undefined cells. Note that allocation never fails: we model an - infinite memory. *) - -Remark succ_nextblock_pos: - forall m, Zsucc m.(nextblock) > 0. -Proof. intro. generalize (nextblock_pos m). omega. Qed. - -Definition alloc (m: mem) (lo hi: Z) := - (mkmem (update m.(nextblock) - (empty_block lo hi) - m.(blocks)) - (Zsucc m.(nextblock)) - (succ_nextblock_pos m), - m.(nextblock)). - -(** Freeing a block. Return the updated memory state where the given - block address has been invalidated: future reads and writes to this - address will fail. Note that we make no attempt to return the block - to an allocation pool: the given block address will never be allocated - later. *) - -Definition free (m: mem) (b: block) := - mkmem (update b - (empty_block 0 0) - m.(blocks)) - m.(nextblock) - m.(nextblock_pos). - -(** Freeing of a list of blocks. *) - -Definition free_list (m:mem) (l:list block) := - List.fold_right (fun b m => free m b) m l. - -(** Return the low and high bounds for the given block address. - Those bounds are 0 for freed or not yet allocated address. *) - -Definition low_bound (m: mem) (b: block) := - low (m.(blocks) b). -Definition high_bound (m: mem) (b: block) := - high (m.(blocks) b). - -(** A block address is valid if it was previously allocated. It remains valid - even after being freed. *) - -Definition valid_block (m: mem) (b: block) := - b < m.(nextblock). - -(** Reading and writing [N] adjacent locations in a [contentmap]. - - We define two functions and prove some of their properties: - - [getN n ofs m] returns the datum at offset [ofs] in block contents [m] - after checking that the contents of offsets [ofs+1] to [ofs+n+1] - are [Cont]. - - [setN n ofs v m] updates the block contents [m], storing the content [v] - at offset [ofs] and the content [Cont] at offsets [ofs+1] to [ofs+n+1]. - *) - -Fixpoint check_cont (n: nat) (p: Z) (m: contentmap) {struct n} : bool := - match n with - | O => true - | S n1 => - match m p with - | Cont => check_cont n1 (p + 1) m - | _ => false - end - end. - -Definition eq_nat: forall (p q: nat), {p=q} + {p<>q}. -Proof. decide equality. Defined. - -Definition getN (n: nat) (p: Z) (m: contentmap) : val := - match m p with - | Datum n' v => - if eq_nat n n' && check_cont n (p + 1) m then v else Vundef - | _ => - Vundef - end. - -Fixpoint set_cont (n: nat) (p: Z) (m: contentmap) {struct n} : contentmap := - match n with - | O => m - | S n1 => update p Cont (set_cont n1 (p + 1) m) - end. - -Definition setN (n: nat) (p: Z) (v: val) (m: contentmap) : contentmap := - update p (Datum n v) (set_cont n (p + 1) m). - -Lemma check_cont_spec: - forall n m p, - if check_cont n p m - then (forall q, p <= q < p + Z_of_nat n -> m q = Cont) - else (exists q, p <= q < p + Z_of_nat n /\ m q <> Cont). -Proof. - induction n; intros. - simpl. intros; omegaContradiction. - simpl check_cont. repeat rewrite inj_S. caseEq (m p); intros. - exists p; split. omega. congruence. - exists p; split. omega. congruence. - generalize (IHn m (p + 1)). case (check_cont n (p + 1) m). - intros. assert (p = q \/ p + 1 <= q < p + Zsucc (Z_of_nat n)) by omega. - elim H2; intro. congruence. apply H0; omega. - intros [q [A B]]. exists q; split. omega. auto. -Qed. - -Lemma check_cont_true: - forall n m p, - (forall q, p <= q < p + Z_of_nat n -> m q = Cont) -> - check_cont n p m = true. -Proof. - intros. generalize (check_cont_spec n m p). - destruct (check_cont n p m). auto. - intros [q [A B]]. elim B; auto. -Qed. - -Lemma check_cont_false: - forall n m p q, - p <= q < p + Z_of_nat n -> m q <> Cont -> - check_cont n p m = false. -Proof. - intros. generalize (check_cont_spec n m p). - destruct (check_cont n p m). - intros. elim H0; auto. - auto. -Qed. - -Lemma set_cont_inside: - forall n p m q, - p <= q < p + Z_of_nat n -> - (set_cont n p m) q = Cont. -Proof. - induction n; intros. - unfold Z_of_nat in H. omegaContradiction. - simpl. - assert (p = q \/ p + 1 <= q < (p + 1) + Z_of_nat n). - rewrite inj_S in H. omega. - elim H0; intro. - subst q. apply update_s. - rewrite update_o. apply IHn. auto. - red; intro; subst q. omega. -Qed. - -Lemma set_cont_outside: - forall n p m q, - q < p \/ p + Z_of_nat n <= q -> - (set_cont n p m) q = m q. -Proof. - induction n; intros. - simpl. auto. - simpl. rewrite inj_S in H. - rewrite update_o. apply IHn. omega. omega. -Qed. - -Lemma getN_setN_same: - forall n p v m, - getN n p (setN n p v m) = v. -Proof. - intros. unfold getN, setN. rewrite update_s. - rewrite check_cont_true. unfold proj_sumbool. - rewrite dec_eq_true. auto. - intros. rewrite update_o. apply set_cont_inside. auto. - omega. -Qed. - -Lemma getN_setN_other: - forall n1 n2 p1 p2 v m, - p1 + Z_of_nat n1 < p2 \/ p2 + Z_of_nat n2 < p1 -> - getN n2 p2 (setN n1 p1 v m) = getN n2 p2 m. -Proof. - intros. unfold getN, setN. - generalize (check_cont_spec n2 m (p2 + 1)); - destruct (check_cont n2 (p2 + 1) m); intros. - rewrite check_cont_true. - rewrite update_o. rewrite set_cont_outside. auto. - omega. omega. - intros. rewrite update_o. rewrite set_cont_outside. auto. - omega. omega. - destruct H0 as [q [A B]]. - rewrite (check_cont_false n2 (update p1 (Datum n1 v) (set_cont n1 (p1 + 1) m)) (p2 + 1) q). - rewrite update_o. rewrite set_cont_outside. auto. - omega. omega. omega. - rewrite update_o. rewrite set_cont_outside. auto. - omega. omega. -Qed. - -Lemma getN_setN_overlap: - forall n1 n2 p1 p2 v m, - p1 <> p2 -> - p1 + Z_of_nat n1 >= p2 -> p2 + Z_of_nat n2 >= p1 -> - getN n2 p2 (setN n1 p1 v m) = Vundef. -Proof. - intros. unfold getN, setN. - rewrite update_o; auto. - destruct (zlt p2 p1). - (* [p1] belongs to [[p2, p2 + n2 - 1]], - therefore [check_cont n2 (p2 + 1) ...] is false. *) - rewrite (check_cont_false n2 (update p1 (Datum n1 v) (set_cont n1 (p1 + 1) m)) (p2 + 1) p1). - destruct (set_cont n1 (p1 + 1) m p2); auto. - destruct (eq_nat n2 n); auto. - omega. - rewrite update_s. congruence. - (* [p2] belongs to [[p1 + 1, p1 + n1 - 1]], - therefore [set_cont n1 (p1 + 1) m p2] is [Cont]. *) - rewrite set_cont_inside. auto. omega. -Qed. - -Lemma getN_setN_mismatch: - forall n1 n2 p v m, - n1 <> n2 -> - getN n2 p (setN n1 p v m) = Vundef. -Proof. - intros. unfold getN, setN. rewrite update_s. - unfold proj_sumbool; rewrite dec_eq_false; simpl. auto. auto. -Qed. - -Lemma getN_setN_characterization: - forall m v n1 p1 n2 p2, - getN n2 p2 (setN n1 p1 v m) = v - \/ getN n2 p2 (setN n1 p1 v m) = getN n2 p2 m - \/ getN n2 p2 (setN n1 p1 v m) = Vundef. -Proof. - intros. destruct (zeq p1 p2). subst p2. - destruct (eq_nat n1 n2). subst n2. - left; apply getN_setN_same. - right; right; apply getN_setN_mismatch; auto. - destruct (zlt (p1 + Z_of_nat n1) p2). - right; left; apply getN_setN_other; auto. - destruct (zlt (p2 + Z_of_nat n2) p1). - right; left; apply getN_setN_other; auto. - right; right; apply getN_setN_overlap; omega. -Qed. - -Lemma getN_init: - forall n p, - getN n p (fun y => Undef) = Vundef. -Proof. - intros. auto. -Qed. - -(** [valid_access m chunk b ofs] holds if a memory access (load or store) - of the given chunk is possible in [m] at address [b, ofs]. - This means: -- The block [b] is valid. -- The range of bytes accessed is within the bounds of [b]. -- The offset [ofs] is aligned. -*) - -Inductive valid_access (m: mem) (chunk: memory_chunk) (b: block) (ofs: Z) : Prop := - | valid_access_intro: - valid_block m b -> - low_bound m b <= ofs -> - ofs + size_chunk chunk <= high_bound m b -> - (align_chunk chunk | ofs) -> - valid_access m chunk b ofs. - -(** The following function checks whether accessing the given memory chunk - at the given offset in the given block respects the bounds of the block. *) - -Definition in_bounds (m: mem) (chunk: memory_chunk) (b: block) (ofs: Z) : - {valid_access m chunk b ofs} + {~valid_access m chunk b ofs}. -Proof. - intros. - destruct (zlt b m.(nextblock)). - destruct (zle (low_bound m b) ofs). - destruct (zle (ofs + size_chunk chunk) (high_bound m b)). - destruct (Zdivide_dec (align_chunk chunk) ofs (align_chunk_pos chunk)). - left; constructor; auto. - right; red; intro V; inv V; contradiction. - right; red; intro V; inv V; omega. - right; red; intro V; inv V; omega. - right; red; intro V; inv V; contradiction. -Defined. - -Lemma in_bounds_true: - forall m chunk b ofs (A: Type) (a1 a2: A), - valid_access m chunk b ofs -> - (if in_bounds m chunk b ofs then a1 else a2) = a1. -Proof. - intros. destruct (in_bounds m chunk b ofs). auto. contradiction. -Qed. - -(** [valid_pointer] holds if the given block address is valid and the - given offset falls within the bounds of the corresponding block. *) - -Definition valid_pointer (m: mem) (b: block) (ofs: Z) : bool := - zlt b m.(nextblock) && - zle (low_bound m b) ofs && - zlt ofs (high_bound m b). - -(** [load chunk m b ofs] perform a read in memory state [m], at address - [b] and offset [ofs]. [None] is returned if the address is invalid - or the memory access is out of bounds. *) - -Definition load (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z) - : option val := - if in_bounds m chunk b ofs then - Some (Val.load_result chunk - (getN (pred_size_chunk chunk) ofs (contents (blocks m b)))) - else - None. - -Lemma load_inv: - forall chunk m b ofs v, - load chunk m b ofs = Some v -> - valid_access m chunk b ofs /\ - v = Val.load_result chunk - (getN (pred_size_chunk chunk) ofs (contents (blocks m b))). -Proof. - intros until v; unfold load. - destruct (in_bounds m chunk b ofs); intros. - split. auto. congruence. - congruence. -Qed. - -(** [loadv chunk m addr] is similar, but the address and offset are given - as a single value [addr], which must be a pointer value. *) - -Definition loadv (chunk: memory_chunk) (m: mem) (addr: val) : option val := - match addr with - | Vptr b ofs => load chunk m b (Int.signed ofs) - | _ => None - end. - -(* The memory state [m] after a store of value [v] at offset [ofs] - in block [b]. *) - -Definition unchecked_store - (chunk: memory_chunk) (m: mem) (b: block) - (ofs: Z) (v: val) : mem := - let c := m.(blocks) b in - mkmem - (update b - (mkblock c.(low) c.(high) - (setN (pred_size_chunk chunk) ofs v c.(contents))) - m.(blocks)) - m.(nextblock) - m.(nextblock_pos). - -(** [store chunk m b ofs v] perform a write in memory state [m]. - Value [v] is stored at address [b] and offset [ofs]. - Return the updated memory store, or [None] if the address is invalid - or the memory access is out of bounds. *) - -Definition store (chunk: memory_chunk) (m: mem) (b: block) - (ofs: Z) (v: val) : option mem := - if in_bounds m chunk b ofs - then Some(unchecked_store chunk m b ofs v) - else None. - -Lemma store_inv: - forall chunk m b ofs v m', - store chunk m b ofs v = Some m' -> - valid_access m chunk b ofs /\ - m' = unchecked_store chunk m b ofs v. -Proof. - intros until m'; unfold store. - destruct (in_bounds m chunk b ofs); intros. - split. auto. congruence. - congruence. -Qed. - -(** [storev chunk m addr v] is similar, but the address and offset are given - as a single value [addr], which must be a pointer value. *) - -Definition storev (chunk: memory_chunk) (m: mem) (addr v: val) : option mem := - match addr with - | Vptr b ofs => store chunk m b (Int.signed ofs) v - | _ => None - end. - -(** Build a block filled with the given initialization data. *) - -Fixpoint contents_init_data (pos: Z) (id: list init_data) {struct id}: contentmap := - match id with - | nil => (fun y => Undef) - | Init_int8 n :: id' => - setN 0%nat pos (Vint n) (contents_init_data (pos + 1) id') - | Init_int16 n :: id' => - setN 1%nat pos (Vint n) (contents_init_data (pos + 1) id') - | Init_int32 n :: id' => - setN 3%nat pos (Vint n) (contents_init_data (pos + 1) id') - | Init_float32 f :: id' => - setN 3%nat pos (Vfloat f) (contents_init_data (pos + 1) id') - | Init_float64 f :: id' => - setN 7%nat pos (Vfloat f) (contents_init_data (pos + 1) id') - | Init_space n :: id' => - contents_init_data (pos + Zmax n 0) id' - | Init_addrof s n :: id' => - (* Not handled properly yet *) - contents_init_data (pos + 4) id' - end. - -Definition size_init_data (id: init_data) : Z := - match id with - | Init_int8 _ => 1 - | Init_int16 _ => 2 - | Init_int32 _ => 4 - | Init_float32 _ => 4 - | Init_float64 _ => 8 - | Init_space n => Zmax n 0 - | Init_addrof _ _ => 4 - end. - -Definition size_init_data_list (id: list init_data): Z := - List.fold_right (fun id sz => size_init_data id + sz) 0 id. - -Remark size_init_data_list_pos: - forall id, size_init_data_list id >= 0. -Proof. - induction id; simpl. - omega. - assert (size_init_data a >= 0). destruct a; simpl; try omega. - generalize (Zmax2 z 0). omega. omega. -Qed. - -Definition block_init_data (id: list init_data) : block_contents := - mkblock 0 (size_init_data_list id) (contents_init_data 0 id). - -Definition alloc_init_data (m: mem) (id: list init_data) : mem * block := - (mkmem (update m.(nextblock) - (block_init_data id) - m.(blocks)) - (Zsucc m.(nextblock)) - (succ_nextblock_pos m), - m.(nextblock)). - -Remark block_init_data_empty: - block_init_data nil = empty_block 0 0. -Proof. - auto. -Qed. - -(** * Properties of the memory operations *) - -(** ** Properties related to block validity *) - -Lemma valid_not_valid_diff: - forall m b b', valid_block m b -> ~(valid_block m b') -> b <> b'. -Proof. - intros; red; intros. subst b'. contradiction. -Qed. - -Lemma valid_access_valid_block: - forall m chunk b ofs, - valid_access m chunk b ofs -> valid_block m b. -Proof. - intros. inv H; auto. -Qed. - -Lemma valid_access_aligned: - forall m chunk b ofs, - valid_access m chunk b ofs -> (align_chunk chunk | ofs). -Proof. - intros. inv H; auto. -Qed. - -Lemma valid_access_compat: - forall m chunk1 chunk2 b ofs, - size_chunk chunk1 = size_chunk chunk2 -> - valid_access m chunk1 b ofs -> - valid_access m chunk2 b ofs. -Proof. - intros. inv H0. rewrite H in H3. constructor; auto. - rewrite <- (align_chunk_compat _ _ H). auto. -Qed. - -Hint Resolve valid_not_valid_diff valid_access_valid_block valid_access_aligned: mem. - -(** ** Properties related to [load] *) - -Theorem valid_access_load: - forall m chunk b ofs, - valid_access m chunk b ofs -> - exists v, load chunk m b ofs = Some v. -Proof. - intros. econstructor. unfold load. rewrite in_bounds_true; auto. -Qed. - -Theorem load_valid_access: - forall m chunk b ofs v, - load chunk m b ofs = Some v -> - valid_access m chunk b ofs. -Proof. - intros. generalize (load_inv _ _ _ _ _ H). tauto. -Qed. - -Hint Resolve load_valid_access valid_access_load. - -(** ** Properties related to [store] *) - -Lemma valid_access_store: - forall m1 chunk b ofs v, - valid_access m1 chunk b ofs -> - exists m2, store chunk m1 b ofs v = Some m2. -Proof. - intros. econstructor. unfold store. rewrite in_bounds_true; auto. -Qed. - -Hint Resolve valid_access_store: mem. - -Section STORE. -Variable chunk: memory_chunk. -Variable m1: mem. -Variable b: block. -Variable ofs: Z. -Variable v: val. -Variable m2: mem. -Hypothesis STORE: store chunk m1 b ofs v = Some m2. - -Lemma low_bound_store: - forall b', low_bound m2 b' = low_bound m1 b'. -Proof. - intro. elim (store_inv _ _ _ _ _ _ STORE); intros. - subst m2. unfold low_bound, unchecked_store; simpl. - unfold update. destruct (zeq b' b); auto. subst b'; auto. -Qed. - -Lemma high_bound_store: - forall b', high_bound m2 b' = high_bound m1 b'. -Proof. - intro. elim (store_inv _ _ _ _ _ _ STORE); intros. - subst m2. unfold high_bound, unchecked_store; simpl. - unfold update. destruct (zeq b' b); auto. subst b'; auto. -Qed. - -Lemma nextblock_store: - nextblock m2 = nextblock m1. -Proof. - intros. elim (store_inv _ _ _ _ _ _ STORE); intros. - subst m2; reflexivity. -Qed. - -Lemma store_valid_block_1: - forall b', valid_block m1 b' -> valid_block m2 b'. -Proof. - unfold valid_block; intros. rewrite nextblock_store; auto. -Qed. - -Lemma store_valid_block_2: - forall b', valid_block m2 b' -> valid_block m1 b'. -Proof. - unfold valid_block; intros. rewrite nextblock_store in H; auto. -Qed. - -Hint Resolve store_valid_block_1 store_valid_block_2: mem. - -Lemma store_valid_access_1: - forall chunk' b' ofs', - valid_access m1 chunk' b' ofs' -> valid_access m2 chunk' b' ofs'. -Proof. - intros. inv H. constructor; auto with mem. - rewrite low_bound_store; auto. - rewrite high_bound_store; auto. -Qed. - -Lemma store_valid_access_2: - forall chunk' b' ofs', - valid_access m2 chunk' b' ofs' -> valid_access m1 chunk' b' ofs'. -Proof. - intros. inv H. constructor; auto with mem. - rewrite low_bound_store in H1; auto. - rewrite high_bound_store in H2; auto. -Qed. - -Lemma store_valid_access_3: - valid_access m1 chunk b ofs. -Proof. - elim (store_inv _ _ _ _ _ _ STORE); intros. auto. -Qed. - -Hint Resolve store_valid_access_1 store_valid_access_2 - store_valid_access_3: mem. - -Theorem load_store_similar: - forall chunk', - size_chunk chunk' = size_chunk chunk -> - load chunk' m2 b ofs = Some (Val.load_result chunk' v). -Proof. - intros. destruct (store_inv _ _ _ _ _ _ STORE). - unfold load. rewrite in_bounds_true. - decEq. decEq. rewrite H1. unfold unchecked_store; simpl. - rewrite update_s. simpl. - replace (pred_size_chunk chunk) with (pred_size_chunk chunk'). - apply getN_setN_same. - repeat rewrite size_chunk_pred in H. omega. - apply store_valid_access_1. - inv H0. constructor; auto. congruence. - rewrite (align_chunk_compat _ _ H). auto. -Qed. - -Theorem load_store_same: - load chunk m2 b ofs = Some (Val.load_result chunk v). -Proof. - eapply load_store_similar; eauto. -Qed. - -Theorem load_store_other: - forall chunk' b' ofs', - b' <> b - \/ ofs' + size_chunk chunk' <= ofs - \/ ofs + size_chunk chunk <= ofs' -> - load chunk' m2 b' ofs' = load chunk' m1 b' ofs'. -Proof. - intros. destruct (store_inv _ _ _ _ _ _ STORE). - unfold load. destruct (in_bounds m1 chunk' b' ofs'). - rewrite in_bounds_true. decEq. decEq. - rewrite H1; unfold unchecked_store; simpl. - unfold update. destruct (zeq b' b). subst b'. - simpl. repeat rewrite size_chunk_pred in H. - apply getN_setN_other. elim H; intro. congruence. omega. - auto. - eauto with mem. - destruct (in_bounds m2 chunk' b' ofs'); auto. - elim n. eauto with mem. -Qed. - -Theorem load_store_overlap: - forall chunk' ofs' v', - load chunk' m2 b ofs' = Some v' -> - ofs' <> ofs -> - ofs' + size_chunk chunk' > ofs -> - ofs + size_chunk chunk > ofs' -> - v' = Vundef. -Proof. - intros. destruct (store_inv _ _ _ _ _ _ STORE). - destruct (load_inv _ _ _ _ _ H). rewrite H6. - rewrite H4. unfold unchecked_store. simpl. rewrite update_s. - simpl. rewrite getN_setN_overlap. - destruct chunk'; reflexivity. - auto. rewrite size_chunk_pred in H2. omega. - rewrite size_chunk_pred in H1. omega. -Qed. - -Theorem load_store_overlap': - forall chunk' ofs', - valid_access m1 chunk' b ofs' -> - ofs' <> ofs -> - ofs' + size_chunk chunk' > ofs -> - ofs + size_chunk chunk > ofs' -> - load chunk' m2 b ofs' = Some Vundef. -Proof. - intros. - assert (exists v', load chunk' m2 b ofs' = Some v'). - eauto with mem. - destruct H3 as [v' LOAD]. rewrite LOAD. decEq. - eapply load_store_overlap; eauto. -Qed. - -Theorem load_store_mismatch: - forall chunk' v', - load chunk' m2 b ofs = Some v' -> - size_chunk chunk' <> size_chunk chunk -> - v' = Vundef. -Proof. - intros. destruct (store_inv _ _ _ _ _ _ STORE). - destruct (load_inv _ _ _ _ _ H). rewrite H4. - rewrite H2. unfold unchecked_store. simpl. rewrite update_s. - simpl. rewrite getN_setN_mismatch. - destruct chunk'; reflexivity. - repeat rewrite size_chunk_pred in H0; omega. -Qed. - -Theorem load_store_mismatch': - forall chunk', - valid_access m1 chunk' b ofs -> - size_chunk chunk' <> size_chunk chunk -> - load chunk' m2 b ofs = Some Vundef. -Proof. - intros. - assert (exists v', load chunk' m2 b ofs = Some v'). - eauto with mem. - destruct H1 as [v' LOAD]. rewrite LOAD. decEq. - eapply load_store_mismatch; eauto. -Qed. - -Inductive load_store_cases - (chunk1: memory_chunk) (b1: block) (ofs1: Z) - (chunk2: memory_chunk) (b2: block) (ofs2: Z) : Type := - | lsc_similar: - b1 = b2 -> ofs1 = ofs2 -> size_chunk chunk1 = size_chunk chunk2 -> - load_store_cases chunk1 b1 ofs1 chunk2 b2 ofs2 - | lsc_other: - b1 <> b2 \/ ofs2 + size_chunk chunk2 <= ofs1 \/ ofs1 + size_chunk chunk1 <= ofs2 -> - load_store_cases chunk1 b1 ofs1 chunk2 b2 ofs2 - | lsc_overlap: - b1 = b2 -> ofs1 <> ofs2 -> ofs2 + size_chunk chunk2 > ofs1 -> ofs1 + size_chunk chunk1 > ofs2 -> - load_store_cases chunk1 b1 ofs1 chunk2 b2 ofs2 - | lsc_mismatch: - b1 = b2 -> ofs1 = ofs2 -> size_chunk chunk1 <> size_chunk chunk2 -> - load_store_cases chunk1 b1 ofs1 chunk2 b2 ofs2. - -Definition load_store_classification: - forall (chunk1: memory_chunk) (b1: block) (ofs1: Z) - (chunk2: memory_chunk) (b2: block) (ofs2: Z), - load_store_cases chunk1 b1 ofs1 chunk2 b2 ofs2. -Proof. - intros. destruct (eq_block b1 b2). - destruct (zeq ofs1 ofs2). - destruct (zeq (size_chunk chunk1) (size_chunk chunk2)). - apply lsc_similar; auto. - apply lsc_mismatch; auto. - destruct (zle (ofs2 + size_chunk chunk2) ofs1). - apply lsc_other. tauto. - destruct (zle (ofs1 + size_chunk chunk1) ofs2). - apply lsc_other. tauto. - apply lsc_overlap; auto. - apply lsc_other; tauto. -Qed. - -Theorem load_store_characterization: - forall chunk' b' ofs', - valid_access m1 chunk' b' ofs' -> - load chunk' m2 b' ofs' = - match load_store_classification chunk b ofs chunk' b' ofs' with - | lsc_similar _ _ _ => Some (Val.load_result chunk' v) - | lsc_other _ => load chunk' m1 b' ofs' - | lsc_overlap _ _ _ _ => Some Vundef - | lsc_mismatch _ _ _ => Some Vundef - end. -Proof. - intros. - assert (exists v', load chunk' m2 b' ofs' = Some v') by eauto with mem. - destruct H0 as [v' LOAD]. - destruct (load_store_classification chunk b ofs chunk' b' ofs'). - subst b' ofs'. apply load_store_similar; auto. - apply load_store_other; intuition. - subst b'. rewrite LOAD. decEq. - eapply load_store_overlap; eauto. - subst b' ofs'. rewrite LOAD. decEq. - eapply load_store_mismatch; eauto. -Qed. - -End STORE. - -Hint Resolve store_valid_block_1 store_valid_block_2: mem. -Hint Resolve store_valid_access_1 store_valid_access_2 - store_valid_access_3: mem. - -(** ** Properties related to [alloc]. *) - -Section ALLOC. - -Variable m1: mem. -Variables lo hi: Z. -Variable m2: mem. -Variable b: block. -Hypothesis ALLOC: alloc m1 lo hi = (m2, b). - -Lemma nextblock_alloc: - nextblock m2 = Zsucc (nextblock m1). -Proof. - injection ALLOC; intros. rewrite <- H0; auto. -Qed. - -Lemma alloc_result: - b = nextblock m1. -Proof. - injection ALLOC; auto. -Qed. - -Lemma valid_block_alloc: - forall b', valid_block m1 b' -> valid_block m2 b'. -Proof. - unfold valid_block; intros. rewrite nextblock_alloc. omega. -Qed. - -Lemma fresh_block_alloc: - ~(valid_block m1 b). -Proof. - unfold valid_block. rewrite alloc_result. omega. -Qed. - -Lemma valid_new_block: - valid_block m2 b. -Proof. - unfold valid_block. rewrite alloc_result. rewrite nextblock_alloc. omega. -Qed. - -Hint Resolve valid_block_alloc fresh_block_alloc valid_new_block: mem. - -Lemma valid_block_alloc_inv: - forall b', valid_block m2 b' -> b' = b \/ valid_block m1 b'. -Proof. - unfold valid_block; intros. - rewrite nextblock_alloc in H. rewrite alloc_result. - unfold block; omega. -Qed. - -Lemma low_bound_alloc: - forall b', low_bound m2 b' = if zeq b' b then lo else low_bound m1 b'. -Proof. - intros. injection ALLOC; intros. rewrite <- H0; unfold low_bound; simpl. - unfold update. rewrite H. destruct (zeq b' b); auto. -Qed. - -Lemma low_bound_alloc_same: - low_bound m2 b = lo. -Proof. - rewrite low_bound_alloc. apply zeq_true. -Qed. - -Lemma low_bound_alloc_other: - forall b', valid_block m1 b' -> low_bound m2 b' = low_bound m1 b'. -Proof. - intros; rewrite low_bound_alloc. - apply zeq_false. eauto with mem. -Qed. - -Lemma high_bound_alloc: - forall b', high_bound m2 b' = if zeq b' b then hi else high_bound m1 b'. -Proof. - intros. injection ALLOC; intros. rewrite <- H0; unfold high_bound; simpl. - unfold update. rewrite H. destruct (zeq b' b); auto. -Qed. - -Lemma high_bound_alloc_same: - high_bound m2 b = hi. -Proof. - rewrite high_bound_alloc. apply zeq_true. -Qed. - -Lemma high_bound_alloc_other: - forall b', valid_block m1 b' -> high_bound m2 b' = high_bound m1 b'. -Proof. - intros; rewrite high_bound_alloc. - apply zeq_false. eauto with mem. -Qed. - -Lemma valid_access_alloc_other: - forall chunk b' ofs, - valid_access m1 chunk b' ofs -> - valid_access m2 chunk b' ofs. -Proof. - intros. inv H. constructor; auto with mem. - rewrite low_bound_alloc_other; auto. - rewrite high_bound_alloc_other; auto. -Qed. - -Lemma valid_access_alloc_same: - forall chunk ofs, - lo <= ofs -> ofs + size_chunk chunk <= hi -> (align_chunk chunk | ofs) -> - valid_access m2 chunk b ofs. -Proof. - intros. constructor; auto with mem. - rewrite low_bound_alloc_same; auto. - rewrite high_bound_alloc_same; auto. -Qed. - -Hint Resolve valid_access_alloc_other valid_access_alloc_same: mem. - -Lemma valid_access_alloc_inv: - forall chunk b' ofs, - valid_access m2 chunk b' ofs -> - valid_access m1 chunk b' ofs \/ - (b' = b /\ lo <= ofs /\ ofs + size_chunk chunk <= hi /\ (align_chunk chunk | ofs)). -Proof. - intros. inv H. - elim (valid_block_alloc_inv _ H0); intro. - subst b'. rewrite low_bound_alloc_same in H1. - rewrite high_bound_alloc_same in H2. - right. tauto. - left. constructor; auto. - rewrite low_bound_alloc_other in H1; auto. - rewrite high_bound_alloc_other in H2; auto. -Qed. - -Theorem load_alloc_unchanged: - forall chunk b' ofs, - valid_block m1 b' -> - load chunk m2 b' ofs = load chunk m1 b' ofs. -Proof. - intros. unfold load. - destruct (in_bounds m2 chunk b' ofs). - elim (valid_access_alloc_inv _ _ _ v). intro. - rewrite in_bounds_true; auto. - injection ALLOC; intros. rewrite <- H2; simpl. - rewrite update_o. auto. rewrite H1. apply sym_not_equal. eauto with mem. - intros [A [B C]]. subst b'. elimtype False. eauto with mem. - destruct (in_bounds m1 chunk b' ofs). - elim n; eauto with mem. - auto. -Qed. - -Theorem load_alloc_other: - forall chunk b' ofs v, - load chunk m1 b' ofs = Some v -> - load chunk m2 b' ofs = Some v. -Proof. - intros. rewrite <- H. apply load_alloc_unchanged. eauto with mem. -Qed. - -Theorem load_alloc_same: - forall chunk ofs v, - load chunk m2 b ofs = Some v -> - v = Vundef. -Proof. - intros. destruct (load_inv _ _ _ _ _ H). rewrite H1. - injection ALLOC; intros. rewrite <- H3; simpl. - rewrite <- H2. rewrite update_s. - simpl. rewrite getN_init. destruct chunk; auto. -Qed. - -Theorem load_alloc_same': - forall chunk ofs, - lo <= ofs -> ofs + size_chunk chunk <= hi -> (align_chunk chunk | ofs) -> - load chunk m2 b ofs = Some Vundef. -Proof. - intros. assert (exists v, load chunk m2 b ofs = Some v). - apply valid_access_load. constructor; eauto with mem. - rewrite low_bound_alloc_same. auto. - rewrite high_bound_alloc_same. auto. - destruct H2 as [v LOAD]. rewrite LOAD. decEq. - eapply load_alloc_same; eauto. -Qed. - -End ALLOC. - -Hint Resolve valid_block_alloc fresh_block_alloc valid_new_block: mem. -Hint Resolve valid_access_alloc_other valid_access_alloc_same: mem. -Hint Resolve load_alloc_unchanged: mem. - -(** ** Properties related to [free]. *) - -Section FREE. - -Variable m: mem. -Variable bf: block. - -Lemma valid_block_free_1: - forall b, valid_block m b -> valid_block (free m bf) b. -Proof. - unfold valid_block, free; intros; simpl; auto. -Qed. - -Lemma valid_block_free_2: - forall b, valid_block (free m bf) b -> valid_block m b. -Proof. - unfold valid_block, free; intros; simpl in *; auto. -Qed. - -Hint Resolve valid_block_free_1 valid_block_free_2: mem. - -Lemma low_bound_free: - forall b, b <> bf -> low_bound (free m bf) b = low_bound m b. -Proof. - intros. unfold low_bound, free; simpl. - rewrite update_o; auto. -Qed. - -Lemma high_bound_free: - forall b, b <> bf -> high_bound (free m bf) b = high_bound m b. -Proof. - intros. unfold high_bound, free; simpl. - rewrite update_o; auto. -Qed. - -Lemma low_bound_free_same: - forall m b, low_bound (free m b) b = 0. -Proof. - intros. unfold low_bound; simpl. rewrite update_s. simpl; omega. -Qed. - -Lemma high_bound_free_same: - forall m b, high_bound (free m b) b = 0. -Proof. - intros. unfold high_bound; simpl. rewrite update_s. simpl; omega. -Qed. - -Lemma valid_access_free_1: - forall chunk b ofs, - valid_access m chunk b ofs -> b <> bf -> - valid_access (free m bf) chunk b ofs. -Proof. - intros. inv H. constructor; auto with mem. - rewrite low_bound_free; auto. rewrite high_bound_free; auto. -Qed. - -Lemma valid_access_free_2: - forall chunk ofs, ~(valid_access (free m bf) chunk bf ofs). -Proof. - intros; red; intros. inv H. - unfold free, low_bound in H1; simpl in H1. rewrite update_s in H1. simpl in H1. - unfold free, high_bound in H2; simpl in H2. rewrite update_s in H2. simpl in H2. - generalize (size_chunk_pos chunk). omega. -Qed. - -Hint Resolve valid_access_free_1 valid_access_free_2: mem. - -Lemma valid_access_free_inv: - forall chunk b ofs, - valid_access (free m bf) chunk b ofs -> - valid_access m chunk b ofs /\ b <> bf. -Proof. - intros. destruct (eq_block b bf). subst b. - elim (valid_access_free_2 _ _ H). - inv H. rewrite low_bound_free in H1; auto. - rewrite high_bound_free in H2; auto. - split; auto. constructor; auto with mem. -Qed. - -Theorem load_free: - forall chunk b ofs, - b <> bf -> - load chunk (free m bf) b ofs = load chunk m b ofs. -Proof. - intros. unfold load. - destruct (in_bounds m chunk b ofs). - rewrite in_bounds_true; auto with mem. - unfold free; simpl. rewrite update_o; auto. - destruct (in_bounds (free m bf) chunk b ofs); auto. - elim n. elim (valid_access_free_inv _ _ _ v); auto. -Qed. - -End FREE. - -(** ** Properties related to [free_list] *) - -Lemma valid_block_free_list_1: - forall bl m b, valid_block m b -> valid_block (free_list m bl) b. -Proof. - induction bl; simpl; intros. auto. - apply valid_block_free_1; auto. -Qed. - -Lemma valid_block_free_list_2: - forall bl m b, valid_block (free_list m bl) b -> valid_block m b. -Proof. - induction bl; simpl; intros. auto. - apply IHbl. apply valid_block_free_2 with a; auto. -Qed. - -Lemma valid_access_free_list: - forall chunk b ofs m bl, - valid_access m chunk b ofs -> ~In b bl -> - valid_access (free_list m bl) chunk b ofs. -Proof. - induction bl; simpl; intros. auto. - apply valid_access_free_1. apply IHbl. auto. intuition. intuition congruence. -Qed. - -Lemma valid_access_free_list_inv: - forall chunk b ofs m bl, - valid_access (free_list m bl) chunk b ofs -> - ~In b bl /\ valid_access m chunk b ofs. -Proof. - induction bl; simpl; intros. - tauto. - elim (valid_access_free_inv _ _ _ _ _ H); intros. - elim (IHbl H0); intros. - intuition congruence. -Qed. - -(** ** Properties related to pointer validity *) - -Lemma valid_pointer_valid_access: - forall m b ofs, - valid_pointer m b ofs = true <-> valid_access m Mint8unsigned b ofs. -Proof. - unfold valid_pointer; intros; split; intros. - destruct (andb_prop _ _ H). destruct (andb_prop _ _ H0). - constructor. red. eapply proj_sumbool_true; eauto. - eapply proj_sumbool_true; eauto. - change (size_chunk Mint8unsigned) with 1. - generalize (proj_sumbool_true _ H1). omega. - simpl. apply Zone_divide. - inv H. unfold proj_sumbool. - rewrite zlt_true; auto. rewrite zle_true; auto. - change (size_chunk Mint8unsigned) with 1 in H2. - rewrite zlt_true. auto. omega. -Qed. - -Theorem valid_pointer_alloc: - forall (m1 m2: mem) (lo hi: Z) (b b': block) (ofs: Z), - alloc m1 lo hi = (m2, b') -> - valid_pointer m1 b ofs = true -> - valid_pointer m2 b ofs = true. -Proof. - intros. rewrite valid_pointer_valid_access in H0. - rewrite valid_pointer_valid_access. - eauto with mem. -Qed. - -Theorem valid_pointer_store: - forall (chunk: memory_chunk) (m1 m2: mem) (b b': block) (ofs ofs': Z) (v: val), - store chunk m1 b' ofs' v = Some m2 -> - valid_pointer m1 b ofs = true -> valid_pointer m2 b ofs = true. -Proof. - intros. rewrite valid_pointer_valid_access in H0. - rewrite valid_pointer_valid_access. - eauto with mem. -Qed. - -(** * Generic injections between memory states. *) - -Section GENERIC_INJECT. - -Definition meminj : Type := block -> option (block * Z). - -Variable val_inj: meminj -> val -> val -> Prop. - -Hypothesis val_inj_undef: - forall mi, val_inj mi Vundef Vundef. - -Definition mem_inj (mi: meminj) (m1 m2: mem) := - forall chunk b1 ofs v1 b2 delta, - mi b1 = Some(b2, delta) -> - load chunk m1 b1 ofs = Some v1 -> - exists v2, load chunk m2 b2 (ofs + delta) = Some v2 /\ val_inj mi v1 v2. - -Lemma valid_access_inj: - forall mi m1 m2 chunk b1 ofs b2 delta, - mi b1 = Some(b2, delta) -> - mem_inj mi m1 m2 -> - valid_access m1 chunk b1 ofs -> - valid_access m2 chunk b2 (ofs + delta). -Proof. - intros. - assert (exists v1, load chunk m1 b1 ofs = Some v1) by eauto with mem. - destruct H2 as [v1 LOAD1]. - destruct (H0 _ _ _ _ _ _ H LOAD1) as [v2 [LOAD2 VCP]]. - eauto with mem. -Qed. - -Hint Resolve valid_access_inj: mem. - -Lemma store_unmapped_inj: - forall mi m1 m2 b ofs v chunk m1', - mem_inj mi m1 m2 -> - mi b = None -> - store chunk m1 b ofs v = Some m1' -> - mem_inj mi m1' m2. -Proof. - intros; red; intros. - assert (load chunk0 m1 b1 ofs0 = Some v1). - rewrite <- H3; symmetry. eapply load_store_other; eauto. - left. congruence. - eapply H; eauto. -Qed. - -Lemma store_outside_inj: - forall mi m1 m2 chunk b ofs v m2', - mem_inj mi m1 m2 -> - (forall b' delta, - mi b' = Some(b, delta) -> - high_bound m1 b' + delta <= ofs - \/ ofs + size_chunk chunk <= low_bound m1 b' + delta) -> - store chunk m2 b ofs v = Some m2' -> - mem_inj mi m1 m2'. -Proof. - intros; red; intros. - exploit H; eauto. intros [v2 [LOAD2 VINJ]]. - exists v2; split; auto. - rewrite <- LOAD2. eapply load_store_other; eauto. - destruct (eq_block b2 b). subst b2. - right. generalize (H0 _ _ H2); intro. - assert (valid_access m1 chunk0 b1 ofs0) by eauto with mem. - inv H5. omega. auto. -Qed. - -Definition meminj_no_overlap (mi: meminj) (m: mem) : Prop := - forall b1 b1' delta1 b2 b2' delta2, - b1 <> b2 -> - mi b1 = Some (b1', delta1) -> - mi b2 = Some (b2', delta2) -> - b1' <> b2' - \/ low_bound m b1 >= high_bound m b1 - \/ low_bound m b2 >= high_bound m b2 - \/ high_bound m b1 + delta1 <= low_bound m b2 + delta2 - \/ high_bound m b2 + delta2 <= low_bound m b1 + delta1. - -Lemma store_mapped_inj: - forall mi m1 m2 b1 ofs b2 delta v1 v2 chunk m1', - mem_inj mi m1 m2 -> - meminj_no_overlap mi m1 -> - mi b1 = Some(b2, delta) -> - store chunk m1 b1 ofs v1 = Some m1' -> - (forall chunk', size_chunk chunk' = size_chunk chunk -> - val_inj mi (Val.load_result chunk' v1) (Val.load_result chunk' v2)) -> - exists m2', - store chunk m2 b2 (ofs + delta) v2 = Some m2' /\ mem_inj mi m1' m2'. -Proof. - intros. - assert (exists m2', store chunk m2 b2 (ofs + delta) v2 = Some m2') by eauto with mem. - destruct H4 as [m2' STORE2]. - exists m2'; split. auto. - red. intros chunk' b1' ofs' v b2' delta' CP LOAD1. - assert (valid_access m1 chunk' b1' ofs') by eauto with mem. - generalize (load_store_characterization _ _ _ _ _ _ H2 _ _ _ H4). - destruct (load_store_classification chunk b1 ofs chunk' b1' ofs'); - intro. - (* similar *) - subst b1' ofs'. - rewrite CP in H1. inv H1. - rewrite LOAD1 in H5. inv H5. - exists (Val.load_result chunk' v2). split. - eapply load_store_similar; eauto. - auto. - (* disjoint *) - rewrite LOAD1 in H5. - destruct (H _ _ _ _ _ _ CP (sym_equal H5)) as [v2' [LOAD2 VCP]]. - exists v2'. split; auto. - rewrite <- LOAD2. eapply load_store_other; eauto. - destruct (eq_block b1 b1'). subst b1'. - rewrite CP in H1; inv H1. - right. elim o; [congruence | omega]. - assert (valid_access m1 chunk b1 ofs) by eauto with mem. - generalize (H0 _ _ _ _ _ _ n H1 CP). intros [A | [A | [A | A]]]. - auto. - inv H6. generalize (size_chunk_pos chunk). intro. omegaContradiction. - inv H4. generalize (size_chunk_pos chunk'). intro. omegaContradiction. - right. inv H4. inv H6. omega. - (* overlapping *) - subst b1'. rewrite CP in H1; inv H1. - assert (exists v2', load chunk' m2' b2 (ofs' + delta) = Some v2') by eauto with mem. - destruct H1 as [v2' LOAD2']. - assert (v2' = Vundef). eapply load_store_overlap; eauto. - omega. omega. omega. - exists v2'; split. auto. - replace v with Vundef by congruence. subst v2'. apply val_inj_undef. - (* mismatch *) - subst b1' ofs'. rewrite CP in H1; inv H1. - assert (exists v2', load chunk' m2' b2 (ofs + delta) = Some v2') by eauto with mem. - destruct H1 as [v2' LOAD2']. - assert (v2' = Vundef). eapply load_store_mismatch; eauto. - exists v2'; split. auto. - replace v with Vundef by congruence. subst v2'. apply val_inj_undef. -Qed. - -Definition inj_offset_aligned (delta: Z) (size: Z) : Prop := - forall chunk, size_chunk chunk <= size -> (align_chunk chunk | delta). - -Lemma alloc_parallel_inj: - forall mi m1 m2 lo1 hi1 m1' b1 lo2 hi2 m2' b2 delta, - mem_inj mi m1 m2 -> - alloc m1 lo1 hi1 = (m1', b1) -> - alloc m2 lo2 hi2 = (m2', b2) -> - mi b1 = Some(b2, delta) -> - lo2 <= lo1 + delta -> hi1 + delta <= hi2 -> - inj_offset_aligned delta (hi1 - lo1) -> - mem_inj mi m1' m2'. -Proof. - intros; red; intros. - exploit (valid_access_alloc_inv m1); eauto with mem. - intros [A | [A [B [C D]]]]. - assert (load chunk m1 b0 ofs = Some v1). - rewrite <- H7. symmetry. eapply load_alloc_unchanged; eauto with mem. - exploit H; eauto. intros [v2 [LOAD2 VINJ]]. - exists v2; split. - rewrite <- LOAD2. eapply load_alloc_unchanged; eauto with mem. - auto. - subst b0. rewrite H2 in H6. inversion H6. subst b3 delta0. - assert (v1 = Vundef). eapply load_alloc_same with (m1 := m1); eauto. - subst v1. - assert (exists v2, load chunk m2' b2 (ofs + delta) = Some v2). - apply valid_access_load. - eapply valid_access_alloc_same; eauto. omega. omega. - apply Zdivide_plus_r; auto. apply H5. omega. - destruct H8 as [v2 LOAD2]. - assert (v2 = Vundef). eapply load_alloc_same with (m1 := m2); eauto. - subst v2. - exists Vundef; split. auto. apply val_inj_undef. -Qed. - -Lemma alloc_right_inj: - forall mi m1 m2 lo hi b2 m2', - mem_inj mi m1 m2 -> - alloc m2 lo hi = (m2', b2) -> - mem_inj mi m1 m2'. -Proof. - intros; red; intros. - exploit H; eauto. intros [v2 [LOAD2 VINJ]]. - exists v2; split; auto. - assert (valid_block m2 b0). - apply valid_access_valid_block with chunk (ofs + delta). - eauto with mem. - rewrite <- LOAD2. eapply load_alloc_unchanged; eauto. -Qed. - -Hypothesis val_inj_undef_any: - forall mi v, val_inj mi Vundef v. - -Lemma alloc_left_unmapped_inj: - forall mi m1 m2 lo hi b1 m1', - mem_inj mi m1 m2 -> - alloc m1 lo hi = (m1', b1) -> - mi b1 = None -> - mem_inj mi m1' m2. -Proof. - intros; red; intros. - exploit (valid_access_alloc_inv m1); eauto with mem. - intros [A | [A [B C]]]. - eapply H; eauto. - rewrite <- H3. symmetry. eapply load_alloc_unchanged; eauto with mem. - subst b0. congruence. -Qed. - -Lemma alloc_left_mapped_inj: - forall mi m1 m2 lo hi b1 m1' b2 delta, - mem_inj mi m1 m2 -> - alloc m1 lo hi = (m1', b1) -> - mi b1 = Some(b2, delta) -> - valid_block m2 b2 -> - low_bound m2 b2 <= lo + delta -> hi + delta <= high_bound m2 b2 -> - inj_offset_aligned delta (hi - lo) -> - mem_inj mi m1' m2. -Proof. - intros; red; intros. - exploit (valid_access_alloc_inv m1); eauto with mem. - intros [A | [A [B [C D]]]]. - eapply H; eauto. - rewrite <- H7. symmetry. eapply load_alloc_unchanged; eauto with mem. - subst b0. rewrite H1 in H6. inversion H6. subst b3 delta0. - assert (v1 = Vundef). eapply load_alloc_same with (m1 := m1); eauto. - subst v1. - assert (exists v2, load chunk m2 b2 (ofs + delta) = Some v2). - apply valid_access_load. constructor. auto. omega. omega. - apply Zdivide_plus_r; auto. apply H5. omega. - destruct H8 as [v2 LOAD2]. exists v2; split. auto. - apply val_inj_undef_any. -Qed. - -Lemma free_parallel_inj: - forall mi m1 m2 b1 b2 delta, - mem_inj mi m1 m2 -> - mi b1 = Some(b2, delta) -> - (forall b delta', mi b = Some(b2, delta') -> b = b1) -> - mem_inj mi (free m1 b1) (free m2 b2). -Proof. - intros; red; intros. - exploit valid_access_free_inv; eauto with mem. intros [A B]. - assert (load chunk m1 b0 ofs = Some v1). - rewrite <- H3. symmetry. apply load_free. auto. - exploit H; eauto. intros [v2 [LOAD2 INJ]]. - exists v2; split. - rewrite <- LOAD2. apply load_free. - red; intro; subst b3. elim B. eauto. - auto. -Qed. - -Lemma free_left_inj: - forall mi m1 m2 b1, - mem_inj mi m1 m2 -> - mem_inj mi (free m1 b1) m2. -Proof. - intros; red; intros. - exploit valid_access_free_inv; eauto with mem. intros [A B]. - eapply H; eauto with mem. - rewrite <- H1; symmetry; eapply load_free; eauto. -Qed. - -Lemma free_list_left_inj: - forall mi bl m1 m2, - mem_inj mi m1 m2 -> - mem_inj mi (free_list m1 bl) m2. -Proof. - induction bl; intros; simpl. - auto. - apply free_left_inj. auto. -Qed. - -Lemma free_right_inj: - forall mi m1 m2 b2, - mem_inj mi m1 m2 -> - (forall b1 delta chunk ofs, - mi b1 = Some(b2, delta) -> ~(valid_access m1 chunk b1 ofs)) -> - mem_inj mi m1 (free m2 b2). -Proof. - intros; red; intros. - assert (b0 <> b2). - red; intro; subst b0. elim (H0 b1 delta chunk ofs H1). - eauto with mem. - exploit H; eauto. intros [v2 [LOAD2 INJ]]. - exists v2; split; auto. - rewrite <- LOAD2. apply load_free. auto. -Qed. - -Lemma valid_pointer_inj: - forall mi m1 m2 b1 ofs b2 delta, - mi b1 = Some(b2, delta) -> - mem_inj mi m1 m2 -> - valid_pointer m1 b1 ofs = true -> - valid_pointer m2 b2 (ofs + delta) = true. -Proof. - intros. rewrite valid_pointer_valid_access in H1. - rewrite valid_pointer_valid_access. eauto with mem. -Qed. - -End GENERIC_INJECT. - -(** ** Store extensions *) - -(** A store [m2] extends a store [m1] if [m2] can be obtained from [m1] - by increasing the sizes of the memory blocks of [m1] (decreasing - the low bounds, increasing the high bounds), while still keeping the - same contents for block offsets that are valid in [m1]. *) - -Definition inject_id : meminj := fun b => Some(b, 0). - -Definition val_inj_id (mi: meminj) (v1 v2: val) : Prop := v1 = v2. - -Definition extends (m1 m2: mem) := - nextblock m1 = nextblock m2 /\ mem_inj val_inj_id inject_id m1 m2. - -Theorem extends_refl: - forall (m: mem), extends m m. -Proof. - intros; split. auto. - red; unfold inject_id; intros. inv H. - exists v1; split. replace (ofs + 0) with ofs by omega. auto. - unfold val_inj_id; auto. -Qed. - -Theorem alloc_extends: - forall (m1 m2 m1' m2': mem) (lo1 hi1 lo2 hi2: Z) (b1 b2: block), - extends m1 m2 -> - lo2 <= lo1 -> hi1 <= hi2 -> - alloc m1 lo1 hi1 = (m1', b1) -> - alloc m2 lo2 hi2 = (m2', b2) -> - b1 = b2 /\ extends m1' m2'. -Proof. - intros. destruct H. - assert (b1 = b2). - transitivity (nextblock m1). eapply alloc_result; eauto. - symmetry. rewrite H. eapply alloc_result; eauto. - subst b2. split. auto. split. - rewrite (nextblock_alloc _ _ _ _ _ H2). - rewrite (nextblock_alloc _ _ _ _ _ H3). - congruence. - eapply alloc_parallel_inj; eauto. - unfold val_inj_id; auto. - unfold inject_id; eauto. - omega. omega. - red; intros. apply Zdivide_0. -Qed. - -Theorem free_extends: - forall (m1 m2: mem) (b: block), - extends m1 m2 -> - extends (free m1 b) (free m2 b). -Proof. - intros. destruct H. split. - simpl; auto. - eapply free_parallel_inj; eauto. - unfold inject_id. eauto. - unfold inject_id; intros. congruence. -Qed. - -Theorem load_extends: - forall (chunk: memory_chunk) (m1 m2: mem) (b: block) (ofs: Z) (v: val), - extends m1 m2 -> - load chunk m1 b ofs = Some v -> - load chunk m2 b ofs = Some v. -Proof. - intros. destruct H. - exploit H1; eauto. unfold inject_id. eauto. - unfold val_inj_id. intros [v2 [LOAD EQ]]. - replace (ofs + 0) with ofs in LOAD by omega. congruence. -Qed. - -Theorem store_within_extends: - forall (chunk: memory_chunk) (m1 m2 m1': mem) (b: block) (ofs: Z) (v: val), - extends m1 m2 -> - store chunk m1 b ofs v = Some m1' -> - exists m2', store chunk m2 b ofs v = Some m2' /\ extends m1' m2'. -Proof. - intros. destruct H. - exploit store_mapped_inj; eauto. - unfold val_inj_id; eauto. - unfold meminj_no_overlap, inject_id; intros. - inv H3. inv H4. auto. - unfold inject_id; eauto. - unfold val_inj_id; intros. eauto. - intros [m2' [STORE MINJ]]. - exists m2'; split. - replace (ofs + 0) with ofs in STORE by omega. auto. - split. - rewrite (nextblock_store _ _ _ _ _ _ H0). - rewrite (nextblock_store _ _ _ _ _ _ STORE). - auto. - auto. -Qed. - -Theorem store_outside_extends: - forall (chunk: memory_chunk) (m1 m2 m2': mem) (b: block) (ofs: Z) (v: val), - extends m1 m2 -> - ofs + size_chunk chunk <= low_bound m1 b \/ high_bound m1 b <= ofs -> - store chunk m2 b ofs v = Some m2' -> - extends m1 m2'. -Proof. - intros. destruct H. split. - rewrite (nextblock_store _ _ _ _ _ _ H1). auto. - eapply store_outside_inj; eauto. - unfold inject_id; intros. inv H3. omega. -Qed. - -Theorem valid_pointer_extends: - forall m1 m2 b ofs, - extends m1 m2 -> valid_pointer m1 b ofs = true -> - valid_pointer m2 b ofs = true. -Proof. - intros. destruct H. - replace ofs with (ofs + 0) by omega. - apply valid_pointer_inj with val_inj_id inject_id m1 b; auto. -Qed. - -(** * The ``less defined than'' relation over memory states *) - -(** A memory state [m1] is less defined than [m2] if, for all addresses, - the value [v1] read in [m1] at this address is less defined than - the value [v2] read in [m2], that is, either [v1 = v2] or [v1 = Vundef]. *) - -Definition val_inj_lessdef (mi: meminj) (v1 v2: val) : Prop := - Val.lessdef v1 v2. - -Definition lessdef (m1 m2: mem) : Prop := - nextblock m1 = nextblock m2 /\ - mem_inj val_inj_lessdef inject_id m1 m2. - -Lemma lessdef_refl: - forall m, lessdef m m. -Proof. - intros; split. auto. - red; intros. unfold inject_id in H. inv H. - exists v1; split. replace (ofs + 0) with ofs by omega. auto. - red. constructor. -Qed. - -Lemma load_lessdef: - forall m1 m2 chunk b ofs v1, - lessdef m1 m2 -> load chunk m1 b ofs = Some v1 -> - exists v2, load chunk m2 b ofs = Some v2 /\ Val.lessdef v1 v2. -Proof. - intros. destruct H. - exploit H1; eauto. unfold inject_id. eauto. - intros [v2 [LOAD INJ]]. exists v2; split. - replace ofs with (ofs + 0) by omega. auto. - auto. -Qed. - -Lemma loadv_lessdef: - forall m1 m2 chunk addr1 addr2 v1, - lessdef m1 m2 -> Val.lessdef addr1 addr2 -> - loadv chunk m1 addr1 = Some v1 -> - exists v2, loadv chunk m2 addr2 = Some v2 /\ Val.lessdef v1 v2. -Proof. - intros. inv H0. - destruct addr2; simpl in *; try discriminate. - eapply load_lessdef; eauto. - simpl in H1; discriminate. -Qed. - -Lemma store_lessdef: - forall m1 m2 chunk b ofs v1 v2 m1', - lessdef m1 m2 -> Val.lessdef v1 v2 -> - store chunk m1 b ofs v1 = Some m1' -> - exists m2', store chunk m2 b ofs v2 = Some m2' /\ lessdef m1' m2'. -Proof. - intros. destruct H. - exploit store_mapped_inj; eauto. - unfold val_inj_lessdef; intros; constructor. - red; unfold inject_id; intros. inv H4. inv H5. auto. - unfold inject_id; eauto. - unfold val_inj_lessdef; intros. - apply Val.load_result_lessdef. eexact H0. - intros [m2' [STORE MINJ]]. - exists m2'; split. replace ofs with (ofs + 0) by omega. auto. - split. - rewrite (nextblock_store _ _ _ _ _ _ H1). - rewrite (nextblock_store _ _ _ _ _ _ STORE). - auto. - auto. -Qed. - -Lemma storev_lessdef: - forall m1 m2 chunk addr1 v1 addr2 v2 m1', - lessdef m1 m2 -> Val.lessdef addr1 addr2 -> Val.lessdef v1 v2 -> - storev chunk m1 addr1 v1 = Some m1' -> - exists m2', storev chunk m2 addr2 v2 = Some m2' /\ lessdef m1' m2'. -Proof. - intros. inv H0. - destruct addr2; simpl in H2; try discriminate. - simpl. eapply store_lessdef; eauto. - discriminate. -Qed. - -Lemma alloc_lessdef: - forall m1 m2 lo hi b1 m1' b2 m2', - lessdef m1 m2 -> alloc m1 lo hi = (m1', b1) -> alloc m2 lo hi = (m2', b2) -> - b1 = b2 /\ lessdef m1' m2'. -Proof. - intros. destruct H. - assert (b1 = b2). - transitivity (nextblock m1). eapply alloc_result; eauto. - symmetry. rewrite H. eapply alloc_result; eauto. - subst b2. split. auto. split. - rewrite (nextblock_alloc _ _ _ _ _ H0). - rewrite (nextblock_alloc _ _ _ _ _ H1). - congruence. - eapply alloc_parallel_inj; eauto. - unfold val_inj_lessdef; auto. - unfold inject_id; eauto. - omega. omega. - red; intros. apply Zdivide_0. -Qed. - -Lemma free_lessdef: - forall m1 m2 b, lessdef m1 m2 -> lessdef (free m1 b) (free m2 b). -Proof. - intros. destruct H. split. - simpl; auto. - eapply free_parallel_inj; eauto. - unfold inject_id. eauto. - unfold inject_id; intros. congruence. -Qed. - -Lemma free_left_lessdef: - forall m1 m2 b, - lessdef m1 m2 -> lessdef (free m1 b) m2. -Proof. - intros. destruct H. split. - rewrite <- H. auto. - apply free_left_inj; auto. -Qed. - -Lemma free_right_lessdef: - forall m1 m2 b, - lessdef m1 m2 -> low_bound m1 b >= high_bound m1 b -> - lessdef m1 (free m2 b). -Proof. - intros. destruct H. unfold lessdef. split. - rewrite H. auto. - apply free_right_inj; auto. intros. unfold inject_id in H2. inv H2. - red; intro. inv H2. generalize (size_chunk_pos chunk); intro. omega. -Qed. - -Lemma valid_block_lessdef: - forall m1 m2 b, lessdef m1 m2 -> valid_block m1 b -> valid_block m2 b. -Proof. - unfold valid_block. intros. destruct H. rewrite <- H; auto. -Qed. - -Lemma valid_pointer_lessdef: - forall m1 m2 b ofs, - lessdef m1 m2 -> valid_pointer m1 b ofs = true -> valid_pointer m2 b ofs = true. -Proof. - intros. destruct H. - replace ofs with (ofs + 0) by omega. - apply valid_pointer_inj with val_inj_lessdef inject_id m1 b; auto. -Qed. - -(** ** Memory injections *) - -(** A memory injection [f] is a function from addresses to either [None] - or [Some] of an address and an offset. It defines a correspondence - between the blocks of two memory states [m1] and [m2]: -- if [f b = None], the block [b] of [m1] has no equivalent in [m2]; -- if [f b = Some(b', ofs)], the block [b] of [m2] corresponds to - a sub-block at offset [ofs] of the block [b'] in [m2]. -*) - -(** A memory injection defines a relation between values that is the - identity relation, except for pointer values which are shifted - as prescribed by the memory injection. *) - -Inductive val_inject (mi: meminj): val -> val -> Prop := - | val_inject_int: - forall i, val_inject mi (Vint i) (Vint i) - | val_inject_float: - forall f, val_inject mi (Vfloat f) (Vfloat f) - | val_inject_ptr: - forall b1 ofs1 b2 ofs2 x, - mi b1 = Some (b2, x) -> - ofs2 = Int.add ofs1 (Int.repr x) -> - val_inject mi (Vptr b1 ofs1) (Vptr b2 ofs2) - | val_inject_undef: forall v, - val_inject mi Vundef v. - -Hint Resolve val_inject_int val_inject_float val_inject_ptr - val_inject_undef. - -Inductive val_list_inject (mi: meminj): list val -> list val-> Prop:= - | val_nil_inject : - val_list_inject mi nil nil - | val_cons_inject : forall v v' vl vl' , - val_inject mi v v' -> val_list_inject mi vl vl'-> - val_list_inject mi (v :: vl) (v' :: vl'). - -Hint Resolve val_nil_inject val_cons_inject. - -(** A memory state [m1] injects into another memory state [m2] via the - memory injection [f] if the following conditions hold: -- loads in [m1] must have matching loads in [m2] in the sense - of the [mem_inj] predicate; -- unallocated blocks in [m1] must be mapped to [None] by [f]; -- if [f b = Some(b', delta)], [b'] must be valid in [m2]; -- distinct blocks in [m1] are mapped to non-overlapping sub-blocks in [m2]; -- the sizes of [m2]'s blocks are representable with signed machine integers; -- the offsets [delta] are representable with signed machine integers. -*) - -Record mem_inject (f: meminj) (m1 m2: mem) : Prop := - mk_mem_inject { - mi_inj: - mem_inj val_inject f m1 m2; - mi_freeblocks: - forall b, ~(valid_block m1 b) -> f b = None; - mi_mappedblocks: - forall b b' delta, f b = Some(b', delta) -> valid_block m2 b'; - mi_no_overlap: - meminj_no_overlap f m1; - mi_range_1: - forall b b' delta, - f b = Some(b', delta) -> - Int.min_signed <= delta <= Int.max_signed; - mi_range_2: - forall b b' delta, - f b = Some(b', delta) -> - delta = 0 \/ (Int.min_signed <= low_bound m2 b' /\ high_bound m2 b' <= Int.max_signed) - }. - - -(** The following lemmas establish the absence of machine integer overflow - during address computations. *) - -Lemma address_inject: - forall f m1 m2 chunk b1 ofs1 b2 delta, - mem_inject f m1 m2 -> - valid_access m1 chunk b1 (Int.signed ofs1) -> - f b1 = Some (b2, delta) -> - Int.signed (Int.add ofs1 (Int.repr delta)) = Int.signed ofs1 + delta. -Proof. - intros. inversion H. - elim (mi_range_4 _ _ _ H1); intro. - (* delta = 0 *) - subst delta. change (Int.repr 0) with Int.zero. - rewrite Int.add_zero. omega. - (* delta <> 0 *) - rewrite Int.add_signed. - repeat rewrite Int.signed_repr. auto. - eauto. - assert (valid_access m2 chunk b2 (Int.signed ofs1 + delta)). - eapply valid_access_inj; eauto. - inv H3. generalize (size_chunk_pos chunk); omega. - eauto. -Qed. - -Lemma valid_pointer_inject_no_overflow: - forall f m1 m2 b ofs b' x, - mem_inject f m1 m2 -> - valid_pointer m1 b (Int.signed ofs) = true -> - f b = Some(b', x) -> - Int.min_signed <= Int.signed ofs + Int.signed (Int.repr x) <= Int.max_signed. -Proof. - intros. inv H. rewrite valid_pointer_valid_access in H0. - assert (valid_access m2 Mint8unsigned b' (Int.signed ofs + x)). - eapply valid_access_inj; eauto. - inv H. change (size_chunk Mint8unsigned) with 1 in H4. - rewrite Int.signed_repr; eauto. - exploit mi_range_4; eauto. intros [A | [A B]]. - subst x. rewrite Zplus_0_r. apply Int.signed_range. - omega. -Qed. - -Lemma valid_pointer_inject: - forall f m1 m2 b ofs b' ofs', - mem_inject f m1 m2 -> - valid_pointer m1 b (Int.signed ofs) = true -> - val_inject f (Vptr b ofs) (Vptr b' ofs') -> - valid_pointer m2 b' (Int.signed ofs') = true. -Proof. - intros. inv H1. - exploit valid_pointer_inject_no_overflow; eauto. intro NOOV. - inv H. rewrite Int.add_signed. rewrite Int.signed_repr; auto. - rewrite Int.signed_repr; eauto. - eapply valid_pointer_inj; eauto. -Qed. - -Lemma different_pointers_inject: - forall f m m' b1 ofs1 b2 ofs2 b1' delta1 b2' delta2, - mem_inject f m m' -> - b1 <> b2 -> - valid_pointer m b1 (Int.signed ofs1) = true -> - valid_pointer m b2 (Int.signed ofs2) = true -> - f b1 = Some (b1', delta1) -> - f b2 = Some (b2', delta2) -> - b1' <> b2' \/ - Int.signed (Int.add ofs1 (Int.repr delta1)) <> - Int.signed (Int.add ofs2 (Int.repr delta2)). -Proof. - intros. - rewrite valid_pointer_valid_access in H1. - rewrite valid_pointer_valid_access in H2. - rewrite (address_inject _ _ _ _ _ _ _ _ H H1 H3). - rewrite (address_inject _ _ _ _ _ _ _ _ H H2 H4). - inv H1. simpl in H7. inv H2. simpl in H10. - exploit (mi_no_overlap _ _ _ H); eauto. - intros [A | [A | [A | [A | A]]]]. - auto. omegaContradiction. omegaContradiction. - right. omega. right. omega. -Qed. - -(** Relation between injections and loads. *) - -Lemma load_inject: - forall f m1 m2 chunk b1 ofs b2 delta v1, - mem_inject f m1 m2 -> - load chunk m1 b1 ofs = Some v1 -> - f b1 = Some (b2, delta) -> - exists v2, load chunk m2 b2 (ofs + delta) = Some v2 /\ val_inject f v1 v2. -Proof. - intros. inversion H. - eapply mi_inj0; eauto. -Qed. - -Lemma loadv_inject: - forall f m1 m2 chunk a1 a2 v1, - mem_inject f m1 m2 -> - loadv chunk m1 a1 = Some v1 -> - val_inject f a1 a2 -> - exists v2, loadv chunk m2 a2 = Some v2 /\ val_inject f v1 v2. -Proof. - intros. inv H1; simpl in H0; try discriminate. - exploit load_inject; eauto. intros [v2 [LOAD INJ]]. - exists v2; split; auto. simpl. - replace (Int.signed (Int.add ofs1 (Int.repr x))) - with (Int.signed ofs1 + x). - auto. symmetry. eapply address_inject; eauto with mem. -Qed. - -(** Relation between injections and stores. *) - -Inductive val_content_inject (f: meminj): memory_chunk -> val -> val -> Prop := - | val_content_inject_base: - forall chunk v1 v2, - val_inject f v1 v2 -> - val_content_inject f chunk v1 v2 - | val_content_inject_8: - forall chunk n1 n2, - chunk = Mint8unsigned \/ chunk = Mint8signed -> - Int.zero_ext 8 n1 = Int.zero_ext 8 n2 -> - val_content_inject f chunk (Vint n1) (Vint n2) - | val_content_inject_16: - forall chunk n1 n2, - chunk = Mint16unsigned \/ chunk = Mint16signed -> - Int.zero_ext 16 n1 = Int.zero_ext 16 n2 -> - val_content_inject f chunk (Vint n1) (Vint n2) - | val_content_inject_32: - forall f1 f2, - Float.singleoffloat f1 = Float.singleoffloat f2 -> - val_content_inject f Mfloat32 (Vfloat f1) (Vfloat f2). - -Hint Resolve val_content_inject_base. - -Lemma load_result_inject: - forall f chunk v1 v2 chunk', - val_content_inject f chunk v1 v2 -> - size_chunk chunk = size_chunk chunk' -> - val_inject f (Val.load_result chunk' v1) (Val.load_result chunk' v2). -Proof. - intros. inv H; simpl. - inv H1; destruct chunk'; simpl; econstructor; eauto. - - elim H1; intro; subst chunk; - destruct chunk'; simpl in H0; try discriminate; simpl. - replace (Int.sign_ext 8 n1) with (Int.sign_ext 8 n2). - constructor. apply Int.sign_ext_equal_if_zero_equal; auto. compute; auto. - rewrite H2. constructor. - replace (Int.sign_ext 8 n1) with (Int.sign_ext 8 n2). - constructor. apply Int.sign_ext_equal_if_zero_equal; auto. compute; auto. - rewrite H2. constructor. - - elim H1; intro; subst chunk; - destruct chunk'; simpl in H0; try discriminate; simpl. - replace (Int.sign_ext 16 n1) with (Int.sign_ext 16 n2). - constructor. apply Int.sign_ext_equal_if_zero_equal; auto. compute; auto. - rewrite H2. constructor. - replace (Int.sign_ext 16 n1) with (Int.sign_ext 16 n2). - constructor. apply Int.sign_ext_equal_if_zero_equal; auto. compute; auto. - rewrite H2. constructor. - - destruct chunk'; simpl in H0; try discriminate; simpl. - constructor. rewrite H1; constructor. -Qed. - -Lemma store_mapped_inject_1 : - forall f chunk m1 b1 ofs v1 n1 m2 b2 delta v2, - mem_inject f m1 m2 -> - store chunk m1 b1 ofs v1 = Some n1 -> - f b1 = Some (b2, delta) -> - val_content_inject f chunk v1 v2 -> - exists n2, - store chunk m2 b2 (ofs + delta) v2 = Some n2 - /\ mem_inject f n1 n2. -Proof. - intros. inversion H. - exploit store_mapped_inj; eauto. - intros; constructor. - intros. apply load_result_inject with chunk; eauto. - intros [n2 [STORE MINJ]]. - exists n2; split. auto. constructor. - (* inj *) - auto. - (* freeblocks *) - intros. apply mi_freeblocks0. red; intro. elim H3. eauto with mem. - (* mappedblocks *) - intros. eauto with mem. - (* no_overlap *) - red; intros. - repeat rewrite (low_bound_store _ _ _ _ _ _ H0). - repeat rewrite (high_bound_store _ _ _ _ _ _ H0). - eapply mi_no_overlap0; eauto. - (* range *) - auto. - intros. - repeat rewrite (low_bound_store _ _ _ _ _ _ STORE). - repeat rewrite (high_bound_store _ _ _ _ _ _ STORE). - eapply mi_range_4; eauto. -Qed. - -Lemma store_mapped_inject: - forall f chunk m1 b1 ofs v1 n1 m2 b2 delta v2, - mem_inject f m1 m2 -> - store chunk m1 b1 ofs v1 = Some n1 -> - f b1 = Some (b2, delta) -> - val_inject f v1 v2 -> - exists n2, - store chunk m2 b2 (ofs + delta) v2 = Some n2 - /\ mem_inject f n1 n2. -Proof. - intros. eapply store_mapped_inject_1; eauto. -Qed. - -Lemma store_unmapped_inject: - forall f chunk m1 b1 ofs v1 n1 m2, - mem_inject f m1 m2 -> - store chunk m1 b1 ofs v1 = Some n1 -> - f b1 = None -> - mem_inject f n1 m2. -Proof. - intros. inversion H. - constructor. - (* inj *) - eapply store_unmapped_inj; eauto. - (* freeblocks *) - intros. apply mi_freeblocks0. red; intros; elim H2; eauto with mem. - (* mappedblocks *) - intros. eapply mi_mappedblocks0; eauto with mem. - (* no_overlap *) - red; intros. - repeat rewrite (low_bound_store _ _ _ _ _ _ H0). - repeat rewrite (high_bound_store _ _ _ _ _ _ H0). - eapply mi_no_overlap0; eauto. - (* range *) - auto. auto. -Qed. - -Lemma storev_mapped_inject_1: - forall f chunk m1 a1 v1 n1 m2 a2 v2, - mem_inject f m1 m2 -> - storev chunk m1 a1 v1 = Some n1 -> - val_inject f a1 a2 -> - val_content_inject f chunk v1 v2 -> - exists n2, - storev chunk m2 a2 v2 = Some n2 /\ mem_inject f n1 n2. -Proof. - intros. inv H1; simpl in H0; try discriminate. - simpl. replace (Int.signed (Int.add ofs1 (Int.repr x))) - with (Int.signed ofs1 + x). - eapply store_mapped_inject_1; eauto. - symmetry. eapply address_inject; eauto with mem. -Qed. - -Lemma storev_mapped_inject: - forall f chunk m1 a1 v1 n1 m2 a2 v2, - mem_inject f m1 m2 -> - storev chunk m1 a1 v1 = Some n1 -> - val_inject f a1 a2 -> - val_inject f v1 v2 -> - exists n2, - storev chunk m2 a2 v2 = Some n2 /\ mem_inject f n1 n2. -Proof. - intros. eapply storev_mapped_inject_1; eauto. -Qed. - -(** Relation between injections and [free] *) - -Lemma meminj_no_overlap_free: - forall mi m b, - meminj_no_overlap mi m -> - meminj_no_overlap mi (free m b). -Proof. - intros; red; intros. - assert (low_bound (free m b) b >= high_bound (free m b) b). - rewrite low_bound_free_same; rewrite high_bound_free_same; auto. - omega. - destruct (eq_block b1 b); destruct (eq_block b2 b); subst; auto. - repeat (rewrite low_bound_free; auto). - repeat (rewrite high_bound_free; auto). -Qed. - -Lemma meminj_no_overlap_free_list: - forall mi m bl, - meminj_no_overlap mi m -> - meminj_no_overlap mi (free_list m bl). -Proof. - induction bl; simpl; intros. auto. - apply meminj_no_overlap_free. auto. -Qed. - -Lemma free_inject: - forall f m1 m2 l b, - (forall b1 delta, f b1 = Some(b, delta) -> In b1 l) -> - mem_inject f m1 m2 -> - mem_inject f (free_list m1 l) (free m2 b). -Proof. - intros. inversion H0. constructor. - (* inj *) - apply free_right_inj. apply free_list_left_inj. auto. - intros; red; intros. - elim (valid_access_free_list_inv _ _ _ _ _ H2); intros. - elim H3; eauto. - (* freeblocks *) - intros. apply mi_freeblocks0. red; intro; elim H1. - apply valid_block_free_list_1; auto. - (* mappedblocks *) - intros. apply valid_block_free_1. eauto. - (* overlap *) - apply meminj_no_overlap_free_list; auto. - (* range *) - auto. - intros. destruct (eq_block b' b). subst b'. - rewrite low_bound_free_same; rewrite high_bound_free_same. - right; compute; intuition congruence. - rewrite low_bound_free; auto. rewrite high_bound_free; auto. - eauto. -Qed. - -(** Monotonicity properties of memory injections. *) - -Definition inject_incr (f1 f2: meminj) : Prop := - forall b, f1 b = f2 b \/ f1 b = None. - -Lemma inject_incr_refl : - forall f , inject_incr f f . -Proof. unfold inject_incr . intros. left . auto . Qed. - -Lemma inject_incr_trans : - forall f1 f2 f3, - inject_incr f1 f2 -> inject_incr f2 f3 -> inject_incr f1 f3 . -Proof . - unfold inject_incr; intros. - generalize (H b); generalize (H0 b); intuition congruence. -Qed. - -Lemma val_inject_incr: - forall f1 f2 v v', - inject_incr f1 f2 -> - val_inject f1 v v' -> - val_inject f2 v v'. -Proof. - intros. inversion H0. - constructor. - constructor. - elim (H b1); intro. - apply val_inject_ptr with x. congruence. auto. - congruence. - constructor. -Qed. - -Lemma val_list_inject_incr: - forall f1 f2 vl vl' , - inject_incr f1 f2 -> val_list_inject f1 vl vl' -> - val_list_inject f2 vl vl'. -Proof. - induction vl; intros; inv H0. auto. - constructor. eapply val_inject_incr; eauto. auto. -Qed. - -Hint Resolve inject_incr_refl val_inject_incr val_list_inject_incr. - -(** Properties of injections and allocations. *) - -Definition extend_inject - (b: block) (x: option (block * Z)) (f: meminj) : meminj := - fun (b': block) => if zeq b' b then x else f b'. - -Lemma extend_inject_incr: - forall f b x, - f b = None -> - inject_incr f (extend_inject b x f). -Proof. - intros; red; intros. unfold extend_inject. - destruct (zeq b0 b). subst b0; auto. auto. -Qed. - -Lemma alloc_right_inject: - forall f m1 m2 lo hi m2' b, - mem_inject f m1 m2 -> - alloc m2 lo hi = (m2', b) -> - mem_inject f m1 m2'. -Proof. - intros. inversion H. constructor. - eapply alloc_right_inj; eauto. - auto. - intros. eauto with mem. - auto. - auto. - intros. replace (low_bound m2' b') with (low_bound m2 b'). - replace (high_bound m2' b') with (high_bound m2 b'). - eauto. - symmetry. eapply high_bound_alloc_other; eauto. - symmetry. eapply low_bound_alloc_other; eauto. -Qed. - -Lemma alloc_unmapped_inject: - forall f m1 m2 lo hi m1' b, - mem_inject f m1 m2 -> - alloc m1 lo hi = (m1', b) -> - mem_inject (extend_inject b None f) m1' m2 /\ - inject_incr f (extend_inject b None f). -Proof. - intros. inversion H. - assert (inject_incr f (extend_inject b None f)). - apply extend_inject_incr. apply mi_freeblocks0. eauto with mem. - split; auto. constructor. - (* inj *) - eapply alloc_left_unmapped_inj; eauto. - red; intros. unfold extend_inject in H2. - destruct (zeq b1 b). congruence. - exploit mi_inj0; eauto. intros [v2 [LOAD VINJ]]. - exists v2; split. auto. - apply val_inject_incr with f; auto. - unfold extend_inject. apply zeq_true. - (* freeblocks *) - intros. unfold extend_inject. destruct (zeq b0 b). auto. - apply mi_freeblocks0; red; intro. elim H2. eauto with mem. - (* mappedblocks *) - intros. unfold extend_inject in H2. destruct (zeq b0 b). - discriminate. eauto. - (* overlap *) - red; unfold extend_inject, update; intros. - repeat rewrite (low_bound_alloc _ _ _ _ _ H0). - repeat rewrite (high_bound_alloc _ _ _ _ _ H0). - destruct (zeq b1 b); try discriminate. - destruct (zeq b2 b); try discriminate. - eauto. - (* range *) - unfold extend_inject; intros. - destruct (zeq b0 b). discriminate. eauto. - unfold extend_inject; intros. - destruct (zeq b0 b). discriminate. eauto. -Qed. - -Lemma alloc_mapped_inject: - forall f m1 m2 lo hi m1' b b' ofs, - mem_inject f m1 m2 -> - alloc m1 lo hi = (m1', b) -> - valid_block m2 b' -> - Int.min_signed <= ofs <= Int.max_signed -> - Int.min_signed <= low_bound m2 b' -> - high_bound m2 b' <= Int.max_signed -> - low_bound m2 b' <= lo + ofs -> - hi + ofs <= high_bound m2 b' -> - inj_offset_aligned ofs (hi-lo) -> - (forall b0 ofs0, - f b0 = Some (b', ofs0) -> - high_bound m1 b0 + ofs0 <= lo + ofs \/ - hi + ofs <= low_bound m1 b0 + ofs0) -> - mem_inject (extend_inject b (Some (b', ofs)) f) m1' m2 /\ - inject_incr f (extend_inject b (Some (b', ofs)) f). -Proof. - intros. inversion H. - assert (inject_incr f (extend_inject b (Some (b', ofs)) f)). - apply extend_inject_incr. apply mi_freeblocks0. eauto with mem. - split; auto. - constructor. - (* inj *) - eapply alloc_left_mapped_inj; eauto. - red; intros. unfold extend_inject in H10. - rewrite zeq_false in H10. - exploit mi_inj0; eauto. intros [v2 [LOAD VINJ]]. - exists v2; split. auto. eapply val_inject_incr; eauto. - eauto with mem. - unfold extend_inject. apply zeq_true. - (* freeblocks *) - intros. unfold extend_inject. rewrite zeq_false. - apply mi_freeblocks0. red; intro. elim H10; eauto with mem. - apply sym_not_equal; eauto with mem. - (* mappedblocks *) - unfold extend_inject; intros. - destruct (zeq b0 b). inv H10. auto. eauto. - (* overlap *) - red; unfold extend_inject, update; intros. - repeat rewrite (low_bound_alloc _ _ _ _ _ H0). - repeat rewrite (high_bound_alloc _ _ _ _ _ H0). - destruct (zeq b1 b); [inv H11|idtac]; - (destruct (zeq b2 b); [inv H12|idtac]). - congruence. - destruct (zeq b1' b2'). subst b2'. generalize (H8 _ _ H12). tauto. auto. - destruct (zeq b1' b2'). subst b2'. generalize (H8 _ _ H11). tauto. auto. - eauto. - (* range *) - unfold extend_inject; intros. - destruct (zeq b0 b). inv H10. auto. eauto. - unfold extend_inject; intros. - destruct (zeq b0 b). inv H10. auto. eauto. -Qed. - -Lemma alloc_parallel_inject: - forall f m1 m2 lo hi m1' m2' b1 b2, - mem_inject f m1 m2 -> - alloc m1 lo hi = (m1', b1) -> - alloc m2 lo hi = (m2', b2) -> - Int.min_signed <= lo -> hi <= Int.max_signed -> - mem_inject (extend_inject b1 (Some(b2, 0)) f) m1' m2' /\ - inject_incr f (extend_inject b1 (Some(b2, 0)) f). -Proof. - intros. - eapply alloc_mapped_inject; eauto. - eapply alloc_right_inject; eauto. - eauto with mem. - compute; intuition congruence. - rewrite (low_bound_alloc_same _ _ _ _ _ H1). auto. - rewrite (high_bound_alloc_same _ _ _ _ _ H1). auto. - rewrite (low_bound_alloc_same _ _ _ _ _ H1). omega. - rewrite (high_bound_alloc_same _ _ _ _ _ H1). omega. - red; intros. apply Zdivide_0. - intros. elimtype False. inv H. - exploit mi_mappedblocks0; eauto. - change (~ valid_block m2 b2). eauto with mem. -Qed. - -Definition meminj_init (m: mem) : meminj := - fun (b: block) => if zlt b m.(nextblock) then Some(b, 0) else None. - -Definition mem_inject_neutral (m: mem) : Prop := - forall f chunk b ofs v, - load chunk m b ofs = Some v -> val_inject f v v. - -Lemma init_inject: - forall m, - mem_inject_neutral m -> - mem_inject (meminj_init m) m m. -Proof. - intros; constructor. - (* inj *) - red; intros. unfold meminj_init in H0. - destruct (zlt b1 (nextblock m)); inversion H0. - subst b2 delta. exists v1; split. - rewrite Zplus_0_r. auto. eapply H; eauto. - (* free blocks *) - unfold valid_block, meminj_init; intros. - apply zlt_false. omega. - (* mapped blocks *) - unfold valid_block, meminj_init; intros. - destruct (zlt b (nextblock m)); inversion H0. subst b'; auto. - (* overlap *) - red; unfold meminj_init; intros. - destruct (zlt b1 (nextblock m)); inversion H1. - destruct (zlt b2 (nextblock m)); inversion H2. - left; congruence. - (* range *) - unfold meminj_init; intros. - destruct (zlt b (nextblock m)); inversion H0. subst delta. - compute; intuition congruence. - unfold meminj_init; intros. - destruct (zlt b (nextblock m)); inversion H0. subst delta. - auto. -Qed. - -Remark getN_setN_inject: - forall f m v n1 p1 n2 p2, - val_inject f (getN n2 p2 m) (getN n2 p2 m) -> - val_inject f v v -> - val_inject f (getN n2 p2 (setN n1 p1 v m)) - (getN n2 p2 (setN n1 p1 v m)). -Proof. - intros. - destruct (getN_setN_characterization m v n1 p1 n2 p2) - as [A | [A | A]]; rewrite A; auto. -Qed. - -Remark getN_contents_init_data_inject: - forall f n ofs id pos, - val_inject f (getN n ofs (contents_init_data pos id)) - (getN n ofs (contents_init_data pos id)). -Proof. - induction id; simpl; intros. - repeat rewrite getN_init. constructor. - destruct a; auto; apply getN_setN_inject; auto. -Qed. - -Lemma alloc_init_data_neutral: - forall m id m' b, - mem_inject_neutral m -> - alloc_init_data m id = (m', b) -> - mem_inject_neutral m'. -Proof. - intros. injection H0; intros A B. - red; intros. - exploit load_inv; eauto. intros [C D]. - rewrite <- B in D; simpl in D. rewrite A in D. - unfold update in D. destruct (zeq b0 b). - subst b0. rewrite D. simpl. - apply load_result_inject with chunk. constructor. - apply getN_contents_init_data_inject. auto. - apply H with chunk b0 ofs. unfold load. - rewrite in_bounds_true. congruence. - inversion C. constructor. - generalize H2. unfold valid_block. rewrite <- B; simpl. - rewrite A. unfold block in n; intros. omega. - replace (low_bound m b0) with (low_bound m' b0). auto. - unfold low_bound; rewrite <- B; simpl; rewrite A. rewrite update_o; auto. - replace (high_bound m b0) with (high_bound m' b0). auto. - unfold high_bound; rewrite <- B; simpl; rewrite A. rewrite update_o; auto. - auto. -Qed. - -(** ** Memory shifting *) - -(** A special case of memory injection where blocks are not coalesced: - each source block injects in a distinct target block. *) - -Definition memshift := block -> option Z. - -Definition meminj_of_shift (mi: memshift) : meminj := - fun b => match mi b with None => None | Some x => Some (b, x) end. - -Definition val_shift (mi: memshift) (v1 v2: val): Prop := - val_inject (meminj_of_shift mi) v1 v2. - -Record mem_shift (f: memshift) (m1 m2: mem) : Prop := - mk_mem_shift { - ms_inj: - mem_inj val_inject (meminj_of_shift f) m1 m2; - ms_samedomain: - nextblock m1 = nextblock m2; - ms_domain: - forall b, match f b with Some _ => b < nextblock m1 | None => b >= nextblock m1 end; - ms_range_1: - forall b delta, - f b = Some delta -> - Int.min_signed <= delta <= Int.max_signed; - ms_range_2: - forall b delta, - f b = Some delta -> - Int.min_signed <= low_bound m2 b /\ high_bound m2 b <= Int.max_signed - }. - -(** The following lemmas establish the absence of machine integer overflow - during address computations. *) - -Lemma address_shift: - forall f m1 m2 chunk b ofs1 delta, - mem_shift f m1 m2 -> - valid_access m1 chunk b (Int.signed ofs1) -> - f b = Some delta -> - Int.signed (Int.add ofs1 (Int.repr delta)) = Int.signed ofs1 + delta. -Proof. - intros. inversion H. - elim (ms_range_4 _ _ H1); intros. - rewrite Int.add_signed. - repeat rewrite Int.signed_repr. auto. - eauto. - assert (valid_access m2 chunk b (Int.signed ofs1 + delta)). - eapply valid_access_inj with (mi := meminj_of_shift f); eauto. - unfold meminj_of_shift. rewrite H1; auto. - inv H4. generalize (size_chunk_pos chunk); omega. - eauto. -Qed. - -Lemma valid_pointer_shift_no_overflow: - forall f m1 m2 b ofs x, - mem_shift f m1 m2 -> - valid_pointer m1 b (Int.signed ofs) = true -> - f b = Some x -> - Int.min_signed <= Int.signed ofs + Int.signed (Int.repr x) <= Int.max_signed. -Proof. - intros. inv H. rewrite valid_pointer_valid_access in H0. - assert (valid_access m2 Mint8unsigned b (Int.signed ofs + x)). - eapply valid_access_inj with (mi := meminj_of_shift f); eauto. - unfold meminj_of_shift. rewrite H1; auto. - inv H. change (size_chunk Mint8unsigned) with 1 in H4. - rewrite Int.signed_repr; eauto. - exploit ms_range_4; eauto. intros [A B]. omega. -Qed. - -Lemma valid_pointer_shift: - forall f m1 m2 b ofs b' ofs', - mem_shift f m1 m2 -> - valid_pointer m1 b (Int.signed ofs) = true -> - val_shift f (Vptr b ofs) (Vptr b' ofs') -> - valid_pointer m2 b' (Int.signed ofs') = true. -Proof. - intros. unfold val_shift in H1. inv H1. - assert (f b = Some x). - unfold meminj_of_shift in H5. destruct (f b); congruence. - exploit valid_pointer_shift_no_overflow; eauto. intro NOOV. - inv H. rewrite Int.add_signed. rewrite Int.signed_repr; auto. - rewrite Int.signed_repr; eauto. - eapply valid_pointer_inj; eauto. -Qed. - -(** Relation between shifts and loads. *) - -Lemma load_shift: - forall f m1 m2 chunk b ofs delta v1, - mem_shift f m1 m2 -> - load chunk m1 b ofs = Some v1 -> - f b = Some delta -> - exists v2, load chunk m2 b (ofs + delta) = Some v2 /\ val_shift f v1 v2. -Proof. - intros. inversion H. - unfold val_shift. eapply ms_inj0; eauto. - unfold meminj_of_shift; rewrite H1; auto. -Qed. - -Lemma loadv_shift: - forall f m1 m2 chunk a1 a2 v1, - mem_shift f m1 m2 -> - loadv chunk m1 a1 = Some v1 -> - val_shift f a1 a2 -> - exists v2, loadv chunk m2 a2 = Some v2 /\ val_shift f v1 v2. -Proof. - intros. unfold val_shift in H1. inv H1; simpl in H0; try discriminate. - generalize H2. unfold meminj_of_shift. caseEq (f b1); intros; inv H3. - exploit load_shift; eauto. intros [v2 [LOAD INJ]]. - exists v2; split; auto. simpl. - replace (Int.signed (Int.add ofs1 (Int.repr x))) - with (Int.signed ofs1 + x). - auto. symmetry. eapply address_shift; eauto with mem. -Qed. - -(** Relation between shifts and stores. *) - -Lemma store_within_shift: - forall f chunk m1 b ofs v1 n1 m2 delta v2, - mem_shift f m1 m2 -> - store chunk m1 b ofs v1 = Some n1 -> - f b = Some delta -> - val_shift f v1 v2 -> - exists n2, - store chunk m2 b (ofs + delta) v2 = Some n2 - /\ mem_shift f n1 n2. -Proof. - intros. inversion H. - exploit store_mapped_inj; eauto. - intros; constructor. - red. intros until delta2. unfold meminj_of_shift. - destruct (f b1). destruct (f b2). intros. inv H4. inv H5. auto. - congruence. congruence. - unfold meminj_of_shift. rewrite H1. auto. - intros. apply load_result_inject with chunk; eauto. - unfold val_shift in H2. eauto. - intros [n2 [STORE MINJ]]. - exists n2; split. auto. constructor. - (* inj *) - auto. - (* samedomain *) - rewrite (nextblock_store _ _ _ _ _ _ H0). - rewrite (nextblock_store _ _ _ _ _ _ STORE). - auto. - (* domain *) - rewrite (nextblock_store _ _ _ _ _ _ H0). auto. - (* range *) - auto. - intros. - repeat rewrite (low_bound_store _ _ _ _ _ _ STORE). - repeat rewrite (high_bound_store _ _ _ _ _ _ STORE). - eapply ms_range_4; eauto. -Qed. - -Lemma store_outside_shift: - forall f chunk m1 b ofs m2 v m2' delta, - mem_shift f m1 m2 -> - f b = Some delta -> - high_bound m1 b + delta <= ofs - \/ ofs + size_chunk chunk <= low_bound m1 b + delta -> - store chunk m2 b ofs v = Some m2' -> - mem_shift f m1 m2'. -Proof. - intros. inversion H. constructor. - (* inj *) - eapply store_outside_inj; eauto. - unfold meminj_of_shift. intros b' d'. caseEq (f b'); intros; inv H4. - congruence. - (* samedomain *) - rewrite (nextblock_store _ _ _ _ _ _ H2). - auto. - (* domain *) - auto. - (* range *) - auto. - intros. - repeat rewrite (low_bound_store _ _ _ _ _ _ H2). - repeat rewrite (high_bound_store _ _ _ _ _ _ H2). - eapply ms_range_4; eauto. -Qed. - -Lemma storev_shift: - forall f chunk m1 a1 v1 n1 m2 a2 v2, - mem_shift f m1 m2 -> - storev chunk m1 a1 v1 = Some n1 -> - val_shift f a1 a2 -> - val_shift f v1 v2 -> - exists n2, - storev chunk m2 a2 v2 = Some n2 /\ mem_shift f n1 n2. -Proof. - intros. unfold val_shift in H1. inv H1; simpl in H0; try discriminate. - generalize H3. unfold meminj_of_shift. caseEq (f b1); intros; inv H4. - exploit store_within_shift; eauto. intros [n2 [A B]]. - exists n2; split; auto. - unfold storev. - replace (Int.signed (Int.add ofs1 (Int.repr x))) - with (Int.signed ofs1 + x). - auto. symmetry. eapply address_shift; eauto with mem. -Qed. - -(** Relation between shifts and [free]. *) - -Lemma free_shift: - forall f m1 m2 b, - mem_shift f m1 m2 -> - mem_shift f (free m1 b) (free m2 b). -Proof. - intros. inv H. constructor. - (* inj *) - apply free_right_inj. apply free_left_inj; auto. - intros until ofs. unfold meminj_of_shift. caseEq (f b1); intros; inv H0. - apply valid_access_free_2. - (* samedomain *) - simpl. auto. - (* domain *) - simpl. auto. - (* range *) - auto. - intros. destruct (eq_block b0 b). - subst b0. rewrite low_bound_free_same. rewrite high_bound_free_same. - vm_compute; intuition congruence. - rewrite low_bound_free; auto. rewrite high_bound_free; auto. eauto. -Qed. - -(** Relation between shifts and allocation. *) - -Definition shift_incr (f1 f2: memshift) : Prop := - forall b, f1 b = f2 b \/ f1 b = None. - -Remark shift_incr_inject_incr: - forall f1 f2, - shift_incr f1 f2 -> inject_incr (meminj_of_shift f1) (meminj_of_shift f2). -Proof. - intros. unfold meminj_of_shift. red. intros. - elim (H b); intro. rewrite H0. auto. rewrite H0. auto. -Qed. - -Lemma val_shift_incr: - forall f1 f2 v1 v2, - shift_incr f1 f2 -> val_shift f1 v1 v2 -> val_shift f2 v1 v2. -Proof. - unfold val_shift; intros. - apply val_inject_incr with (meminj_of_shift f1). - apply shift_incr_inject_incr. auto. auto. -Qed. - -(*** -Remark mem_inj_incr: - forall f1 f2 m1 m2, - inject_incr f1 f2 -> mem_inj val_inject f1 m1 m2 -> mem_inj val_inject f2 m1 m2. -Proof. - intros; red; intros. - destruct (H b1). rewrite <- H3 in H1. - exploit H0; eauto. intros [v2 [A B]]. - exists v2; split. auto. apply val_inject_incr with f1; auto. - congruence. -***) - -Lemma alloc_shift: - forall f m1 m2 lo1 hi1 m1' b delta lo2 hi2, - mem_shift f m1 m2 -> - alloc m1 lo1 hi1 = (m1', b) -> - lo2 <= lo1 + delta -> hi1 + delta <= hi2 -> - Int.min_signed <= delta <= Int.max_signed -> - Int.min_signed <= lo2 -> hi2 <= Int.max_signed -> - inj_offset_aligned delta (hi1-lo1) -> - exists f', exists m2', - alloc m2 lo2 hi2 = (m2', b) - /\ mem_shift f' m1' m2' - /\ shift_incr f f' - /\ f' b = Some delta. -Proof. - intros. inv H. caseEq (alloc m2 lo2 hi2). intros m2' b' ALLOC2. - assert (b' = b). - rewrite (alloc_result _ _ _ _ _ H0). - rewrite (alloc_result _ _ _ _ _ ALLOC2). - auto. - subst b'. - assert (f b = None). - generalize (ms_domain0 b). - rewrite (alloc_result _ _ _ _ _ H0). - destruct (f (nextblock m1)). - intros. omegaContradiction. - auto. - set (f' := fun (b': block) => if zeq b' b then Some delta else f b'). - assert (shift_incr f f'). - red; unfold f'; intros. - destruct (zeq b0 b); auto. - subst b0. auto. - exists f'; exists m2'. - split. auto. - (* mem_shift *) - split. constructor. - (* inj *) - assert (mem_inj val_inject (meminj_of_shift f') m1 m2). - red; intros. - assert (meminj_of_shift f b1 = Some (b2, delta0)). - rewrite <- H8. unfold meminj_of_shift, f'. - destruct (zeq b1 b); auto. - subst b1. - assert (valid_block m1 b) by eauto with mem. - assert (~valid_block m1 b) by eauto with mem. - contradiction. - exploit ms_inj0; eauto. intros [v2 [A B]]. - exists v2; split; auto. - apply val_inject_incr with (meminj_of_shift f). - apply shift_incr_inject_incr. auto. auto. - eapply alloc_parallel_inj; eauto. - unfold meminj_of_shift, f'. rewrite zeq_true. auto. - (* samedomain *) - rewrite (nextblock_alloc _ _ _ _ _ H0). - rewrite (nextblock_alloc _ _ _ _ _ ALLOC2). - congruence. - (* domain *) - intros. unfold f'. - rewrite (nextblock_alloc _ _ _ _ _ H0). - rewrite (alloc_result _ _ _ _ _ H0). - destruct (zeq b0 (nextblock m1)). omega. - generalize (ms_domain0 b0). destruct (f b0); omega. - (* range *) - unfold f'; intros. destruct (zeq b0 b). congruence. eauto. - unfold f'; intros. - rewrite (low_bound_alloc _ _ _ _ _ ALLOC2). - rewrite (high_bound_alloc _ _ _ _ _ ALLOC2). - destruct (zeq b0 b). auto. eauto. - (* shift_incr *) - split. auto. - (* f' b = delta *) - unfold f'. apply zeq_true. -Qed. - -(** ** Relation between signed and unsigned loads and stores *) - -(** Target processors do not distinguish between signed and unsigned - stores of 8- and 16-bit quantities. We show these are equivalent. *) - -(** Signed 8- and 16-bit stores can be performed like unsigned stores. *) - -Remark in_bounds_equiv: - forall chunk1 chunk2 m b ofs (A: Type) (a1 a2: A), - size_chunk chunk1 = size_chunk chunk2 -> - (if in_bounds m chunk1 b ofs then a1 else a2) = - (if in_bounds m chunk2 b ofs then a1 else a2). -Proof. - intros. destruct (in_bounds m chunk1 b ofs). - rewrite in_bounds_true. auto. eapply valid_access_compat; eauto. - destruct (in_bounds m chunk2 b ofs); auto. - elim n. eapply valid_access_compat with (chunk1 := chunk2); eauto. -Qed. - -Lemma storev_8_signed_unsigned: - forall m a v, - storev Mint8signed m a v = storev Mint8unsigned m a v. -Proof. - intros. unfold storev. destruct a; auto. - unfold store. rewrite (in_bounds_equiv Mint8signed Mint8unsigned). - auto. auto. -Qed. - -Lemma storev_16_signed_unsigned: - forall m a v, - storev Mint16signed m a v = storev Mint16unsigned m a v. -Proof. - intros. unfold storev. destruct a; auto. - unfold store. rewrite (in_bounds_equiv Mint16signed Mint16unsigned). - auto. auto. -Qed. - -(** Likewise, some target processors (e.g. the PowerPC) do not have - a ``load 8-bit signed integer'' instruction. - We show that it can be synthesized as a ``load 8-bit unsigned integer'' - followed by a sign extension. *) - -Lemma loadv_8_signed_unsigned: - forall m a, - loadv Mint8signed m a = option_map (Val.sign_ext 8) (loadv Mint8unsigned m a). -Proof. - intros. unfold Mem.loadv. destruct a; try reflexivity. - unfold load. rewrite (in_bounds_equiv Mint8signed Mint8unsigned). - destruct (in_bounds m Mint8unsigned b (Int.signed i)); auto. - simpl. - destruct (getN 0 (Int.signed i) (contents (blocks m b))); auto. - simpl. rewrite Int.sign_ext_zero_ext. auto. compute; auto. - auto. -Qed. - diff --git a/common/Memdata.v b/common/Memdata.v new file mode 100644 index 0000000..2c5fdb6 --- /dev/null +++ b/common/Memdata.v @@ -0,0 +1,1058 @@ +Require Import Coqlib. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. + +(** * Properties of memory chunks *) + +(** Memory reads and writes are performed by quantities called memory chunks, + encoding the type, size and signedness of the chunk being addressed. + The following functions extract the size information from a chunk. *) + +Definition size_chunk (chunk: memory_chunk) : Z := + match chunk with + | Mint8signed => 1 + | Mint8unsigned => 1 + | Mint16signed => 2 + | Mint16unsigned => 2 + | Mint32 => 4 + | Mfloat32 => 4 + | Mfloat64 => 8 + end. + +Lemma size_chunk_pos: + forall chunk, size_chunk chunk > 0. +Proof. + intros. destruct chunk; simpl; omega. +Qed. + +Definition size_chunk_nat (chunk: memory_chunk) : nat := + nat_of_Z(size_chunk chunk). + +Lemma size_chunk_conv: + forall chunk, size_chunk chunk = Z_of_nat (size_chunk_nat chunk). +Proof. + intros. destruct chunk; reflexivity. +Qed. + +Lemma size_chunk_nat_pos: + forall chunk, exists n, size_chunk_nat chunk = S n. +Proof. + intros. + generalize (size_chunk_pos chunk). rewrite size_chunk_conv. + destruct (size_chunk_nat chunk). + simpl; intros; omegaContradiction. + intros; exists n; auto. +Qed. + +(** Memory reads and writes must respect alignment constraints: + the byte offset of the location being addressed should be an exact + multiple of the natural alignment for the chunk being addressed. + This natural alignment is defined by the following + [align_chunk] function. Some target architectures + (e.g. the PowerPC) have no alignment constraints, which we could + reflect by taking [align_chunk chunk = 1]. However, other architectures + have stronger alignment requirements. The following definition is + appropriate for PowerPC and ARM. *) + +Definition align_chunk (chunk: memory_chunk) : Z := + match chunk with + | Mint8signed => 1 + | Mint8unsigned => 1 + | Mint16signed => 2 + | Mint16unsigned => 2 + | _ => 4 + end. + +Lemma align_chunk_pos: + forall chunk, align_chunk chunk > 0. +Proof. + intro. destruct chunk; simpl; omega. +Qed. + +Lemma align_size_chunk_divides: + forall chunk, (align_chunk chunk | size_chunk chunk). +Proof. + intros. destruct chunk; simpl; try apply Zdivide_refl. exists 2; auto. +Qed. + +Lemma align_chunk_compat: + forall chunk1 chunk2, + size_chunk chunk1 = size_chunk chunk2 -> align_chunk chunk1 = align_chunk chunk2. +Proof. + intros chunk1 chunk2. + destruct chunk1; destruct chunk2; simpl; congruence. +Qed. + +(** The type (integer/pointer or float) of a chunk. *) + +Definition type_of_chunk (c: memory_chunk) : typ := + match c with + | Mint8signed => Tint + | Mint8unsigned => Tint + | Mint16signed => Tint + | Mint16unsigned => Tint + | Mint32 => Tint + | Mfloat32 => Tfloat + | Mfloat64 => Tfloat + end. + +(** * Memory values *) + +(** A ``memory value'' is a byte-sized quantity that describes the current + content of a memory cell. It can be either: +- a concrete 8-bit integer; +- a byte-sized fragment of an opaque pointer; +- the special constant [Undef] that represents uninitialized memory. +*) + +(** Values stored in memory cells. *) + +Inductive memval: Type := + | Undef: memval + | Byte: byte -> memval + | Pointer: block -> int -> nat -> memval. + +(** * Encoding and decoding integers *) + +(** We define functions to convert between integers and lists of bytes + according to a given memory chunk. *) + +Parameter big_endian: bool. + +Definition rev_if_le (l: list byte) : list byte := + if big_endian then l else List.rev l. + +Lemma rev_if_le_involutive: + forall l, rev_if_le (rev_if_le l) = l. +Proof. + intros; unfold rev_if_le; destruct big_endian. + auto. + apply List.rev_involutive. +Qed. + +Lemma rev_if_le_length: + forall l, length (rev_if_le l) = length l. +Proof. + intros; unfold rev_if_le; destruct big_endian. + auto. + apply List.rev_length. +Qed. + +Definition encode_int (c: memory_chunk) (x: int) : list byte := + let n := Int.unsigned x in + rev_if_le (match c with + | Mint8signed | Mint8unsigned => + Byte.repr n :: nil + | Mint16signed | Mint16unsigned => + Byte.repr (n/256) :: Byte.repr n :: nil + | Mint32 => + Byte.repr (n/16777216) :: Byte.repr (n/65536) :: Byte.repr (n/256) :: Byte.repr n :: nil + | Mfloat32 => + Byte.zero :: Byte.zero :: Byte.zero :: Byte.zero :: nil + | Mfloat64 => + Byte.zero :: Byte.zero :: Byte.zero :: Byte.zero :: + Byte.zero :: Byte.zero :: Byte.zero :: Byte.zero :: nil + end). + +Definition decode_int (c: memory_chunk) (b: list byte) : int := + match c, rev_if_le b with + | Mint8signed, b1 :: nil => + Int.sign_ext 8 (Int.repr (Byte.unsigned b1)) + | Mint8unsigned, b1 :: nil => + Int.repr (Byte.unsigned b1) + | Mint16signed, b1 :: b2 :: nil => + Int.sign_ext 16 (Int.repr (Byte.unsigned b1 * 256 + Byte.unsigned b2)) + | Mint16unsigned, b1 :: b2 :: nil => + Int.repr (Byte.unsigned b1 * 256 + Byte.unsigned b2) + | Mint32, b1 :: b2 :: b3 :: b4 :: nil => + Int.repr (Byte.unsigned b1 * 16777216 + Byte.unsigned b2 * 65536 + + Byte.unsigned b3 * 256 + Byte.unsigned b4) + | _, _ => Int.zero + end. + +Lemma encode_int_length: + forall chunk n, length(encode_int chunk n) = size_chunk_nat chunk. +Proof. + intros. unfold encode_int. rewrite rev_if_le_length. + destruct chunk; reflexivity. +Qed. + +Lemma decode_encode_int8unsigned: forall n, + decode_int Mint8unsigned (encode_int Mint8unsigned n) = Int.zero_ext 8 n. +Proof. + intros. unfold decode_int, encode_int. rewrite rev_if_le_involutive. + simpl. auto. +Qed. + +Lemma decode_encode_int8signed: forall n, + decode_int Mint8signed (encode_int Mint8signed n) = Int.sign_ext 8 n. +Proof. + intros. unfold decode_int, encode_int. rewrite rev_if_le_involutive. simpl. + change (Int.repr (Int.unsigned n mod Byte.modulus)) + with (Int.zero_ext 8 n). + apply Int.sign_ext_zero_ext. compute; auto. +Qed. + +Remark recombine_16: + forall x, + (x / 256) mod Byte.modulus * 256 + x mod Byte.modulus = x mod (two_p 16). +Proof. + intros. symmetry. apply (Zmod_recombine x 256 256); omega. +Qed. + +Lemma decode_encode_int16unsigned: forall n, + decode_int Mint16unsigned (encode_int Mint16unsigned n) = Int.zero_ext 16 n. +Proof. + intros. unfold decode_int, encode_int. rewrite rev_if_le_involutive. simpl. + rewrite recombine_16. auto. +Qed. + +Lemma decode_encode_int16signed: forall n, + decode_int Mint16signed (encode_int Mint16signed n) = Int.sign_ext 16 n. +Proof. + intros. unfold decode_int, encode_int. rewrite rev_if_le_involutive. simpl. + rewrite recombine_16. + fold (Int.zero_ext 16 n). apply Int.sign_ext_zero_ext. compute; auto. +Qed. + +Remark recombine_32: + forall x, + (x / 16777216) mod Byte.modulus * 16777216 + + (x / 65536) mod Byte.modulus * 65536 + + (x / 256) mod Byte.modulus * 256 + + x mod Byte.modulus = + x mod Int.modulus. +Proof. + intros. change Byte.modulus with 256. + exploit (Zmod_recombine x 65536 65536). omega. omega. intro EQ1. + exploit (Zmod_recombine x 256 256). omega. omega. + change (256 * 256) with 65536. intro EQ2. + exploit (Zmod_recombine (x/65536) 256 256). omega. omega. + rewrite Zdiv_Zdiv. change (65536*256) with 16777216. change (256 * 256) with 65536. + intro EQ3. + change Int.modulus with (65536 * 65536). + rewrite EQ1. rewrite EQ2. rewrite EQ3. omega. + omega. omega. +Qed. + +Lemma decode_encode_int32: forall n, + decode_int Mint32 (encode_int Mint32 n) = n. +Proof. + intros. unfold decode_int, encode_int. rewrite rev_if_le_involutive. simpl. + rewrite recombine_32. + transitivity (Int.repr (Int.unsigned n)). 2: apply Int.repr_unsigned. + apply Int.eqm_samerepr. apply Int.eqm_sym. red. apply Int.eqmod_mod. + apply Int.modulus_pos. +Qed. + +Lemma encode_int8_signed_unsigned: forall n, + encode_int Mint8signed n = encode_int Mint8unsigned n. +Proof. + intros; reflexivity. +Qed. + +Remark encode_8_mod: + forall x y, + Int.eqmod (two_p 8) (Int.unsigned x) (Int.unsigned y) -> + encode_int Mint8unsigned x = encode_int Mint8unsigned y. +Proof. + intros. unfold encode_int. decEq. decEq. apply Byte.eqm_samerepr. exact H. +Qed. + +Lemma encode_int8_zero_ext: + forall x, + encode_int Mint8unsigned (Int.zero_ext 8 x) = encode_int Mint8unsigned x. +Proof. + intros. apply encode_8_mod. apply Int.eqmod_sym. + apply Int.eqmod_two_p_zero_ext. compute; auto. +Qed. + +Lemma encode_int8_sign_ext: + forall x, + encode_int Mint8signed (Int.sign_ext 8 x) = encode_int Mint8signed x. +Proof. + intros. repeat rewrite encode_int8_signed_unsigned. + apply encode_8_mod. apply Int.eqmod_sym. + apply Int.eqmod_two_p_sign_ext. compute; auto. +Qed. + +Lemma encode_int16_signed_unsigned: forall n, + encode_int Mint16signed n = encode_int Mint16unsigned n. +Proof. + intros; reflexivity. +Qed. + +Remark encode_16_mod: + forall x y, + Int.eqmod (two_p 16) (Int.unsigned x) (Int.unsigned y) -> + encode_int Mint16unsigned x = encode_int Mint16unsigned y. +Proof. + intros. unfold encode_int. decEq. + set (x' := Int.unsigned x) in *. + set (y' := Int.unsigned y) in *. + assert (Int.eqmod (two_p 8) x' y'). + eapply Int.eqmod_divides; eauto. exists (two_p 8); auto. + assert (Int.eqmod (two_p 8) (x' / 256) (y' / 256)). + destruct H as [k EQ]. + exists k. rewrite EQ. + replace (k * two_p 16) with ((k * two_p 8) * two_p 8). + rewrite Zplus_comm. rewrite Z_div_plus. omega. + omega. rewrite <- Zmult_assoc. auto. + decEq. apply Byte.eqm_samerepr. exact H1. + decEq. apply Byte.eqm_samerepr. exact H0. +Qed. + +Lemma encode_int16_zero_ext: + forall x, + encode_int Mint16unsigned (Int.zero_ext 16 x) = encode_int Mint16unsigned x. +Proof. + intros. apply encode_16_mod. apply Int.eqmod_sym. + apply (Int.eqmod_two_p_zero_ext 16). compute; auto. +Qed. + +Lemma encode_int16_sign_ext: + forall x, + encode_int Mint16signed (Int.sign_ext 16 x) = encode_int Mint16signed x. +Proof. + intros. repeat rewrite encode_int16_signed_unsigned. + apply encode_16_mod. apply Int.eqmod_sym. + apply Int.eqmod_two_p_sign_ext. compute; auto. +Qed. + +Lemma decode_int8_zero_ext: + forall l, + Int.zero_ext 8 (decode_int Mint8unsigned l) = decode_int Mint8unsigned l. +Proof. + intros; simpl. destruct (rev_if_le l); auto. destruct l0; auto. + unfold Int.zero_ext. decEq. + generalize (Byte.unsigned_range i). intro. + rewrite Int.unsigned_repr. apply Zmod_small. assumption. + assert (Byte.modulus < Int.max_unsigned). vm_compute. auto. + omega. +Qed. + +Lemma decode_int8_sign_ext: + forall l, + Int.sign_ext 8 (decode_int Mint8signed l) = decode_int Mint8signed l. +Proof. + intros; simpl. destruct (rev_if_le l); auto. destruct l0; auto. + rewrite Int.sign_ext_idem. auto. vm_compute; auto. +Qed. + +Lemma decode_int16_zero_ext: + forall l, + Int.zero_ext 16 (decode_int Mint16unsigned l) = decode_int Mint16unsigned l. +Proof. + intros; simpl. destruct (rev_if_le l); auto. destruct l0; auto. destruct l0; auto. + unfold Int.zero_ext. decEq. + generalize (Byte.unsigned_range i) (Byte.unsigned_range i0). + change Byte.modulus with 256. intros. + assert (0 <= Byte.unsigned i * 256 + Byte.unsigned i0 < 65536). omega. + rewrite Int.unsigned_repr. apply Zmod_small. assumption. + assert (65536 < Int.max_unsigned). vm_compute. auto. + omega. +Qed. + +Lemma decode_int16_sign_ext: + forall l, + Int.sign_ext 16 (decode_int Mint16signed l) = decode_int Mint16signed l. +Proof. + intros; simpl. destruct (rev_if_le l); auto. destruct l0; auto. destruct l0; auto. + rewrite Int.sign_ext_idem. auto. vm_compute; auto. +Qed. + +Lemma decode_int8_signed_unsigned: + forall l, + decode_int Mint8signed l = Int.sign_ext 8 (decode_int Mint8unsigned l). +Proof. + unfold decode_int; intros. destruct (rev_if_le l); auto. destruct l0; auto. +Qed. + +Lemma decode_int16_signed_unsigned: + forall l, + decode_int Mint16signed l = Int.sign_ext 16 (decode_int Mint16unsigned l). +Proof. + unfold decode_int; intros. destruct (rev_if_le l); auto. + destruct l0; auto. destruct l0; auto. +Qed. + +(** * Encoding and decoding floats *) + +Parameter encode_float: memory_chunk -> float -> list byte. +Parameter decode_float: memory_chunk -> list byte -> float. + +Axiom encode_float_length: + forall chunk n, length(encode_float chunk n) = size_chunk_nat chunk. + +(* More realistic: + decode_float Mfloat32 (encode_float Mfloat32 (Float.singleoffloat n)) = + Float.singleoffloat n +*) +Axiom decode_encode_float32: forall n, + decode_float Mfloat32 (encode_float Mfloat32 n) = Float.singleoffloat n. +Axiom decode_encode_float64: forall n, + decode_float Mfloat64 (encode_float Mfloat64 n) = n. + +Axiom encode_float32_singleoffloat: forall n, + encode_float Mfloat32 (Float.singleoffloat n) = encode_float Mfloat32 n. + +Axiom encode_float8_signed_unsigned: forall n, + encode_float Mint8signed n = encode_float Mint8unsigned n. +Axiom encode_float16_signed_unsigned: forall n, + encode_float Mint16signed n = encode_float Mint16unsigned n. + +Axiom encode_float32_cast: + forall f, + encode_float Mfloat32 (Float.singleoffloat f) = encode_float Mfloat32 f. + +Axiom decode_float32_cast: + forall l, + Float.singleoffloat (decode_float Mfloat32 l) = decode_float Mfloat32 l. + +(** * Encoding and decoding values *) + +Definition inj_bytes (bl: list byte) : list memval := + List.map Byte bl. + +Fixpoint proj_bytes (vl: list memval) : option (list byte) := + match vl with + | nil => Some nil + | Byte b :: vl' => + match proj_bytes vl' with None => None | Some bl => Some(b :: bl) end + | _ => None + end. + +Remark length_inj_bytes: + forall bl, length (inj_bytes bl) = length bl. +Proof. + intros. apply List.map_length. +Qed. + +Remark proj_inj_bytes: + forall bl, proj_bytes (inj_bytes bl) = Some bl. +Proof. + induction bl; simpl. auto. rewrite IHbl. auto. +Qed. + +Lemma inj_proj_bytes: + forall cl bl, proj_bytes cl = Some bl -> cl = inj_bytes bl. +Proof. + induction cl; simpl; intros. + inv H; auto. + destruct a; try congruence. destruct (proj_bytes cl); inv H. + simpl. decEq. auto. +Qed. + +Fixpoint inj_pointer (n: nat) (b: block) (ofs: int) {struct n}: list memval := + match n with + | O => nil + | S m => Pointer b ofs m :: inj_pointer m b ofs + end. + +Fixpoint check_pointer (n: nat) (b: block) (ofs: int) (vl: list memval) + {struct n} : bool := + match n, vl with + | O, nil => true + | S m, Pointer b' ofs' m' :: vl' => + eq_block b b' && Int.eq_dec ofs ofs' && beq_nat m m' && check_pointer m b ofs vl' + | _, _ => false + end. + +Definition proj_pointer (vl: list memval) : val := + match vl with + | Pointer b ofs n :: vl' => + if check_pointer (size_chunk_nat Mint32) b ofs vl + then Vptr b ofs + else Vundef + | _ => Vundef + end. + +Definition encode_val (chunk: memory_chunk) (v: val) : list memval := + match v, chunk with + | Vptr b ofs, Mint32 => inj_pointer (size_chunk_nat Mint32) b ofs + | Vint n, _ => inj_bytes (encode_int chunk n) + | Vfloat f, _ => inj_bytes (encode_float chunk f) + | _, _ => list_repeat (size_chunk_nat chunk) Undef + end. + +Definition decode_val (chunk: memory_chunk) (vl: list memval) : val := + match proj_bytes vl with + | Some bl => + match chunk with + | Mint8signed | Mint8unsigned + | Mint16signed | Mint16unsigned | Mint32 => + Vint(decode_int chunk bl) + | Mfloat32 | Mfloat64 => + Vfloat(decode_float chunk bl) + end + | None => + match chunk with + | Mint32 => proj_pointer vl + | _ => Vundef + end + end. + +(* +Lemma inj_pointer_length: + forall b ofs n, List.length(inj_pointer n b ofs) = n. +Proof. + induction n; simpl; congruence. +Qed. +*) + +Lemma encode_val_length: + forall chunk v, length(encode_val chunk v) = size_chunk_nat chunk. +Proof. + intros. destruct v; simpl. + apply length_list_repeat. + rewrite length_inj_bytes. apply encode_int_length. + rewrite length_inj_bytes. apply encode_float_length. + destruct chunk; try (apply length_list_repeat). reflexivity. +Qed. + +Lemma check_inj_pointer: + forall b ofs n, check_pointer n b ofs (inj_pointer n b ofs) = true. +Proof. + induction n; simpl. auto. + unfold proj_sumbool. rewrite dec_eq_true. rewrite dec_eq_true. + rewrite <- beq_nat_refl. simpl; auto. +Qed. + +Definition decode_encode_val (v1: val) (chunk1 chunk2: memory_chunk) (v2: val) : Prop := + match v1, chunk1, chunk2 with + | Vundef, _, _ => v2 = Vundef + | Vint n, Mint8signed, Mint8signed => v2 = Vint(Int.sign_ext 8 n) + | Vint n, Mint8unsigned, Mint8signed => v2 = Vint(Int.sign_ext 8 n) + | Vint n, Mint8signed, Mint8unsigned => v2 = Vint(Int.zero_ext 8 n) + | Vint n, Mint8unsigned, Mint8unsigned => v2 = Vint(Int.zero_ext 8 n) + | Vint n, Mint16signed, Mint16signed => v2 = Vint(Int.sign_ext 16 n) + | Vint n, Mint16unsigned, Mint16signed => v2 = Vint(Int.sign_ext 16 n) + | Vint n, Mint16signed, Mint16unsigned => v2 = Vint(Int.zero_ext 16 n) + | Vint n, Mint16unsigned, Mint16unsigned => v2 = Vint(Int.zero_ext 16 n) + | Vint n, Mint32, Mint32 => v2 = Vint n + | Vint n, Mint32, Mfloat32 => v2 = Vfloat(decode_float Mfloat32 (encode_int Mint32 n)) + | Vint n, _, _ => True (* nothing interesting to say about v2 *) + | Vptr b ofs, Mint32, Mint32 => v2 = Vptr b ofs + | Vptr b ofs, _, _ => v2 = Vundef + | Vfloat f, Mfloat32, Mfloat32 => v2 = Vfloat(Float.singleoffloat f) + | Vfloat f, Mfloat32, Mint32 => v2 = Vint(decode_int Mint32 (encode_float Mfloat32 f)) + | Vfloat f, Mfloat64, Mfloat64 => v2 = Vfloat f + | Vfloat f, _, _ => True (* nothing interesting to say about v2 *) + end. + +Lemma decode_encode_val_general: + forall v chunk1 chunk2, + decode_encode_val v chunk1 chunk2 (decode_val chunk2 (encode_val chunk1 v)). +Proof. + intros. destruct v. +(* Vundef *) + simpl. destruct (size_chunk_nat_pos chunk1) as [psz EQ]. + rewrite EQ. simpl. + unfold decode_val. simpl. destruct chunk2; auto. +(* Vint *) + simpl. + destruct chunk1; auto; destruct chunk2; auto; unfold decode_val; + rewrite proj_inj_bytes. + rewrite decode_encode_int8signed. auto. + rewrite encode_int8_signed_unsigned. rewrite decode_encode_int8unsigned. auto. + rewrite <- encode_int8_signed_unsigned. rewrite decode_encode_int8signed. auto. + rewrite decode_encode_int8unsigned. auto. + rewrite decode_encode_int16signed. auto. + rewrite encode_int16_signed_unsigned. rewrite decode_encode_int16unsigned. auto. + rewrite <- encode_int16_signed_unsigned. rewrite decode_encode_int16signed. auto. + rewrite decode_encode_int16unsigned. auto. + rewrite decode_encode_int32. auto. + auto. +(* Vfloat *) + unfold decode_val, encode_val, decode_encode_val; + destruct chunk1; auto; destruct chunk2; auto; unfold decode_val; + rewrite proj_inj_bytes. + auto. + rewrite decode_encode_float32. auto. + rewrite decode_encode_float64. auto. +(* Vptr *) + unfold decode_val, encode_val, decode_encode_val; + destruct chunk1; auto; destruct chunk2; auto. + simpl. generalize (check_inj_pointer b i (size_chunk_nat Mint32)). + simpl. intro. rewrite H. auto. +Qed. + +Lemma decode_encode_val_similar: + forall v1 chunk1 chunk2 v2, + type_of_chunk chunk1 = type_of_chunk chunk2 -> + size_chunk chunk1 = size_chunk chunk2 -> + Val.has_type v1 (type_of_chunk chunk1) -> + decode_encode_val v1 chunk1 chunk2 v2 -> + v2 = Val.load_result chunk2 v1. +Proof. + intros. + destruct v1. + simpl in *. destruct chunk2; simpl; auto. + red in H1. + destruct chunk1; simpl in H1; try contradiction; + destruct chunk2; simpl in *; discriminate || auto. + red in H1. + destruct chunk1; simpl in H1; try contradiction; + destruct chunk2; simpl in *; discriminate || auto. + red in H1. + destruct chunk1; simpl in H1; try contradiction; + destruct chunk2; simpl in *; discriminate || auto. +Qed. + +Lemma decode_val_type: + forall chunk cl, + Val.has_type (decode_val chunk cl) (type_of_chunk chunk). +Proof. + intros. unfold decode_val. + destruct (proj_bytes cl). + destruct chunk; simpl; auto. + destruct chunk; simpl; auto. + unfold proj_pointer. destruct cl; try (exact I). + destruct m; try (exact I). + destruct (check_pointer (size_chunk_nat Mint32) b i (Pointer b i n :: cl)); + exact I. +Qed. + +Lemma encode_val_int8_signed_unsigned: + forall v, encode_val Mint8signed v = encode_val Mint8unsigned v. +Proof. + intros. destruct v; simpl; auto. rewrite encode_float8_signed_unsigned; auto. +Qed. + +Lemma encode_val_int16_signed_unsigned: + forall v, encode_val Mint16signed v = encode_val Mint16unsigned v. +Proof. + intros. destruct v; simpl; auto. rewrite encode_float16_signed_unsigned; auto. +Qed. + +Lemma encode_val_int8_zero_ext: + forall n, encode_val Mint8unsigned (Vint (Int.zero_ext 8 n)) = encode_val Mint8unsigned (Vint n). +Proof. + intros; unfold encode_val. rewrite encode_int8_zero_ext. auto. +Qed. + +Lemma encode_val_int8_sign_ext: + forall n, encode_val Mint8signed (Vint (Int.sign_ext 8 n)) = encode_val Mint8signed (Vint n). +Proof. + intros; unfold encode_val. rewrite encode_int8_sign_ext. auto. +Qed. + +Lemma encode_val_int16_zero_ext: + forall n, encode_val Mint16unsigned (Vint (Int.zero_ext 16 n)) = encode_val Mint16unsigned (Vint n). +Proof. + intros; unfold encode_val. rewrite encode_int16_zero_ext. auto. +Qed. + +Lemma encode_val_int16_sign_ext: + forall n, encode_val Mint16signed (Vint (Int.sign_ext 16 n)) = encode_val Mint16signed (Vint n). +Proof. + intros; unfold encode_val. rewrite encode_int16_sign_ext. auto. +Qed. + +Lemma decode_val_int_inv: + forall chunk cl n, + decode_val chunk cl = Vint n -> + type_of_chunk chunk = Tint /\ + exists bytes, proj_bytes cl = Some bytes /\ n = decode_int chunk bytes. +Proof. + intros until n. unfold decode_val. destruct (proj_bytes cl). +Opaque decode_int. + destruct chunk; intro EQ; inv EQ; split; auto; exists l; auto. + destruct chunk; try congruence. unfold proj_pointer. + destruct cl; try congruence. destruct m; try congruence. + destruct (check_pointer (size_chunk_nat Mint32) b i (Pointer b i n0 :: cl)); + congruence. +Qed. + +Lemma decode_val_float_inv: + forall chunk cl f, + decode_val chunk cl = Vfloat f -> + type_of_chunk chunk = Tfloat /\ + exists bytes, proj_bytes cl = Some bytes /\ f = decode_float chunk bytes. +Proof. + intros until f. unfold decode_val. destruct (proj_bytes cl). + destruct chunk; intro EQ; inv EQ; split; auto; exists l; auto. + destruct chunk; try congruence. unfold proj_pointer. + destruct cl; try congruence. destruct m; try congruence. + destruct (check_pointer (size_chunk_nat Mint32) b i (Pointer b i n :: cl)); + congruence. +Qed. + +Lemma decode_val_cast: + forall chunk l, + let v := decode_val chunk l in + match chunk with + | Mint8signed => v = Val.sign_ext 8 v + | Mint8unsigned => v = Val.zero_ext 8 v + | Mint16signed => v = Val.sign_ext 16 v + | Mint16unsigned => v = Val.zero_ext 16 v + | Mfloat32 => v = Val.singleoffloat v + | _ => True + end. +Proof. + unfold decode_val; intros; destruct chunk; auto; destruct (proj_bytes l); auto. + unfold Val.sign_ext. decEq. symmetry. apply decode_int8_sign_ext. + unfold Val.zero_ext. decEq. symmetry. apply decode_int8_zero_ext. + unfold Val.sign_ext. decEq. symmetry. apply decode_int16_sign_ext. + unfold Val.zero_ext. decEq. symmetry. apply decode_int16_zero_ext. + unfold Val.singleoffloat. decEq. symmetry. apply decode_float32_cast. +Qed. + +(** Pointers cannot be forged. *) + +Definition memval_valid_first (mv: memval) : Prop := + match mv with + | Pointer b ofs n => n = pred (size_chunk_nat Mint32) + | _ => True + end. + +Definition memval_valid_cont (mv: memval) : Prop := + match mv with + | Pointer b ofs n => n <> pred (size_chunk_nat Mint32) + | _ => True + end. + +Inductive encoding_shape: list memval -> Prop := + | encoding_shape_intro: forall mv1 mvl, + memval_valid_first mv1 -> + (forall mv, In mv mvl -> memval_valid_cont mv) -> + encoding_shape (mv1 :: mvl). + +Lemma encode_val_shape: + forall chunk v, encoding_shape (encode_val chunk v). +Proof. + intros. + destruct (size_chunk_nat_pos chunk) as [sz1 EQ]. + assert (encoding_shape (list_repeat (size_chunk_nat chunk) Undef)). + rewrite EQ; simpl; constructor. exact I. + intros. replace mv with Undef. exact I. symmetry; eapply in_list_repeat; eauto. + assert (forall bl, length bl = size_chunk_nat chunk -> + encoding_shape (inj_bytes bl)). + intros. destruct bl; simpl in *. congruence. + constructor. exact I. unfold inj_bytes. intros. + exploit list_in_map_inv; eauto. intros [x [A B]]. subst mv. exact I. + destruct v; simpl. + auto. + apply H0. apply encode_int_length. + apply H0. apply encode_float_length. + destruct chunk; auto. + constructor. red. auto. + simpl; intros. intuition; subst mv; red; simpl; congruence. +Qed. + +Lemma check_pointer_inv: + forall b ofs n mv, + check_pointer n b ofs mv = true -> mv = inj_pointer n b ofs. +Proof. + induction n; destruct mv; simpl. + auto. + congruence. + congruence. + destruct m; try congruence. intro. + destruct (andb_prop _ _ H). destruct (andb_prop _ _ H0). + destruct (andb_prop _ _ H2). + decEq. decEq. symmetry; eapply proj_sumbool_true; eauto. + symmetry; eapply proj_sumbool_true; eauto. + symmetry; apply beq_nat_true; auto. + auto. +Qed. + +Inductive decoding_shape: list memval -> Prop := + | decoding_shape_intro: forall mv1 mvl, + memval_valid_first mv1 -> mv1 <> Undef -> + (forall mv, In mv mvl -> memval_valid_cont mv /\ mv <> Undef) -> + decoding_shape (mv1 :: mvl). + +Lemma decode_val_shape: + forall chunk mvl, + List.length mvl = size_chunk_nat chunk -> + decode_val chunk mvl = Vundef \/ decoding_shape mvl. +Proof. + intros. destruct (size_chunk_nat_pos chunk) as [sz EQ]. + unfold decode_val. + caseEq (proj_bytes mvl). + intros bl PROJ. right. exploit inj_proj_bytes; eauto. intros. subst mvl. + destruct bl; simpl in H. congruence. simpl. constructor. + red; auto. congruence. + unfold inj_bytes; intros. exploit list_in_map_inv; eauto. intros [b [A B]]. + subst mv. split. red; auto. congruence. + intros. destruct chunk; auto. unfold proj_pointer. + destruct mvl; auto. destruct m; auto. + caseEq (check_pointer (size_chunk_nat Mint32) b i (Pointer b i n :: mvl)); auto. + intros. right. exploit check_pointer_inv; eauto. simpl; intros; inv H2. + constructor. red. auto. congruence. + simpl; intros. intuition; subst mv; simpl; congruence. +Qed. + +Lemma encode_val_pointer_inv: + forall chunk v b ofs n mvl, + encode_val chunk v = Pointer b ofs n :: mvl -> + chunk = Mint32 /\ v = Vptr b ofs /\ mvl = inj_pointer (pred (size_chunk_nat Mint32)) b ofs. +Proof. + intros until mvl. + destruct (size_chunk_nat_pos chunk) as [sz SZ]. + unfold encode_val. rewrite SZ. destruct v. + simpl. congruence. + generalize (encode_int_length chunk i). destruct (encode_int chunk i); simpl; congruence. + generalize (encode_float_length chunk f). destruct (encode_float chunk f); simpl; congruence. + destruct chunk; try (simpl; congruence). + simpl. intuition congruence. +Qed. + +Lemma decode_val_pointer_inv: + forall chunk mvl b ofs, + decode_val chunk mvl = Vptr b ofs -> + chunk = Mint32 /\ mvl = inj_pointer (size_chunk_nat Mint32) b ofs. +Proof. + intros until ofs; unfold decode_val. + destruct (proj_bytes mvl). + destruct chunk; congruence. + destruct chunk; try congruence. + unfold proj_pointer. destruct mvl. congruence. destruct m; try congruence. + case_eq (check_pointer (size_chunk_nat Mint32) b0 i (Pointer b0 i n :: mvl)); intros. + inv H0. split; auto. apply check_pointer_inv; auto. + congruence. +Qed. + +Inductive pointer_encoding_shape: list memval -> Prop := + | pointer_encoding_shape_intro: forall mv1 mvl, + ~memval_valid_cont mv1 -> + (forall mv, In mv mvl -> ~memval_valid_first mv) -> + pointer_encoding_shape (mv1 :: mvl). + +Lemma encode_pointer_shape: + forall b ofs, pointer_encoding_shape (encode_val Mint32 (Vptr b ofs)). +Proof. + intros. simpl. constructor. + unfold memval_valid_cont. red; intro. elim H. auto. + unfold memval_valid_first. simpl; intros; intuition; subst mv; congruence. +Qed. + +Lemma decode_pointer_shape: + forall chunk mvl b ofs, + decode_val chunk mvl = Vptr b ofs -> + chunk = Mint32 /\ pointer_encoding_shape mvl. +Proof. + intros. exploit decode_val_pointer_inv; eauto. intros [A B]. + split; auto. subst mvl. apply encode_pointer_shape. +Qed. + +(* +Lemma proj_bytes_none: + forall mv, + match mv with Byte _ => False | _ => True end -> + forall mvl, + In mv mvl -> + proj_bytes mvl = None. +Proof. + induction mvl; simpl; intros. + elim H0. + destruct a; auto. destruct H0. subst mv. contradiction. + rewrite (IHmvl H0); auto. +Qed. + +Lemma decode_val_undef: + forall chunk mv mv1 mvl, + match mv with + | Pointer b ofs n => n = pred (size_chunk_nat Mint32) + | Undef => True + | _ => False + end -> + In mv mvl -> + decode_val chunk (mv1 :: mvl) = Vundef. +Proof. + intros. unfold decode_val. + replace (proj_bytes (mv1 :: mvl)) with (@None (list byte)). + destruct chunk; auto. unfold proj_pointer. destruct mv1; auto. + case_eq (check_pointer (size_chunk_nat Mint32) b i (Pointer b i n :: mvl)); intros. + exploit check_pointer_inv; eauto. simpl. intros. inv H2. + simpl in H0. intuition; subst mv; simpl in H; congruence. + auto. + symmetry. apply proj_bytes_none with mv. + destruct mv; tauto. auto with coqlib. +Qed. + +*) + +(** * Compatibility with memory injections *) + +(** Relating two memory values according to a memory injection. *) + +Inductive memval_inject (f: meminj): memval -> memval -> Prop := + | memval_inject_byte: + forall n, memval_inject f (Byte n) (Byte n) + | memval_inject_ptr: + forall b1 ofs1 b2 ofs2 delta n, + f b1 = Some (b2, delta) -> + ofs2 = Int.add ofs1 (Int.repr delta) -> + memval_inject f (Pointer b1 ofs1 n) (Pointer b2 ofs2 n) + | memval_inject_undef: + forall mv, memval_inject f Undef mv. + +Lemma memval_inject_incr: + forall f f' v1 v2, memval_inject f v1 v2 -> inject_incr f f' -> memval_inject f' v1 v2. +Proof. + intros. inv H; econstructor. rewrite (H0 _ _ _ H1). reflexivity. auto. +Qed. + +(** [decode_val], applied to lists of memory values that are pairwise + related by [memval_inject], returns values that are related by [val_inject]. *) + +Lemma proj_bytes_inject: + forall f vl vl', + list_forall2 (memval_inject f) vl vl' -> + forall bl, + proj_bytes vl = Some bl -> + proj_bytes vl' = Some bl. +Proof. + induction 1; simpl. congruence. + inv H; try congruence. + destruct (proj_bytes al); intros. + inv H. rewrite (IHlist_forall2 l); auto. + congruence. +Qed. + +Lemma check_pointer_inject: + forall f vl vl', + list_forall2 (memval_inject f) vl vl' -> + forall n b ofs b' delta, + check_pointer n b ofs vl = true -> + f b = Some(b', delta) -> + check_pointer n b' (Int.add ofs (Int.repr delta)) vl' = true. +Proof. + induction 1; intros; destruct n; simpl in *; auto. + inv H; auto. + destruct (andb_prop _ _ H1). destruct (andb_prop _ _ H). + destruct (andb_prop _ _ H5). + assert (n = n0) by (apply beq_nat_true; auto). + assert (b = b0) by (eapply proj_sumbool_true; eauto). + assert (ofs = ofs1) by (eapply proj_sumbool_true; eauto). + subst. rewrite H3 in H2; inv H2. + unfold proj_sumbool. rewrite dec_eq_true. rewrite dec_eq_true. + rewrite <- beq_nat_refl. simpl. eauto. + congruence. +Qed. + +Lemma proj_pointer_inject: + forall f vl1 vl2, + list_forall2 (memval_inject f) vl1 vl2 -> + val_inject f (proj_pointer vl1) (proj_pointer vl2). +Proof. + intros. unfold proj_pointer. + inversion H; subst. auto. inversion H0; subst; auto. + case_eq (check_pointer (size_chunk_nat Mint32) b0 ofs1 (Pointer b0 ofs1 n :: al)); intros. + exploit check_pointer_inject. eexact H. eauto. eauto. + intro. rewrite H4. econstructor; eauto. + constructor. +Qed. + +Lemma proj_bytes_not_inject: + forall f vl vl', + list_forall2 (memval_inject f) vl vl' -> + proj_bytes vl = None -> proj_bytes vl' <> None -> In Undef vl. +Proof. + induction 1; simpl; intros. + congruence. + inv H; try congruence. + right. apply IHlist_forall2. + destruct (proj_bytes al); congruence. + destruct (proj_bytes bl); congruence. + auto. +Qed. + +Lemma check_pointer_undef: + forall n b ofs vl, + In Undef vl -> check_pointer n b ofs vl = false. +Proof. + induction n; intros; simpl. + destruct vl. elim H. auto. + destruct vl. auto. + destruct m; auto. simpl in H; destruct H. congruence. + rewrite IHn; auto. apply andb_false_r. +Qed. + +Lemma proj_pointer_undef: + forall vl, In Undef vl -> proj_pointer vl = Vundef. +Proof. + intros; unfold proj_pointer. + destruct vl; auto. destruct m; auto. + rewrite check_pointer_undef. auto. auto. +Qed. + +Theorem decode_val_inject: + forall f vl1 vl2 chunk, + list_forall2 (memval_inject f) vl1 vl2 -> + val_inject f (decode_val chunk vl1) (decode_val chunk vl2). +Proof. + intros. unfold decode_val. + case_eq (proj_bytes vl1); intros. + exploit proj_bytes_inject; eauto. intros. rewrite H1. + destruct chunk; constructor. + destruct chunk; auto. + case_eq (proj_bytes vl2); intros. + rewrite proj_pointer_undef. auto. eapply proj_bytes_not_inject; eauto. congruence. + apply proj_pointer_inject; auto. +Qed. + +(** Symmetrically, [encode_val], applied to values related by [val_inject], + returns lists of memory values that are pairwise + related by [memval_inject]. *) + +Lemma inj_bytes_inject: + forall f bl, list_forall2 (memval_inject f) (inj_bytes bl) (inj_bytes bl). +Proof. + induction bl; constructor; auto. constructor. +Qed. + +Lemma repeat_Undef_inject_any: + forall f vl, + list_forall2 (memval_inject f) (list_repeat (length vl) Undef) vl. +Proof. + induction vl; simpl; constructor; auto. constructor. +Qed. + +Lemma repeat_Undef_inject_self: + forall f n, + list_forall2 (memval_inject f) (list_repeat n Undef) (list_repeat n Undef). +Proof. + induction n; simpl; constructor; auto. constructor. +Qed. + +Theorem encode_val_inject: + forall f v1 v2 chunk, + val_inject f v1 v2 -> + list_forall2 (memval_inject f) (encode_val chunk v1) (encode_val chunk v2). +Proof. + intros. inv H; simpl. + apply inj_bytes_inject. + apply inj_bytes_inject. + destruct chunk; try apply repeat_Undef_inject_self. + unfold inj_pointer; simpl; repeat econstructor; auto. + replace (size_chunk_nat chunk) with (length (encode_val chunk v2)). + apply repeat_Undef_inject_any. apply encode_val_length. +Qed. + +(** The identity injection has interesting properties. *) + +Definition inject_id : meminj := fun b => Some(b, 0). + +Lemma val_inject_id: + forall v1 v2, + val_inject inject_id v1 v2 <-> Val.lessdef v1 v2. +Proof. + intros; split; intros. + inv H. constructor. constructor. + unfold inject_id in H0. inv H0. rewrite Int.add_zero. constructor. + constructor. + inv H. destruct v2; econstructor. unfold inject_id; reflexivity. rewrite Int.add_zero; auto. + constructor. +Qed. + +Lemma memval_inject_id: + forall mv, memval_inject inject_id mv mv. +Proof. + destruct mv; econstructor. unfold inject_id; reflexivity. rewrite Int.add_zero; auto. +Qed. + diff --git a/common/Memdataaux.ml b/common/Memdataaux.ml new file mode 100644 index 0000000..3a39428 --- /dev/null +++ b/common/Memdataaux.ml @@ -0,0 +1,68 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +open Camlcoq +open Integers +open AST + +let big_endian = + match Configuration.arch with + | "powerpc" -> true + | "arm" -> false + | _ -> assert false + +let encode_float chunk f = + match chunk with + | Mint8unsigned | Mint8signed -> + [Byte.zero] + | Mint16unsigned | Mint16signed -> + [Byte.zero; Byte.zero] + | Mint32 -> + [Byte.zero; Byte.zero; Byte.zero; Byte.zero] + | Mfloat32 -> + let bits = Int32.bits_of_float f in + let byte n = + coqint_of_camlint + (Int32.logand (Int32.shift_right_logical bits n) 0xFFl) in + if big_endian then + [byte 24; byte 16; byte 8; byte 0] + else + [byte 0; byte 8; byte 16; byte 24] + | Mfloat64 -> + let bits = Int64.bits_of_float f in + let byte n = + coqint_of_camlint + (Int64.to_int32 + (Int64.logand (Int64.shift_right_logical bits n) 0xFFL)) in + if big_endian then + [byte 56; byte 48; byte 40; byte 32; byte 24; byte 16; byte 8; byte 0] + else + [byte 0; byte 8; byte 16; byte 24; byte 32; byte 40; byte 48; byte 56] + +let decode_float chunk bytes = + match chunk with + | Mfloat32 -> + let combine accu b = + Int32.logor (Int32.shift_left accu 8) (camlint_of_coqint b) in + Int32.float_of_bits + (List.fold_left combine 0l + (if big_endian then bytes else List.rev bytes)) + | Mfloat64 -> + let combine accu b = + Int64.logor (Int64.shift_left accu 8) + (Int64.of_int32 (camlint_of_coqint b)) in + Int64.float_of_bits + (List.fold_left combine 0L + (if big_endian then bytes else List.rev bytes)) + | _ -> + 0.0 (* unspecified *) + diff --git a/common/Memory.v b/common/Memory.v new file mode 100644 index 0000000..3092021 --- /dev/null +++ b/common/Memory.v @@ -0,0 +1,2844 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* Sandrine Blazy, ENSIIE and INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** This file develops the memory model that is used in the dynamic + semantics of all the languages used in the compiler. + It defines a type [mem] of memory states, the following 4 basic + operations over memory states, and their properties: +- [load]: read a memory chunk at a given address; +- [store]: store a memory chunk at a given address; +- [alloc]: allocate a fresh memory block; +- [free]: invalidate a memory block. +*) + +Require Import Coqlib. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Export Memdata. +Require Export Memtype. + +Definition update (A: Type) (x: Z) (v: A) (f: Z -> A) : Z -> A := + fun y => if zeq y x then v else f y. + +Implicit Arguments update [A]. + +Lemma update_s: + forall (A: Type) (x: Z) (v: A) (f: Z -> A), + update x v f x = v. +Proof. + intros; unfold update. apply zeq_true. +Qed. + +Lemma update_o: + forall (A: Type) (x: Z) (v: A) (f: Z -> A) (y: Z), + x <> y -> update x v f y = f y. +Proof. + intros; unfold update. apply zeq_false; auto. +Qed. + +Module Mem : MEM. + +Record mem_ : Type := mkmem { + contents: block -> Z -> memval; + access: block -> Z -> bool; + bound: block -> Z * Z; + next: block; + next_pos: next > 0; + next_noaccess: forall b ofs, b >= next -> access b ofs = false; + bound_noaccess: forall b ofs, ofs < fst(bound b) \/ ofs >= snd(bound b) -> access b ofs = false +}. + +Definition mem := mem_. + +(** * Validity of blocks and accesses *) + +(** A block address is valid if it was previously allocated. It remains valid + even after being freed. *) + +Definition nextblock (m: mem) : block := m.(next). + +Theorem nextblock_pos: + forall m, nextblock m > 0. +Proof next_pos. + +Definition valid_block (m: mem) (b: block) := + b < nextblock m. + +Theorem valid_not_valid_diff: + forall m b b', valid_block m b -> ~(valid_block m b') -> b <> b'. +Proof. + intros; red; intros. subst b'. contradiction. +Qed. + +Hint Local Resolve valid_not_valid_diff: mem. + +(** Permissions *) + +Definition perm (m: mem) (b: block) (ofs: Z) (p: permission) : Prop := + m.(access) b ofs = true. + +Theorem perm_implies: + forall m b ofs p1 p2, perm m b ofs p1 -> perm_order p1 p2 -> perm m b ofs p2. +Proof. + unfold perm; auto. +Qed. + +Hint Local Resolve perm_implies: mem. + +Theorem perm_valid_block: + forall m b ofs p, perm m b ofs p -> valid_block m b. +Proof. + unfold perm; intros. + destruct (zlt b m.(next)). + auto. + assert (access m b ofs = false). eapply next_noaccess; eauto. + congruence. +Qed. + +Hint Local Resolve perm_valid_block: mem. + +Theorem perm_dec: + forall m b ofs p, {perm m b ofs p} + {~ perm m b ofs p}. +Proof. + unfold perm; intros. + destruct (access m b ofs). left; auto. right; congruence. +Qed. + +Definition range_perm (m: mem) (b: block) (lo hi: Z) (p: permission) : Prop := + forall ofs, lo <= ofs < hi -> perm m b ofs p. + +Theorem range_perm_implies: + forall m b lo hi p1 p2, + range_perm m b lo hi p1 -> perm_order p1 p2 -> range_perm m b lo hi p2. +Proof. + unfold range_perm; intros; eauto with mem. +Qed. + +Hint Local Resolve range_perm_implies: mem. + +Lemma range_perm_dec: + forall m b lo hi p, {range_perm m b lo hi p} + {~ range_perm m b lo hi p}. +Proof. + intros. + assert (forall n, 0 <= n -> + {range_perm m b lo (lo + n) p} + {~ range_perm m b lo (lo + n) p}). + apply natlike_rec2. + left. red; intros. omegaContradiction. + intros. destruct H0. + destruct (perm_dec m b (lo + z) p). + left. red; intros. destruct (zeq ofs (lo + z)). congruence. apply r. omega. + right; red; intro. elim n. apply H0. omega. + right; red; intro. elim n. red; intros. apply H0. omega. + destruct (zlt lo hi). + replace hi with (lo + (hi - lo)) by omega. apply H. omega. + left; red; intros. omegaContradiction. +Qed. + +(** [valid_access m chunk b ofs p] holds if a memory access + of the given chunk is possible in [m] at address [b, ofs] + with permissions [p]. + This means: +- The range of bytes accessed all have permission [p]. +- The offset [ofs] is aligned. +*) + +Definition valid_access (m: mem) (chunk: memory_chunk) (b: block) (ofs: Z) (p: permission): Prop := + range_perm m b ofs (ofs + size_chunk chunk) p + /\ (align_chunk chunk | ofs). + +Theorem valid_access_writable_any: + forall m chunk b ofs p, + valid_access m chunk b ofs Writable -> + valid_access m chunk b ofs p. +Proof. + intros. inv H. constructor; auto with mem. +Qed. + +Theorem valid_access_implies: + forall m chunk b ofs p1 p2, + valid_access m chunk b ofs p1 -> perm_order p1 p2 -> + valid_access m chunk b ofs p2. +Proof. + intros. inv H. constructor; eauto with mem. +Qed. + +Hint Local Resolve valid_access_implies: mem. + +Theorem valid_access_valid_block: + forall m chunk b ofs, + valid_access m chunk b ofs Nonempty -> + valid_block m b. +Proof. + intros. destruct H. + assert (perm m b ofs Nonempty). + apply H. generalize (size_chunk_pos chunk). omega. + eauto with mem. +Qed. + +Hint Local Resolve valid_access_valid_block: mem. + +Lemma valid_access_perm: + forall m chunk b ofs p, + valid_access m chunk b ofs p -> + perm m b ofs p. +Proof. + intros. destruct H. apply H. generalize (size_chunk_pos chunk). omega. +Qed. + +Lemma valid_access_compat: + forall m chunk1 chunk2 b ofs p, + size_chunk chunk1 = size_chunk chunk2 -> + valid_access m chunk1 b ofs p-> + valid_access m chunk2 b ofs p. +Proof. + intros. inv H0. rewrite H in H1. constructor; auto. + rewrite <- (align_chunk_compat _ _ H). auto. +Qed. + +Lemma valid_access_dec: + forall m chunk b ofs p, + {valid_access m chunk b ofs p} + {~ valid_access m chunk b ofs p}. +Proof. + intros. + destruct (range_perm_dec m b ofs (ofs + size_chunk chunk) p). + destruct (Zdivide_dec (align_chunk chunk) ofs (align_chunk_pos chunk)). + left; constructor; auto. + right; red; intro V; inv V; contradiction. + right; red; intro V; inv V; contradiction. +Qed. + +(** [valid_pointer] is a boolean-valued function that says whether + the byte at the given location is nonempty. *) + +Definition valid_pointer (m: mem) (b: block) (ofs: Z): bool := + perm_dec m b ofs Nonempty. + +Theorem valid_pointer_nonempty_perm: + forall m b ofs, + valid_pointer m b ofs = true <-> perm m b ofs Nonempty. +Proof. + intros. unfold valid_pointer. + destruct (perm_dec m b ofs Nonempty); simpl; + intuition congruence. +Qed. + +Theorem valid_pointer_valid_access: + forall m b ofs, + valid_pointer m b ofs = true <-> valid_access m Mint8unsigned b ofs Nonempty. +Proof. + intros. rewrite valid_pointer_nonempty_perm. + split; intros. + split. simpl; red; intros. replace ofs0 with ofs by omega. auto. + simpl. apply Zone_divide. + destruct H. apply H. simpl. omega. +Qed. + +(** Bounds *) + +(** Each block has a low bound and a high bound, determined at allocation time + and invariant afterward. The crucial properties of bounds is + that any offset below the low bound or above the high bound is + empty. *) + +Definition bounds (m: mem) (b: block) : Z*Z := m.(bound) b. + +Notation low_bound m b := (fst(bounds m b)). +Notation high_bound m b := (snd(bounds m b)). + +Theorem perm_in_bounds: + forall m b ofs p, perm m b ofs p -> low_bound m b <= ofs < high_bound m b. +Proof. + unfold perm, bounds. intros. + destruct (zlt ofs (fst (bound m b))). + exploit bound_noaccess. left; eauto. congruence. + destruct (zlt ofs (snd (bound m b))). + omega. + exploit bound_noaccess. right; eauto. congruence. +Qed. + +Theorem range_perm_in_bounds: + forall m b lo hi p, + range_perm m b lo hi p -> lo < hi -> low_bound m b <= lo /\ hi <= high_bound m b. +Proof. + intros. split. + exploit (perm_in_bounds m b lo p). apply H. omega. omega. + exploit (perm_in_bounds m b (hi-1) p). apply H. omega. omega. +Qed. + +Theorem valid_access_in_bounds: + forall m chunk b ofs p, + valid_access m chunk b ofs p -> + low_bound m b <= ofs /\ ofs + size_chunk chunk <= high_bound m b. +Proof. + intros. inv H. apply range_perm_in_bounds with p; auto. + generalize (size_chunk_pos chunk). omega. +Qed. + +Hint Local Resolve perm_in_bounds range_perm_in_bounds valid_access_in_bounds. + +(** * Store operations *) + +(** The initial store *) + +Program Definition empty: mem := + mkmem (fun b ofs => Undef) + (fun b ofs => false) + (fun b => (0,0)) + 1 _ _ _. +Next Obligation. + omega. +Qed. + +Definition nullptr: block := 0. + +(** Allocation of a fresh block with the given bounds. Return an updated + memory state and the address of the fresh block, which initially contains + undefined cells. Note that allocation never fails: we model an + infinite memory. *) + +Program Definition alloc (m: mem) (lo hi: Z) := + (mkmem (update m.(next) + (fun ofs => Undef) + m.(contents)) + (update m.(next) + (fun ofs => zle lo ofs && zlt ofs hi) + m.(access)) + (update m.(next) (lo, hi) m.(bound)) + (Zsucc m.(next)) + _ _ _, + m.(next)). +Next Obligation. + generalize (next_pos m). omega. +Qed. +Next Obligation. + rewrite update_o. apply next_noaccess. omega. omega. +Qed. +Next Obligation. + unfold update in *. destruct (zeq b (next m)). + simpl in H. destruct H. + unfold proj_sumbool. rewrite zle_false. auto. omega. + unfold proj_sumbool. rewrite zlt_false. apply andb_false_r. auto. + apply bound_noaccess. auto. +Qed. + +(** Freeing a block between the given bounds. + Return the updated memory state where the given range of the given block + has been invalidated: future reads and writes to this + range will fail. Requires write permission on the given range. *) + +Program Definition unchecked_free (m: mem) (b: block) (lo hi: Z): mem := + mkmem m.(contents) + (update b + (fun ofs => if zle lo ofs && zlt ofs hi then false else m.(access) b ofs) + m.(access)) + m.(bound) + m.(next) _ _ _. +Next Obligation. + apply next_pos. +Qed. +Next Obligation. + unfold update. destruct (zeq b0 b). subst b0. + destruct (zle lo ofs); simpl; auto. + destruct (zlt ofs hi); simpl; auto. + apply next_noaccess; auto. + apply next_noaccess; auto. + apply next_noaccess; auto. +Qed. +Next Obligation. + unfold update. destruct (zeq b0 b). subst b0. + destruct (zle lo ofs); simpl; auto. + destruct (zlt ofs hi); simpl; auto. + apply bound_noaccess; auto. + apply bound_noaccess; auto. + apply bound_noaccess; auto. +Qed. + +Definition free (m: mem) (b: block) (lo hi: Z): option mem := + if range_perm_dec m b lo hi Freeable + then Some(unchecked_free m b lo hi) + else None. + +Fixpoint free_list (m: mem) (l: list (block * Z * Z)) {struct l}: option mem := + match l with + | nil => Some m + | (b, lo, hi) :: l' => + match free m b lo hi with + | None => None + | Some m' => free_list m' l' + end + end. + +(** Memory reads. *) + +(** Reading N adjacent bytes in a block content. *) + +Fixpoint getN (n: nat) (p: Z) (c: Z -> memval) {struct n}: list memval := + match n with + | O => nil + | S n' => c p :: getN n' (p + 1) c + end. + +(** [load chunk m b ofs] perform a read in memory state [m], at address + [b] and offset [ofs]. It returns the value of the memory chunk + at that address. [None] is returned if the accessed bytes + are not readable. *) + +Definition load (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z): option val := + if valid_access_dec m chunk b ofs Readable + then Some(decode_val chunk (getN (size_chunk_nat chunk) ofs (m.(contents) b))) + else None. + +(** [loadv chunk m addr] is similar, but the address and offset are given + as a single value [addr], which must be a pointer value. *) + +Definition loadv (chunk: memory_chunk) (m: mem) (addr: val) : option val := + match addr with + | Vptr b ofs => load chunk m b (Int.signed ofs) + | _ => None + end. + +(** [loadbytes m b ofs n] reads [n] consecutive bytes starting at + location [(b, ofs)]. Returns [None] if the accessed locations are + not readable or do not contain bytes. *) + +Definition loadbytes (m: mem) (b: block) (ofs n: Z): option (list byte) := + if range_perm_dec m b ofs (ofs + n) Readable + then proj_bytes (getN (nat_of_Z n) ofs (m.(contents) b)) + else None. + +(** Memory stores. *) + +(** Writing N adjacent bytes in a block content. *) + +Fixpoint setN (vl: list memval) (p: Z) (c: Z -> memval) {struct vl}: Z -> memval := + match vl with + | nil => c + | v :: vl' => setN vl' (p + 1) (update p v c) + end. + +Definition unchecked_store (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z) (v: val): mem := + mkmem (update b + (setN (encode_val chunk v) ofs (m.(contents) b)) + m.(contents)) + m.(access) + m.(bound) + m.(next) + m.(next_pos) + m.(next_noaccess) + m.(bound_noaccess). + +(** [store chunk m b ofs v] perform a write in memory state [m]. + Value [v] is stored at address [b] and offset [ofs]. + Return the updated memory store, or [None] if the accessed bytes + are not writable. *) + +Definition store (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z) (v: val): option mem := + if valid_access_dec m chunk b ofs Writable + then Some(unchecked_store chunk m b ofs v) + else None. + +(** [storev chunk m addr v] is similar, but the address and offset are given + as a single value [addr], which must be a pointer value. *) + +Definition storev (chunk: memory_chunk) (m: mem) (addr v: val) : option mem := + match addr with + | Vptr b ofs => store chunk m b (Int.signed ofs) v + | _ => None + end. + +(** * Properties of the memory operations *) + +(** Properties of the empty store. *) + +Theorem nextblock_empty: nextblock empty = 1. +Proof. reflexivity. Qed. + +Theorem perm_empty: forall b ofs p, ~perm empty b ofs p. +Proof. + intros. unfold perm, empty; simpl. congruence. +Qed. + +Theorem valid_access_empty: forall chunk b ofs p, ~valid_access empty chunk b ofs p. +Proof. + intros. red; intros. elim (perm_empty b ofs p). apply H. + generalize (size_chunk_pos chunk); omega. +Qed. + +(** ** Properties related to [load] *) + +Theorem valid_access_load: + forall m chunk b ofs, + valid_access m chunk b ofs Readable -> + exists v, load chunk m b ofs = Some v. +Proof. + intros. econstructor. unfold load. rewrite pred_dec_true; eauto. +Qed. + +Theorem load_valid_access: + forall m chunk b ofs v, + load chunk m b ofs = Some v -> + valid_access m chunk b ofs Readable. +Proof. + intros until v. unfold load. + destruct (valid_access_dec m chunk b ofs Readable); intros. + auto. + congruence. +Qed. + +Lemma load_result: + forall chunk m b ofs v, + load chunk m b ofs = Some v -> + v = decode_val chunk (getN (size_chunk_nat chunk) ofs (m.(contents) b)). +Proof. + intros until v. unfold load. + destruct (valid_access_dec m chunk b ofs Readable); intros. + congruence. + congruence. +Qed. + +Hint Local Resolve load_valid_access valid_access_load: mem. + +Theorem load_type: + forall m chunk b ofs v, + load chunk m b ofs = Some v -> + Val.has_type v (type_of_chunk chunk). +Proof. + intros. exploit load_result; eauto; intros. rewrite H0. + apply decode_val_type. +Qed. + +Theorem load_cast: + forall m chunk b ofs v, + load chunk m b ofs = Some v -> + match chunk with + | Mint8signed => v = Val.sign_ext 8 v + | Mint8unsigned => v = Val.zero_ext 8 v + | Mint16signed => v = Val.sign_ext 16 v + | Mint16unsigned => v = Val.zero_ext 16 v + | Mfloat32 => v = Val.singleoffloat v + | _ => True + end. +Proof. + intros. exploit load_result; eauto. + set (l := getN (size_chunk_nat chunk) ofs (contents m b)). + intros. subst v. apply decode_val_cast. +Qed. + +Theorem load_int8_signed_unsigned: + forall m b ofs, + load Mint8signed m b ofs = option_map (Val.sign_ext 8) (load Mint8unsigned m b ofs). +Proof. + intros. unfold load. + change (size_chunk_nat Mint8signed) with (size_chunk_nat Mint8unsigned). + set (cl := getN (size_chunk_nat Mint8unsigned) ofs (contents m b)). + destruct (valid_access_dec m Mint8signed b ofs Readable). + rewrite pred_dec_true; auto. unfold decode_val. + destruct (proj_bytes cl); auto. rewrite decode_int8_signed_unsigned. auto. + rewrite pred_dec_false; auto. +Qed. + +Theorem load_int16_signed_unsigned: + forall m b ofs, + load Mint16signed m b ofs = option_map (Val.sign_ext 16) (load Mint16unsigned m b ofs). +Proof. + intros. unfold load. + change (size_chunk_nat Mint16signed) with (size_chunk_nat Mint16unsigned). + set (cl := getN (size_chunk_nat Mint16unsigned) ofs (contents m b)). + destruct (valid_access_dec m Mint16signed b ofs Readable). + rewrite pred_dec_true; auto. unfold decode_val. + destruct (proj_bytes cl); auto. rewrite decode_int16_signed_unsigned. auto. + rewrite pred_dec_false; auto. +Qed. + +Theorem loadbytes_load: + forall chunk m b ofs bytes, + loadbytes m b ofs (size_chunk chunk) = Some bytes -> + (align_chunk chunk | ofs) -> + load chunk m b ofs = + Some(match type_of_chunk chunk with + | Tint => Vint(decode_int chunk bytes) + | Tfloat => Vfloat(decode_float chunk bytes) + end). +Proof. + unfold loadbytes, load; intros. + destruct (range_perm_dec m b ofs (ofs + size_chunk chunk) Readable); + try congruence. + rewrite pred_dec_true. decEq. unfold size_chunk_nat. + unfold decode_val; rewrite H. destruct chunk; auto. + split; auto. +Qed. + +Theorem load_int_loadbytes: + forall chunk m b ofs n, + load chunk m b ofs = Some(Vint n) -> + type_of_chunk chunk = Tint /\ + exists bytes, loadbytes m b ofs (size_chunk chunk) = Some bytes + /\ n = decode_int chunk bytes. +Proof. + intros. exploit load_valid_access; eauto. intros [A B]. + exploit decode_val_int_inv. symmetry. eapply load_result; eauto. + intros [C [bytes [D E]]]. + split. auto. exists bytes; split. + unfold loadbytes. rewrite pred_dec_true; auto. auto. +Qed. + +Theorem load_float_loadbytes: + forall chunk m b ofs f, + load chunk m b ofs = Some(Vfloat f) -> + type_of_chunk chunk = Tfloat /\ + exists bytes, loadbytes m b ofs (size_chunk chunk) = Some bytes + /\ f = decode_float chunk bytes. +Proof. + intros. exploit load_valid_access; eauto. intros [A B]. + exploit decode_val_float_inv. symmetry. eapply load_result; eauto. + intros [C [bytes [D E]]]. + split. auto. exists bytes; split. + unfold loadbytes. rewrite pred_dec_true; auto. auto. +Qed. + +Lemma getN_length: + forall c n p, length (getN n p c) = n. +Proof. + induction n; simpl; intros. auto. decEq; auto. +Qed. + +Theorem loadbytes_length: + forall m b ofs n bytes, + loadbytes m b ofs n = Some bytes -> + length bytes = nat_of_Z n. +Proof. + unfold loadbytes; intros. + destruct (range_perm_dec m b ofs (ofs + n) Readable); try congruence. + exploit inj_proj_bytes; eauto. intros. + transitivity (length (inj_bytes bytes)). + symmetry. unfold inj_bytes. apply List.map_length. + rewrite <- H0. apply getN_length. +Qed. + +Lemma getN_concat: + forall c n1 n2 p, + getN (n1 + n2)%nat p c = getN n1 p c ++ getN n2 (p + Z_of_nat n1) c. +Proof. + induction n1; intros. + simpl. decEq. omega. + rewrite inj_S. simpl. decEq. + replace (p + Zsucc (Z_of_nat n1)) with ((p + 1) + Z_of_nat n1) by omega. + auto. +Qed. + +Theorem loadbytes_concat: + forall m b ofs n1 n2 bytes1 bytes2, + loadbytes m b ofs n1 = Some bytes1 -> + loadbytes m b (ofs + n1) n2 = Some bytes2 -> + n1 >= 0 -> n2 >= 0 -> + loadbytes m b ofs (n1 + n2) = Some(bytes1 ++ bytes2). +Proof. + unfold loadbytes; intros. + destruct (range_perm_dec m b ofs (ofs + n1) Readable); try congruence. + destruct (range_perm_dec m b (ofs + n1) (ofs + n1 + n2) Readable); try congruence. + rewrite pred_dec_true. rewrite nat_of_Z_plus; auto. + rewrite getN_concat. rewrite nat_of_Z_eq; auto. + rewrite (inj_proj_bytes _ _ H). rewrite (inj_proj_bytes _ _ H0). + unfold inj_bytes. rewrite <- List.map_app. apply proj_inj_bytes. + red; intros. + assert (ofs0 < ofs + n1 \/ ofs0 >= ofs + n1) by omega. + destruct H4. apply r; omega. apply r0; omega. +Qed. + +Theorem loadbytes_split: + forall m b ofs n1 n2 bytes, + loadbytes m b ofs (n1 + n2) = Some bytes -> + n1 >= 0 -> n2 >= 0 -> + exists bytes1, exists bytes2, + loadbytes m b ofs n1 = Some bytes1 + /\ loadbytes m b (ofs + n1) n2 = Some bytes2 + /\ bytes = bytes1 ++ bytes2. +Proof. + unfold loadbytes; intros. + destruct (range_perm_dec m b ofs (ofs + (n1 + n2)) Readable); + try congruence. + rewrite nat_of_Z_plus in H; auto. rewrite getN_concat in H. + rewrite nat_of_Z_eq in H; auto. + repeat rewrite pred_dec_true. + exploit inj_proj_bytes; eauto. unfold inj_bytes. intros. + exploit list_append_map_inv; eauto. intros [l1 [l2 [P [Q R]]]]. + exists l1; exists l2; intuition. + rewrite <- P. apply proj_inj_bytes. + rewrite <- Q. apply proj_inj_bytes. + red; intros; apply r; omega. + red; intros; apply r; omega. +Qed. + +(** ** Properties related to [store] *) + +Theorem valid_access_store: + forall m1 chunk b ofs v, + valid_access m1 chunk b ofs Writable -> + { m2: mem | store chunk m1 b ofs v = Some m2 }. +Proof. + intros. econstructor. unfold store. rewrite pred_dec_true; auto. +Qed. + +Hint Local Resolve valid_access_store: mem. + +Section STORE. +Variable chunk: memory_chunk. +Variable m1: mem. +Variable b: block. +Variable ofs: Z. +Variable v: val. +Variable m2: mem. +Hypothesis STORE: store chunk m1 b ofs v = Some m2. + +Lemma store_result: + m2 = unchecked_store chunk m1 b ofs v. +Proof. + unfold store in STORE. + destruct (valid_access_dec m1 chunk b ofs Writable). + congruence. + congruence. +Qed. + +Theorem perm_store_1: + forall b' ofs' p, perm m1 b' ofs' p -> perm m2 b' ofs' p. +Proof. + intros. rewrite store_result. exact H. +Qed. + +Theorem perm_store_2: + forall b' ofs' p, perm m2 b' ofs' p -> perm m1 b' ofs' p. +Proof. + intros. rewrite store_result in H. exact H. +Qed. + +Hint Local Resolve perm_store_1 perm_store_2: mem. + +Theorem nextblock_store: + nextblock m2 = nextblock m1. +Proof. + intros. rewrite store_result. reflexivity. +Qed. + +Theorem store_valid_block_1: + forall b', valid_block m1 b' -> valid_block m2 b'. +Proof. + unfold valid_block; intros. rewrite nextblock_store; auto. +Qed. + +Theorem store_valid_block_2: + forall b', valid_block m2 b' -> valid_block m1 b'. +Proof. + unfold valid_block; intros. rewrite nextblock_store in H; auto. +Qed. + +Hint Local Resolve store_valid_block_1 store_valid_block_2: mem. + +Theorem store_valid_access_1: + forall chunk' b' ofs' p, + valid_access m1 chunk' b' ofs' p -> valid_access m2 chunk' b' ofs' p. +Proof. + intros. inv H. constructor; try red; auto with mem. +Qed. + +Theorem store_valid_access_2: + forall chunk' b' ofs' p, + valid_access m2 chunk' b' ofs' p -> valid_access m1 chunk' b' ofs' p. +Proof. + intros. inv H. constructor; try red; auto with mem. +Qed. + +Theorem store_valid_access_3: + valid_access m1 chunk b ofs Writable. +Proof. + unfold store in STORE. destruct (valid_access_dec m1 chunk b ofs Writable). + auto. + congruence. +Qed. + +Hint Local Resolve store_valid_access_1 store_valid_access_2 + store_valid_access_3: mem. + +Theorem bounds_store: + forall b', bounds m2 b' = bounds m1 b'. +Proof. + intros. rewrite store_result. simpl. auto. +Qed. + +Remark setN_other: + forall vl c p q, + (forall r, p <= r < p + Z_of_nat (length vl) -> r <> q) -> + setN vl p c q = c q. +Proof. + induction vl; intros; simpl. + auto. + simpl length in H. rewrite inj_S in H. + transitivity (update p a c q). + apply IHvl. intros. apply H. omega. + apply update_o. apply H. omega. +Qed. + +Remark setN_outside: + forall vl c p q, + q < p \/ q >= p + Z_of_nat (length vl) -> + setN vl p c q = c q. +Proof. + intros. apply setN_other. + intros. omega. +Qed. + +Remark getN_setN_same: + forall vl p c, + getN (length vl) p (setN vl p c) = vl. +Proof. + induction vl; intros; simpl. + auto. + decEq. + rewrite setN_outside. apply update_s. omega. + apply IHvl. +Qed. + +Remark getN_setN_outside: + forall vl q c n p, + p + Z_of_nat n <= q \/ q + Z_of_nat (length vl) <= p -> + getN n p (setN vl q c) = getN n p c. +Proof. + induction n; intros; simpl. + auto. + rewrite inj_S in H. decEq. + apply setN_outside. omega. + apply IHn. omega. +Qed. + +Theorem load_store_similar: + forall chunk', + size_chunk chunk' = size_chunk chunk -> + exists v', load chunk' m2 b ofs = Some v' /\ decode_encode_val v chunk chunk' v'. +Proof. + intros. + exploit (valid_access_load m2 chunk'). + eapply valid_access_compat. symmetry; eauto. eauto with mem. + intros [v' LOAD]. + exists v'; split; auto. + exploit load_result; eauto. intros B. + rewrite B. rewrite store_result; simpl. + rewrite update_s. + replace (size_chunk_nat chunk') with (length (encode_val chunk v)). + rewrite getN_setN_same. apply decode_encode_val_general. + rewrite encode_val_length. repeat rewrite size_chunk_conv in H. + apply inj_eq_rev; auto. +Qed. + +Theorem load_store_same: + Val.has_type v (type_of_chunk chunk) -> + load chunk m2 b ofs = Some (Val.load_result chunk v). +Proof. + intros. + destruct (load_store_similar chunk) as [v' [A B]]. auto. + rewrite A. decEq. eapply decode_encode_val_similar; eauto. +Qed. + +Theorem load_store_other: + forall chunk' b' ofs', + b' <> b + \/ ofs' + size_chunk chunk' <= ofs + \/ ofs + size_chunk chunk <= ofs' -> + load chunk' m2 b' ofs' = load chunk' m1 b' ofs'. +Proof. + intros. unfold load. + destruct (valid_access_dec m1 chunk' b' ofs' Readable). + rewrite pred_dec_true. + decEq. decEq. rewrite store_result; unfold unchecked_store; simpl. + unfold update. destruct (zeq b' b). subst b'. + apply getN_setN_outside. rewrite encode_val_length. repeat rewrite <- size_chunk_conv. + intuition. + auto. + eauto with mem. + rewrite pred_dec_false. auto. + eauto with mem. +Qed. + +Theorem loadbytes_store_same: + loadbytes m2 b ofs (size_chunk chunk) = + match v with + | Vundef => None + | Vint n => Some(encode_int chunk n) + | Vfloat n => Some(encode_float chunk n) + | Vptr _ _ => None + end. +Proof. + intros. + assert (valid_access m2 chunk b ofs Readable) by eauto with mem. + unfold loadbytes. rewrite pred_dec_true. rewrite store_result; simpl. + rewrite update_s. + replace (nat_of_Z (size_chunk chunk)) + with (length (encode_val chunk v)). + rewrite getN_setN_same. + destruct (size_chunk_nat_pos chunk) as [sz1 EQ]. + unfold encode_val; destruct v. + rewrite EQ; auto. + apply proj_inj_bytes. + apply proj_inj_bytes. + rewrite EQ; destruct chunk; auto. + apply encode_val_length. + apply H. +Qed. + +Theorem loadbytes_store_other: + forall b' ofs' n, + b' <> b + \/ n <= 0 + \/ ofs' + n <= ofs + \/ ofs + size_chunk chunk <= ofs' -> + loadbytes m2 b' ofs' n = loadbytes m1 b' ofs' n. +Proof. + intros. unfold loadbytes. + destruct (range_perm_dec m1 b' ofs' (ofs' + n) Readable). + rewrite pred_dec_true. + decEq. rewrite store_result; unfold unchecked_store; simpl. + unfold update. destruct (zeq b' b). subst b'. + destruct H. congruence. + destruct (zle n 0). + rewrite (nat_of_Z_neg _ z). auto. + destruct H. omegaContradiction. + apply getN_setN_outside. rewrite encode_val_length. rewrite <- size_chunk_conv. + rewrite nat_of_Z_eq. auto. omega. + auto. + red; intros. eauto with mem. + rewrite pred_dec_false. auto. + red; intro; elim n0; red; intros; eauto with mem. +Qed. + +Lemma setN_property: + forall (P: memval -> Prop) vl p q c, + (forall v, In v vl -> P v) -> + p <= q < p + Z_of_nat (length vl) -> + P(setN vl p c q). +Proof. + induction vl; intros. + simpl in H0. omegaContradiction. + simpl length in H0. rewrite inj_S in H0. simpl. + destruct (zeq p q). subst q. rewrite setN_outside. rewrite update_s. + auto with coqlib. omega. + apply IHvl. auto with coqlib. omega. +Qed. + +Lemma getN_in: + forall c q n p, + p <= q < p + Z_of_nat n -> + In (c q) (getN n p c). +Proof. + induction n; intros. + simpl in H; omegaContradiction. + rewrite inj_S in H. simpl. destruct (zeq p q). + subst q. auto. + right. apply IHn. omega. +Qed. + +Theorem load_pointer_store: + forall chunk' b' ofs' v_b v_o, + load chunk' m2 b' ofs' = Some(Vptr v_b v_o) -> + (chunk = Mint32 /\ v = Vptr v_b v_o /\ chunk' = Mint32 /\ b' = b /\ ofs' = ofs) + \/ (b' <> b \/ ofs' + size_chunk chunk' <= ofs \/ ofs + size_chunk chunk <= ofs'). +Proof. + intros. exploit load_result; eauto. rewrite store_result; simpl. + unfold update. destruct (zeq b' b); auto. subst b'. intro DEC. + destruct (zle (ofs' + size_chunk chunk') ofs); auto. + destruct (zle (ofs + size_chunk chunk) ofs'); auto. + destruct (size_chunk_nat_pos chunk) as [sz SZ]. + destruct (size_chunk_nat_pos chunk') as [sz' SZ']. + exploit decode_pointer_shape; eauto. intros [CHUNK' PSHAPE]. clear CHUNK'. + generalize (encode_val_shape chunk v). intro VSHAPE. + set (c := contents m1 b) in *. + set (c' := setN (encode_val chunk v) ofs c) in *. + destruct (zeq ofs ofs'). + +(* 1. ofs = ofs': must be same chunks and same value *) + subst ofs'. inv VSHAPE. + exploit decode_val_pointer_inv; eauto. intros [A B]. + subst chunk'. simpl in B. inv B. + generalize H4. unfold c'. rewrite <- H0. simpl. + rewrite setN_outside; try omega. rewrite update_s. intros. + exploit (encode_val_pointer_inv chunk v v_b v_o). + rewrite <- H0. subst mv1. eauto. intros [C [D E]]. + left; auto. + + destruct (zlt ofs ofs'). + +(* 2. ofs < ofs': + + ofs ofs' ofs+|chunk| + [-------------------] write + [-------------------] read + + The byte at ofs' satisfies memval_valid_cont (consequence of write). + For the read to return a pointer, it must satisfy ~memval_valid_cont. +*) + elimtype False. + assert (~memval_valid_cont (c' ofs')). + rewrite SZ' in PSHAPE. simpl in PSHAPE. inv PSHAPE. auto. + assert (memval_valid_cont (c' ofs')). + inv VSHAPE. unfold c'. rewrite <- H1. simpl. + apply setN_property. auto. + assert (length mvl = sz). + generalize (encode_val_length chunk v). rewrite <- H1. rewrite SZ. + simpl; congruence. + rewrite H4. rewrite size_chunk_conv in z0. omega. + contradiction. + +(* 3. ofs > ofs': + + ofs' ofs ofs'+|chunk'| + [-------------------] write + [----------------] read + + The byte at ofs satisfies memval_valid_first (consequence of write). + For the read to return a pointer, it must satisfy ~memval_valid_first. +*) + elimtype False. + assert (memval_valid_first (c' ofs)). + inv VSHAPE. unfold c'. rewrite <- H0. simpl. + rewrite setN_outside. rewrite update_s. auto. omega. + assert (~memval_valid_first (c' ofs)). + rewrite SZ' in PSHAPE. simpl in PSHAPE. inv PSHAPE. + apply H4. apply getN_in. rewrite size_chunk_conv in z. + rewrite SZ' in z. rewrite inj_S in z. omega. + contradiction. +Qed. + +End STORE. + +Hint Local Resolve perm_store_1 perm_store_2: mem. +Hint Local Resolve store_valid_block_1 store_valid_block_2: mem. +Hint Local Resolve store_valid_access_1 store_valid_access_2 + store_valid_access_3: mem. + +Theorem load_store_pointer_overlap: + forall chunk m1 b ofs v_b v_o m2 chunk' ofs' v, + store chunk m1 b ofs (Vptr v_b v_o) = Some m2 -> + load chunk' m2 b ofs' = Some v -> + ofs' <> ofs -> + ofs' + size_chunk chunk' > ofs -> + ofs + size_chunk chunk > ofs' -> + v = Vundef. +Proof. + intros. + exploit store_result; eauto. intro ST. + exploit load_result; eauto. intro LD. + rewrite LD; clear LD. +Opaque encode_val. + rewrite ST; simpl. + rewrite update_s. + set (c := contents m1 b). + set (c' := setN (encode_val chunk (Vptr v_b v_o)) ofs c). + destruct (decode_val_shape chunk' (getN (size_chunk_nat chunk') ofs' c')) + as [OK | VSHAPE]. + apply getN_length. + exact OK. + elimtype False. + destruct (size_chunk_nat_pos chunk) as [sz SZ]. + destruct (size_chunk_nat_pos chunk') as [sz' SZ']. + assert (ENC: encode_val chunk (Vptr v_b v_o) = list_repeat (size_chunk_nat chunk) Undef + \/ pointer_encoding_shape (encode_val chunk (Vptr v_b v_o))). + destruct chunk; try (left; reflexivity). + right. apply encode_pointer_shape. + assert (GET: getN (size_chunk_nat chunk) ofs c' = encode_val chunk (Vptr v_b v_o)). + unfold c'. rewrite <- (encode_val_length chunk (Vptr v_b v_o)). + apply getN_setN_same. + destruct (zlt ofs ofs'). + +(* ofs < ofs': + + ofs ofs' ofs+|chunk| + [-------------------] write + [-------------------] read + + The byte at ofs' is Undef or not memval_valid_first (because write of pointer). + The byte at ofs' must be memval_valid_first and not Undef (otherwise load returns Vundef). +*) + assert (memval_valid_first (c' ofs') /\ c' ofs' <> Undef). + rewrite SZ' in VSHAPE. simpl in VSHAPE. inv VSHAPE. auto. + assert (~memval_valid_first (c' ofs') \/ c' ofs' = Undef). + unfold c'. destruct ENC. + right. apply setN_property. rewrite H5. intros. eapply in_list_repeat; eauto. + rewrite encode_val_length. rewrite <- size_chunk_conv. omega. + left. revert H5. rewrite <- GET. rewrite SZ. simpl. intros. inv H5. + apply setN_property. apply H9. rewrite getN_length. + rewrite size_chunk_conv in H3. rewrite SZ in H3. rewrite inj_S in H3. omega. + intuition. + +(* ofs > ofs': + + ofs' ofs ofs'+|chunk'| + [-------------------] write + [----------------] read + + The byte at ofs is Undef or not memval_valid_cont (because write of pointer). + The byte at ofs must be memval_valid_cont and not Undef (otherwise load returns Vundef). +*) + assert (memval_valid_cont (c' ofs) /\ c' ofs <> Undef). + rewrite SZ' in VSHAPE. simpl in VSHAPE. inv VSHAPE. + apply H8. apply getN_in. rewrite size_chunk_conv in H2. + rewrite SZ' in H2. rewrite inj_S in H2. omega. + assert (~memval_valid_cont (c' ofs) \/ c' ofs = Undef). + elim ENC. + rewrite <- GET. rewrite SZ. simpl. intros. right; congruence. + rewrite <- GET. rewrite SZ. simpl. intros. inv H5. auto. + intuition. +Qed. + +Theorem load_store_pointer_mismatch: + forall chunk m1 b ofs v_b v_o m2 chunk' v, + store chunk m1 b ofs (Vptr v_b v_o) = Some m2 -> + load chunk' m2 b ofs = Some v -> + chunk <> Mint32 \/ chunk' <> Mint32 -> + v = Vundef. +Proof. + intros. + exploit store_result; eauto. intro ST. + exploit load_result; eauto. intro LD. + rewrite LD; clear LD. +Opaque encode_val. + rewrite ST; simpl. + rewrite update_s. + set (c1 := contents m1 b). + set (e := encode_val chunk (Vptr v_b v_o)). + destruct (size_chunk_nat_pos chunk) as [sz SZ]. + destruct (size_chunk_nat_pos chunk') as [sz' SZ']. + assert (match e with + | Undef :: _ => True + | Pointer _ _ _ :: _ => chunk = Mint32 + | _ => False + end). +Transparent encode_val. + unfold e, encode_val. rewrite SZ. destruct chunk; simpl; auto. + destruct e as [ | e1 el]. contradiction. + rewrite SZ'. simpl. rewrite setN_outside. rewrite update_s. + destruct e1; try contradiction. + destruct chunk'; auto. + destruct chunk'; auto. intuition. + omega. +Qed. + +Lemma store_similar_chunks: + forall chunk1 chunk2 v1 v2 m b ofs, + encode_val chunk1 v1 = encode_val chunk2 v2 -> + store chunk1 m b ofs v1 = store chunk2 m b ofs v2. +Proof. + intros. unfold store. + assert (size_chunk chunk1 = size_chunk chunk2). + repeat rewrite size_chunk_conv. + rewrite <- (encode_val_length chunk1 v1). + rewrite <- (encode_val_length chunk2 v2). + congruence. + unfold store. + destruct (valid_access_dec m chunk1 b ofs Writable). + rewrite pred_dec_true. unfold unchecked_store. congruence. + eapply valid_access_compat; eauto. + rewrite pred_dec_false; auto. + red; intro; elim n. apply valid_access_compat with chunk2; auto. +Qed. + +Theorem store_signed_unsigned_8: + forall m b ofs v, + store Mint8signed m b ofs v = store Mint8unsigned m b ofs v. +Proof. intros. apply store_similar_chunks. apply encode_val_int8_signed_unsigned. Qed. + +Theorem store_signed_unsigned_16: + forall m b ofs v, + store Mint16signed m b ofs v = store Mint16unsigned m b ofs v. +Proof. intros. apply store_similar_chunks. apply encode_val_int16_signed_unsigned. Qed. + +Theorem store_int8_zero_ext: + forall m b ofs n, + store Mint8unsigned m b ofs (Vint (Int.zero_ext 8 n)) = + store Mint8unsigned m b ofs (Vint n). +Proof. intros. apply store_similar_chunks. apply encode_val_int8_zero_ext. Qed. + +Theorem store_int8_sign_ext: + forall m b ofs n, + store Mint8signed m b ofs (Vint (Int.sign_ext 8 n)) = + store Mint8signed m b ofs (Vint n). +Proof. intros. apply store_similar_chunks. apply encode_val_int8_sign_ext. Qed. + +Theorem store_int16_zero_ext: + forall m b ofs n, + store Mint16unsigned m b ofs (Vint (Int.zero_ext 16 n)) = + store Mint16unsigned m b ofs (Vint n). +Proof. intros. apply store_similar_chunks. apply encode_val_int16_zero_ext. Qed. + +Theorem store_int16_sign_ext: + forall m b ofs n, + store Mint16signed m b ofs (Vint (Int.sign_ext 16 n)) = + store Mint16signed m b ofs (Vint n). +Proof. intros. apply store_similar_chunks. apply encode_val_int16_sign_ext. Qed. + +Theorem store_float32_truncate: + forall m b ofs n, + store Mfloat32 m b ofs (Vfloat (Float.singleoffloat n)) = + store Mfloat32 m b ofs (Vfloat n). +Proof. intros. apply store_similar_chunks. simpl. decEq. apply encode_float32_cast. Qed. + +(** ** Properties related to [alloc]. *) + +Section ALLOC. + +Variable m1: mem. +Variables lo hi: Z. +Variable m2: mem. +Variable b: block. +Hypothesis ALLOC: alloc m1 lo hi = (m2, b). + +Theorem nextblock_alloc: + nextblock m2 = Zsucc (nextblock m1). +Proof. + injection ALLOC; intros. rewrite <- H0; auto. +Qed. + +Theorem alloc_result: + b = nextblock m1. +Proof. + injection ALLOC; auto. +Qed. + +Theorem valid_block_alloc: + forall b', valid_block m1 b' -> valid_block m2 b'. +Proof. + unfold valid_block; intros. rewrite nextblock_alloc. omega. +Qed. + +Theorem fresh_block_alloc: + ~(valid_block m1 b). +Proof. + unfold valid_block. rewrite alloc_result. omega. +Qed. + +Theorem valid_new_block: + valid_block m2 b. +Proof. + unfold valid_block. rewrite alloc_result. rewrite nextblock_alloc. omega. +Qed. + +Hint Local Resolve valid_block_alloc fresh_block_alloc valid_new_block: mem. + +Theorem valid_block_alloc_inv: + forall b', valid_block m2 b' -> b' = b \/ valid_block m1 b'. +Proof. + unfold valid_block; intros. + rewrite nextblock_alloc in H. rewrite alloc_result. + unfold block; omega. +Qed. + +Theorem perm_alloc_1: + forall b' ofs p, perm m1 b' ofs p -> perm m2 b' ofs p. +Proof. + unfold perm; intros. injection ALLOC; intros. rewrite <- H1; simpl. + subst b. unfold update. destruct (zeq b' (next m1)); auto. + assert (access m1 b' ofs = false). apply next_noaccess. omega. congruence. +Qed. + +Theorem perm_alloc_2: + forall ofs, lo <= ofs < hi -> perm m2 b ofs Writable. +Proof. + unfold perm; intros. injection ALLOC; intros. rewrite <- H1; simpl. + subst b. rewrite update_s. unfold proj_sumbool. rewrite zle_true. + rewrite zlt_true. auto. omega. omega. +Qed. + +Theorem perm_alloc_3: + forall ofs p, ofs < lo \/ hi <= ofs -> ~perm m2 b ofs p. +Proof. + unfold perm; intros. injection ALLOC; intros. rewrite <- H1; simpl. + subst b. rewrite update_s. unfold proj_sumbool. + destruct H. rewrite zle_false. simpl. congruence. omega. + rewrite zlt_false. rewrite andb_false_r. congruence. omega. +Qed. + +Hint Local Resolve perm_alloc_1 perm_alloc_2 perm_alloc_3: mem. + +Theorem perm_alloc_inv: + forall b' ofs p, + perm m2 b' ofs p -> + if zeq b' b then lo <= ofs < hi else perm m1 b' ofs p. +Proof. + intros until p; unfold perm. inv ALLOC. simpl. + unfold update. destruct (zeq b' (next m1)); intros. + destruct (andb_prop _ _ H). + split; eapply proj_sumbool_true; eauto. + auto. +Qed. + +Theorem valid_access_alloc_other: + forall chunk b' ofs p, + valid_access m1 chunk b' ofs p -> + valid_access m2 chunk b' ofs p. +Proof. + intros. inv H. constructor; auto with mem. + red; auto with mem. +Qed. + +Theorem valid_access_alloc_same: + forall chunk ofs, + lo <= ofs -> ofs + size_chunk chunk <= hi -> (align_chunk chunk | ofs) -> + valid_access m2 chunk b ofs Writable. +Proof. + intros. constructor; auto with mem. + red; intros. apply perm_alloc_2. omega. +Qed. + +Hint Local Resolve valid_access_alloc_other valid_access_alloc_same: mem. + +Theorem valid_access_alloc_inv: + forall chunk b' ofs p, + valid_access m2 chunk b' ofs p -> + if eq_block b' b + then lo <= ofs /\ ofs + size_chunk chunk <= hi /\ (align_chunk chunk | ofs) + else valid_access m1 chunk b' ofs p. +Proof. + intros. inv H. + generalize (size_chunk_pos chunk); intro. + unfold eq_block. destruct (zeq b' b). subst b'. + assert (perm m2 b ofs p). apply H0. omega. + assert (perm m2 b (ofs + size_chunk chunk - 1) p). apply H0. omega. + exploit perm_alloc_inv. eexact H2. rewrite zeq_true. intro. + exploit perm_alloc_inv. eexact H3. rewrite zeq_true. intro. + intuition omega. + split; auto. red; intros. + exploit perm_alloc_inv. apply H0. eauto. rewrite zeq_false; auto. +Qed. + +Theorem bounds_alloc: + forall b', bounds m2 b' = if eq_block b' b then (lo, hi) else bounds m1 b'. +Proof. + injection ALLOC; intros. rewrite <- H; rewrite <- H0; simpl. + unfold update. auto. +Qed. + +Theorem bounds_alloc_same: + bounds m2 b = (lo, hi). +Proof. + rewrite bounds_alloc. apply dec_eq_true. +Qed. + +Theorem bounds_alloc_other: + forall b', b' <> b -> bounds m2 b' = bounds m1 b'. +Proof. + intros. rewrite bounds_alloc. apply dec_eq_false. auto. +Qed. + +Theorem load_alloc_unchanged: + forall chunk b' ofs, + valid_block m1 b' -> + load chunk m2 b' ofs = load chunk m1 b' ofs. +Proof. + intros. unfold load. + destruct (valid_access_dec m2 chunk b' ofs Readable). + exploit valid_access_alloc_inv; eauto. destruct (eq_block b' b); intros. + subst b'. elimtype False. eauto with mem. + rewrite pred_dec_true; auto. + injection ALLOC; intros. rewrite <- H2; simpl. + rewrite update_o. auto. rewrite H1. apply sym_not_equal; eauto with mem. + rewrite pred_dec_false. auto. + eauto with mem. +Qed. + +Theorem load_alloc_other: + forall chunk b' ofs v, + load chunk m1 b' ofs = Some v -> + load chunk m2 b' ofs = Some v. +Proof. + intros. rewrite <- H. apply load_alloc_unchanged. eauto with mem. +Qed. + +Theorem load_alloc_same: + forall chunk ofs v, + load chunk m2 b ofs = Some v -> + v = Vundef. +Proof. + intros. exploit load_result; eauto. intro. rewrite H0. + injection ALLOC; intros. rewrite <- H2; simpl. rewrite <- H1. + rewrite update_s. destruct chunk; reflexivity. +Qed. + +Theorem load_alloc_same': + forall chunk ofs, + lo <= ofs -> ofs + size_chunk chunk <= hi -> (align_chunk chunk | ofs) -> + load chunk m2 b ofs = Some Vundef. +Proof. + intros. assert (exists v, load chunk m2 b ofs = Some v). + apply valid_access_load. constructor; auto. + red; intros. eapply perm_implies. apply perm_alloc_2. omega. auto with mem. + destruct H2 as [v LOAD]. rewrite LOAD. decEq. + eapply load_alloc_same; eauto. +Qed. + +End ALLOC. + +Hint Local Resolve valid_block_alloc fresh_block_alloc valid_new_block: mem. +Hint Local Resolve valid_access_alloc_other valid_access_alloc_same: mem. + +(** ** Properties related to [free]. *) + +Theorem range_perm_free: + forall m1 b lo hi, + range_perm m1 b lo hi Freeable -> + { m2: mem | free m1 b lo hi = Some m2 }. +Proof. + intros; unfold free. rewrite pred_dec_true; auto. econstructor; eauto. +Qed. + +Section FREE. + +Variable m1: mem. +Variable bf: block. +Variables lo hi: Z. +Variable m2: mem. +Hypothesis FREE: free m1 bf lo hi = Some m2. + +Theorem free_range_perm: + range_perm m1 bf lo hi Writable. +Proof. + unfold free in FREE. destruct (range_perm_dec m1 bf lo hi Freeable). + auto. congruence. +Qed. + +Lemma free_result: + m2 = unchecked_free m1 bf lo hi. +Proof. + unfold free in FREE. destruct (range_perm_dec m1 bf lo hi Freeable). + congruence. congruence. +Qed. + +Theorem nextblock_free: + nextblock m2 = nextblock m1. +Proof. + rewrite free_result; reflexivity. +Qed. + +Theorem valid_block_free_1: + forall b, valid_block m1 b -> valid_block m2 b. +Proof. + intros. rewrite free_result. assumption. +Qed. + +Theorem valid_block_free_2: + forall b, valid_block m2 b -> valid_block m1 b. +Proof. + intros. rewrite free_result in H. assumption. +Qed. + +Hint Local Resolve valid_block_free_1 valid_block_free_2: mem. + +Theorem perm_free_1: + forall b ofs p, + b <> bf \/ ofs < lo \/ hi <= ofs -> + perm m1 b ofs p -> + perm m2 b ofs p. +Proof. + intros. rewrite free_result. unfold perm, unchecked_free; simpl. + unfold update. destruct (zeq b bf). subst b. + destruct (zle lo ofs); simpl. + destruct (zlt ofs hi); simpl. + elimtype False; intuition. + auto. auto. + auto. +Qed. + +Theorem perm_free_2: + forall ofs p, lo <= ofs < hi -> ~ perm m2 bf ofs p. +Proof. + intros. rewrite free_result. unfold perm, unchecked_free; simpl. + rewrite update_s. unfold proj_sumbool. rewrite zle_true. rewrite zlt_true. + simpl. congruence. omega. omega. +Qed. + +Theorem perm_free_3: + forall b ofs p, + perm m2 b ofs p -> perm m1 b ofs p. +Proof. + intros until p. rewrite free_result. unfold perm, unchecked_free; simpl. + unfold update. destruct (zeq b bf). subst b. + destruct (zle lo ofs); simpl. + destruct (zlt ofs hi); simpl. + congruence. auto. auto. + auto. +Qed. + +Theorem valid_access_free_1: + forall chunk b ofs p, + valid_access m1 chunk b ofs p -> + b <> bf \/ lo >= hi \/ ofs + size_chunk chunk <= lo \/ hi <= ofs -> + valid_access m2 chunk b ofs p. +Proof. + intros. inv H. constructor; auto with mem. + red; intros. eapply perm_free_1; eauto. + destruct (zlt lo hi). intuition. right. omega. +Qed. + +Theorem valid_access_free_2: + forall chunk ofs p, + lo < hi -> ofs + size_chunk chunk > lo -> ofs < hi -> + ~(valid_access m2 chunk bf ofs p). +Proof. + intros; red; intros. inv H2. + generalize (size_chunk_pos chunk); intros. + destruct (zlt ofs lo). + elim (perm_free_2 lo p). + omega. apply H3. omega. + elim (perm_free_2 ofs p). + omega. apply H3. omega. +Qed. + +Theorem valid_access_free_inv_1: + forall chunk b ofs p, + valid_access m2 chunk b ofs p -> + valid_access m1 chunk b ofs p. +Proof. + intros. destruct H. split; auto. + red; intros. generalize (H ofs0 H1). + rewrite free_result. unfold perm, unchecked_free; simpl. + unfold update. destruct (zeq b bf). subst b. + destruct (zle lo ofs0); simpl. + destruct (zlt ofs0 hi); simpl. + congruence. auto. auto. auto. +Qed. + +Theorem valid_access_free_inv_2: + forall chunk ofs p, + valid_access m2 chunk bf ofs p -> + lo >= hi \/ ofs + size_chunk chunk <= lo \/ hi <= ofs. +Proof. + intros. + destruct (zlt lo hi); auto. + destruct (zle (ofs + size_chunk chunk) lo); auto. + destruct (zle hi ofs); auto. + elim (valid_access_free_2 chunk ofs p); auto. omega. +Qed. + +Theorem bounds_free: + forall b, bounds m2 b = bounds m1 b. +Proof. + intros. rewrite free_result; simpl. auto. +Qed. + +Theorem load_free: + forall chunk b ofs, + b <> bf \/ lo >= hi \/ ofs + size_chunk chunk <= lo \/ hi <= ofs -> + load chunk m2 b ofs = load chunk m1 b ofs. +Proof. + intros. unfold load. + destruct (valid_access_dec m2 chunk b ofs Readable). + rewrite pred_dec_true. + rewrite free_result; auto. + apply valid_access_free_inv_1; auto. + rewrite pred_dec_false; auto. + red; intro; elim n. eapply valid_access_free_1; eauto. +Qed. + +End FREE. + +Hint Local Resolve valid_block_free_1 valid_block_free_2 + perm_free_1 perm_free_2 perm_free_3 + valid_access_free_1 valid_access_free_inv_1: mem. + +(** * Generic injections *) + +(** A memory state [m1] generically injects into another memory state [m2] via the + memory injection [f] if the following conditions hold: +- each access in [m2] that corresponds to a valid access in [m1] + is itself valid; +- the memory value associated in [m1] to an accessible address + must inject into [m2]'s memory value at the corersponding address. +*) + +Record mem_inj (f: meminj) (m1 m2: mem) : Prop := + mk_mem_inj { + mi_access: + forall b1 b2 delta chunk ofs p, + f b1 = Some(b2, delta) -> + valid_access m1 chunk b1 ofs p -> + valid_access m2 chunk b2 (ofs + delta) p; + mi_memval: + forall b1 ofs b2 delta, + f b1 = Some(b2, delta) -> + perm m1 b1 ofs Nonempty -> + memval_inject f (m1.(contents) b1 ofs) (m2.(contents) b2 (ofs + delta)) + }. + +(** Preservation of permissions *) + +Lemma perm_inj: + forall f m1 m2 b1 ofs p b2 delta, + mem_inj f m1 m2 -> + perm m1 b1 ofs p -> + f b1 = Some(b2, delta) -> + perm m2 b2 (ofs + delta) p. +Proof. + intros. + assert (valid_access m1 Mint8unsigned b1 ofs p). + split. red; intros. simpl in H2. replace ofs0 with ofs by omega. auto. + simpl. apply Zone_divide. + exploit mi_access; eauto. intros [A B]. + apply A. simpl; omega. +Qed. + +(** Preservation of loads. *) + +Lemma getN_inj: + forall f m1 m2 b1 b2 delta, + mem_inj f m1 m2 -> + f b1 = Some(b2, delta) -> + forall n ofs, + range_perm m1 b1 ofs (ofs + Z_of_nat n) Readable -> + list_forall2 (memval_inject f) + (getN n ofs (m1.(contents) b1)) + (getN n (ofs + delta) (m2.(contents) b2)). +Proof. + induction n; intros; simpl. + constructor. + rewrite inj_S in H1. + constructor. + eapply mi_memval; eauto. apply H1. omega. + replace (ofs + delta + 1) with ((ofs + 1) + delta) by omega. + apply IHn. red; intros; apply H1; omega. +Qed. + +Lemma load_inj: + forall f m1 m2 chunk b1 ofs b2 delta v1, + mem_inj f m1 m2 -> + load chunk m1 b1 ofs = Some v1 -> + f b1 = Some (b2, delta) -> + exists v2, load chunk m2 b2 (ofs + delta) = Some v2 /\ val_inject f v1 v2. +Proof. + intros. + exists (decode_val chunk (getN (size_chunk_nat chunk) (ofs + delta) (m2.(contents) b2))). + split. unfold load. apply pred_dec_true. + eapply mi_access; eauto with mem. + exploit load_result; eauto. intro. rewrite H2. + apply decode_val_inject. apply getN_inj; auto. + rewrite <- size_chunk_conv. exploit load_valid_access; eauto. intros [A B]. auto. +Qed. + +(** Preservation of stores. *) + +Lemma setN_inj: + forall (access: Z -> Prop) delta f vl1 vl2, + list_forall2 (memval_inject f) vl1 vl2 -> + forall p c1 c2, + (forall q, access q -> memval_inject f (c1 q) (c2 (q + delta))) -> + (forall q, access q -> memval_inject f ((setN vl1 p c1) q) + ((setN vl2 (p + delta) c2) (q + delta))). +Proof. + induction 1; intros; simpl. + auto. + replace (p + delta + 1) with ((p + 1) + delta) by omega. + apply IHlist_forall2; auto. + intros. unfold update at 1. destruct (zeq q0 p). subst q0. + rewrite update_s. auto. + rewrite update_o. auto. omega. +Qed. + +Definition meminj_no_overlap (f: meminj) (m: mem) : Prop := + forall b1 b1' delta1 b2 b2' delta2, + b1 <> b2 -> + f b1 = Some (b1', delta1) -> + f b2 = Some (b2', delta2) -> + b1' <> b2' +(* + \/ low_bound m b1 >= high_bound m b1 + \/ low_bound m b2 >= high_bound m b2 *) + \/ high_bound m b1 + delta1 <= low_bound m b2 + delta2 + \/ high_bound m b2 + delta2 <= low_bound m b1 + delta1. + +Lemma meminj_no_overlap_perm: + forall f m b1 b1' delta1 b2 b2' delta2 ofs1 ofs2, + meminj_no_overlap f m -> + b1 <> b2 -> + f b1 = Some (b1', delta1) -> + f b2 = Some (b2', delta2) -> + perm m b1 ofs1 Nonempty -> + perm m b2 ofs2 Nonempty -> + b1' <> b2' \/ ofs1 + delta1 <> ofs2 + delta2. +Proof. + intros. exploit H; eauto. intro. + exploit perm_in_bounds. eexact H3. intro. + exploit perm_in_bounds. eexact H4. intro. + destruct H5. auto. right. omega. +Qed. + +Lemma store_mapped_inj: + forall f chunk m1 b1 ofs v1 n1 m2 b2 delta v2, + mem_inj f m1 m2 -> + store chunk m1 b1 ofs v1 = Some n1 -> + meminj_no_overlap f m1 -> + f b1 = Some (b2, delta) -> + val_inject f v1 v2 -> + exists n2, + store chunk m2 b2 (ofs + delta) v2 = Some n2 + /\ mem_inj f n1 n2. +Proof. + intros. inversion H. + assert (valid_access m2 chunk b2 (ofs + delta) Writable). + eapply mi_access0; eauto with mem. + destruct (valid_access_store _ _ _ _ v2 H4) as [n2 STORE]. + exists n2; split. eauto. + constructor. +(* access *) + eauto with mem. +(* contents *) + intros. + assert (perm m1 b0 ofs0 Readable). eapply perm_store_2; eauto. + rewrite (store_result _ _ _ _ _ _ H0). + rewrite (store_result _ _ _ _ _ _ STORE). + unfold unchecked_store; simpl. unfold update. + destruct (zeq b0 b1). subst b0. + (* block = b1, block = b2 *) + assert (b3 = b2) by congruence. subst b3. + assert (delta0 = delta) by congruence. subst delta0. + rewrite zeq_true. + apply setN_inj with (access := fun ofs => perm m1 b1 ofs Nonempty). + apply encode_val_inject; auto. auto. auto. + destruct (zeq b3 b2). subst b3. + (* block <> b1, block = b2 *) + rewrite setN_other. auto. + rewrite encode_val_length. rewrite <- size_chunk_conv. intros. + assert (b2 <> b2 \/ ofs0 + delta0 <> (r - delta) + delta). + eapply meminj_no_overlap_perm; eauto. + exploit store_valid_access_3. eexact H0. intros [A B]. + eapply perm_implies. apply A. omega. auto with mem. + destruct H9. congruence. omega. + (* block <> b1, block <> b2 *) + eauto. +Qed. + +Lemma store_unmapped_inj: + forall f chunk m1 b1 ofs v1 n1 m2, + mem_inj f m1 m2 -> + store chunk m1 b1 ofs v1 = Some n1 -> + f b1 = None -> + mem_inj f n1 m2. +Proof. + intros. inversion H. + constructor. +(* access *) + eauto with mem. +(* contents *) + intros. + rewrite (store_result _ _ _ _ _ _ H0). + unfold unchecked_store; simpl. rewrite update_o. eauto with mem. + congruence. +Qed. + +Lemma store_outside_inj: + forall f m1 m2 chunk b ofs v m2', + mem_inj f m1 m2 -> + (forall b' delta ofs', + f b' = Some(b, delta) -> + perm m1 b' ofs' Nonempty -> + ofs' + delta < ofs \/ ofs' + delta >= ofs + size_chunk chunk) -> + store chunk m2 b ofs v = Some m2' -> + mem_inj f m1 m2'. +Proof. + intros. inversion H. constructor. +(* access *) + eauto with mem. +(* contents *) + intros. + rewrite (store_result _ _ _ _ _ _ H1). + unfold unchecked_store; simpl. unfold update. destruct (zeq b2 b). subst b2. + rewrite setN_outside. auto. + rewrite encode_val_length. rewrite <- size_chunk_conv. + eapply H0; eauto. + eauto with mem. +Qed. + +(** Preservation of allocations *) + +Lemma alloc_right_inj: + forall f m1 m2 lo hi b2 m2', + mem_inj f m1 m2 -> + alloc m2 lo hi = (m2', b2) -> + mem_inj f m1 m2'. +Proof. + intros. injection H0. intros NEXT MEM. + inversion H. constructor. +(* access *) + intros. eauto with mem. +(* contents *) + intros. + assert (valid_access m2 Mint8unsigned b0 (ofs + delta) Nonempty). + eapply mi_access0; eauto. + split. simpl. red; intros. assert (ofs0 = ofs) by omega. congruence. + simpl. apply Zone_divide. + assert (valid_block m2 b0) by eauto with mem. + rewrite <- MEM; simpl. rewrite update_o. eauto with mem. + rewrite NEXT. apply sym_not_equal. eauto with mem. +Qed. + +Lemma alloc_left_unmapped_inj: + forall f m1 m2 lo hi m1' b1, + mem_inj f m1 m2 -> + alloc m1 lo hi = (m1', b1) -> + f b1 = None -> + mem_inj f m1' m2. +Proof. + intros. inversion H. constructor. +(* access *) + unfold update; intros. + exploit valid_access_alloc_inv; eauto. unfold eq_block. intros. + destruct (zeq b0 b1). congruence. eauto. +(* contents *) + injection H0; intros NEXT MEM. intros. + rewrite <- MEM; simpl. rewrite NEXT. unfold update. + exploit perm_alloc_inv; eauto. intros. + destruct (zeq b0 b1). constructor. eauto. +Qed. + +Definition inj_offset_aligned (delta: Z) (size: Z) : Prop := + forall chunk, size_chunk chunk <= size -> (align_chunk chunk | delta). + +Lemma alloc_left_mapped_inj: + forall f m1 m2 lo hi m1' b1 b2 delta, + mem_inj f m1 m2 -> + alloc m1 lo hi = (m1', b1) -> + valid_block m2 b2 -> + inj_offset_aligned delta (hi-lo) -> + (forall ofs p, lo <= ofs < hi -> perm m2 b2 (ofs + delta) p) -> + f b1 = Some(b2, delta) -> + mem_inj f m1' m2. +Proof. + intros. inversion H. constructor. +(* access *) + intros. + exploit valid_access_alloc_inv; eauto. unfold eq_block. intros. + destruct (zeq b0 b1). subst b0. rewrite H4 in H5. inversion H5; clear H5; subst b3 delta0. + split. red; intros. + replace ofs0 with ((ofs0 - delta) + delta) by omega. + apply H3. omega. + destruct H6. apply Zdivide_plus_r. auto. apply H2. omega. + eauto. +(* contents *) + injection H0; intros NEXT MEM. + intros. rewrite <- MEM; simpl. rewrite NEXT. unfold update. + exploit perm_alloc_inv; eauto. intros. + destruct (zeq b0 b1). constructor. eauto. +Qed. + +Lemma free_left_inj: + forall f m1 m2 b lo hi m1', + mem_inj f m1 m2 -> + free m1 b lo hi = Some m1' -> + mem_inj f m1' m2. +Proof. + intros. exploit free_result; eauto. intro FREE. inversion H. constructor. +(* access *) + intros. eauto with mem. +(* contents *) + intros. rewrite FREE; simpl. eauto with mem. +Qed. + +Lemma free_right_inj: + forall f m1 m2 b lo hi m2', + mem_inj f m1 m2 -> + free m2 b lo hi = Some m2' -> + (forall b1 delta ofs p, + f b1 = Some(b, delta) -> perm m1 b1 ofs p -> + lo <= ofs + delta < hi -> False) -> + mem_inj f m1 m2'. +Proof. + intros. exploit free_result; eauto. intro FREE. inversion H. constructor. +(* access *) + intros. exploit mi_access0; eauto. intros [RG AL]. split; auto. + red; intros. eapply perm_free_1; eauto. + destruct (zeq b2 b); auto. subst b. right. + destruct (zlt ofs0 lo); auto. destruct (zle hi ofs0); auto. + elimtype False. eapply H1 with (ofs := ofs0 - delta). eauto. + apply H3. omega. omega. +(* contents *) + intros. rewrite FREE; simpl. eauto. +Qed. + +(** * Memory extensions *) + +(** A store [m2] extends a store [m1] if [m2] can be obtained from [m1] + by increasing the sizes of the memory blocks of [m1] (decreasing + the low bounds, increasing the high bounds), and replacing some of + the [Vundef] values stored in [m1] by more defined values stored + in [m2] at the same locations. *) + +Record extends_ (m1 m2: mem) : Prop := + mk_extends { + mext_next: nextblock m1 = nextblock m2; + mext_inj: mem_inj inject_id m1 m2 +(* + mext_bounds: forall b, low_bound m2 b <= low_bound m1 b /\ high_bound m1 b <= high_bound m2 b +*) + }. + +Definition extends := extends_. + +Theorem extends_refl: + forall m, extends m m. +Proof. + intros. constructor. auto. constructor. + intros. unfold inject_id in H; inv H. replace (ofs + 0) with ofs by omega. auto. + intros. unfold inject_id in H; inv H. replace (ofs + 0) with ofs by omega. + apply memval_inject_id. +(* intros. omega. *) +Qed. + +Theorem load_extends: + forall chunk m1 m2 b ofs v1, + extends m1 m2 -> + load chunk m1 b ofs = Some v1 -> + exists v2, load chunk m2 b ofs = Some v2 /\ Val.lessdef v1 v2. +Proof. + intros. inv H. exploit load_inj; eauto. unfold inject_id; reflexivity. + intros [v2 [A B]]. exists v2; split. + replace (ofs + 0) with ofs in A by omega. auto. + rewrite val_inject_id in B. auto. +Qed. + +Theorem loadv_extends: + forall chunk m1 m2 addr1 addr2 v1, + extends m1 m2 -> + loadv chunk m1 addr1 = Some v1 -> + Val.lessdef addr1 addr2 -> + exists v2, loadv chunk m2 addr2 = Some v2 /\ Val.lessdef v1 v2. +Proof. + unfold loadv; intros. inv H1. + destruct addr2; try congruence. eapply load_extends; eauto. + congruence. +Qed. + +Theorem store_within_extends: + forall chunk m1 m2 b ofs v1 m1' v2, + extends m1 m2 -> + store chunk m1 b ofs v1 = Some m1' -> + Val.lessdef v1 v2 -> + exists m2', + store chunk m2 b ofs v2 = Some m2' + /\ extends m1' m2'. +Proof. + intros. inversion H. + exploit store_mapped_inj; eauto. + unfold inject_id; red; intros. inv H3; inv H4. auto. + unfold inject_id; reflexivity. + rewrite val_inject_id. eauto. + intros [m2' [A B]]. + exists m2'; split. + replace (ofs + 0) with ofs in A by omega. auto. + split; auto. + rewrite (nextblock_store _ _ _ _ _ _ H0). + rewrite (nextblock_store _ _ _ _ _ _ A). + auto. +(* + intros. + rewrite (bounds_store _ _ _ _ _ _ H0). + rewrite (bounds_store _ _ _ _ _ _ A). + auto. +*) +Qed. + +Theorem store_outside_extends: + forall chunk m1 m2 b ofs v m2', + extends m1 m2 -> + store chunk m2 b ofs v = Some m2' -> + ofs + size_chunk chunk <= low_bound m1 b \/ high_bound m1 b <= ofs -> + extends m1 m2'. +Proof. + intros. inversion H. constructor. + rewrite (nextblock_store _ _ _ _ _ _ H0). auto. + eapply store_outside_inj; eauto. + unfold inject_id; intros. inv H2. + exploit perm_in_bounds; eauto. omega. +(* + intros. + rewrite (bounds_store _ _ _ _ _ _ H0). auto. +*) +Qed. + +Theorem storev_extends: + forall chunk m1 m2 addr1 v1 m1' addr2 v2, + extends m1 m2 -> + storev chunk m1 addr1 v1 = Some m1' -> + Val.lessdef addr1 addr2 -> + Val.lessdef v1 v2 -> + exists m2', + storev chunk m2 addr2 v2 = Some m2' + /\ extends m1' m2'. +Proof. + unfold storev; intros. inv H1. + destruct addr2; try congruence. eapply store_within_extends; eauto. + congruence. +Qed. + +Theorem alloc_extends: + forall m1 m2 lo1 hi1 b m1' lo2 hi2, + extends m1 m2 -> + alloc m1 lo1 hi1 = (m1', b) -> + lo2 <= lo1 -> hi1 <= hi2 -> + exists m2', + alloc m2 lo2 hi2 = (m2', b) + /\ extends m1' m2'. +Proof. + intros. inv H. + case_eq (alloc m2 lo2 hi2); intros m2' b' ALLOC. + assert (b' = b). + rewrite (alloc_result _ _ _ _ _ H0). + rewrite (alloc_result _ _ _ _ _ ALLOC). + auto. + subst b'. + exists m2'; split; auto. + constructor. + rewrite (nextblock_alloc _ _ _ _ _ H0). + rewrite (nextblock_alloc _ _ _ _ _ ALLOC). + congruence. + eapply alloc_left_mapped_inj with (m1 := m1) (m2 := m2') (b2 := b) (delta := 0); eauto. + eapply alloc_right_inj; eauto. + eauto with mem. + red. intros. apply Zdivide_0. + intros. eapply perm_alloc_2; eauto. omega. +(* + intros. destruct (zeq b0 b). subst b0. + rewrite (bounds_alloc_same _ _ _ _ _ H0). + rewrite (bounds_alloc_same _ _ _ _ _ ALLOC). + simpl. auto. + rewrite (bounds_alloc_other _ _ _ _ _ H0); auto. + rewrite (bounds_alloc_other _ _ _ _ _ ALLOC); auto. +*) +Qed. + +Theorem free_left_extends: + forall m1 m2 b lo hi m1', + extends m1 m2 -> + free m1 b lo hi = Some m1' -> + extends m1' m2. +Proof. + intros. inv H. constructor. + rewrite (nextblock_free _ _ _ _ _ H0). auto. + eapply free_left_inj; eauto. +(* + intros. rewrite (bounds_free _ _ _ _ _ H0). auto. +*) +Qed. + +Theorem free_right_extends: + forall m1 m2 b lo hi m2', + extends m1 m2 -> + free m2 b lo hi = Some m2' -> + (forall ofs p, lo <= ofs < hi -> ~perm m1 b ofs p) -> + extends m1 m2'. +Proof. + intros. inv H. constructor. + rewrite (nextblock_free _ _ _ _ _ H0). auto. + eapply free_right_inj; eauto. + unfold inject_id; intros. inv H. + elim (H1 ofs p); auto. omega. +(* + intros. rewrite (bounds_free _ _ _ _ _ H0). auto. +*) +Qed. + +Theorem free_parallel_extends: + forall m1 m2 b lo hi m1', + extends m1 m2 -> + free m1 b lo hi = Some m1' -> + exists m2', + free m2 b lo hi = Some m2' + /\ extends m1' m2'. +Proof. + intros. inversion H. + assert ({ m2': mem | free m2 b lo hi = Some m2' }). + apply range_perm_free. red; intros. + replace ofs with (ofs + 0) by omega. + eapply perm_inj with (b1 := b); eauto. + eapply free_range_perm; eauto. + destruct X as [m2' FREE]. exists m2'; split; auto. + inv H. constructor. + rewrite (nextblock_free _ _ _ _ _ H0). + rewrite (nextblock_free _ _ _ _ _ FREE). auto. + eapply free_right_inj with (m1 := m1'); eauto. + eapply free_left_inj; eauto. + unfold inject_id; intros. inv H. + assert (~perm m1' b ofs p). eapply perm_free_2; eauto. omega. + contradiction. +(* + intros. + rewrite (bounds_free _ _ _ _ _ H0). + rewrite (bounds_free _ _ _ _ _ FREE). + auto. +*) +Qed. + +Theorem valid_block_extends: + forall m1 m2 b, + extends m1 m2 -> + (valid_block m1 b <-> valid_block m2 b). +Proof. + intros. inv H. unfold valid_block. rewrite mext_next0. omega. +Qed. + +Theorem perm_extends: + forall m1 m2 b ofs p, + extends m1 m2 -> perm m1 b ofs p -> perm m2 b ofs p. +Proof. + intros. inv H. replace ofs with (ofs + 0) by omega. + eapply perm_inj; eauto. +Qed. + +Theorem valid_access_extends: + forall m1 m2 chunk b ofs p, + extends m1 m2 -> valid_access m1 chunk b ofs p -> valid_access m2 chunk b ofs p. +Proof. + intros. inv H. replace ofs with (ofs + 0) by omega. + eapply mi_access; eauto. auto. +Qed. + +(* +Theorem bounds_extends: + forall m1 m2 b, + extends m1 m2 -> low_bound m2 b <= low_bound m1 b /\ high_bound m1 b <= high_bound m2 b. +Proof. + intros. inv H. auto. +Qed. +*) + +(** * Memory injections *) + +(** A memory state [m1] injects into another memory state [m2] via the + memory injection [f] if the following conditions hold: +- each access in [m2] that corresponds to a valid access in [m1] + is itself valid; +- the memory value associated in [m1] to an accessible address + must inject into [m2]'s memory value at the corersponding address; +- unallocated blocks in [m1] must be mapped to [None] by [f]; +- if [f b = Some(b', delta)], [b'] must be valid in [m2]; +- distinct blocks in [m1] are mapped to non-overlapping sub-blocks in [m2]; +- the sizes of [m2]'s blocks are representable with signed machine integers; +- the offsets [delta] are representable with signed machine integers. +*) + +Record inject_ (f: meminj) (m1 m2: mem) : Prop := + mk_inject { + mi_inj: + mem_inj f m1 m2; + mi_freeblocks: + forall b, ~(valid_block m1 b) -> f b = None; + mi_mappedblocks: + forall b b' delta, f b = Some(b', delta) -> valid_block m2 b'; + mi_no_overlap: + meminj_no_overlap f m1; + mi_range_offset: + forall b b' delta, + f b = Some(b', delta) -> + Int.min_signed <= delta <= Int.max_signed; + mi_range_block: + forall b b' delta, + f b = Some(b', delta) -> + delta = 0 \/ + (Int.min_signed <= low_bound m2 b' /\ high_bound m2 b' <= Int.max_signed) + }. + +Definition inject := inject_. + +Hint Local Resolve mi_mappedblocks mi_range_offset: mem. + +(** Preservation of access validity and pointer validity *) + +Theorem valid_block_inject_1: + forall f m1 m2 b1 b2 delta, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + valid_block m1 b1. +Proof. + intros. inv H. destruct (zlt b1 (nextblock m1)). auto. + assert (f b1 = None). eapply mi_freeblocks; eauto. congruence. +Qed. + +Theorem valid_block_inject_2: + forall f m1 m2 b1 b2 delta, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + valid_block m2 b2. +Proof. + intros. eapply mi_mappedblocks; eauto. +Qed. + +Hint Local Resolve valid_block_inject_1 valid_block_inject_2: mem. + +Theorem perm_inject: + forall f m1 m2 b1 b2 delta ofs p, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + perm m1 b1 ofs p -> perm m2 b2 (ofs + delta) p. +Proof. + intros. inv H0. eapply perm_inj; eauto. +Qed. + +Theorem valid_access_inject: + forall f m1 m2 chunk b1 ofs b2 delta p, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + valid_access m1 chunk b1 ofs p -> + valid_access m2 chunk b2 (ofs + delta) p. +Proof. + intros. eapply mi_access; eauto. apply mi_inj; auto. +Qed. + +Theorem valid_pointer_inject: + forall f m1 m2 b1 ofs b2 delta, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + valid_pointer m1 b1 ofs = true -> + valid_pointer m2 b2 (ofs + delta) = true. +Proof. + intros. + rewrite valid_pointer_valid_access in H1. + rewrite valid_pointer_valid_access. + eapply valid_access_inject; eauto. +Qed. + +(** The following lemmas establish the absence of machine integer overflow + during address computations. *) + +Lemma address_inject: + forall f m1 m2 b1 ofs1 b2 delta, + inject f m1 m2 -> + perm m1 b1 (Int.signed ofs1) Nonempty -> + f b1 = Some (b2, delta) -> + Int.signed (Int.add ofs1 (Int.repr delta)) = Int.signed ofs1 + delta. +Proof. + intros. + exploit perm_inject; eauto. intro A. + exploit perm_in_bounds. eexact A. intros [B C]. + exploit mi_range_block; eauto. intros [D | [E F]]. + subst delta. rewrite Int.add_zero. omega. + rewrite Int.add_signed. + repeat rewrite Int.signed_repr. auto. + eapply mi_range_offset; eauto. + omega. + eapply mi_range_offset; eauto. +Qed. + +Lemma address_inject': + forall f m1 m2 chunk b1 ofs1 b2 delta, + inject f m1 m2 -> + valid_access m1 chunk b1 (Int.signed ofs1) Nonempty -> + f b1 = Some (b2, delta) -> + Int.signed (Int.add ofs1 (Int.repr delta)) = Int.signed ofs1 + delta. +Proof. + intros. destruct H0. eapply address_inject; eauto. + apply H0. generalize (size_chunk_pos chunk). omega. +Qed. + +Theorem valid_pointer_inject_no_overflow: + forall f m1 m2 b ofs b' x, + inject f m1 m2 -> + valid_pointer m1 b (Int.signed ofs) = true -> + f b = Some(b', x) -> + Int.min_signed <= Int.signed ofs + Int.signed (Int.repr x) <= Int.max_signed. +Proof. + intros. rewrite valid_pointer_valid_access in H0. + exploit address_inject'; eauto. intros. + rewrite Int.signed_repr; eauto. + rewrite <- H2. apply Int.signed_range. + eapply mi_range_offset; eauto. +Qed. + +Theorem valid_pointer_inject_val: + forall f m1 m2 b ofs b' ofs', + inject f m1 m2 -> + valid_pointer m1 b (Int.signed ofs) = true -> + val_inject f (Vptr b ofs) (Vptr b' ofs') -> + valid_pointer m2 b' (Int.signed ofs') = true. +Proof. + intros. inv H1. + exploit valid_pointer_inject_no_overflow; eauto. intro NOOV. + rewrite Int.add_signed. rewrite Int.signed_repr; auto. + rewrite Int.signed_repr. + eapply valid_pointer_inject; eauto. + eapply mi_range_offset; eauto. +Qed. + +Theorem inject_no_overlap: + forall f m1 m2 b1 b2 b1' b2' delta1 delta2 ofs1 ofs2, + inject f m1 m2 -> + b1 <> b2 -> + f b1 = Some (b1', delta1) -> + f b2 = Some (b2', delta2) -> + perm m1 b1 ofs1 Nonempty -> + perm m1 b2 ofs2 Nonempty -> + b1' <> b2' \/ ofs1 + delta1 <> ofs2 + delta2. +Proof. + intros. inv H. eapply meminj_no_overlap_perm; eauto. +Qed. + +Theorem different_pointers_inject: + forall f m m' b1 ofs1 b2 ofs2 b1' delta1 b2' delta2, + inject f m m' -> + b1 <> b2 -> + valid_pointer m b1 (Int.signed ofs1) = true -> + valid_pointer m b2 (Int.signed ofs2) = true -> + f b1 = Some (b1', delta1) -> + f b2 = Some (b2', delta2) -> + b1' <> b2' \/ + Int.signed (Int.add ofs1 (Int.repr delta1)) <> + Int.signed (Int.add ofs2 (Int.repr delta2)). +Proof. + intros. + rewrite valid_pointer_valid_access in H1. + rewrite valid_pointer_valid_access in H2. + rewrite (address_inject' _ _ _ _ _ _ _ _ H H1 H3). + rewrite (address_inject' _ _ _ _ _ _ _ _ H H2 H4). + inv H1. simpl in H5. inv H2. simpl in H1. + eapply meminj_no_overlap_perm. + eapply mi_no_overlap; eauto. eauto. eauto. eauto. + apply (H5 (Int.signed ofs1)). omega. + apply (H1 (Int.signed ofs2)). omega. +Qed. + +(** Preservation of loads *) + +Theorem load_inject: + forall f m1 m2 chunk b1 ofs b2 delta v1, + inject f m1 m2 -> + load chunk m1 b1 ofs = Some v1 -> + f b1 = Some (b2, delta) -> + exists v2, load chunk m2 b2 (ofs + delta) = Some v2 /\ val_inject f v1 v2. +Proof. + intros. inv H. eapply load_inj; eauto. +Qed. + +Theorem loadv_inject: + forall f m1 m2 chunk a1 a2 v1, + inject f m1 m2 -> + loadv chunk m1 a1 = Some v1 -> + val_inject f a1 a2 -> + exists v2, loadv chunk m2 a2 = Some v2 /\ val_inject f v1 v2. +Proof. + intros. inv H1; simpl in H0; try discriminate. + exploit load_inject; eauto. intros [v2 [LOAD INJ]]. + exists v2; split; auto. simpl. + replace (Int.signed (Int.add ofs1 (Int.repr delta))) + with (Int.signed ofs1 + delta). + auto. symmetry. eapply address_inject'; eauto with mem. +Qed. + +(** Preservation of stores *) + +Theorem store_mapped_inject: + forall f chunk m1 b1 ofs v1 n1 m2 b2 delta v2, + inject f m1 m2 -> + store chunk m1 b1 ofs v1 = Some n1 -> + f b1 = Some (b2, delta) -> + val_inject f v1 v2 -> + exists n2, + store chunk m2 b2 (ofs + delta) v2 = Some n2 + /\ inject f n1 n2. +Proof. + intros. inversion H. + exploit store_mapped_inj; eauto. intros [n2 [STORE MI]]. + exists n2; split. eauto. constructor. +(* inj *) + auto. +(* freeblocks *) + eauto with mem. +(* mappedblocks *) + eauto with mem. +(* no overlap *) + red; intros. + repeat rewrite (bounds_store _ _ _ _ _ _ H0). + eauto. +(* range offset *) + eauto. +(* range blocks *) + intros. rewrite (bounds_store _ _ _ _ _ _ STORE). eauto. +Qed. + +Theorem store_unmapped_inject: + forall f chunk m1 b1 ofs v1 n1 m2, + inject f m1 m2 -> + store chunk m1 b1 ofs v1 = Some n1 -> + f b1 = None -> + inject f n1 m2. +Proof. + intros. inversion H. + constructor. +(* inj *) + eapply store_unmapped_inj; eauto. +(* freeblocks *) + eauto with mem. +(* mappedblocks *) + eauto with mem. +(* no overlap *) + red; intros. + repeat rewrite (bounds_store _ _ _ _ _ _ H0). + eauto. +(* range offset *) + eauto. +(* range blocks *) + auto. +Qed. + +Theorem store_outside_inject: + forall f m1 m2 chunk b ofs v m2', + inject f m1 m2 -> + (forall b' delta, + f b' = Some(b, delta) -> + high_bound m1 b' + delta <= ofs + \/ ofs + size_chunk chunk <= low_bound m1 b' + delta) -> + store chunk m2 b ofs v = Some m2' -> + inject f m1 m2'. +Proof. + intros. inversion H. constructor. +(* inj *) + eapply store_outside_inj; eauto. + intros. exploit perm_in_bounds; eauto. intro. + exploit H0; eauto. intro. omega. +(* freeblocks *) + auto. +(* mappedblocks *) + eauto with mem. +(* no overlap *) + auto. +(* range offset *) + auto. +(* rang blocks *) + intros. rewrite (bounds_store _ _ _ _ _ _ H1). eauto. +Qed. + +Theorem storev_mapped_inject: + forall f chunk m1 a1 v1 n1 m2 a2 v2, + inject f m1 m2 -> + storev chunk m1 a1 v1 = Some n1 -> + val_inject f a1 a2 -> + val_inject f v1 v2 -> + exists n2, + storev chunk m2 a2 v2 = Some n2 /\ inject f n1 n2. +Proof. + intros. inv H1; simpl in H0; try discriminate. + simpl. replace (Int.signed (Int.add ofs1 (Int.repr delta))) + with (Int.signed ofs1 + delta). + eapply store_mapped_inject; eauto. + symmetry. eapply address_inject'; eauto with mem. +Qed. + +(* Preservation of allocations *) + +Theorem alloc_right_inject: + forall f m1 m2 lo hi b2 m2', + inject f m1 m2 -> + alloc m2 lo hi = (m2', b2) -> + inject f m1 m2'. +Proof. + intros. injection H0. intros NEXT MEM. + inversion H. constructor. +(* inj *) + eapply alloc_right_inj; eauto. +(* freeblocks *) + auto. +(* mappedblocks *) + eauto with mem. +(* no overlap *) + auto. +(* range offset *) + auto. +(* range block *) + intros. rewrite (bounds_alloc_other _ _ _ _ _ H0). eauto. + eapply valid_not_valid_diff; eauto with mem. +Qed. + +Theorem alloc_left_unmapped_inject: + forall f m1 m2 lo hi m1' b1, + inject f m1 m2 -> + alloc m1 lo hi = (m1', b1) -> + exists f', + inject f' m1' m2 + /\ inject_incr f f' + /\ f' b1 = None + /\ (forall b, b <> b1 -> f' b = f b). +Proof. + intros. inversion H. + assert (inject_incr f (update b1 None f)). + red; unfold update; intros. destruct (zeq b b1). subst b. + assert (f b1 = None). eauto with mem. congruence. + auto. + assert (mem_inj (update b1 None f) m1 m2). + inversion mi_inj0; constructor; eauto with mem. + unfold update; intros. destruct (zeq b0 b1). congruence. eauto. + unfold update; intros. destruct (zeq b0 b1). congruence. + apply memval_inject_incr with f; auto. + exists (update b1 None f); split. constructor. +(* inj *) + eapply alloc_left_unmapped_inj; eauto. apply update_s. +(* freeblocks *) + intros. unfold update. destruct (zeq b b1). auto. + apply mi_freeblocks0. red; intro; elim H3. eauto with mem. +(* mappedblocks *) + unfold update; intros. destruct (zeq b b1). congruence. eauto. +(* no overlap *) + unfold update; red; intros. + destruct (zeq b0 b1); destruct (zeq b2 b1); try congruence. + repeat rewrite (bounds_alloc_other _ _ _ _ _ H0); eauto. +(* range offset *) + unfold update; intros. + destruct (zeq b b1). congruence. eauto. +(* range block *) + unfold update; intros. + destruct (zeq b b1). congruence. eauto. +(* incr *) + split. auto. +(* image *) + split. apply update_s. +(* incr *) + intros; apply update_o; auto. +Qed. + +Theorem alloc_left_mapped_inject: + forall f m1 m2 lo hi m1' b1 b2 delta, + inject f m1 m2 -> + alloc m1 lo hi = (m1', b1) -> + valid_block m2 b2 -> + Int.min_signed <= delta <= Int.max_signed -> + delta = 0 \/ Int.min_signed <= low_bound m2 b2 /\ high_bound m2 b2 <= Int.max_signed -> + (forall ofs p, lo <= ofs < hi -> perm m2 b2 (ofs + delta) p) -> + inj_offset_aligned delta (hi-lo) -> + (forall b ofs, + f b = Some (b2, ofs) -> + high_bound m1 b + ofs <= lo + delta \/ + hi + delta <= low_bound m1 b + ofs) -> + exists f', + inject f' m1' m2 + /\ inject_incr f f' + /\ f' b1 = Some(b2, delta) + /\ (forall b, b <> b1 -> f' b = f b). +Proof. + intros. inversion H. + assert (inject_incr f (update b1 (Some(b2, delta)) f)). + red; unfold update; intros. destruct (zeq b b1). subst b. + assert (f b1 = None). eauto with mem. congruence. + auto. + assert (mem_inj (update b1 (Some(b2, delta)) f) m1 m2). + inversion mi_inj0; constructor; eauto with mem. + unfold update; intros. destruct (zeq b0 b1). + inv H8. elim (fresh_block_alloc _ _ _ _ _ H0). eauto with mem. + eauto. + unfold update; intros. destruct (zeq b0 b1). + inv H8. elim (fresh_block_alloc _ _ _ _ _ H0). eauto with mem. + apply memval_inject_incr with f; auto. + exists (update b1 (Some(b2, delta)) f). split. constructor. +(* inj *) + eapply alloc_left_mapped_inj; eauto. apply update_s. +(* freeblocks *) + unfold update; intros. destruct (zeq b b1). subst b. + elim H9. eauto with mem. + eauto with mem. +(* mappedblocks *) + unfold update; intros. destruct (zeq b b1). inv H9. auto. + eauto. +(* overlap *) + unfold update; red; intros. + repeat rewrite (bounds_alloc _ _ _ _ _ H0). unfold eq_block. + destruct (zeq b0 b1); destruct (zeq b3 b1); simpl. + inv H10; inv H11. congruence. + inv H10. destruct (zeq b1' b2'); auto. subst b2'. + right. generalize (H6 _ _ H11). tauto. + inv H11. destruct (zeq b1' b2'); auto. subst b2'. + right. eapply H6; eauto. + eauto. +(* range offset *) + unfold update; intros. destruct (zeq b b1). inv H9. auto. eauto. +(* range block *) + unfold update; intros. destruct (zeq b b1). inv H9. auto. eauto. +(* incr *) + split. auto. +(* image of b1 *) + split. apply update_s. +(* image of others *) + intros. apply update_o; auto. +Qed. + +Theorem alloc_parallel_inject: + forall f m1 m2 lo1 hi1 m1' b1 lo2 hi2, + inject f m1 m2 -> + alloc m1 lo1 hi1 = (m1', b1) -> + lo2 <= lo1 -> hi1 <= hi2 -> + exists f', exists m2', exists b2, + alloc m2 lo2 hi2 = (m2', b2) + /\ inject f' m1' m2' + /\ inject_incr f f' + /\ f' b1 = Some(b2, 0) + /\ (forall b, b <> b1 -> f' b = f b). +Proof. + intros. + case_eq (alloc m2 lo2 hi2). intros m2' b2 ALLOC. + exploit alloc_left_mapped_inject. + eapply alloc_right_inject; eauto. + eauto. + instantiate (1 := b2). eauto with mem. + instantiate (1 := 0). generalize Int.min_signed_neg Int.max_signed_pos; omega. + auto. + intros. eapply perm_alloc_2; eauto. omega. + red; intros. apply Zdivide_0. + intros. elimtype False. apply (valid_not_valid_diff m2 b2 b2); eauto with mem. + intros [f' [A [B [C D]]]]. + exists f'; exists m2'; exists b2; auto. +Qed. + +(** Preservation of [free] operations *) + +Lemma free_left_inject: + forall f m1 m2 b lo hi m1', + inject f m1 m2 -> + free m1 b lo hi = Some m1' -> + inject f m1' m2. +Proof. + intros. inversion H. constructor. +(* inj *) + eapply free_left_inj; eauto. +(* freeblocks *) + eauto with mem. +(* mappedblocks *) + auto. +(* no overlap *) + red; intros. repeat rewrite (bounds_free _ _ _ _ _ H0). eauto. +(* range offset *) + auto. +(* range block *) + auto. +Qed. + +Lemma free_list_left_inject: + forall f m2 l m1 m1', + inject f m1 m2 -> + free_list m1 l = Some m1' -> + inject f m1' m2. +Proof. + induction l; simpl; intros. + inv H0. auto. + destruct a as [[b lo] hi]. generalize H0. case_eq (free m1 b lo hi); intros. + apply IHl with m; auto. eapply free_left_inject; eauto. + congruence. +Qed. + +Lemma free_right_inject: + forall f m1 m2 b lo hi m2', + inject f m1 m2 -> + free m2 b lo hi = Some m2' -> + (forall b1 delta ofs p, + f b1 = Some(b, delta) -> perm m1 b1 ofs p -> + lo <= ofs + delta < hi -> False) -> + inject f m1 m2'. +Proof. + intros. inversion H. constructor. +(* inj *) + eapply free_right_inj; eauto. +(* freeblocks *) + auto. +(* mappedblocks *) + eauto with mem. +(* no overlap *) + auto. +(* range offset *) + auto. +(* range blocks *) + intros. rewrite (bounds_free _ _ _ _ _ H0). eauto. +Qed. + +Lemma perm_free_list: + forall l m m' b ofs p, + free_list m l = Some m' -> + perm m' b ofs p -> + perm m b ofs p /\ + (forall lo hi, In (b, lo, hi) l -> lo <= ofs < hi -> False). +Proof. + induction l; intros until p; simpl. + intros. inv H. split; auto. + destruct a as [[b1 lo1] hi1]. + case_eq (free m b1 lo1 hi1); intros; try congruence. + exploit IHl; eauto. intros [A B]. + split. eauto with mem. + intros. destruct H2. inv H2. + elim (perm_free_2 _ _ _ _ _ H ofs p). auto. auto. + eauto. +Qed. + +Theorem free_inject: + forall f m1 l m1' m2 b lo hi m2', + inject f m1 m2 -> + free_list m1 l = Some m1' -> + free m2 b lo hi = Some m2' -> + (forall b1 delta ofs p, + f b1 = Some(b, delta) -> + perm m1 b1 ofs p -> lo <= ofs + delta < hi -> + exists lo1, exists hi1, In (b1, lo1, hi1) l /\ lo1 <= ofs < hi1) -> + inject f m1' m2'. +Proof. + intros. + eapply free_right_inject; eauto. + eapply free_list_left_inject; eauto. + intros. exploit perm_free_list; eauto. intros [A B]. + exploit H2; eauto. intros [lo1 [hi1 [C D]]]. eauto. +Qed. + +(* +Theorem free_inject': + forall f m1 l m1' m2 b lo hi m2', + inject f m1 m2 -> + free_list m1 l = Some m1' -> + free m2 b lo hi = Some m2' -> + (forall b1 delta, + f b1 = Some(b, delta) -> In (b1, low_bound m1 b1, high_bound m1 b1) l) -> + inject f m1' m2'. +Proof. + intros. eapply free_inject; eauto. + intros. exists (low_bound m1 b1); exists (high_bound m1 b1). + split. eauto. apply perm_in_bounds with p. auto. +Qed. +*) + +(** Injecting a memory into itself. *) + +Definition flat_inj (thr: block) : meminj := + fun (b: block) => if zlt b thr then Some(b, 0) else None. + +Definition inject_neutral (thr: block) (m: mem) := + mem_inj (flat_inj thr) m m. + +Remark flat_inj_no_overlap: + forall thr m, meminj_no_overlap (flat_inj thr) m. +Proof. + unfold flat_inj; intros; red; intros. + destruct (zlt b1 thr); inversion H0; subst. + destruct (zlt b2 thr); inversion H1; subst. + auto. +Qed. + +Theorem neutral_inject: + forall m, inject_neutral (nextblock m) m -> inject (flat_inj (nextblock m)) m m. +Proof. + intros. constructor. +(* meminj *) + auto. +(* freeblocks *) + unfold flat_inj, valid_block; intros. + apply zlt_false. omega. +(* mappedblocks *) + unfold flat_inj, valid_block; intros. + destruct (zlt b (nextblock m)); inversion H0; subst. auto. +(* no overlap *) + apply flat_inj_no_overlap. +(* range *) + unfold flat_inj; intros. + destruct (zlt b (nextblock m)); inv H0. + generalize Int.min_signed_neg Int.max_signed_pos; omega. +(* range *) + unfold flat_inj; intros. + destruct (zlt b (nextblock m)); inv H0. auto. +Qed. + +Theorem empty_inject_neutral: + forall thr, inject_neutral thr empty. +Proof. + intros; red; constructor. +(* access *) + unfold flat_inj; intros. destruct (zlt b1 thr); inv H. + replace (ofs + 0) with ofs by omega; auto. +(* contents *) + intros; simpl; constructor. +Qed. + +Theorem alloc_inject_neutral: + forall thr m lo hi b m', + alloc m lo hi = (m', b) -> + inject_neutral thr m -> + nextblock m < thr -> + inject_neutral thr m'. +Proof. + intros; red. + eapply alloc_left_mapped_inj with (m1 := m) (b2 := b) (delta := 0). + eapply alloc_right_inj; eauto. eauto. eauto with mem. + red. intros. apply Zdivide_0. + intros. eapply perm_alloc_2; eauto. omega. + unfold flat_inj. apply zlt_true. + rewrite (alloc_result _ _ _ _ _ H). auto. +Qed. + +Theorem store_inject_neutral: + forall chunk m b ofs v m' thr, + store chunk m b ofs v = Some m' -> + inject_neutral thr m -> + b < thr -> + val_inject (flat_inj thr) v v -> + inject_neutral thr m'. +Proof. + intros; red. + exploit store_mapped_inj. eauto. eauto. apply flat_inj_no_overlap. + unfold flat_inj. apply zlt_true; auto. eauto. + replace (ofs + 0) with ofs by omega. + intros [m'' [A B]]. congruence. +Qed. + +End Mem. + +Notation mem := Mem.mem. + +Hint Resolve + Mem.valid_not_valid_diff + Mem.perm_implies + Mem.perm_valid_block + Mem.range_perm_implies + Mem.valid_access_implies + Mem.valid_access_valid_block + Mem.valid_access_perm + Mem.valid_access_load + Mem.load_valid_access + Mem.valid_access_store + Mem.perm_store_1 + Mem.perm_store_2 + Mem.nextblock_store + Mem.store_valid_block_1 + Mem.store_valid_block_2 + Mem.store_valid_access_1 + Mem.store_valid_access_2 + Mem.store_valid_access_3 + Mem.nextblock_alloc + Mem.alloc_result + Mem.valid_block_alloc + Mem.fresh_block_alloc + Mem.valid_new_block + Mem.perm_alloc_1 + Mem.perm_alloc_2 + Mem.perm_alloc_3 + Mem.perm_alloc_inv + Mem.valid_access_alloc_other + Mem.valid_access_alloc_same + Mem.valid_access_alloc_inv + Mem.range_perm_free + Mem.free_range_perm + Mem.nextblock_free + Mem.valid_block_free_1 + Mem.valid_block_free_2 + Mem.perm_free_1 + Mem.perm_free_2 + Mem.perm_free_3 + Mem.valid_access_free_1 + Mem.valid_access_free_2 + Mem.valid_access_free_inv_1 + Mem.valid_access_free_inv_2 +: mem. diff --git a/common/Memtype.v b/common/Memtype.v new file mode 100644 index 0000000..cfbe511 --- /dev/null +++ b/common/Memtype.v @@ -0,0 +1,989 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** This file defines the interface for the memory model that + is used in the dynamic semantics of all the languages used in the compiler. + It defines a type [mem] of memory states, the following 4 basic + operations over memory states, and their properties: +- [load]: read a memory chunk at a given address; +- [store]: store a memory chunk at a given address; +- [alloc]: allocate a fresh memory block; +- [free]: invalidate a memory block. +*) + +Require Import Coqlib. +Require Import AST. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Import Memdata. + +(** Memory states are accessed by addresses [b, ofs]: pairs of a block + identifier [b] and a byte offset [ofs] within that block. + Each address is in one of the following five states: +- Freeable (exclusive access): all operations permitted +- Writable: load, store and pointer comparison operations are permitted, + but freeing is not. +- Readable: only load and pointer comparison operations are permitted. +- Nonempty: valid, but only pointer comparisons are permitted. +- Empty: not yet allocated or previously freed; no operation permitted. + +The first four cases are represented by the following type of permissions. +Being empty is represented by the absence of any permission. +*) + +Inductive permission: Type := + | Freeable: permission + | Writable: permission + | Readable: permission + | Nonempty: permission. + +(** In the list, each permission implies the other permissions further down the + list. We reflect this fact by the following order over permissions. *) + +Inductive perm_order: permission -> permission -> Prop := + | perm_F_any: forall p, perm_order Freeable p + | perm_W_R: perm_order Writable Readable + | perm_any_N: forall p, perm_order p Nonempty. + +Hint Constructors perm_order: mem. + +Module Type MEM. + +(** The abstract type of memory states. *) +Parameter mem: Type. + +Definition nullptr: block := 0. + +(** * Operations on memory states *) + +(** [empty] is the initial memory state. *) +Parameter empty: mem. + +(** [alloc m lo hi] allocates a fresh block of size [hi - lo] bytes. + Valid offsets in this block are between [lo] included and [hi] excluded. + These offsets are writable in the returned memory state. + This block is not initialized: its contents are initially undefined. + Returns a pair [(m', b)] of the updated memory state [m'] and + the identifier [b] of the newly-allocated block. + Note that [alloc] never fails: we are modeling an infinite memory. *) +Parameter alloc: forall (m: mem) (lo hi: Z), mem * block. + +(** [free m b lo hi] frees (deallocates) the range of offsets from [lo] + included to [hi] excluded in block [b]. Returns the updated memory + state, or [None] if the freed addresses are not writable. *) +Parameter free: forall (m: mem) (b: block) (lo hi: Z), option mem. + +(** [load chunk m b ofs] reads a memory quantity [chunk] from + addresses [b, ofs] to [b, ofs + size_chunk chunk - 1] in memory state + [m]. Returns the value read, or [None] if the accessed addresses + are not readable. *) +Parameter load: forall (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z), option val. + +(** [store chunk m b ofs v] writes value [v] as memory quantity [chunk] + from addresses [b, ofs] to [b, ofs + size_chunk chunk - 1] in memory state + [m]. Returns the updated memory state, or [None] if the accessed addresses + are not writable. *) +Parameter store: forall (chunk: memory_chunk) (m: mem) (b: block) (ofs: Z) (v: val), option mem. + +(** [loadv] and [storev] are variants of [load] and [store] where + the address being accessed is passed as a value (of the [Vptr] kind). *) + +Definition loadv (chunk: memory_chunk) (m: mem) (addr: val) : option val := + match addr with + | Vptr b ofs => load chunk m b (Int.signed ofs) + | _ => None + end. + +Definition storev (chunk: memory_chunk) (m: mem) (addr v: val) : option mem := + match addr with + | Vptr b ofs => store chunk m b (Int.signed ofs) v + | _ => None + end. + +(** [loadbytes m b ofs n] reads and returns the byte-level representation of + the values contained at offsets [ofs] to [ofs + n - 1] within block [b] + in memory state [m]. These values must be integers or floats. + [None] is returned if the accessed addresses are not readable + or contain undefined or pointer values. *) +Parameter loadbytes: forall (m: mem) (b: block) (ofs n: Z), option (list byte). + +(** [free_list] frees all the given (block, lo, hi) triples. *) +Fixpoint free_list (m: mem) (l: list (block * Z * Z)) {struct l}: option mem := + match l with + | nil => Some m + | (b, lo, hi) :: l' => + match free m b lo hi with + | None => None + | Some m' => free_list m' l' + end + end. + +(** * Permissions, block validity, access validity, and bounds *) + +(** The next block of a memory state is the block identifier for the + next allocation. It increases by one at each allocation. + Block identifiers below [nextblock] are said to be valid, meaning + that they have been allocated previously. Block identifiers above + [nextblock] are fresh or invalid, i.e. not yet allocated. Note that + a block identifier remains valid after a [free] operation over this + block. *) + +Parameter nextblock: mem -> block. +Axiom nextblock_pos: + forall m, nextblock m > 0. + +Definition valid_block (m: mem) (b: block) := + b < nextblock m. +Axiom valid_not_valid_diff: + forall m b b', valid_block m b -> ~(valid_block m b') -> b <> b'. + +(** [perm m b ofs p] holds if the address [b, ofs] in memory state [m] + has permission [p]: one of writable, readable, and nonempty. + If the address is empty, [perm m b ofs p] is false for all values of [p]. *) +Parameter perm: forall (m: mem) (b: block) (ofs: Z) (p: permission), Prop. + +(** Logical implications between permissions *) + +Axiom perm_implies: + forall m b ofs p1 p2, perm m b ofs p1 -> perm_order p1 p2 -> perm m b ofs p2. + +(** Having a (nonempty) permission implies that the block is valid. + In other words, invalid blocks, not yet allocated, are all empty. *) +Axiom perm_valid_block: + forall m b ofs p, perm m b ofs p -> valid_block m b. + +(* Unused? +(** The [Mem.perm] predicate is decidable. *) +Axiom perm_dec: + forall m b ofs p, {perm m b ofs p} + {~ perm m b ofs p}. +*) + +(** [range_perm m b lo hi p] holds iff the addresses [b, lo] to [b, hi-1] + all have permission [p]. *) +Definition range_perm (m: mem) (b: block) (lo hi: Z) (p: permission) : Prop := + forall ofs, lo <= ofs < hi -> perm m b ofs p. + +Axiom range_perm_implies: + forall m b lo hi p1 p2, + range_perm m b lo hi p1 -> perm_order p1 p2 -> range_perm m b lo hi p2. + +(** An access to a memory quantity [chunk] at address [b, ofs] with + permission [p] is valid in [m] if the accessed addresses all have + permission [p] and moreover the offset is properly aligned. *) +Definition valid_access (m: mem) (chunk: memory_chunk) (b: block) (ofs: Z) (p: permission): Prop := + range_perm m b ofs (ofs + size_chunk chunk) p + /\ (align_chunk chunk | ofs). + +Axiom valid_access_implies: + forall m chunk b ofs p1 p2, + valid_access m chunk b ofs p1 -> perm_order p1 p2 -> + valid_access m chunk b ofs p2. + +Axiom valid_access_valid_block: + forall m chunk b ofs, + valid_access m chunk b ofs Nonempty -> + valid_block m b. + +Axiom valid_access_perm: + forall m chunk b ofs p, + valid_access m chunk b ofs p -> + perm m b ofs p. + +(** [valid_pointer m b ofs] returns [true] if the address [b, ofs] + is nonempty in [m] and [false] if it is empty. *) + +Parameter valid_pointer: forall (m: mem) (b: block) (ofs: Z), bool. + +Axiom valid_pointer_nonempty_perm: + forall m b ofs, + valid_pointer m b ofs = true <-> perm m b ofs Nonempty. +Axiom valid_pointer_valid_access: + forall m b ofs, + valid_pointer m b ofs = true <-> valid_access m Mint8unsigned b ofs Nonempty. + +(** Each block has associated low and high bounds. These are the bounds + that were given when the block was allocated. *) + +Parameter bounds: forall (m: mem) (b: block), Z*Z. + +Notation low_bound m b := (fst(bounds m b)). +Notation high_bound m b := (snd(bounds m b)). + +(** The crucial properties of bounds is that any offset below the low + bound or above the high bound is empty. *) + +Axiom perm_in_bounds: + forall m b ofs p, perm m b ofs p -> low_bound m b <= ofs < high_bound m b. + +Axiom range_perm_in_bounds: + forall m b lo hi p, + range_perm m b lo hi p -> lo < hi -> + low_bound m b <= lo /\ hi <= high_bound m b. + +Axiom valid_access_in_bounds: + forall m chunk b ofs p, + valid_access m chunk b ofs p -> + low_bound m b <= ofs /\ ofs + size_chunk chunk <= high_bound m b. + +(** * Properties of the memory operations *) + +(** ** Properties of the initial memory state. *) + +Axiom nextblock_empty: nextblock empty = 1. +Axiom perm_empty: forall b ofs p, ~perm empty b ofs p. +Axiom valid_access_empty: + forall chunk b ofs p, ~valid_access empty chunk b ofs p. + +(** ** Properties of [load]. *) + +(** A load succeeds if and only if the access is valid for reading *) +Axiom valid_access_load: + forall m chunk b ofs, + valid_access m chunk b ofs Readable -> + exists v, load chunk m b ofs = Some v. +Axiom load_valid_access: + forall m chunk b ofs v, + load chunk m b ofs = Some v -> + valid_access m chunk b ofs Readable. + +(** The value returned by [load] belongs to the type of the memory quantity + accessed: [Vundef], [Vint] or [Vptr] for an integer quantity, + [Vundef] or [Vfloat] for a float quantity. *) +Axiom load_type: + forall m chunk b ofs v, + load chunk m b ofs = Some v -> + Val.has_type v (type_of_chunk chunk). + +(** For a small integer or float type, the value returned by [load] + is invariant under the corresponding cast. *) +Axiom load_cast: + forall m chunk b ofs v, + load chunk m b ofs = Some v -> + match chunk with + | Mint8signed => v = Val.sign_ext 8 v + | Mint8unsigned => v = Val.zero_ext 8 v + | Mint16signed => v = Val.sign_ext 16 v + | Mint16unsigned => v = Val.zero_ext 16 v + | Mfloat32 => v = Val.singleoffloat v + | _ => True + end. + +Axiom load_int8_signed_unsigned: + forall m b ofs, + load Mint8signed m b ofs = option_map (Val.sign_ext 8) (load Mint8unsigned m b ofs). + +Axiom load_int16_signed_unsigned: + forall m b ofs, + load Mint16signed m b ofs = option_map (Val.sign_ext 16) (load Mint16unsigned m b ofs). + + +(** ** Properties of [loadbytes]. *) + +(** If [loadbytes] succeeds, the corresponding [load] succeeds and + returns a [Vint] or [Vfloat] value that is determined by the + bytes read by [loadbytes]. *) +Axiom loadbytes_load: + forall chunk m b ofs bytes, + loadbytes m b ofs (size_chunk chunk) = Some bytes -> + (align_chunk chunk | ofs) -> + load chunk m b ofs = + Some(match type_of_chunk chunk with + | Tint => Vint(decode_int chunk bytes) + | Tfloat => Vfloat(decode_float chunk bytes) + end). + +(** Conversely, if [load] returns an int or a float, the corresponding + [loadbytes] succeeds and returns a list of bytes which decodes into the + result of [load]. *) +Axiom load_int_loadbytes: + forall chunk m b ofs n, + load chunk m b ofs = Some(Vint n) -> + type_of_chunk chunk = Tint /\ + exists bytes, loadbytes m b ofs (size_chunk chunk) = Some bytes + /\ n = decode_int chunk bytes. + +Axiom load_float_loadbytes: + forall chunk m b ofs f, + load chunk m b ofs = Some(Vfloat f) -> + type_of_chunk chunk = Tfloat /\ + exists bytes, loadbytes m b ofs (size_chunk chunk) = Some bytes + /\ f = decode_float chunk bytes. + + +(** [loadbytes] returns a list of length [n] (the number of bytes read). *) +Axiom loadbytes_length: + forall m b ofs n bytes, + loadbytes m b ofs n = Some bytes -> + length bytes = nat_of_Z n. + +(** Composing or decomposing [loadbytes] operations at adjacent addresses. *) +Axiom loadbytes_concat: + forall m b ofs n1 n2 bytes1 bytes2, + loadbytes m b ofs n1 = Some bytes1 -> + loadbytes m b (ofs + n1) n2 = Some bytes2 -> + n1 >= 0 -> n2 >= 0 -> + loadbytes m b ofs (n1 + n2) = Some(bytes1 ++ bytes2). +Axiom loadbytes_split: + forall m b ofs n1 n2 bytes, + loadbytes m b ofs (n1 + n2) = Some bytes -> + n1 >= 0 -> n2 >= 0 -> + exists bytes1, exists bytes2, + loadbytes m b ofs n1 = Some bytes1 + /\ loadbytes m b (ofs + n1) n2 = Some bytes2 + /\ bytes = bytes1 ++ bytes2. + +(** ** Properties of [store]. *) + +(** [store] preserves block validity, permissions, access validity, and bounds. + Moreover, a [store] succeeds if and only if the corresponding access + is valid for writing. *) + +Axiom nextblock_store: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + nextblock m2 = nextblock m1. +Axiom store_valid_block_1: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall b', valid_block m1 b' -> valid_block m2 b'. +Axiom store_valid_block_2: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall b', valid_block m2 b' -> valid_block m1 b'. + +Axiom perm_store_1: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall b' ofs' p, perm m1 b' ofs' p -> perm m2 b' ofs' p. +Axiom perm_store_2: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall b' ofs' p, perm m2 b' ofs' p -> perm m1 b' ofs' p. + +Axiom valid_access_store: + forall m1 chunk b ofs v, + valid_access m1 chunk b ofs Writable -> + { m2: mem | store chunk m1 b ofs v = Some m2 }. +Axiom store_valid_access_1: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall chunk' b' ofs' p, + valid_access m1 chunk' b' ofs' p -> valid_access m2 chunk' b' ofs' p. +Axiom store_valid_access_2: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall chunk' b' ofs' p, + valid_access m2 chunk' b' ofs' p -> valid_access m1 chunk' b' ofs' p. +Axiom store_valid_access_3: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + valid_access m1 chunk b ofs Writable. + +Axiom bounds_store: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall b', bounds m2 b' = bounds m1 b'. + +(** Load-store properties. *) + +Axiom load_store_similar: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall chunk', + size_chunk chunk' = size_chunk chunk -> + exists v', load chunk' m2 b ofs = Some v' /\ decode_encode_val v chunk chunk' v'. + +Axiom load_store_same: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + Val.has_type v (type_of_chunk chunk) -> + load chunk m2 b ofs = Some (Val.load_result chunk v). + +Axiom load_store_other: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall chunk' b' ofs', + b' <> b + \/ ofs' + size_chunk chunk' <= ofs + \/ ofs + size_chunk chunk <= ofs' -> + load chunk' m2 b' ofs' = load chunk' m1 b' ofs'. + +(** Integrity of pointer values. *) + +Axiom load_store_pointer_overlap: + forall chunk m1 b ofs v_b v_o m2 chunk' ofs' v, + store chunk m1 b ofs (Vptr v_b v_o) = Some m2 -> + load chunk' m2 b ofs' = Some v -> + ofs' <> ofs -> + ofs' + size_chunk chunk' > ofs -> + ofs + size_chunk chunk > ofs' -> + v = Vundef. +Axiom load_store_pointer_mismatch: + forall chunk m1 b ofs v_b v_o m2 chunk' v, + store chunk m1 b ofs (Vptr v_b v_o) = Some m2 -> + load chunk' m2 b ofs = Some v -> + chunk <> Mint32 \/ chunk' <> Mint32 -> + v = Vundef. +Axiom load_pointer_store: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall chunk' b' ofs' v_b v_o, + load chunk' m2 b' ofs' = Some(Vptr v_b v_o) -> + (chunk = Mint32 /\ v = Vptr v_b v_o /\ chunk' = Mint32 /\ b' = b /\ ofs' = ofs) + \/ (b' <> b \/ ofs' + size_chunk chunk' <= ofs \/ ofs + size_chunk chunk <= ofs'). + +(** Load-store properties for [loadbytes]. *) + +Axiom loadbytes_store_same: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + loadbytes m2 b ofs (size_chunk chunk) = + match v with + | Vundef => None + | Vint n => Some(encode_int chunk n) + | Vfloat n => Some(encode_float chunk n) + | Vptr _ _ => None + end. +Axiom loadbytes_store_other: + forall chunk m1 b ofs v m2, store chunk m1 b ofs v = Some m2 -> + forall b' ofs' n, + b' <> b \/ n <= 0 \/ ofs' + n <= ofs \/ ofs + size_chunk chunk <= ofs' -> + loadbytes m2 b' ofs' n = loadbytes m1 b' ofs' n. + +(** [store] is insensitive to the signedness or the high bits of + small integer quantities. *) + +Axiom store_signed_unsigned_8: + forall m b ofs v, + store Mint8signed m b ofs v = store Mint8unsigned m b ofs v. +Axiom store_signed_unsigned_16: + forall m b ofs v, + store Mint16signed m b ofs v = store Mint16unsigned m b ofs v. +Axiom store_int8_zero_ext: + forall m b ofs n, + store Mint8unsigned m b ofs (Vint (Int.zero_ext 8 n)) = + store Mint8unsigned m b ofs (Vint n). +Axiom store_int8_sign_ext: + forall m b ofs n, + store Mint8signed m b ofs (Vint (Int.sign_ext 8 n)) = + store Mint8signed m b ofs (Vint n). +Axiom store_int16_zero_ext: + forall m b ofs n, + store Mint16unsigned m b ofs (Vint (Int.zero_ext 16 n)) = + store Mint16unsigned m b ofs (Vint n). +Axiom store_int16_sign_ext: + forall m b ofs n, + store Mint16signed m b ofs (Vint (Int.sign_ext 16 n)) = + store Mint16signed m b ofs (Vint n). +Axiom store_float32_truncate: + forall m b ofs n, + store Mfloat32 m b ofs (Vfloat (Float.singleoffloat n)) = + store Mfloat32 m b ofs (Vfloat n). + +(** ** Properties of [alloc]. *) + +(** The identifier of the freshly allocated block is the next block + of the initial memory state. *) + +Axiom alloc_result: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + b = nextblock m1. + +(** Effect of [alloc] on block validity. *) + +Axiom nextblock_alloc: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + nextblock m2 = Zsucc (nextblock m1). + +Axiom valid_block_alloc: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall b', valid_block m1 b' -> valid_block m2 b'. +Axiom fresh_block_alloc: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + ~(valid_block m1 b). +Axiom valid_new_block: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + valid_block m2 b. +Axiom valid_block_alloc_inv: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall b', valid_block m2 b' -> b' = b \/ valid_block m1 b'. + +(** Effect of [alloc] on permissions. *) + +Axiom perm_alloc_1: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall b' ofs p, perm m1 b' ofs p -> perm m2 b' ofs p. +Axiom perm_alloc_2: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall ofs, lo <= ofs < hi -> perm m2 b ofs Freeable. +Axiom perm_alloc_3: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall ofs p, ofs < lo \/ hi <= ofs -> ~perm m2 b ofs p. +Axiom perm_alloc_inv: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall b' ofs p, + perm m2 b' ofs p -> + if zeq b' b then lo <= ofs < hi else perm m1 b' ofs p. + +(** Effect of [alloc] on access validity. *) + +Axiom valid_access_alloc_other: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall chunk b' ofs p, + valid_access m1 chunk b' ofs p -> + valid_access m2 chunk b' ofs p. +Axiom valid_access_alloc_same: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall chunk ofs, + lo <= ofs -> ofs + size_chunk chunk <= hi -> (align_chunk chunk | ofs) -> + valid_access m2 chunk b ofs Freeable. +Axiom valid_access_alloc_inv: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall chunk b' ofs p, + valid_access m2 chunk b' ofs p -> + if eq_block b' b + then lo <= ofs /\ ofs + size_chunk chunk <= hi /\ (align_chunk chunk | ofs) + else valid_access m1 chunk b' ofs p. + +(** Effect of [alloc] on bounds. *) + +Axiom bounds_alloc: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall b', bounds m2 b' = if eq_block b' b then (lo, hi) else bounds m1 b'. + +Axiom bounds_alloc_same: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + bounds m2 b = (lo, hi). + +Axiom bounds_alloc_other: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall b', b' <> b -> bounds m2 b' = bounds m1 b'. + +(** Load-alloc properties. *) + +Axiom load_alloc_unchanged: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall chunk b' ofs, + valid_block m1 b' -> + load chunk m2 b' ofs = load chunk m1 b' ofs. +Axiom load_alloc_other: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall chunk b' ofs v, + load chunk m1 b' ofs = Some v -> + load chunk m2 b' ofs = Some v. +Axiom load_alloc_same: + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall chunk ofs v, + load chunk m2 b ofs = Some v -> + v = Vundef. +Axiom load_alloc_same': + forall m1 lo hi m2 b, alloc m1 lo hi = (m2, b) -> + forall chunk ofs, + lo <= ofs -> ofs + size_chunk chunk <= hi -> (align_chunk chunk | ofs) -> + load chunk m2 b ofs = Some Vundef. + +(** ** Properties of [free]. *) + +(** [free] succeeds if and only if the correspond range of addresses + has [Freeable] permission. *) + +Axiom range_perm_free: + forall m1 b lo hi, + range_perm m1 b lo hi Freeable -> + { m2: mem | free m1 b lo hi = Some m2 }. +Axiom free_range_perm: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + range_perm m1 bf lo hi Freeable. + +(** Block validity is preserved by [free]. *) + +Axiom nextblock_free: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + nextblock m2 = nextblock m1. +Axiom valid_block_free_1: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall b, valid_block m1 b -> valid_block m2 b. +Axiom valid_block_free_2: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall b, valid_block m2 b -> valid_block m1 b. + +(** Effect of [free] on permissions. *) + +Axiom perm_free_1: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall b ofs p, + b <> bf \/ ofs < lo \/ hi <= ofs -> + perm m1 b ofs p -> + perm m2 b ofs p. +Axiom perm_free_2: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall ofs p, lo <= ofs < hi -> ~ perm m2 bf ofs p. +Axiom perm_free_3: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall b ofs p, + perm m2 b ofs p -> perm m1 b ofs p. + +(** Effect of [free] on access validity. *) + +Axiom valid_access_free_1: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall chunk b ofs p, + valid_access m1 chunk b ofs p -> + b <> bf \/ lo >= hi \/ ofs + size_chunk chunk <= lo \/ hi <= ofs -> + valid_access m2 chunk b ofs p. +Axiom valid_access_free_2: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall chunk ofs p, + lo < hi -> ofs + size_chunk chunk > lo -> ofs < hi -> + ~(valid_access m2 chunk bf ofs p). +Axiom valid_access_free_inv_1: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall chunk b ofs p, + valid_access m2 chunk b ofs p -> + valid_access m1 chunk b ofs p. +Axiom valid_access_free_inv_2: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall chunk ofs p, + valid_access m2 chunk bf ofs p -> + lo >= hi \/ ofs + size_chunk chunk <= lo \/ hi <= ofs. + +(** [free] preserves bounds. *) + +Axiom bounds_free: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall b, bounds m2 b = bounds m1 b. + +(** Load-free properties *) + +Axiom load_free: + forall m1 bf lo hi m2, free m1 bf lo hi = Some m2 -> + forall chunk b ofs, + b <> bf \/ lo >= hi \/ ofs + size_chunk chunk <= lo \/ hi <= ofs -> + load chunk m2 b ofs = load chunk m1 b ofs. + +(** * Relating two memory states. *) + +(** ** Memory extensions *) + +(** A store [m2] extends a store [m1] if [m2] can be obtained from [m1] + by relaxing the permissions of [m1] (for instance, allocating larger + blocks) and replacing some of the [Vundef] values stored in [m1] by + more defined values stored in [m2] at the same addresses. *) + +Parameter extends: mem -> mem -> Prop. + +Axiom extends_refl: + forall m, extends m m. + +Axiom load_extends: + forall chunk m1 m2 b ofs v1, + extends m1 m2 -> + load chunk m1 b ofs = Some v1 -> + exists v2, load chunk m2 b ofs = Some v2 /\ Val.lessdef v1 v2. + +Axiom loadv_extends: + forall chunk m1 m2 addr1 addr2 v1, + extends m1 m2 -> + loadv chunk m1 addr1 = Some v1 -> + Val.lessdef addr1 addr2 -> + exists v2, loadv chunk m2 addr2 = Some v2 /\ Val.lessdef v1 v2. + +Axiom store_within_extends: + forall chunk m1 m2 b ofs v1 m1' v2, + extends m1 m2 -> + store chunk m1 b ofs v1 = Some m1' -> + Val.lessdef v1 v2 -> + exists m2', + store chunk m2 b ofs v2 = Some m2' + /\ extends m1' m2'. + +Axiom store_outside_extends: + forall chunk m1 m2 b ofs v m2', + extends m1 m2 -> + store chunk m2 b ofs v = Some m2' -> + ofs + size_chunk chunk <= low_bound m1 b \/ high_bound m1 b <= ofs -> + extends m1 m2'. + +Axiom storev_extends: + forall chunk m1 m2 addr1 v1 m1' addr2 v2, + extends m1 m2 -> + storev chunk m1 addr1 v1 = Some m1' -> + Val.lessdef addr1 addr2 -> + Val.lessdef v1 v2 -> + exists m2', + storev chunk m2 addr2 v2 = Some m2' + /\ extends m1' m2'. + +Axiom alloc_extends: + forall m1 m2 lo1 hi1 b m1' lo2 hi2, + extends m1 m2 -> + alloc m1 lo1 hi1 = (m1', b) -> + lo2 <= lo1 -> hi1 <= hi2 -> + exists m2', + alloc m2 lo2 hi2 = (m2', b) + /\ extends m1' m2'. + +Axiom free_left_extends: + forall m1 m2 b lo hi m1', + extends m1 m2 -> + free m1 b lo hi = Some m1' -> + extends m1' m2. + +Axiom free_right_extends: + forall m1 m2 b lo hi m2', + extends m1 m2 -> + free m2 b lo hi = Some m2' -> + (forall ofs p, lo <= ofs < hi -> ~perm m1 b ofs p) -> + extends m1 m2'. + +Axiom free_parallel_extends: + forall m1 m2 b lo hi m1', + extends m1 m2 -> + free m1 b lo hi = Some m1' -> + exists m2', + free m2 b lo hi = Some m2' + /\ extends m1' m2'. + +Axiom valid_block_extends: + forall m1 m2 b, + extends m1 m2 -> + (valid_block m1 b <-> valid_block m2 b). +Axiom perm_extends: + forall m1 m2 b ofs p, + extends m1 m2 -> perm m1 b ofs p -> perm m2 b ofs p. +Axiom valid_access_extends: + forall m1 m2 chunk b ofs p, + extends m1 m2 -> valid_access m1 chunk b ofs p -> valid_access m2 chunk b ofs p. + +(** * Memory injections *) + +(** A memory injection [f] is a function from addresses to either [None] + or [Some] of an address and an offset. It defines a correspondence + between the blocks of two memory states [m1] and [m2]: +- if [f b = None], the block [b] of [m1] has no equivalent in [m2]; +- if [f b = Some(b', ofs)], the block [b] of [m2] corresponds to + a sub-block at offset [ofs] of the block [b'] in [m2]. + +A memory injection [f] defines a relation [val_inject] between values +that is the identity for integer and float values, and relocates pointer +values as prescribed by [f]. (See module [Values].) + +Likewise, a memory injection [f] defines a relation between memory states +that we now axiomatize. *) + +Parameter inject: meminj -> mem -> mem -> Prop. + +Axiom valid_block_inject_1: + forall f m1 m2 b1 b2 delta, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + valid_block m1 b1. + +Axiom valid_block_inject_2: + forall f m1 m2 b1 b2 delta, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + valid_block m2 b2. + +Axiom perm_inject: + forall f m1 m2 b1 b2 delta ofs p, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + perm m1 b1 ofs p -> perm m2 b2 (ofs + delta) p. + +Axiom valid_access_inject: + forall f m1 m2 chunk b1 ofs b2 delta p, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + valid_access m1 chunk b1 ofs p -> + valid_access m2 chunk b2 (ofs + delta) p. + +Axiom valid_pointer_inject: + forall f m1 m2 b1 ofs b2 delta, + f b1 = Some(b2, delta) -> + inject f m1 m2 -> + valid_pointer m1 b1 ofs = true -> + valid_pointer m2 b2 (ofs + delta) = true. + +Axiom address_inject: + forall f m1 m2 b1 ofs1 b2 delta, + inject f m1 m2 -> + perm m1 b1 (Int.signed ofs1) Nonempty -> + f b1 = Some (b2, delta) -> + Int.signed (Int.add ofs1 (Int.repr delta)) = Int.signed ofs1 + delta. + +Axiom valid_pointer_inject_no_overflow: + forall f m1 m2 b ofs b' x, + inject f m1 m2 -> + valid_pointer m1 b (Int.signed ofs) = true -> + f b = Some(b', x) -> + Int.min_signed <= Int.signed ofs + Int.signed (Int.repr x) <= Int.max_signed. + +Axiom valid_pointer_inject_val: + forall f m1 m2 b ofs b' ofs', + inject f m1 m2 -> + valid_pointer m1 b (Int.signed ofs) = true -> + val_inject f (Vptr b ofs) (Vptr b' ofs') -> + valid_pointer m2 b' (Int.signed ofs') = true. + +Axiom inject_no_overlap: + forall f m1 m2 b1 b2 b1' b2' delta1 delta2 ofs1 ofs2, + inject f m1 m2 -> + b1 <> b2 -> + f b1 = Some (b1', delta1) -> + f b2 = Some (b2', delta2) -> + perm m1 b1 ofs1 Nonempty -> + perm m1 b2 ofs2 Nonempty -> + b1' <> b2' \/ ofs1 + delta1 <> ofs2 + delta2. + +Axiom different_pointers_inject: + forall f m m' b1 ofs1 b2 ofs2 b1' delta1 b2' delta2, + inject f m m' -> + b1 <> b2 -> + valid_pointer m b1 (Int.signed ofs1) = true -> + valid_pointer m b2 (Int.signed ofs2) = true -> + f b1 = Some (b1', delta1) -> + f b2 = Some (b2', delta2) -> + b1' <> b2' \/ + Int.signed (Int.add ofs1 (Int.repr delta1)) <> + Int.signed (Int.add ofs2 (Int.repr delta2)). + +Axiom load_inject: + forall f m1 m2 chunk b1 ofs b2 delta v1, + inject f m1 m2 -> + load chunk m1 b1 ofs = Some v1 -> + f b1 = Some (b2, delta) -> + exists v2, load chunk m2 b2 (ofs + delta) = Some v2 /\ val_inject f v1 v2. + +Axiom loadv_inject: + forall f m1 m2 chunk a1 a2 v1, + inject f m1 m2 -> + loadv chunk m1 a1 = Some v1 -> + val_inject f a1 a2 -> + exists v2, loadv chunk m2 a2 = Some v2 /\ val_inject f v1 v2. + +Axiom store_mapped_inject: + forall f chunk m1 b1 ofs v1 n1 m2 b2 delta v2, + inject f m1 m2 -> + store chunk m1 b1 ofs v1 = Some n1 -> + f b1 = Some (b2, delta) -> + val_inject f v1 v2 -> + exists n2, + store chunk m2 b2 (ofs + delta) v2 = Some n2 + /\ inject f n1 n2. + +Axiom store_unmapped_inject: + forall f chunk m1 b1 ofs v1 n1 m2, + inject f m1 m2 -> + store chunk m1 b1 ofs v1 = Some n1 -> + f b1 = None -> + inject f n1 m2. + +Axiom store_outside_inject: + forall f m1 m2 chunk b ofs v m2', + inject f m1 m2 -> + (forall b' delta, + f b' = Some(b, delta) -> + high_bound m1 b' + delta <= ofs + \/ ofs + size_chunk chunk <= low_bound m1 b' + delta) -> + store chunk m2 b ofs v = Some m2' -> + inject f m1 m2'. + +Axiom storev_mapped_inject: + forall f chunk m1 a1 v1 n1 m2 a2 v2, + inject f m1 m2 -> + storev chunk m1 a1 v1 = Some n1 -> + val_inject f a1 a2 -> + val_inject f v1 v2 -> + exists n2, + storev chunk m2 a2 v2 = Some n2 /\ inject f n1 n2. + +Axiom alloc_right_inject: + forall f m1 m2 lo hi b2 m2', + inject f m1 m2 -> + alloc m2 lo hi = (m2', b2) -> + inject f m1 m2'. + +Axiom alloc_left_unmapped_inject: + forall f m1 m2 lo hi m1' b1, + inject f m1 m2 -> + alloc m1 lo hi = (m1', b1) -> + exists f', + inject f' m1' m2 + /\ inject_incr f f' + /\ f' b1 = None + /\ (forall b, b <> b1 -> f' b = f b). + +Definition inj_offset_aligned (delta: Z) (size: Z) : Prop := + forall chunk, size_chunk chunk <= size -> (align_chunk chunk | delta). + +Axiom alloc_left_mapped_inject: + forall f m1 m2 lo hi m1' b1 b2 delta, + inject f m1 m2 -> + alloc m1 lo hi = (m1', b1) -> + valid_block m2 b2 -> + Int.min_signed <= delta <= Int.max_signed -> + delta = 0 \/ Int.min_signed <= low_bound m2 b2 /\ high_bound m2 b2 <= Int.max_signed -> + (forall ofs p, lo <= ofs < hi -> perm m2 b2 (ofs + delta) p) -> + inj_offset_aligned delta (hi-lo) -> + (forall b ofs, + f b = Some (b2, ofs) -> + high_bound m1 b + ofs <= lo + delta \/ + hi + delta <= low_bound m1 b + ofs) -> + exists f', + inject f' m1' m2 + /\ inject_incr f f' + /\ f' b1 = Some(b2, delta) + /\ (forall b, b <> b1 -> f' b = f b). + +Axiom alloc_parallel_inject: + forall f m1 m2 lo1 hi1 m1' b1 lo2 hi2, + inject f m1 m2 -> + alloc m1 lo1 hi1 = (m1', b1) -> + lo2 <= lo1 -> hi1 <= hi2 -> + exists f', exists m2', exists b2, + alloc m2 lo2 hi2 = (m2', b2) + /\ inject f' m1' m2' + /\ inject_incr f f' + /\ f' b1 = Some(b2, 0) + /\ (forall b, b <> b1 -> f' b = f b). + +Axiom free_inject: + forall f m1 l m1' m2 b lo hi m2', + inject f m1 m2 -> + free_list m1 l = Some m1' -> + free m2 b lo hi = Some m2' -> + (forall b1 delta ofs p, + f b1 = Some(b, delta) -> perm m1 b1 ofs p -> lo <= ofs + delta < hi -> + exists lo1, exists hi1, In (b1, lo1, hi1) l /\ lo1 <= ofs < hi1) -> + inject f m1' m2'. + +(** Memory states that inject into themselves. *) + +Definition flat_inj (thr: block) : meminj := + fun (b: block) => if zlt b thr then Some(b, 0) else None. + +Parameter inject_neutral: forall (thr: block) (m: mem), Prop. + +Axiom neutral_inject: + forall m, inject_neutral (nextblock m) m -> + inject (flat_inj (nextblock m)) m m. + +Axiom empty_inject_neutral: + forall thr, inject_neutral thr empty. + +Axiom alloc_inject_neutral: + forall thr m lo hi b m', + alloc m lo hi = (m', b) -> + inject_neutral thr m -> + nextblock m < thr -> + inject_neutral thr m'. + +Axiom store_inject_neutral: + forall chunk m b ofs v m' thr, + store chunk m b ofs v = Some m' -> + inject_neutral thr m -> + b < thr -> + val_inject (flat_inj thr) v v -> + inject_neutral thr m'. + +End MEM. diff --git a/common/Values.v b/common/Values.v index 19a8077..056cffb 100644 --- a/common/Values.v +++ b/common/Values.v @@ -46,6 +46,8 @@ Definition Vmone: val := Vint Int.mone. Definition Vtrue: val := Vint Int.one. Definition Vfalse: val := Vint Int.zero. +(** * Operations over values *) + (** The module [Val] defines a number of arithmetic and logical operations over type [val]. Most of these operations are straightforward extensions of the corresponding integer or floating-point operations. *) @@ -984,3 +986,82 @@ Proof. Qed. End Val. + +(** * Values and memory injections *) + +(** A memory injection [f] is a function from addresses to either [None] + or [Some] of an address and an offset. It defines a correspondence + between the blocks of two memory states [m1] and [m2]: +- if [f b = None], the block [b] of [m1] has no equivalent in [m2]; +- if [f b = Some(b', ofs)], the block [b] of [m2] corresponds to + a sub-block at offset [ofs] of the block [b'] in [m2]. +*) + +Definition meminj : Type := block -> option (block * Z). + +(** A memory injection defines a relation between values that is the + identity relation, except for pointer values which are shifted + as prescribed by the memory injection. Moreover, [Vundef] values + inject into any other value. *) + +Inductive val_inject (mi: meminj): val -> val -> Prop := + | val_inject_int: + forall i, val_inject mi (Vint i) (Vint i) + | val_inject_float: + forall f, val_inject mi (Vfloat f) (Vfloat f) + | val_inject_ptr: + forall b1 ofs1 b2 ofs2 delta, + mi b1 = Some (b2, delta) -> + ofs2 = Int.add ofs1 (Int.repr delta) -> + val_inject mi (Vptr b1 ofs1) (Vptr b2 ofs2) + | val_inject_undef: forall v, + val_inject mi Vundef v. + +Hint Resolve val_inject_int val_inject_float val_inject_ptr + val_inject_undef. + +Inductive val_list_inject (mi: meminj): list val -> list val-> Prop:= + | val_nil_inject : + val_list_inject mi nil nil + | val_cons_inject : forall v v' vl vl' , + val_inject mi v v' -> val_list_inject mi vl vl'-> + val_list_inject mi (v :: vl) (v' :: vl'). + +Hint Resolve val_nil_inject val_cons_inject. + +(** Monotone evolution of a memory injection. *) + +Definition inject_incr (f1 f2: meminj) : Prop := + forall b b' delta, f1 b = Some(b', delta) -> f2 b = Some(b', delta). + +Lemma inject_incr_refl : + forall f , inject_incr f f . +Proof. unfold inject_incr. auto. Qed. + +Lemma inject_incr_trans : + forall f1 f2 f3, + inject_incr f1 f2 -> inject_incr f2 f3 -> inject_incr f1 f3 . +Proof . + unfold inject_incr; intros. eauto. +Qed. + +Lemma val_inject_incr: + forall f1 f2 v v', + inject_incr f1 f2 -> + val_inject f1 v v' -> + val_inject f2 v v'. +Proof. + intros. inv H0; eauto. +Qed. + +Lemma val_list_inject_incr: + forall f1 f2 vl vl' , + inject_incr f1 f2 -> val_list_inject f1 vl vl' -> + val_list_inject f2 vl vl'. +Proof. + induction vl; intros; inv H0. auto. + constructor. eapply val_inject_incr; eauto. auto. +Qed. + +Hint Resolve inject_incr_refl val_inject_incr val_list_inject_incr. + -- cgit v1.2.3