Library Events


Observable events, execution traces, and semantics of external calls.

Require Import Coqlib.
Require Intv.
Require Import AST.
Require Import Integers.
Require Import Values.
Require Import Memory.
Require Import Globalenvs.

Events and traces


Require comp.
The observable behaviour of programs is stated in terms of 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 memory states nor pointer values (other than pointers to global variables), 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 parameters, and its result.
  • A volatile load from a global memory location, recording the chunk and address being read and the value just read.
  • A volatile store to a global memory location, recording the chunk and address being written and the value stored there.
  • An annotation, recording the text of the annotation and the values of the arguments.
    The values attached to these events are of the following form. As mentioned above, we do not expose pointer values directly. Pointers relative to a global variable are shown with the name of the variable instead of the block identifier.

Inductive eventval: Type :=
  | EVint: int -> eventval

  | EVptr_global: ident -> int -> eventval.

Inductive aevent: Type :=
  | Aevent_syscall: ident -> list eventval -> eventval -> aevent
  | Aevent_vload: memory_chunk -> ident -> int -> eventval -> aevent
  | Aevent_vstore: memory_chunk -> ident -> int -> eventval -> aevent
  | Aevent_annot: ident -> list eventval -> aevent.

Inductive mem_le (m1 m2 : mem) : Prop := Mem_le {
  mem_le_valid_block:
    forall b,
      Mem.valid_block m1 b -> Mem.valid_block m2 b;

External calls cannot increase the max permissions of a valid block. They can decrease the max permissions, e.g. by freeing.
  mem_le_max_perm:
    forall b ofs p,
      Mem.valid_block m1 b -> Mem.perm m2 b ofs Max p -> Mem.perm m1 b ofs Max p;

External call cannot modify memory unless they have Max, Writable permissions.
  mem_le_readonly:
    forall chunk b ofs,
    Mem.valid_block m1 b ->
    (forall ofs', ofs <= ofs' < ofs + size_chunk chunk ->
                          ~(Mem.perm m1 b ofs' Max Writable)) ->
    Mem.load chunk m2 b ofs = Mem.load chunk m1 b ofs
}.

Lemma mem_le_refl: reflexive _ mem_le.
Proof.
  split; tauto.
Qed.

Lemma mem_le_trans: forall x y, mem_le x y -> forall z, mem_le y z -> mem_le x z.
Proof.
  destruct 1; destruct 1; split; eauto.
  intros.
  etransitivity.
  eapply mem_le_readonly1.
  eauto.
  intros.
  intro.
  eapply H0.
  eassumption.
  eauto.
  eauto.
Qed.

Lemma inhabited_val : inhabited val.
Proof.
 constructor.
 exact Vundef.
Qed.

Definition sw : proc.semworld := proc.Semworld
  (ident * signature)
  aevent
  mem_le_refl
  mem_le_trans
  (list val)
  inhabited_val.

Lemma mem_le_store : forall m1 chunk b o v m2,
  Mem.store chunk m1 b o v = Some m2 ->
  mem_le m1 m2.
Proof.
  split.
   intros; eauto using Mem.store_valid_block_1.
   intros; eauto using Mem.perm_store_2.
   intros.
   eapply Mem.load_store_other.
   eassumption.
   destruct (eq_block b0 b); auto.
   subst.
   right.
   destruct (Mem.store_valid_access_3 _ _ _ _ _ _ H).
   unfold Mem.range_perm in H2.
   destruct (zle (ofs + size_chunk chunk0) o); auto.
   destruct (zle ofs o).
    exfalso.
    eapply H1.
    split.
    eassumption.
    omega.
    eapply Mem.perm_cur_max.
    apply H2.
    generalize (size_chunk_pos chunk); omega.
   clear z.
   right.
   destruct (zle (o + size_chunk chunk) ofs); auto.
   exfalso.
  assert (o <= ofs) by omega.
  eapply H1.
  split.
  eapply Zle_refl.
  generalize (size_chunk_pos chunk0); omega.
  eapply Mem.perm_cur_max.
  eapply H2.
  omega.
Qed.

Lemma mem_le_alloc : forall m lo hi m' b',
  Mem.alloc m lo hi = (m', b') ->
  mem_le m m'.
Proof.
  intros.
  split.
   eauto using Mem.valid_block_alloc.
   intros.
   eapply Mem.perm_alloc_4.
   eassumption.
   assumption.
   intro; subst.
   eapply Mem.fresh_block_alloc.
   eassumption.
   assumption.
  intros.
  eauto using Mem.load_alloc_unchanged.
Qed.

Lemma mem_le_free : forall m b lo hi m',
  Mem.free m b lo hi = Some m' ->
  mem_le m m'.
Proof.
  split.
   intros; eauto using Mem.valid_block_free_1.
   intros.
   eauto using Mem.perm_free_3.
  intros.
  eapply Mem.load_free.
  eassumption.
  destruct (zeq b0 b); auto.
  subst.
  right.
  destruct (zle hi lo).
   left; omega.
  right.
  destruct (zle (ofs + size_chunk chunk) lo); auto.
  destruct (zle hi ofs); auto.
  exfalso.
  generalize (size_chunk_pos chunk).
  intro.
  destruct (zle ofs lo).
   eapply H1.
   split.
   eassumption.
   omega.
   eapply Mem.perm_cur_max.
   eapply Mem.perm_implies.
   eapply Mem.free_range_perm.
   eassumption.
   omega.
   constructor.
  clear z0.
  eapply H1.
  split.
  eapply Zle_refl.
  omega.
   eapply Mem.perm_cur_max.
   eapply Mem.perm_implies.
   eapply Mem.free_range_perm.
   eassumption.
   omega.
   constructor.
Qed.

The dynamic semantics for programs collect traces of events. Traces are of two kinds: finite (type trace) or infinite (type traceinf).

Definition event := comp.xevent sw.

Definition trace := list event.

Definition E0 : trace := nil.

Definition Eapp (t1 t2: trace) : trace := t1 ++ t2.

CoInductive traceinf : Type :=
  | Econsinf: event -> traceinf -> traceinf.

Fixpoint Eappinf (t: trace) (T: traceinf) {struct t} : traceinf :=
  match t with
  | nil => T
  | ev :: t' => Econsinf ev (Eappinf t' T)
  end.

Concatenation of traces is written ** in the finite case or *** in the infinite case.

Infix "**" := Eapp (at level 60, right associativity).
Infix "***" := Eappinf (at level 60, right associativity).

Lemma E0_left: forall t, E0 ** t = t.
Proof. auto. Qed.

Lemma E0_right: forall t, t ** E0 = t.
Proof. intros. unfold E0, Eapp. rewrite <- app_nil_end. auto. Qed.

Lemma Eapp_assoc: forall t1 t2 t3, (t1 ** t2) ** t3 = t1 ** (t2 ** t3).
Proof. intros. unfold Eapp, trace. apply app_ass. Qed.

Lemma Eapp_E0_inv: forall t1 t2, t1 ** t2 = E0 -> t1 = E0 /\ t2 = E0.
Proof (@app_eq_nil event).

Lemma E0_left_inf: forall T, E0 *** T = T.
Proof. auto. Qed.

Lemma Eappinf_assoc: forall t1 t2 T, (t1 ** t2) *** T = t1 *** (t2 *** T).
Proof.
  induction t1; intros; simpl. auto. decEq; auto.
Qed.

Hint Rewrite E0_left E0_right Eapp_assoc
             E0_left_inf Eappinf_assoc: trace_rewrite.

Opaque trace E0 Eapp Eappinf.

The following traceEq tactic proves equalities between traces or infinite traces.

Ltac substTraceHyp :=
  match goal with
  | [ H: (@eq trace ?x ?y) |- _ ] =>
       subst x || clear H
  end.

Ltac decomposeTraceEq :=
  match goal with
  | [ |- (_ ** _) = (_ ** _) ] =>
      apply (f_equal2 Eapp); auto; decomposeTraceEq
  | _ =>
      auto
  end.

Ltac traceEq :=
  repeat substTraceHyp; autorewrite with trace_rewrite; decomposeTraceEq.

Bisimilarity between infinite traces.

CoInductive traceinf_sim: traceinf -> traceinf -> Prop :=
  | traceinf_sim_cons: forall e T1 T2,
      traceinf_sim T1 T2 ->
      traceinf_sim (Econsinf e T1) (Econsinf e T2).

Lemma traceinf_sim_refl:
  forall T, traceinf_sim T T.
Proof.
  cofix COINDHYP; intros.
  destruct T. constructor. apply COINDHYP.
Qed.

Lemma traceinf_sim_sym:
  forall T1 T2, traceinf_sim T1 T2 -> traceinf_sim T2 T1.
Proof.
  cofix COINDHYP; intros. inv H; constructor; auto.
Qed.

Lemma traceinf_sim_trans:
  forall T1 T2 T3,
  traceinf_sim T1 T2 -> traceinf_sim T2 T3 -> traceinf_sim T1 T3.
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.

An alternate presentation of infinite traces as infinite concatenations of nonempty finite traces.

CoInductive traceinf': Type :=
  | Econsinf': forall (t: trace) (T: traceinf'), t <> E0 -> traceinf'.

Program Definition split_traceinf' (t: trace) (T: traceinf') (NE: t <> E0): event * traceinf' :=
  match t with
  | nil => _
  | e :: nil => (e, T)
  | e :: t' => (e, Econsinf' t' T _)
  end.
Next Obligation.
  elimtype False. elim NE. auto.
Qed.
Next Obligation.
  red; intro. elim (H e). rewrite H0. auto.
Qed.

CoFixpoint traceinf_of_traceinf' (T': traceinf') : traceinf :=
  match T' with
  | Econsinf' t T'' NOTEMPTY =>
      let (e, tl) := split_traceinf' t T'' NOTEMPTY in
      Econsinf e (traceinf_of_traceinf' tl)
  end.

Remark unroll_traceinf':
  forall T, T = match T with Econsinf' t T' NE => Econsinf' t T' NE end.
Proof.
  intros. destruct T; auto.
Qed.

Remark unroll_traceinf:
  forall T, T = match T with Econsinf t T' => Econsinf t T' end.
Proof.
  intros. destruct T; auto.
Qed.

Lemma traceinf_traceinf'_app:
  forall t T NE,
  traceinf_of_traceinf' (Econsinf' t T NE) = t *** traceinf_of_traceinf' T.
Proof.
  induction t.
  intros. elim NE. auto.
  intros. simpl.
  rewrite (unroll_traceinf (traceinf_of_traceinf' (Econsinf' (a :: t) T NE))).
  simpl. destruct t. auto.
Transparent Eappinf.
  simpl. f_equal. apply IHt.
Qed.

Prefixes of traces.

Definition trace_prefix (t1 t2: trace) :=
  exists t3, t2 = t1 ** t3.

Definition traceinf_prefix (t1: trace) (T2: traceinf) :=
  exists T3, T2 = t1 *** T3.

Lemma trace_prefix_app:
  forall t1 t2 t,
  trace_prefix t1 t2 ->
  trace_prefix (t ** t1) (t ** t2).
Proof.
  intros. destruct H as [t3 EQ]. exists t3. traceEq.
Qed.

Lemma traceinf_prefix_app:
  forall t1 T2 t,
  traceinf_prefix t1 T2 ->
  traceinf_prefix (t ** t1) (t *** T2).
Proof.
  intros. destruct H as [T3 EQ]. exists T3. subst T2. traceEq.
Qed.

Relating values and event values


Notation Event_syscall i l e := (comp.Event (sw := sw) (Aevent_syscall i l e)).
Notation Event_vload c id i ev := (comp.Event (sw := sw) (Aevent_vload c id i ev)).
Notation Event_vstore c id i ev := (comp.Event (sw := sw) (Aevent_vstore c id i ev)).
Notation Event_annot i l := (comp.Event (sw := sw) (Aevent_annot i l)).

for patterns
Notation PEvent_syscall i l e := (comp.Event (Aevent_syscall i l e)).
Notation PEvent_vload c id i ev := (comp.Event (Aevent_vload c id i ev)).
Notation PEvent_vstore c id i ev := (comp.Event (Aevent_vstore c id i ev)).
Notation PEvent_annot i l := (comp.Event (Aevent_annot i l)).

Set Implicit Arguments.

Section EVENTVAL.

Global environment used to translate between global variable names and their block identifiers.
Variables FN F V: Type.
Variable ge: Genv.t FN F V.

Translation between values and event values.

Inductive eventval_match: eventval -> typ -> val -> Prop :=
  | ev_match_int: forall i,
      eventval_match (EVint i) Tint (Vint i)

  | ev_match_ptr: forall id b ofs,
      Genv.find_symbol ge id = Some b ->
      eventval_match (EVptr_global id ofs) Tint (Vptr b ofs).

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).

Some properties of these translation predicates.

Lemma eventval_match_type:
  forall ev ty v,
  eventval_match ev ty v -> Val.has_type v ty.
Proof.
  intros. inv H; constructor.
Qed.

Lemma 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.

Lemma 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; auto.
Qed.

Lemma 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.

Compatibility with memory injections

Variable f: block -> option (block * Z).

Definition meminj_preserves_globals : Prop :=
     (forall id b, Genv.find_symbol ge id = Some b -> f b = Some(b, 0))
  /\ (forall b gv, Genv.find_var_info ge b = Some gv -> f b = Some(b, 0))
  /\ (forall b1 b2 delta gv, Genv.find_var_info ge b2 = Some gv -> f b1 = Some(b2, delta) -> b2 = b1).

Hypothesis glob_pres: meminj_preserves_globals.

Lemma eventval_match_inject:
  forall 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.   destruct glob_pres as [A [B C]].
  exploit A; eauto. intro EQ; rewrite H3 in EQ; inv EQ.
  rewrite Int.add_zero. econstructor; eauto.
Qed.

Lemma eventval_match_inject_2:
  forall ev ty v,
  eventval_match ev ty v -> val_inject f v v.
Proof.
  induction 1. constructor.   destruct glob_pres as [A [B C]].
  exploit A; eauto. intro EQ.
  econstructor; eauto. rewrite Int.add_zero; auto.
Qed.

Lemma 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. inv H; constructor.
  inv H1. constructor. eapply eventval_match_inject; eauto. eauto.
Qed.

Determinism

Lemma 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. congruence.
Qed.

Lemma 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.
  decEq. eapply Genv.genv_vars_inj; eauto.
Qed.

Lemma 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.

Validity

Definition eventval_valid (ev: eventval) : Prop :=
  match ev with
  | EVint _ => True

  | EVptr_global id ofs => exists b, Genv.find_symbol ge id = Some b
  end.

Definition eventval_type (ev: eventval) : typ :=
  match ev with
  | EVint _ => Tint

  | EVptr_global id ofs => Tint
  end.

Lemma eventval_valid_match:
  forall ev ty,
  eventval_valid ev -> eventval_type ev = ty -> exists v, eventval_match ev ty v.
Proof.
  intros. subst ty. destruct ev; simpl in *.
  exists (Vint i); constructor.
  destruct H as [b A]. exists (Vptr b i0); constructor; auto.
Qed.

Lemma eventval_match_valid:
  forall ev ty v,
  eventval_match ev ty v -> eventval_valid ev /\ eventval_type ev = ty.
Proof.
  induction 1; simpl; auto. split; auto. exists b; auto.
Qed.

End EVENTVAL.

Invariance under changes to the global environment

Section EVENTVAL_INV.

Variables FN1 F1 V1 FN2 F2 V2: Type.
Variable ge1: Genv.t FN1 F1 V1.
Variable ge2: Genv.t FN2 F2 V2.

Hypothesis symbols_preserved:
  forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id.

Lemma eventval_match_preserved:
  forall ev ty v,
  eventval_match ge1 ev ty v -> eventval_match ge2 ev ty v.
Proof.
  induction 1; constructor. rewrite symbols_preserved; auto.
Qed.

Lemma eventval_list_match_preserved:
  forall evl tyl vl,
  eventval_list_match ge1 evl tyl vl -> eventval_list_match ge2 evl tyl vl.
Proof.
  induction 1; constructor; auto. eapply eventval_match_preserved; eauto.
Qed.

Lemma eventval_valid_preserved:
  forall ev, eventval_valid ge1 ev -> eventval_valid ge2 ev.
Proof.
  unfold eventval_valid; destruct ev; auto.
  intros [b A]. exists b; rewrite symbols_preserved; auto.
Qed.

End EVENTVAL_INV.

Matching traces.


Section MATCH_TRACES.

Variables FN F V: Type.
Variable ge: Genv.t FN F V.

Matching between traces corresponding to single transitions. Arguments (provided by the program) must be equal. Results (provided by the outside world) can vary as long as they can be converted safely to values.

Inductive match_traces: trace -> trace -> Prop :=
  | match_traces_E0:
      match_traces nil nil
  | match_traces_syscall: forall id args res1 res2,
      eventval_valid ge res1 -> eventval_valid ge res2 -> eventval_type res1 = eventval_type res2 ->
      match_traces (Event_syscall id args res1 :: nil) (Event_syscall id args res2 :: nil)
  | match_traces_vload: forall chunk id ofs res1 res2,
      eventval_valid ge res1 -> eventval_valid ge res2 -> eventval_type res1 = eventval_type res2 ->
      match_traces (Event_vload chunk id ofs res1 :: nil) (Event_vload chunk id ofs res2 :: nil)
  | match_traces_vstore: forall chunk id ofs arg,
      match_traces (Event_vstore chunk id ofs arg :: nil) (Event_vstore chunk id ofs arg :: nil)
  | match_traces_annot: forall id args,
      match_traces (Event_annot id args :: nil) (Event_annot id args :: nil)
  | match_traces_external: forall fs a m r1 m1 r2 m2 (Hle: mem_le m m1 <-> mem_le m m2),
    match_traces (comp.External (sw := sw) fs a m r1 m1 :: nil) (comp.External fs a m r2 m2 :: nil)
.

End MATCH_TRACES.

Invariance by change of global environment

Section MATCH_TRACES_INV.

Variables FN1 F1 V1 FN2 F2 V2: Type.
Variable ge1: Genv.t FN1 F1 V1.
Variable ge2: Genv.t FN2 F2 V2.

Hypothesis symbols_preserved:
  forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id.

Lemma match_traces_preserved:
  forall t1 t2, match_traces ge1 t1 t2 -> match_traces ge2 t1 t2.
Proof.
  induction 1; constructor; auto; eapply eventval_valid_preserved; eauto.
Qed.

End MATCH_TRACES_INV.

An output trace is a trace composed only of output events, that is, events that do not take any result from the outside world.

Definition output_event (ev: event) : Prop :=
  match ev with
  | PEvent_syscall _ _ _ => False
  | PEvent_vload _ _ _ _ => False
  | PEvent_vstore _ _ _ _ => True
  | PEvent_annot _ _ => True
  | comp.External _ _ _ _ _ => False
  end.

Fixpoint output_trace (t: trace) : Prop :=
  match t with
  | nil => True
  | ev :: t' => output_event ev /\ output_trace t'
  end.

Semantics of volatile memory accesses


Definition block_is_volatile (FN F V: Type) (ge: Genv.t FN F V) (b: block) : bool :=
  match Genv.find_var_info ge b with
  | None => false
  | Some gv => gv.(gvar_volatile)
  end.

Inductive volatile_load (FN F V: Type) (ge: Genv.t FN F V):
                   memory_chunk -> mem -> block -> int -> trace -> val -> Prop :=
  | volatile_load_vol: forall chunk m b ofs id ev v,
      block_is_volatile ge b = true ->
      Genv.find_symbol ge id = Some b ->
      eventval_match ge ev (type_of_chunk chunk) v ->
      volatile_load ge chunk m b ofs
                      (Event_vload chunk id ofs ev :: nil)
                      (Val.load_result chunk v)
  | volatile_load_nonvol: forall chunk m b ofs v,
      block_is_volatile ge b = false ->
      Mem.load chunk m b (Int.unsigned ofs) = Some v ->
      volatile_load ge chunk m b ofs E0 v.

Inductive volatile_store (FN F V: Type) (ge: Genv.t FN F V):
                  memory_chunk -> mem -> block -> int -> val -> trace -> mem -> Prop :=
  | volatile_store_vol: forall chunk m b ofs id ev v,
      block_is_volatile ge b = true ->
      Genv.find_symbol ge id = Some b ->
      eventval_match ge ev (type_of_chunk chunk) v ->
      volatile_store ge chunk m b ofs v
                      (Event_vstore chunk id ofs ev :: nil)
                      m
  | volatile_store_nonvol: forall chunk m b ofs v m',
      block_is_volatile ge b = false ->
      Mem.store chunk m b (Int.unsigned ofs) v = Some m' ->
      volatile_store ge chunk m b ofs v E0 m'.

Semantics of external functions

For each external function, its behavior is defined by a predicate relating:
  • the global environment
  • 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).

Definition extcall_sem : Type :=
  forall (FN F V: Type), Genv.t FN F V -> list val -> mem -> trace -> val -> mem -> Prop.

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 k p,
   P b ofs -> Mem.perm m_before b ofs k p -> Mem.perm m_after b ofs k 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 :=
  ~Mem.perm m b ofs Max Nonempty.

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) -> ~Mem.perm m b0 (ofs - delta) Max Nonempty.

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.

Record extcall_properties (sem: extcall_sem)
                          (sg: signature) : Prop := mk_extcall_properties {

The return value of an external call must agree with its signature. No longer true in the context of separate compilation. A substitute for Linear-to-Mach will have to be found.
ec_well_typed: forall F V (ge: Genv.t F V) vargs m1 t vres m2, sem F V ge 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 FN F V (ge: Genv.t FN F V) vargs m1 t vres m2, sem FN F V ge vargs m1 t vres m2 -> List.length vargs = List.length sg.(sig_args);
The semantics is invariant under change of global environment that preserves symbols.
  ec_symbols_preserved:
    forall FN1 F1 V1 (ge1: Genv.t FN1 F1 V1) FN2 F2 V2 (ge2: Genv.t FN2 F2 V2) vargs m1 t vres m2,
    (forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id) ->
    (forall b, block_is_volatile ge2 b = block_is_volatile ge1 b) ->
    sem FN1 F1 V1 ge1 vargs m1 t vres m2 ->
    sem FN2 F2 V2 ge2 vargs m1 t vres m2;

External calls cannot invalidate memory blocks. (Remember that freeing a block does not invalidate its block identifier.)
  ec_valid_block:
    forall FN F V (ge: Genv.t FN F V) vargs m1 t vres m2 b,
    sem FN F V ge vargs m1 t vres m2 ->
    Mem.valid_block m1 b -> Mem.valid_block m2 b;

External calls cannot increase the max permissions of a valid block. They can decrease the max permissions, e.g. by freeing.
  ec_max_perm:
    forall FN F V (ge: Genv.t FN F V) vargs m1 t vres m2 b ofs p,
    sem FN F V ge vargs m1 t vres m2 ->
    Mem.valid_block m1 b -> Mem.perm m2 b ofs Max p -> Mem.perm m1 b ofs Max p;

External call cannot modify memory unless they have Max, Writable permissions.
  ec_readonly:
    forall FN F V (ge: Genv.t FN F V) vargs m1 t vres m2 chunk b ofs,
    sem FN F V ge vargs m1 t vres m2 ->
    Mem.valid_block m1 b ->
    (forall ofs', ofs <= ofs' < ofs + size_chunk chunk ->
                          ~(Mem.perm m1 b ofs' Max Writable)) ->
    Mem.load chunk m2 b ofs = Mem.load chunk m1 b ofs;

External calls must commute with memory extensions, in the following sense. TO DROP for now ec_mem_extends: forall FN F V (ge: Genv.t FN F V) vargs m1 t vres m2 m1' vargs', sem FN F V ge vargs m1 t vres m2 -> Mem.extends m1 m1' -> Val.lessdef_list vargs vargs' -> exists vres', exists m2', sem FN F V ge 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 FN F V (ge: Genv.t FN F V) vargs m1 t vres m2 f m1' vargs', meminj_preserves_globals ge f -> sem FN F V ge vargs m1 t vres m2 -> Mem.inject f m1 m1' -> val_list_inject f vargs vargs' -> exists f', exists vres', exists m2', sem FN F V ge 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 produce at most one event.
  ec_trace_length:
    forall FN F V ge vargs m t vres m',
    sem FN F V ge vargs m t vres m' -> (length t <= 1)%nat;

External calls must be receptive to changes of traces by another, matching trace.
  ec_receptive:
    forall FN F V ge vargs m t1 vres1 m1 t2,
    sem FN F V ge vargs m t1 vres1 m1 -> match_traces ge t1 t2 ->
    exists vres2, exists m2, sem FN F V ge vargs m t2 vres2 m2;

External calls must be deterministic up to matching between traces.
  ec_determ:
    forall FN F V ge vargs m t1 vres1 m1 t2 vres2 m2,
    sem FN F V ge vargs m t1 vres1 m1 -> sem FN F V ge vargs m t2 vres2 m2 ->
    match_traces ge t1 t2 /\ (t1 = t2 -> vres1 = vres2 /\ m1 = m2)
}.

Semantics of volatile loads


Inductive volatile_load_sem (chunk: memory_chunk) (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | volatile_load_sem_intro: forall b ofs m t v,
      volatile_load ge chunk m b ofs t v ->
      volatile_load_sem chunk ge (Vptr b ofs :: nil) m t v m.

Lemma volatile_load_preserved:
  forall FN1 F1 V1 (ge1: Genv.t FN1 F1 V1) FN2 F2 V2 (ge2: Genv.t FN2 F2 V2) chunk m b ofs t v,
  (forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id) ->
  (forall b, block_is_volatile ge2 b = block_is_volatile ge1 b) ->
  volatile_load ge1 chunk m b ofs t v ->
  volatile_load ge2 chunk m b ofs t v.
Proof.
  intros. inv H1; constructor; auto.
  rewrite H0; auto.
  rewrite H; auto.
  eapply eventval_match_preserved; eauto.
  rewrite H0; auto.
Qed.

Lemma volatile_load_extends:
  forall FN F V (ge: Genv.t FN F V) chunk m b ofs t v m',
  volatile_load ge chunk m b ofs t v ->
  Mem.extends m m' ->
  exists v', volatile_load ge chunk m' b ofs t v' /\ Val.lessdef v v'.
Proof.
  intros. inv H.
  econstructor; split; eauto. econstructor; eauto.
  exploit Mem.load_extends; eauto. intros [v' [A B]]. exists v'; split; auto. constructor; auto.
Qed.

Remark meminj_preserves_block_is_volatile:
  forall FN F V (ge: Genv.t FN F V) f b1 b2 delta,
  meminj_preserves_globals ge f ->
  f b1 = Some (b2, delta) ->
  block_is_volatile ge b2 = block_is_volatile ge b1.
Proof.
  intros. destruct H as [A [B C]]. unfold block_is_volatile.
  case_eq (Genv.find_var_info ge b1); intros.
  exploit B; eauto. intro EQ; rewrite H0 in EQ; inv EQ. rewrite H; auto.
  case_eq (Genv.find_var_info ge b2); intros.
  exploit C; eauto. intro EQ. congruence.
  auto.
Qed.

Lemma volatile_load_inject:
  forall FN F V (ge: Genv.t FN F V) f chunk m b ofs t v b' ofs' m',
  meminj_preserves_globals ge f ->
  volatile_load ge chunk m b ofs t v ->
  val_inject f (Vptr b ofs) (Vptr b' ofs') ->
  Mem.inject f m m' ->
  exists v', volatile_load ge chunk m' b' ofs' t v' /\ val_inject f v v'.
Proof.
  intros. inv H0.
  inv H1. exploit (proj1 H); eauto. intros EQ; rewrite H8 in EQ; inv EQ.
  rewrite Int.add_zero. exists (Val.load_result chunk v0); split.
  constructor; auto.
  apply val_load_result_inject. eapply eventval_match_inject_2; eauto.
  exploit Mem.loadv_inject; eauto. simpl; eauto. simpl; intros [v' [A B]]. exists v'; split; auto.
  constructor; auto. rewrite <- H3. inv H1. eapply meminj_preserves_block_is_volatile; eauto.
Qed.

Lemma volatile_load_receptive:
  forall FN F V (ge: Genv.t FN F V) chunk m b ofs t1 t2 v1,
  volatile_load ge chunk m b ofs t1 v1 -> match_traces ge t1 t2 ->
  exists v2, volatile_load ge chunk m b ofs t2 v2.
Proof.
  intros. inv H; inv H0.
  exploit eventval_match_valid; eauto. intros [A B].
  exploit eventval_valid_match. eexact H9. rewrite <- H10; eauto.
  intros [v' EVM]. exists (Val.load_result chunk v'). constructor; auto.
  exists v1; constructor; auto.
Qed.

Lemma volatile_load_ok:
  forall chunk,
  extcall_properties (volatile_load_sem chunk)
                     (mksignature (Tint :: nil) (Some (type_of_chunk chunk))).
Proof.
  intros; constructor; intros.
  inv H1. constructor. eapply volatile_load_preserved; eauto.
  inv H; auto.
  inv H; auto.
  inv H; auto.
  inv H; inv H0; simpl; omega.
  inv H. exploit volatile_load_receptive; eauto. intros [v2 A].
  exists v2; exists m1; constructor; auto.
  inv H; inv H0. inv H1; inv H7; try congruence.
  assert (id = id0) by (eapply Genv.genv_vars_inj; eauto). subst id0.
  exploit eventval_match_valid. eexact H2. intros [V1 T1].
  exploit eventval_match_valid. eexact H4. intros [V2 T2].
  split. constructor; auto. congruence.
  intros EQ; inv EQ.
  assert (v = v0) by (eapply eventval_match_determ_1; eauto). subst v0.
  auto.
  split. constructor. intuition congruence.
Qed.

Inductive volatile_load_global_sem (chunk: memory_chunk) (id: ident) (ofs: int)
                                   (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | volatile_load_global_sem_intro: forall b t v m,
      Genv.find_symbol ge id = Some b ->
      volatile_load ge chunk m b ofs t v ->
      volatile_load_global_sem chunk id ofs ge nil m t v m.

Remark volatile_load_global_charact:
  forall chunk id ofs (FN F V: Type) (ge: Genv.t FN F V) vargs m t vres m',
  volatile_load_global_sem chunk id ofs ge vargs m t vres m' <->
  exists b, Genv.find_symbol ge id = Some b /\ volatile_load_sem chunk ge (Vptr b ofs :: vargs) m t vres m'.
Proof.
  intros; split.
  intros. inv H. exists b; split; auto. constructor; auto.
  intros [b [P Q]]. inv Q. econstructor; eauto.
Qed.

Lemma volatile_load_global_ok:
  forall chunk id ofs,
  extcall_properties (volatile_load_global_sem chunk id ofs)
                     (mksignature nil (Some (type_of_chunk chunk))).
Proof.
  intros; constructor; intros.
  inv H1. econstructor. rewrite H; eauto. eapply volatile_load_preserved; eauto.
  inv H; auto.
  inv H; auto.
  inv H; auto.
  inv H; inv H1; simpl; omega.
  inv H. exploit volatile_load_receptive; eauto. intros [v2 A].
  exists v2; exists m1; econstructor; eauto.
  rewrite volatile_load_global_charact in *.
  destruct H as [b1 [A1 B1]]. destruct H0 as [b2 [A2 B2]].
  rewrite A1 in A2; inv A2.
  eapply ec_determ. eapply volatile_load_ok. eauto. eauto.
Qed.

Semantics of volatile stores


Inductive volatile_store_sem (chunk: memory_chunk) (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | volatile_store_sem_intro: forall b ofs m1 v t m2,
      volatile_store ge chunk m1 b ofs v t m2 ->
      volatile_store_sem chunk ge (Vptr b ofs :: v :: nil) m1 t Vundef m2.

Lemma volatile_store_preserved:
  forall FN1 F1 V1 (ge1: Genv.t FN1 F1 V1) FN2 F2 V2 (ge2: Genv.t FN2 F2 V2) chunk m1 b ofs v t m2,
  (forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id) ->
  (forall b, block_is_volatile ge2 b = block_is_volatile ge1 b) ->
  volatile_store ge1 chunk m1 b ofs v t m2 ->
  volatile_store ge2 chunk m1 b ofs v t m2.
Proof.
  intros. inv H1; constructor; auto.
  rewrite H0; auto.
  rewrite H; auto.
  eapply eventval_match_preserved; eauto.
  rewrite H0; auto.
Qed.

Lemma volatile_store_readonly:
  forall FN F V (ge: Genv.t FN F V) chunk1 m1 b1 ofs1 v t m2 chunk ofs b,
  volatile_store ge chunk1 m1 b1 ofs1 v t m2 ->
  Mem.valid_block m1 b ->
  (forall ofs', ofs <= ofs' < ofs + size_chunk chunk ->
                        ~(Mem.perm m1 b ofs' Max Writable)) ->
  Mem.load chunk m2 b ofs = Mem.load chunk m1 b ofs.
Proof.
  intros. inv H.
  auto.
  eapply Mem.load_store_other; eauto.
  destruct (eq_block b b1); auto. subst b1. right.
  apply (Intv.range_disjoint' (ofs, ofs + size_chunk chunk)
                              (Int.unsigned ofs1, Int.unsigned ofs1 + size_chunk chunk1)).
  red; intros; red; intros.
  elim (H1 x); auto.
  exploit Mem.store_valid_access_3; eauto. intros [A B].
  apply Mem.perm_cur_max. apply A. auto.
  simpl. generalize (size_chunk_pos chunk); omega.
  simpl. generalize (size_chunk_pos chunk1); omega.
Qed.

Lemma volatile_store_extends:
  forall FN F V (ge: Genv.t FN F V) chunk m1 b ofs v t m2 m1' v',
  volatile_store ge chunk m1 b ofs v t m2 ->
  Mem.extends m1 m1' ->
  Val.lessdef v v' ->
  exists m2',
     volatile_store ge chunk m1' b ofs v' t m2'
  /\ Mem.extends m2 m2'
  /\ mem_unchanged_on (loc_out_of_bounds m1) m1' m2'.
Proof.
  intros. inv H.
  econstructor; split. econstructor; eauto. eapply eventval_match_lessdef; eauto.
  split. auto. red; auto.
  exploit Mem.store_within_extends; eauto. intros [m2' [A B]].
  exists m2'; intuition. econstructor; eauto.
  red; split; intros.
  eapply Mem.perm_store_1; eauto.
  rewrite <- H4. eapply Mem.load_store_other; eauto.
  destruct (eq_block b0 b); auto. subst b0; right.
  apply (Intv.range_disjoint' (ofs0, ofs0 + size_chunk chunk0)
                              (Int.unsigned ofs, Int.unsigned ofs + size_chunk chunk)).
  red; intros; red; intros.
  exploit (H x H5). exploit Mem.store_valid_access_3. eexact H3. intros [E G].
  apply Mem.perm_cur_max. apply Mem.perm_implies with Writable; auto with mem.
  auto.
  simpl. generalize (size_chunk_pos chunk0). omega.
  simpl. generalize (size_chunk_pos chunk). omega.
Qed.

Lemma volatile_store_inject:
  forall FN F V (ge: Genv.t FN F V) f chunk m1 b ofs v t m2 m1' b' ofs' v',
  meminj_preserves_globals ge f ->
  volatile_store ge chunk m1 b ofs v t m2 ->
  val_inject f (Vptr b ofs) (Vptr b' ofs') ->
  val_inject f v v' ->
  Mem.inject f m1 m1' ->
  exists m2',
       volatile_store ge chunk m1' b' ofs' v' t m2'
    /\ Mem.inject f m2 m2'
    /\ mem_unchanged_on (loc_unmapped f) m1 m2
    /\ mem_unchanged_on (loc_out_of_reach f m1) m1' m2'.
Proof.
  intros. inv H0.
  inv H1. exploit (proj1 H); eauto. intros EQ; rewrite H9 in EQ; inv EQ.
  rewrite Int.add_zero. exists m1'.
  split. constructor; auto. eapply eventval_match_inject; eauto.
  split. auto. split. red; auto. red; auto.
  assert (Mem.storev chunk m1 (Vptr b ofs) v = Some m2). simpl; auto.
  exploit Mem.storev_mapped_inject; eauto. intros [m2' [A B]].
  inv H1. exists m2'; intuition.
  constructor; auto. rewrite <- H4. eapply meminj_preserves_block_is_volatile; eauto.
  split; intros. eapply Mem.perm_store_1; eauto.
  rewrite <- H6. 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 <- H6. eapply Mem.load_store_other; eauto.
  destruct (eq_block b0 b'); auto. subst b0; right.
  assert (EQ: Int.unsigned (Int.add ofs (Int.repr delta)) = Int.unsigned ofs + delta).
    eapply Mem.address_inject; eauto with mem.
  unfold Mem.storev in A. rewrite EQ in A. rewrite EQ.
  apply (Intv.range_disjoint' (ofs0, ofs0 + size_chunk chunk0)
                              (Int.unsigned ofs + delta, Int.unsigned ofs + delta + size_chunk chunk)).
  red; intros; red; intros. exploit (H1 x H7). eauto.
  exploit Mem.store_valid_access_3. eexact H0. intros [C D].
  apply Mem.perm_cur_max. apply Mem.perm_implies with Writable; auto with mem.
  apply C. red in H8; simpl in H8. omega.
  auto.
  simpl. generalize (size_chunk_pos chunk0). omega.
  simpl. generalize (size_chunk_pos chunk). omega.
Qed.

Lemma volatile_store_receptive:
  forall FN F V (ge: Genv.t FN F V) chunk m b ofs v t1 m1 t2,
  volatile_store ge chunk m b ofs v t1 m1 -> match_traces ge t1 t2 -> t1 = t2.
Proof.
  intros. inv H; inv H0; auto.
Qed.

Lemma volatile_store_ok:
  forall chunk,
  extcall_properties (volatile_store_sem chunk)
                     (mksignature (Tint :: type_of_chunk chunk :: nil) None).
Proof.
  intros; constructor; intros.
  inv H1. constructor. eapply volatile_store_preserved; eauto.
  inv H. inv H1. auto. eauto with mem.
  inv H. inv H2. auto. eauto with mem.
  inv H. eapply volatile_store_readonly; eauto.
  inv H; inv H0; simpl; omega.
  assert (t1 = t2). inv H. eapply volatile_store_receptive; eauto.
  subst t2; exists vres1; exists m1; auto.
  inv H; inv H0. inv H1; inv H8; try congruence.
  assert (id = id0) by (eapply Genv.genv_vars_inj; eauto). subst id0.
  assert (ev = ev0) by (eapply eventval_match_determ_2; eauto). subst ev0.
  split. constructor. auto.
  split. constructor. intuition congruence.
Qed.

Inductive volatile_store_global_sem (chunk: memory_chunk) (id: ident) (ofs: int)
                                   (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | volatile_store_global_sem_intro: forall b m1 v t m2,
      Genv.find_symbol ge id = Some b ->
      volatile_store ge chunk m1 b ofs v t m2 ->
      volatile_store_global_sem chunk id ofs ge (v :: nil) m1 t Vundef m2.

Remark volatile_store_global_charact:
  forall chunk id ofs (FN F V: Type) (ge: Genv.t FN F V) vargs m t vres m',
  volatile_store_global_sem chunk id ofs ge vargs m t vres m' <->
  exists b, Genv.find_symbol ge id = Some b /\ volatile_store_sem chunk ge (Vptr b ofs :: vargs) m t vres m'.
Proof.
  intros; split.
  intros. inv H; exists b; split; auto; econstructor; eauto.
  intros [b [P Q]]. inv Q. econstructor; eauto.
Qed.

Lemma volatile_store_global_ok:
  forall chunk id ofs,
  extcall_properties (volatile_store_global_sem chunk id ofs)
                     (mksignature (type_of_chunk chunk :: nil) None).
Proof.
  intros; constructor; intros.
  inv H1. econstructor. rewrite H; eauto. eapply volatile_store_preserved; eauto.
  inv H. inv H2. auto. eauto with mem.
  inv H. inv H3. auto. eauto with mem.
  inv H. eapply volatile_store_readonly; eauto.
  inv H. inv H1; simpl; omega.
  assert (t1 = t2). inv H. eapply volatile_store_receptive; eauto. subst t2.
  exists vres1; exists m1; congruence.
  rewrite volatile_store_global_charact in *.
  destruct H as [b1 [A1 B1]]. destruct H0 as [b2 [A2 B2]]. rewrite A1 in A2; inv A2.
  eapply ec_determ. eapply volatile_store_ok. eauto. eauto.
Qed.

Semantics of dynamic memory allocation (malloc)


Inductive extcall_malloc_sem (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | extcall_malloc_sem_intro: forall n m m' b m'',
      Mem.alloc m (-4) (Int.unsigned n) = (m', b) ->
      Mem.store Mint32 m' b (-4) (Vint n) = Some m'' ->
      extcall_malloc_sem ge (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.unsigned 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 H1; econstructor; eauto.
  inv H. eauto with mem.
  inv H. exploit Mem.perm_alloc_inv. eauto. eapply Mem.perm_store_2; eauto.
  rewrite zeq_false. auto.
  apply Mem.valid_not_valid_diff with m1; eauto with mem.
  inv H. transitivity (Mem.load chunk m' b ofs).
  eapply Mem.load_store_other; eauto.
  left. apply Mem.valid_not_valid_diff with m1; eauto with mem.
  eapply Mem.load_alloc_unchanged; eauto.
  inv H; simpl; omega.
  assert (t1 = t2). inv H; inv H0; auto. subst t2.
  exists vres1; exists m1; auto.
  inv H; inv H0. split. constructor. intuition congruence.
Qed.

Semantics of dynamic memory deallocation (free)


Inductive extcall_free_sem (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | extcall_free_sem_intro: forall b lo sz m m',
      Mem.load Mint32 m b (Int.unsigned lo - 4) = Some (Vint sz) ->
      Int.unsigned sz > 0 ->
      Mem.free m b (Int.unsigned lo - 4) (Int.unsigned lo + Int.unsigned sz) = Some m' ->
      extcall_free_sem ge (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 H1; econstructor; eauto.
  inv H. eauto with mem.
  inv H. eapply Mem.perm_free_3; eauto.
  inv H. eapply Mem.load_free; eauto.
  destruct (eq_block b b0); auto.
  subst b0. right; right.
  apply (Intv.range_disjoint'
           (ofs, ofs + size_chunk chunk)
           (Int.unsigned lo - 4, Int.unsigned lo + Int.unsigned sz)).
  red; intros; red; intros.
  elim (H1 x). auto. apply Mem.perm_cur_max.
  apply Mem.perm_implies with Freeable; auto with mem.
  exploit Mem.free_range_perm; eauto.
  simpl. generalize (size_chunk_pos chunk); omega.
  simpl. omega.
  inv H; simpl; omega.
  assert (t1 = t2). inv H; inv H0; auto. subst t2.
  exists vres1; exists m1; auto.
  inv H; inv H0. split. constructor. intuition congruence.
Qed.

Semantics of memcpy operations.


Inductive extcall_memcpy_sem (sz al: Z) (FN F V: Type) (ge: Genv.t FN F V): list val -> mem -> trace -> val -> mem -> Prop :=
  | extcall_memcpy_sem_intro: forall bdst odst bsrc osrc m bytes m',
      al = 1 \/ al = 2 \/ al = 4 \/ al = 8 -> sz > 0 ->
      (al | sz) -> (al | Int.unsigned osrc) -> (al | Int.unsigned odst) ->
      bsrc <> bdst \/ Int.unsigned osrc = Int.unsigned odst
                   \/ Int.unsigned osrc + sz <= Int.unsigned odst
                   \/ Int.unsigned odst + sz <= Int.unsigned osrc ->
      Mem.loadbytes m bsrc (Int.unsigned osrc) sz = Some bytes ->
      Mem.storebytes m bdst (Int.unsigned odst) bytes = Some m' ->
      extcall_memcpy_sem sz al ge (Vptr bdst odst :: Vptr bsrc osrc :: nil) m E0 Vundef m'.

Lemma extcall_memcpy_ok:
  forall sz al,
  extcall_properties (extcall_memcpy_sem sz al) (mksignature (Tint :: Tint :: nil) None).
Proof.
  intros. constructor.
  intros. inv H1. econstructor; eauto.
  intros. inv H. eauto with mem.
  intros. inv H. eapply Mem.perm_storebytes_2; eauto.
  intros. inv H. eapply Mem.load_storebytes_other; eauto.
  destruct (eq_block b bdst); auto. subst b. right.
  apply (Intv.range_disjoint'
          (ofs, ofs + size_chunk chunk)
          (Int.unsigned odst, Int.unsigned odst + Z_of_nat (length bytes))).
  red; intros; red; intros. elim (H1 x); auto.
  apply Mem.perm_cur_max.
  eapply Mem.storebytes_range_perm; eauto.
  simpl. generalize (size_chunk_pos chunk); omega.
  simpl. rewrite (Mem.loadbytes_length _ _ _ _ _ H8). rewrite nat_of_Z_eq.
  omega. omega.
  intros; inv H. simpl; omega.
  intros.
  assert (t1 = t2). inv H; inv H0; auto. subst t2.
  exists vres1; exists m1; auto.
  intros; inv H; inv H0. split. constructor. intros; split; congruence.
Qed.

Semantics of system calls.


Inductive extcall_io_sem (name: ident) (sg: signature) (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | extcall_io_sem_intro: forall vargs m args res vres,
      eventval_list_match ge args (sig_args sg) vargs ->
      eventval_match ge res (proj_sig_res sg) vres ->
      extcall_io_sem name sg ge vargs m (Event_syscall name args res :: E0) vres m.

Lemma extcall_io_ok:
  forall name sg,
  extcall_properties (extcall_io_sem name sg) sg.
Proof.
  intros; constructor; intros.
  inv H1. econstructor; eauto.
  eapply eventval_list_match_preserved; eauto.
  eapply eventval_match_preserved; eauto.
  inv H; auto.
  inv H; auto.
  inv H; auto.
  inv H; simpl; omega.
  inv H; inv H0.
  exploit eventval_match_valid; eauto. intros [A B].
  exploit eventval_valid_match. eexact H7. rewrite <- H8; eauto.
  intros [v' EVM]. exists v'; exists m1. econstructor; eauto.
  inv H; inv H0.
  assert (args = args0). eapply eventval_list_match_determ_2; eauto. subst args0.
  exploit eventval_match_valid. eexact H2. intros [V1 T1].
  exploit eventval_match_valid. eexact H3. intros [V2 T2].
  split. constructor; auto. congruence.
  intros EQ; inv EQ. split; auto. eapply eventval_match_determ_1; eauto.
Qed.

Semantics of annotation.


Inductive extcall_annot_sem (text: ident) (targs: list typ) (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | extcall_annot_sem_intro: forall vargs m args,
      eventval_list_match ge args targs vargs ->
      extcall_annot_sem text targs ge vargs m (Event_annot text args :: E0) Vundef m.

Lemma extcall_annot_ok:
  forall text targs,
  extcall_properties (extcall_annot_sem text targs) (mksignature targs None).
Proof.
  intros; constructor; intros.
  inv H1. econstructor; eauto.
  eapply eventval_list_match_preserved; eauto.
  inv H; auto.
  inv H; auto.
  inv H; auto.
  inv H; simpl; omega.
  assert (t1 = t2). inv H; inv H0; auto.
  exists vres1; exists m1; congruence.
  inv H; inv H0.
  assert (args = args0). eapply eventval_list_match_determ_2; eauto. subst args0.
  split. constructor. auto.
Qed.

Inductive extcall_annot_val_sem (text: ident) (targ: typ) (FN F V: Type) (ge: Genv.t FN F V):
              list val -> mem -> trace -> val -> mem -> Prop :=
  | extcall_annot_val_sem_intro: forall varg m arg,
      eventval_match ge arg targ varg ->
      extcall_annot_val_sem text targ ge (varg :: nil) m (Event_annot text (arg :: nil) :: E0) varg m.

Lemma extcall_annot_val_ok:
  forall text targ,
  extcall_properties (extcall_annot_val_sem text targ) (mksignature (targ :: nil) (Some targ)).
Proof.
  intros; constructor; intros.



  inv H1. econstructor; eauto.
  eapply eventval_match_preserved; eauto.

  inv H; auto.

  inv H; auto.

  inv H; auto.


  inv H; simpl; omega.

  assert (t1 = t2). inv H; inv H0; auto. subst t2.
  exists vres1; exists m1; auto.

  inv H; inv H0.
  assert (arg = arg0). eapply eventval_match_determ_2; eauto. subst arg0.
  split. constructor. auto.
Qed.

External functions

For most languages, the functions composing the program are either internal functions, defined within the language, or external functions, defined outside. External functions include system calls but also compiler built-in functions. We define a type for external functions and associated operations.

Inductive pseudo_external_function : Type :=
  | EF_external (name: ident) (sg: signature)
     
A system call or library function. Produces an event in the trace.
  | EF_builtin (name: ident) (sg: signature)
     
A compiler built-in function. Behaves like an external, but can be inlined by the compiler.
  | EF_vload (chunk: memory_chunk)
     
A volatile read operation. If the adress given as first argument points within a volatile global variable, generate an event and return the value found in this event. Otherwise, produce no event and behave like a regular memory load.
  | EF_vstore (chunk: memory_chunk)
     
A volatile store operation. If the adress given as first argument points within a volatile global variable, generate an event. Otherwise, produce no event and behave like a regular memory store.
  | EF_vload_global (chunk: memory_chunk) (id: ident) (ofs: int)
     
A volatile load operation from a global variable. Specialized version of EF_vload.
  | EF_vstore_global (chunk: memory_chunk) (id: ident) (ofs: int)
     
A volatile store operation in a global variable. Specialized version of EF_vstore.
  | EF_malloc
     
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
     
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.
  | EF_memcpy (sz: Z) (al: Z)
     
Block copy, of sz bytes, between addresses that are al-aligned.
  | EF_annot (text: ident) (targs: list typ)
     
A programmer-supplied annotation. Takes zero, one or several arguments, produces an event carrying the text and the values of these arguments, and returns no value.
  | EF_annot_val (text: ident) (targ: typ)
     
Another form of annotation that takes one argument, produces an event carrying the text and the value of this argument, and returns the value of the argument.
  | EF_inline_asm (text: ident).
Inline asm statements. Semantically, treated like an annotation with no parameters (EF_annot text nil). To be used with caution, as it can invalidate the semantic preservation theorem. Generated only if -finline-asm is given.
The type signature of an external function.

Definition pef_sig (ef: pseudo_external_function): signature :=
  match ef with
  | EF_external name sg => sg
  | EF_builtin name sg => sg
  | EF_vload chunk => mksignature (Tint :: nil) (Some (type_of_chunk chunk))
  | EF_vstore chunk => mksignature (Tint :: type_of_chunk chunk :: nil) None
  | EF_vload_global chunk _ _ => mksignature nil (Some (type_of_chunk chunk))
  | EF_vstore_global chunk _ _ => mksignature (type_of_chunk chunk :: nil) None
  | EF_malloc => mksignature (Tint :: nil) (Some Tint)
  | EF_free => mksignature (Tint :: nil) None
  | EF_memcpy sz al => mksignature (Tint :: Tint :: nil) None
  | EF_annot text targs => mksignature targs None
  | EF_annot_val text targ => mksignature (targ :: nil) (Some targ)
  | EF_inline_asm text => mksignature nil None
  end.

Whether an external function must reload its arguments.

Definition pef_reloads (ef: pseudo_external_function) : bool :=
  match ef with
  | EF_annot text targs => false
  | _ => true
  end.

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 pseudo_external_call (ef: pseudo_external_function): extcall_sem :=
  match ef with
  | EF_external name sg => extcall_io_sem name sg
  | EF_builtin name sg => extcall_io_sem name sg
  | EF_vload chunk => volatile_load_sem chunk
  | EF_vstore chunk => volatile_store_sem chunk
  | EF_vload_global chunk id ofs => volatile_load_global_sem chunk id ofs
  | EF_vstore_global chunk id ofs => volatile_store_global_sem chunk id ofs
  | EF_malloc => extcall_malloc_sem
  | EF_free => extcall_free_sem
  | EF_memcpy sz al => extcall_memcpy_sem sz al
  | EF_annot txt targs => extcall_annot_sem txt targs
  | EF_annot_val txt targ=> extcall_annot_val_sem txt targ
  | EF_inline_asm txt => extcall_annot_sem txt nil
  end.

Theorem pseudo_external_call_spec:
  forall ef,
  extcall_properties (pseudo_external_call ef) (pef_sig ef).
Proof.
  intros. unfold pseudo_external_call, pef_sig. destruct ef.
  apply extcall_io_ok.
  apply extcall_io_ok.
  apply volatile_load_ok.
  apply volatile_store_ok.
  apply volatile_load_global_ok.
  apply volatile_store_global_ok.
  apply extcall_malloc_ok.
  apply extcall_free_ok.
  apply extcall_memcpy_ok.
  apply extcall_annot_ok.
  apply extcall_annot_val_ok.
  apply extcall_annot_ok.
Qed.

True external functions


Inductive extcall_true_external_sem (name: ident) (sg: signature) (FN F V: Type) (ge: Genv.t FN F V) :
  list val -> mem -> trace -> val -> mem -> Prop :=
| extcall_true_external_sem_intro:
  forall args
    
    m1 m2 (Hle: mem_le m1 m2) ret,
  extcall_true_external_sem name sg ge args m1 (comp.External (sw := sw) (name, sg) args m1 ret m2 :: nil) ret m2.

Lemma extcall_true_external_ok:
  forall name sg,
  extcall_properties (extcall_true_external_sem name sg) sg.
Proof.
  split.


  inversion 3; subst; econstructor; eauto.

  inversion 1; subst. inversion Hle; subst. eauto.

  inversion 1; subst. inversion Hle; subst. eauto.

  inversion 1; subst. inversion Hle; subst. eauto.

  inversion 1; subst; simpl; omega.

  inversion 1; subst; simpl. inversion 1; subst; simpl.
  esplit. esplit. econstructor. eapply Hle0. assumption.

  inversion 1; subst.
  inversion 1; subst.
  split.
  constructor.
  tauto.
  injection 1; intros; subst; eauto.
Qed.

Inductive external_function : Type :=
  | EF_pseudo (pef: pseudo_external_function)
  | EF_true_external (name: ident) (sg: signature).

The type signature of an external function.

Definition ef_sig (ef: external_function): signature :=
  match ef with
    | EF_pseudo pef => pef_sig pef
    | EF_true_external _ sg => sg
  end.

Whether an external function must reload its arguments.

Definition ef_reloads (ef: external_function) : bool :=
  match ef with
  | EF_pseudo pef => pef_reloads pef
  | _ => true
  end.

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): extcall_sem :=
  match ef with
    | EF_pseudo pef => pseudo_external_call pef
    | EF_true_external name sg => extcall_true_external_sem name sg
  end.

Theorem external_call_spec:
  forall ef,
  extcall_properties (external_call ef) (ef_sig ef).
Proof.
  intros. unfold external_call, ef_sig. destruct ef.
   apply pseudo_external_call_spec.
  apply extcall_true_external_ok.
Qed.

Definition external_call_symbols_preserved_gen ef := ec_symbols_preserved (external_call_spec ef).
Definition external_call_valid_block ef := ec_valid_block (external_call_spec ef).
Definition external_call_max_perm ef := ec_max_perm (external_call_spec ef).
Definition external_call_readonly ef := ec_readonly (external_call_spec ef).
Definition external_call_trace_length ef := ec_trace_length (external_call_spec ef).
Definition external_call_receptive ef := ec_receptive (external_call_spec ef).
Definition external_call_determ ef := ec_determ (external_call_spec ef).

Special cases of external_call_symbols_preserved_gen.

Lemma external_call_symbols_preserved:
  forall ef FN1 F1 F2 FN2 V (ge1: Genv.t FN1 F1 V) (ge2: Genv.t FN2 F2 V) vargs m1 t vres m2,
  external_call ef ge1 vargs m1 t vres m2 ->
  (forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id) ->
  (forall b, Genv.find_var_info ge2 b = Genv.find_var_info ge1 b) ->
  external_call ef ge2 vargs m1 t vres m2.
Proof.
  intros. eapply external_call_symbols_preserved_gen; eauto.
  intros. unfold block_is_volatile. rewrite H1. auto.
Qed.

Require Import Errors.

Lemma external_call_symbols_preserved_2:
  forall ef FN1 F1 V1 FN2 F2 V2 (tvar: V1 -> res V2)
         (ge1: Genv.t FN1 F1 V1) (ge2: Genv.t FN2 F2 V2) vargs m1 t vres m2,
  external_call ef ge1 vargs m1 t vres m2 ->
  (forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id) ->
  (forall b gv1, Genv.find_var_info ge1 b = Some gv1 ->
     exists gv2, Genv.find_var_info ge2 b = Some gv2 /\ transf_globvar tvar gv1 = OK gv2) ->
  (forall b gv2, Genv.find_var_info ge2 b = Some gv2 ->
     exists gv1, Genv.find_var_info ge1 b = Some gv1 /\ transf_globvar tvar gv1 = OK gv2) ->
  external_call ef ge2 vargs m1 t vres m2.
Proof.
  intros. eapply external_call_symbols_preserved_gen; eauto.
  intros. unfold block_is_volatile.
  case_eq (Genv.find_var_info ge1 b); intros.
  exploit H1; eauto. intros [g2 [A B]]. rewrite A. monadInv B. destruct g; auto.
  case_eq (Genv.find_var_info ge2 b); intros.
  exploit H2; eauto. intros [g1 [A B]]. congruence.
  auto.
Qed.

Corollary of external_call_valid_block.

Lemma external_call_nextblock:
  forall ef (FN F V : Type) (ge : Genv.t FN F V) vargs m1 t vres m2,
  external_call ef ge vargs m1 t vres m2 ->
  Mem.nextblock m1 <= Mem.nextblock m2.
Proof.
  intros.
  exploit external_call_valid_block; eauto.
  instantiate (1 := Mem.nextblock m1 - 1). red; omega.
  unfold Mem.valid_block. omega.
Qed.

Corollaries of external_call_determ.

Lemma external_call_match_traces:
  forall ef (FN F V : Type) (ge : Genv.t FN F V) vargs m t1 vres1 m1 t2 vres2 m2,
  external_call ef ge vargs m t1 vres1 m1 ->
  external_call ef ge vargs m t2 vres2 m2 ->
  match_traces ge t1 t2.
Proof.
  intros. exploit external_call_determ. eexact H. eexact H0. tauto.
Qed.

Lemma external_call_deterministic:
  forall ef (FN F V : Type) (ge : Genv.t FN F V) vargs m t vres1 m1 vres2 m2,
  external_call ef ge vargs m t vres1 m1 ->
  external_call ef ge vargs m t vres2 m2 ->
  vres1 = vres2 /\ m1 = m2.
Proof.
  intros. exploit external_call_determ. eexact H. eexact H0. intuition.
Qed.