aboutsummaryrefslogtreecommitdiffstats
path: root/common/Events.v
diff options
context:
space:
mode:
authorBernhard Schommer <bernhardschommer@gmail.com>2015-10-20 13:32:18 +0200
committerBernhard Schommer <bernhardschommer@gmail.com>2015-10-20 13:32:18 +0200
commit4d542bc7eafadb16b845cf05d1eb4988eb55ed0f (patch)
tree1961b41815fc6e392cc0bd2beeb0fb504bc160ce /common/Events.v
parent7a6bb90048db7a254e959b1e3c308bac5fe6c418 (diff)
downloadcompcert-kvx-4d542bc7eafadb16b845cf05d1eb4988eb55ed0f.tar.gz
compcert-kvx-4d542bc7eafadb16b845cf05d1eb4988eb55ed0f.zip
Updated PR by removing whitespaces. Bug 17450.
Diffstat (limited to 'common/Events.v')
-rw-r--r--common/Events.v284
1 files changed, 142 insertions, 142 deletions
diff --git a/common/Events.v b/common/Events.v
index dc38b344..7029a984 100644
--- a/common/Events.v
+++ b/common/Events.v
@@ -31,7 +31,7 @@ Require Import Globalenvs.
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
+ 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:
@@ -40,10 +40,10 @@ Require Import Globalenvs.
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.
+ 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.
+ and address being written and the value stored there.
- An annotation, recording the text of the annotation and the values
of the arguments.
@@ -102,7 +102,7 @@ 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.
@@ -133,7 +133,7 @@ Ltac decomposeTraceEq :=
auto
end.
-Ltac traceEq :=
+Ltac traceEq :=
repeat substTraceHyp; autorewrite with trace_rewrite; decomposeTraceEq.
(** Bisimilarity between infinite traces. *)
@@ -149,7 +149,7 @@ 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.
@@ -157,7 +157,7 @@ Proof.
Qed.
Lemma traceinf_sim_trans:
- forall T1 T2 T3,
+ 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.
@@ -170,11 +170,11 @@ CoInductive traceinf_sim': traceinf -> traceinf -> Prop :=
Lemma traceinf_sim'_sim:
forall T1 T2, traceinf_sim' T1 T2 -> traceinf_sim T1 T2.
Proof.
- cofix COINDHYP; intros. inv H.
+ cofix COINDHYP; intros. inv H.
destruct t. elim H0; auto.
Transparent Eappinf.
Transparent E0.
- simpl.
+ simpl.
destruct t. simpl. constructor. apply COINDHYP; auto.
constructor. apply COINDHYP.
constructor. unfold E0; congruence. auto.
@@ -193,10 +193,10 @@ Program Definition split_traceinf' (t: trace) (T: traceinf') (NE: t <> E0): even
| e :: t' => (e, Econsinf' t' T _)
end.
Next Obligation.
- elimtype False. elim NE. auto.
+ elimtype False. elim NE. auto.
Qed.
Next Obligation.
- red; intro. elim (H e). rewrite H0. auto.
+ red; intro. elim (H e). rewrite H0. auto.
Qed.
CoFixpoint traceinf_of_traceinf' (T': traceinf') : traceinf :=
@@ -224,11 +224,11 @@ Lemma traceinf_traceinf'_app:
Proof.
induction t.
intros. elim NE. auto.
- intros. simpl.
+ intros. simpl.
rewrite (unroll_traceinf (traceinf_of_traceinf' (Econsinf' (a :: t) T NE))).
simpl. destruct t. auto.
Transparent Eappinf.
- simpl. f_equal. apply IHt.
+ simpl. f_equal. apply IHt.
Qed.
(** Prefixes of traces. *)
@@ -244,7 +244,7 @@ Lemma trace_prefix_app:
trace_prefix t1 t2 ->
trace_prefix (t ** t1) (t ** t2).
Proof.
- intros. destruct H as [t3 EQ]. exists t3. traceEq.
+ intros. destruct H as [t3 EQ]. exists t3. traceEq.
Qed.
Lemma traceinf_prefix_app:
@@ -295,7 +295,7 @@ Lemma eventval_match_type:
forall ev ty v,
eventval_match ev ty v -> Val.has_type v ty.
Proof.
- intros. inv H; simpl; auto.
+ intros. inv H; simpl; auto.
Qed.
Lemma eventval_list_match_length:
@@ -331,7 +331,7 @@ 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 Senv.find_symbol_injective; eauto.
+ decEq. eapply Senv.find_symbol_injective; eauto.
Qed.
Lemma eventval_list_match_determ_2:
@@ -370,13 +370,13 @@ Lemma eventval_match_receptive:
Proof.
intros. inv H; destruct ev2; simpl in H2; try discriminate.
- exists (Vint i0); constructor.
-- simpl in H1; exploit Senv.public_symbol_exists; eauto. intros [b FS].
+- simpl in H1; exploit Senv.public_symbol_exists; eauto. intros [b FS].
exists (Vptr b i1); constructor; auto.
- exists (Vlong i0); constructor.
- exists (Vfloat f0); constructor.
- exists (Vsingle f0); constructor; auto.
- exists (Vint i); constructor.
-- simpl in H1. exploit Senv.public_symbol_exists. eexact H1. intros [b' FS].
+- simpl in H1. exploit Senv.public_symbol_exists. eexact H1. intros [b' FS].
exists (Vptr b' i0); constructor; auto.
Qed.
@@ -390,7 +390,7 @@ Lemma eventval_match_same_type:
forall ev1 ty v1 ev2 v2,
eventval_match ev1 ty v1 -> eventval_match ev2 ty v2 -> eventval_type ev1 = eventval_type ev2.
Proof.
- destruct 1; intros EV; inv EV; auto.
+ destruct 1; intros EV; inv EV; auto.
Qed.
End EVENTVAL.
@@ -407,7 +407,7 @@ Hypothesis public_preserved:
Lemma eventval_valid_preserved:
forall ev, eventval_valid ge1 ev -> eventval_valid ge2 ev.
Proof.
- intros. destruct ev; simpl in *; auto. rewrite <- H; auto.
+ intros. destruct ev; simpl in *; auto. rewrite <- H; auto.
Qed.
Hypothesis symbols_preserved:
@@ -418,8 +418,8 @@ Lemma eventval_match_preserved:
eventval_match ge1 ev ty v -> eventval_match ge2 ev ty v.
Proof.
induction 1; constructor; auto.
- rewrite public_preserved; auto.
- rewrite symbols_preserved; auto.
+ rewrite public_preserved; auto.
+ rewrite symbols_preserved; auto.
Qed.
Lemma eventval_list_match_preserved:
@@ -463,12 +463,12 @@ Qed.
Lemma eventval_match_inject_2:
forall ev ty v1,
- eventval_match ge1 ev ty v1 ->
+ eventval_match ge1 ev ty v1 ->
exists v2, eventval_match ge2 ev ty v2 /\ Val.inject f v1 v2.
Proof.
intros. inv H; try (econstructor; split; eauto; constructor; fail).
destruct symb_inj as (A & B & C & D). exploit C; eauto. intros [b2 [EQ FS]].
- exists (Vptr b2 ofs); split. econstructor; eauto.
+ exists (Vptr b2 ofs); split. econstructor; eauto.
econstructor; eauto. rewrite Int.add_zero; auto.
Qed.
@@ -536,7 +536,7 @@ Definition output_event (ev: event) : Prop :=
| Event_vstore _ _ _ _ => True
| Event_annot _ _ => True
end.
-
+
Fixpoint output_trace (t: trace) : Prop :=
match t with
| nil => True
@@ -584,7 +584,7 @@ Inductive volatile_store (ge: Senv.t):
- the trace generated by the call (can be empty).
*)
-Definition extcall_sem : Type :=
+Definition extcall_sem : Type :=
Senv.t -> list val -> mem -> trace -> val -> mem -> Prop.
(** We now specify the expected properties of this predicate. *)
@@ -736,20 +736,20 @@ Lemma volatile_load_inject:
Mem.inject f m m' ->
exists v', volatile_load ge2 chunk m' b' ofs' t v' /\ Val.inject f v v'.
Proof.
- intros until m'; intros SI VL VI MI. generalize SI; intros (A & B & C & D).
+ intros until m'; intros SI VL VI MI. generalize SI; intros (A & B & C & D).
inv VL.
- (* volatile load *)
- inv VI. exploit B; eauto. intros [U V]. subst delta.
+ inv VI. exploit B; eauto. intros [U V]. subst delta.
exploit eventval_match_inject_2; eauto. intros (v2 & X & Y).
rewrite Int.add_zero. exists (Val.load_result chunk v2); split.
constructor; auto.
erewrite D; eauto.
apply Val.load_result_inject. auto.
- (* normal load *)
- exploit Mem.loadv_inject; eauto. simpl; eauto. simpl; intros (v2 & X & Y).
- exists v2; split; auto.
+ exploit Mem.loadv_inject; eauto. simpl; eauto. simpl; intros (v2 & X & Y).
+ exists v2; split; auto.
constructor; auto.
- inv VI. erewrite D; eauto.
+ inv VI. erewrite D; eauto.
Qed.
Lemma volatile_load_receptive:
@@ -765,15 +765,15 @@ Qed.
Lemma volatile_load_ok:
forall chunk,
- extcall_properties (volatile_load_sem chunk)
+ extcall_properties (volatile_load_sem chunk)
(mksignature (Tint :: nil) (Some (type_of_chunk chunk)) cc_default).
Proof.
intros; constructor; intros.
(* well typed *)
-- unfold proj_sig_res; simpl. inv H. inv H0. apply Val.load_result_type.
- eapply Mem.load_type; eauto.
+- unfold proj_sig_res; simpl. inv H. inv H0. apply Val.load_result_type.
+ eapply Mem.load_type; eauto.
(* symbols *)
-- inv H2. constructor. eapply volatile_load_preserved; eauto.
+- inv H2. constructor. eapply volatile_load_preserved; eauto.
(* valid blocks *)
- inv H; auto.
(* max perms *)
@@ -781,12 +781,12 @@ Proof.
(* readonly *)
- inv H. apply Mem.unchanged_on_refl.
(* mem extends *)
-- inv H. inv H1. inv H6. inv H4.
+- inv H. inv H1. inv H6. inv H4.
exploit volatile_load_extends; eauto. intros [v' [A B]].
exists v'; exists m1'; intuition. constructor; auto.
(* mem injects *)
-- inv H0. inv H2. inv H7. inversion H5; subst.
- exploit volatile_load_inject; eauto. intros [v' [A B]].
+- inv H0. inv H2. inv H7. inversion H5; subst.
+ exploit volatile_load_inject; eauto. intros [v' [A B]].
exists f; exists v'; exists m1'; intuition. constructor; auto.
red; intros. congruence.
(* trace length *)
@@ -797,11 +797,11 @@ Proof.
(* determ *)
- inv H; inv H0. inv H1; inv H7; try congruence.
assert (id = id0) by (eapply Senv.find_symbol_injective; eauto). subst id0.
- split. constructor.
+ split. constructor.
eapply eventval_match_valid; eauto.
eapply eventval_match_valid; eauto.
eapply eventval_match_same_type; eauto.
- intros EQ; inv EQ.
+ intros EQ; inv EQ.
assert (v = v0) by (eapply eventval_match_determ_1; eauto). subst v0.
auto.
split. constructor. intuition congruence.
@@ -837,10 +837,10 @@ Lemma volatile_store_readonly:
Proof.
intros. inv H.
apply Mem.unchanged_on_refl.
- eapply Mem.store_unchanged_on; eauto.
- exploit Mem.store_valid_access_3; eauto. intros [P Q].
- intros. unfold loc_not_writable. red; intros. elim H2.
- apply Mem.perm_cur_max. apply P. auto.
+ eapply Mem.store_unchanged_on; eauto.
+ exploit Mem.store_valid_access_3; eauto. intros [P Q].
+ intros. unfold loc_not_writable. red; intros. elim H2.
+ apply Mem.perm_cur_max. apply P. auto.
Qed.
Lemma volatile_store_extends:
@@ -848,19 +848,19 @@ Lemma volatile_store_extends:
volatile_store ge chunk m1 b ofs v t m2 ->
Mem.extends m1 m1' ->
Val.lessdef v v' ->
- exists m2',
+ 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. apply Val.load_result_lessdef; auto.
+ eapply eventval_match_lessdef; eauto. apply Val.load_result_lessdef; auto.
auto with mem.
- exploit Mem.store_within_extends; eauto. intros [m2' [A B]].
exists m2'; intuition.
+ econstructor; eauto.
-+ eapply Mem.store_unchanged_on; eauto.
++ eapply Mem.store_unchanged_on; eauto.
unfold loc_out_of_bounds; intros.
assert (Mem.perm m1 b i Max Nonempty).
{ apply Mem.perm_cur_max. apply Mem.perm_implies with Writable; auto with mem.
@@ -881,13 +881,13 @@ Lemma volatile_store_inject:
/\ Mem.unchanged_on (loc_unmapped f) m1 m2
/\ Mem.unchanged_on (loc_out_of_reach f m1) m1' m2'.
Proof.
- intros until v'; intros SI VS AI VI MI.
- generalize SI; intros (P & Q & R & S).
+ intros until v'; intros SI VS AI VI MI.
+ generalize SI; intros (P & Q & R & S).
inv VS.
- (* volatile store *)
- inv AI. exploit Q; eauto. intros [A B]. subst delta.
+ inv AI. exploit Q; eauto. intros [A B]. subst delta.
rewrite Int.add_zero. exists m1'; split.
- constructor; auto. erewrite S; eauto.
+ constructor; auto. erewrite S; eauto.
eapply eventval_match_inject; eauto. apply Val.load_result_inject. auto.
intuition auto with mem.
- (* normal store *)
@@ -895,17 +895,17 @@ Proof.
assert (Mem.storev chunk m1 (Vptr b ofs) v = Some m2). simpl; auto.
exploit Mem.storev_mapped_inject; eauto. intros [m2' [A B]].
exists m2'; intuition auto.
-+ constructor; auto. erewrite S; eauto.
++ constructor; auto. erewrite S; eauto.
+ eapply Mem.store_unchanged_on; eauto.
unfold loc_unmapped; intros. inv AI; congruence.
-+ eapply Mem.store_unchanged_on; eauto.
- unfold loc_out_of_reach; intros. red; intros. simpl in A.
++ eapply Mem.store_unchanged_on; eauto.
+ unfold loc_out_of_reach; intros. red; intros. simpl in A.
assert (EQ: Int.unsigned (Int.add ofs (Int.repr delta)) = Int.unsigned ofs + delta)
by (eapply Mem.address_inject; eauto with mem).
rewrite EQ in *.
- eelim H3; eauto.
- exploit Mem.store_valid_access_3. eexact H0. intros [X Y].
- apply Mem.perm_cur_max. apply Mem.perm_implies with Writable; auto with mem.
+ eelim H3; eauto.
+ exploit Mem.store_valid_access_3. eexact H0. intros [X Y].
+ apply Mem.perm_cur_max. apply Mem.perm_implies with Writable; auto with mem.
apply X. omega.
Qed.
@@ -913,29 +913,29 @@ Lemma volatile_store_receptive:
forall ge 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.
+ intros. inv H; inv H0; auto.
Qed.
Lemma volatile_store_ok:
forall chunk,
- extcall_properties (volatile_store_sem chunk)
+ extcall_properties (volatile_store_sem chunk)
(mksignature (Tint :: type_of_chunk chunk :: nil) None cc_default).
Proof.
intros; constructor; intros.
(* well typed *)
- unfold proj_sig_res; simpl. inv H; constructor.
(* symbols preserved *)
-- inv H2. constructor. eapply volatile_store_preserved; eauto.
+- inv H2. constructor. eapply volatile_store_preserved; eauto.
(* valid block *)
- inv H. inv H1. auto. eauto with mem.
(* perms *)
-- inv H. inv H2. auto. eauto with mem.
+- inv H. inv H2. auto. eauto with mem.
(* readonly *)
- inv H. eapply volatile_store_readonly; eauto.
(* mem extends*)
- inv H. inv H1. inv H6. inv H7. inv H4.
- exploit volatile_store_extends; eauto. intros [m2' [A [B C]]].
- exists Vundef; exists m2'; intuition. constructor; auto.
+ exploit volatile_store_extends; eauto. intros [m2' [A [B C]]].
+ exists Vundef; exists m2'; intuition. constructor; auto.
(* mem inject *)
- inv H0. inv H2. inv H7. inv H8. inversion H5; subst.
exploit volatile_store_inject; eauto. intros [m2' [A [B [C D]]]].
@@ -963,7 +963,7 @@ Inductive extcall_malloc_sem (ge: Senv.t):
extcall_malloc_sem ge (Vint n :: nil) m E0 (Vptr b Int.zero) m''.
Lemma extcall_malloc_ok:
- extcall_properties extcall_malloc_sem
+ extcall_properties extcall_malloc_sem
(mksignature (Tint :: nil) (Some Tint) cc_default).
Proof.
assert (UNCHANGED:
@@ -978,7 +978,7 @@ Proof.
erewrite Mem.store_mem_contents; eauto. rewrite Maps.PMap.gso by auto.
Local Transparent Mem.alloc. unfold Mem.alloc in H. injection H; intros A B.
rewrite <- B; simpl. rewrite A. rewrite Maps.PMap.gso by auto. auto.
- }
+ }
constructor; intros.
(* well typed *)
@@ -989,34 +989,34 @@ Proof.
- inv H. eauto with mem.
(* perms *)
- inv H. exploit Mem.perm_alloc_inv. eauto. eapply Mem.perm_store_2; eauto.
- rewrite dec_eq_false. auto.
+ rewrite dec_eq_false. auto.
apply Mem.valid_not_valid_diff with m1; eauto with mem.
(* readonly *)
-- inv H. eapply UNCHANGED; eauto.
+- inv H. eapply UNCHANGED; eauto.
(* mem extends *)
-- inv H. inv H1. inv H5. inv H7.
+- 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.
+ 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.
(* mem injects *)
- inv H0. inv H2. inv H6. inv H8.
- exploit Mem.alloc_parallel_inject; eauto. apply Zle_refl. apply Zle_refl.
+ 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.
+ exploit Mem.store_mapped_inject. eexact A. eauto. eauto.
+ instantiate (1 := Vint n). auto.
intros [m2' [E G]].
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 H2. inv H2. eauto with mem.
+ red; intros. destruct (eq_block b1 b).
+ subst b1. rewrite C in H2. inv H2. eauto with mem.
rewrite D in H2 by auto. congruence.
(* trace length *)
- inv H; simpl; omega.
@@ -1024,7 +1024,7 @@ Proof.
- assert (t1 = t2). inv H; inv H0; auto. subst t2.
exists vres1; exists m1; auto.
(* determ *)
-- inv H; inv H0. split. constructor. intuition congruence.
+- inv H; inv H0. split. constructor. intuition congruence.
Qed.
(** ** Semantics of dynamic memory deallocation (free) *)
@@ -1038,7 +1038,7 @@ Inductive extcall_free_sem (ge: Senv.t):
extcall_free_sem ge (Vptr b lo :: nil) m E0 Vundef m'.
Lemma extcall_free_ok:
- extcall_properties extcall_free_sem
+ extcall_properties extcall_free_sem
(mksignature (Tint :: nil) None cc_default).
Proof.
constructor; intros.
@@ -1049,22 +1049,22 @@ Proof.
(* valid block *)
- inv H. eauto with mem.
(* perms *)
-- inv H. eapply Mem.perm_free_3; eauto.
+- inv H. eapply Mem.perm_free_3; eauto.
(* readonly *)
-- inv H. eapply Mem.free_unchanged_on; eauto.
- intros. red; intros. elim H3.
- apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem.
- eapply Mem.free_range_perm; eauto.
+- inv H. eapply Mem.free_unchanged_on; eauto.
+ intros. red; intros. elim H3.
+ apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem.
+ eapply Mem.free_range_perm; eauto.
(* mem extends *)
-- inv H. inv H1. inv H8. inv H6.
- exploit Mem.load_extends; eauto. intros [vsz [A B]]. inv B.
+- 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 Mem.free_unchanged_on; eauto.
- unfold loc_out_of_bounds; intros.
+ eapply Mem.free_unchanged_on; eauto.
+ unfold loc_out_of_bounds; intros.
assert (Mem.perm m1 b i Max Nonempty).
- { apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem.
+ { apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem.
eapply Mem.free_range_perm. eexact H4. eauto. }
tauto.
(* mem inject *)
@@ -1072,24 +1072,24 @@ Proof.
exploit Mem.load_inject; eauto. intros [vsz [A B]]. inv B.
assert (Mem.range_perm m1 b (Int.unsigned lo - 4) (Int.unsigned lo + Int.unsigned sz) Cur Freeable).
eapply Mem.free_range_perm; eauto.
- exploit Mem.address_inject; eauto.
+ exploit Mem.address_inject; eauto.
apply Mem.perm_implies with Freeable; auto with mem.
- apply H0. instantiate (1 := lo). omega.
+ apply H0. instantiate (1 := lo). omega.
intro EQ.
exploit Mem.free_parallel_inject; eauto. intros (m2' & C & D).
exists f, Vundef, m2'; split.
- apply extcall_free_sem_intro with (sz := sz) (m' := m2').
- rewrite EQ. rewrite <- A. f_equal. omega.
+ apply extcall_free_sem_intro with (sz := sz) (m' := m2').
+ rewrite EQ. rewrite <- A. f_equal. omega.
auto.
rewrite ! EQ. rewrite <- C. f_equal; omega.
- split. auto.
+ split. auto.
split. auto.
split. eapply Mem.free_unchanged_on; eauto. unfold loc_unmapped. intros; congruence.
- split. eapply Mem.free_unchanged_on; eauto. unfold loc_out_of_reach.
+ split. eapply Mem.free_unchanged_on; eauto. unfold loc_out_of_reach.
intros. red; intros. eelim H7; eauto.
- apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem.
+ apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable; auto with mem.
apply H0. omega.
- split. auto.
+ split. auto.
red; intros. congruence.
(* trace length *)
- inv H; simpl; omega.
@@ -1122,19 +1122,19 @@ Lemma extcall_memcpy_ok:
Proof.
intros. constructor.
- (* return type *)
- intros. inv H. constructor.
+ intros. inv H. constructor.
- (* change of globalenv *)
intros. inv H2. econstructor; eauto.
- (* valid blocks *)
- intros. inv H. eauto with mem.
+ intros. inv H. eauto with mem.
- (* perms *)
- intros. inv H. eapply Mem.perm_storebytes_2; eauto.
+ intros. inv H. eapply Mem.perm_storebytes_2; eauto.
- (* readonly *)
- intros. inv H. eapply Mem.storebytes_unchanged_on; eauto.
- intros; red; intros. elim H8.
- apply Mem.perm_cur_max. eapply Mem.storebytes_range_perm; eauto.
+ intros. inv H. eapply Mem.storebytes_unchanged_on; eauto.
+ intros; red; intros. elim H8.
+ apply Mem.perm_cur_max. eapply Mem.storebytes_range_perm; eauto.
- (* extensions *)
- intros. inv H.
+ intros. inv H.
inv H1. inv H13. inv H14. inv H10. inv H11.
exploit Mem.loadbytes_length; eauto. intros LEN.
exploit Mem.loadbytes_extends; eauto. intros [bytes2 [A B]].
@@ -1143,28 +1143,28 @@ Proof.
split. econstructor; eauto.
split. constructor.
split. auto.
- eapply Mem.storebytes_unchanged_on; eauto. unfold loc_out_of_bounds; intros.
+ eapply Mem.storebytes_unchanged_on; eauto. unfold loc_out_of_bounds; intros.
assert (Mem.perm m1 bdst i Max Nonempty).
apply Mem.perm_cur_max. apply Mem.perm_implies with Writable; auto with mem.
- eapply Mem.storebytes_range_perm; eauto.
- erewrite list_forall2_length; eauto.
+ eapply Mem.storebytes_range_perm; eauto.
+ erewrite list_forall2_length; eauto.
tauto.
- (* injections *)
intros. inv H0. inv H2. inv H14. inv H15. inv H11. inv H12.
destruct (zeq sz 0).
+ (* special case sz = 0 *)
- assert (bytes = nil).
+ assert (bytes = nil).
{ exploit (Mem.loadbytes_empty m1 bsrc (Int.unsigned osrc) sz). omega. congruence. }
subst.
destruct (Mem.range_perm_storebytes m1' b0 (Int.unsigned (Int.add odst (Int.repr delta0))) nil)
as [m2' SB].
simpl. red; intros; omegaContradiction.
- exists f, Vundef, m2'.
- split. econstructor; eauto.
+ exists f, Vundef, m2'.
+ split. econstructor; eauto.
intros; omegaContradiction.
intros; omegaContradiction.
- right; omega.
- apply Mem.loadbytes_empty. omega.
+ right; omega.
+ apply Mem.loadbytes_empty. omega.
split. auto.
split. eapply Mem.storebytes_empty_inject; eauto.
split. eapply Mem.storebytes_unchanged_on; eauto. unfold loc_unmapped; intros.
@@ -1190,7 +1190,7 @@ Proof.
exploit Mem.loadbytes_inject; eauto. intros [bytes2 [A B]].
exploit Mem.storebytes_mapped_inject; eauto. intros [m2' [C D]].
exists f; exists Vundef; exists m2'.
- split. econstructor; try rewrite EQ1; try rewrite EQ2; eauto.
+ split. econstructor; try rewrite EQ1; try rewrite EQ2; eauto.
intros; eapply Mem.aligned_area_inject with (m := m1); eauto.
intros; eapply Mem.aligned_area_inject with (m := m1); eauto.
eapply Mem.disjoint_or_equal_inject with (m := m1); eauto.
@@ -1201,17 +1201,17 @@ Proof.
split. eapply Mem.storebytes_unchanged_on; eauto. unfold loc_unmapped; intros.
congruence.
split. eapply Mem.storebytes_unchanged_on; eauto. unfold loc_out_of_reach; intros. red; intros.
- eelim H2; eauto.
+ eelim H2; eauto.
apply Mem.perm_cur_max. apply Mem.perm_implies with Writable; auto with mem.
- eapply Mem.storebytes_range_perm; eauto.
- erewrite list_forall2_length; eauto.
+ eapply Mem.storebytes_range_perm; eauto.
+ erewrite list_forall2_length; eauto.
omega.
split. apply inject_incr_refl.
red; intros; congruence.
- (* trace length *)
intros; inv H. simpl; omega.
- (* receptive *)
- intros.
+ intros.
assert (t1 = t2). inv H; inv H0; auto. subst t2.
exists vres1; exists m1; auto.
- (* determ *)
@@ -1235,7 +1235,7 @@ Proof.
(* well typed *)
- inv H. simpl. auto.
(* symbols *)
-- inv H2. econstructor; eauto.
+- inv H2. econstructor; eauto.
eapply eventval_list_match_preserved; eauto.
(* valid blocks *)
- inv H; auto.
@@ -1257,7 +1257,7 @@ Proof.
(* trace length *)
- inv H; simpl; omega.
(* receptive *)
-- assert (t1 = t2). inv H; inv H0; auto.
+- assert (t1 = t2). inv H; inv H0; auto.
exists vres1; exists m1; congruence.
(* determ *)
- inv H; inv H0.
@@ -1280,7 +1280,7 @@ Proof.
(* well typed *)
- inv H. unfold proj_sig_res; simpl. eapply eventval_match_type; eauto.
(* symbols *)
-- inv H2. econstructor; eauto.
+- inv H2. econstructor; eauto.
eapply eventval_match_preserved; eauto.
(* valid blocks *)
- inv H; auto.
@@ -1289,7 +1289,7 @@ Proof.
(* readonly *)
- inv H. apply Mem.unchanged_on_refl.
(* mem extends *)
-- inv H. inv H1. inv H6.
+- inv H. inv H1. inv H6.
exists v2; exists m1'; intuition.
econstructor; eauto.
eapply eventval_match_lessdef; eauto.
@@ -1324,7 +1324,7 @@ Proof.
(* well typed *)
- inv H. simpl. auto.
(* symbols *)
-- inv H2. econstructor; eauto.
+- inv H2. econstructor; eauto.
(* valid blocks *)
- inv H; auto.
(* perms *)
@@ -1386,7 +1386,7 @@ Definition external_call (ef: external_function): extcall_sem :=
| EF_builtin name sg => external_functions_sem name sg
| EF_vload chunk => volatile_load_sem chunk
| EF_vstore chunk => volatile_store_sem chunk
- | EF_malloc => extcall_malloc_sem
+ | 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
@@ -1396,7 +1396,7 @@ Definition external_call (ef: external_function): extcall_sem :=
end.
Theorem external_call_spec:
- forall ef,
+ forall ef,
extcall_properties (external_call ef) (ef_sig ef).
Proof.
intros. unfold external_call, ef_sig; destruct ef.
@@ -1473,7 +1473,7 @@ Lemma external_call_mem_inject:
/\ inject_incr f f'
/\ inject_separated f f' m1 m1'.
Proof.
- intros. destruct H as (A & B & C). eapply external_call_mem_inject_gen with (ge1 := ge); eauto.
+ intros. destruct H as (A & B & C). eapply external_call_mem_inject_gen with (ge1 := ge); eauto.
repeat split; intros.
+ simpl in H3. exploit A; eauto. intros EQ; rewrite EQ in H; inv H. auto.
+ simpl in H3. exploit A; eauto. intros EQ; rewrite EQ in H; inv H. auto.
@@ -1548,29 +1548,29 @@ Inductive external_call'
Lemma decode_longs_lessdef:
forall tyl vl1 vl2, Val.lessdef_list vl1 vl2 -> Val.lessdef_list (decode_longs tyl vl1) (decode_longs tyl vl2).
Proof.
- induction tyl; simpl; intros.
+ induction tyl; simpl; intros.
auto.
- destruct a; inv H; auto. inv H1; auto. constructor; auto. apply Val.longofwords_lessdef; auto.
+ destruct a; inv H; auto. inv H1; auto. constructor; auto. apply Val.longofwords_lessdef; auto.
Qed.
Lemma decode_longs_inject:
forall f tyl vl1 vl2, Val.inject_list f vl1 vl2 -> Val.inject_list f (decode_longs tyl vl1) (decode_longs tyl vl2).
Proof.
- induction tyl; simpl; intros.
+ induction tyl; simpl; intros.
auto.
destruct a; inv H; auto. inv H1; auto. constructor; auto. apply Val.longofwords_inject; auto. Qed.
Lemma encode_long_lessdef:
forall oty v1 v2, Val.lessdef v1 v2 -> Val.lessdef_list (encode_long oty v1) (encode_long oty v2).
Proof.
- intros. destruct oty as [[]|]; simpl; auto.
+ intros. destruct oty as [[]|]; simpl; auto.
constructor. apply Val.hiword_lessdef; auto. constructor. apply Val.loword_lessdef; auto. auto.
Qed.
Lemma encode_long_inject:
forall f oty v1 v2, Val.inject f v1 v2 -> Val.inject_list f (encode_long oty v1) (encode_long oty v2).
Proof.
- intros. destruct oty as [[]|]; simpl; auto.
+ intros. destruct oty as [[]|]; simpl; auto.
constructor. apply Val.hiword_inject; auto. constructor. apply Val.loword_inject; auto. auto.
Qed.
@@ -1589,8 +1589,8 @@ Lemma external_call_well_typed':
external_call' ef ge vargs m1 t vres m2 ->
Val.has_type_list vres (proj_sig_res' (ef_sig ef)).
Proof.
- intros. inv H. apply encode_long_has_type.
- eapply external_call_well_typed; eauto.
+ intros. inv H. apply encode_long_has_type.
+ eapply external_call_well_typed; eauto.
Qed.
Lemma external_call_symbols_preserved':
@@ -1631,9 +1631,9 @@ Lemma external_call_mem_extends':
/\ Mem.extends m2 m2'
/\ Mem.unchanged_on (loc_out_of_bounds m1) m1' m2'.
Proof.
- intros. inv H.
+ intros. inv H.
exploit external_call_mem_extends; eauto.
- eapply decode_longs_lessdef; eauto.
+ eapply decode_longs_lessdef; eauto.
intros (v' & m2' & A & B & C & D).
exists (encode_long (sig_res (ef_sig ef)) v'); exists m2'; intuition.
econstructor; eauto.
@@ -1655,7 +1655,7 @@ Lemma external_call_mem_inject':
/\ inject_incr f f'
/\ inject_separated f f' m1 m1'.
Proof.
- intros. inv H0.
+ intros. inv H0.
exploit external_call_mem_inject; eauto.
eapply decode_longs_inject; eauto.
intros (f' & v' & m2' & A & B & C & D & E & P & Q).
@@ -1670,8 +1670,8 @@ Lemma external_call_determ':
external_call' ef ge vargs m t2 vres2 m2 ->
match_traces ge t1 t2 /\ (t1 = t2 -> vres1 = vres2 /\ m1 = m2).
Proof.
- intros. inv H; inv H0. exploit external_call_determ. eexact H1. eexact H.
- intros [A B]. split. auto. intros. destruct B as [C D]; auto. subst. auto.
+ intros. inv H; inv H0. exploit external_call_determ. eexact H1. eexact H.
+ intros [A B]. split. auto. intros. destruct B as [C D]; auto. subst. auto.
Qed.
Lemma external_call_match_traces':
@@ -1689,7 +1689,7 @@ Lemma external_call_deterministic':
external_call' ef ge vargs m t vres2 m2 ->
vres1 = vres2 /\ m1 = m2.
Proof.
- intros. inv H; inv H0.
+ intros. inv H; inv H0.
exploit external_call_deterministic. eexact H1. eexact H. intros [A B].
split; congruence.
Qed.
@@ -1768,13 +1768,13 @@ Lemma eval_builtin_arg_preserved:
Proof.
assert (EQ: forall id ofs, Senv.symbol_address ge2 id ofs = Senv.symbol_address ge1 id ofs).
{ unfold Senv.symbol_address; simpl; intros. rewrite symbols_preserved; auto. }
- induction 1; eauto with barg. rewrite <- EQ in H; eauto with barg. rewrite <- EQ; eauto with barg.
-Qed.
+ induction 1; eauto with barg. rewrite <- EQ in H; eauto with barg. rewrite <- EQ; eauto with barg.
+Qed.
Lemma eval_builtin_args_preserved:
forall al vl, eval_builtin_args ge1 e sp m al vl -> eval_builtin_args ge2 e sp m al vl.
Proof.
- induction 1; constructor; auto; eapply eval_builtin_arg_preserved; eauto.
+ induction 1; constructor; auto; eapply eval_builtin_arg_preserved; eauto.
Qed.
End EVAL_BUILTIN_ARG_PRESERVED.
@@ -1802,13 +1802,13 @@ Proof.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
- econstructor; eauto with barg.
-- exploit Mem.loadv_extends; eauto. intros (v' & P & Q). exists v'; eauto with barg.
+- exploit Mem.loadv_extends; eauto. intros (v' & P & Q). exists v'; eauto with barg.
- econstructor; eauto with barg.
- exploit Mem.loadv_extends; eauto. intros (v' & P & Q). exists v'; eauto with barg.
- econstructor; eauto with barg.
- destruct IHeval_builtin_arg1 as (vhi' & P & Q).
destruct IHeval_builtin_arg2 as (vlo' & R & S).
- econstructor; split; eauto with barg. apply Val.longofwords_lessdef; auto.
+ econstructor; split; eauto with barg. apply Val.longofwords_lessdef; auto.
Qed.
Lemma eval_builtin_args_lessdef:
@@ -1817,7 +1817,7 @@ Lemma eval_builtin_args_lessdef:
Proof.
induction 1.
- econstructor; split. constructor. auto.
-- exploit eval_builtin_arg_lessdef; eauto. intros (v1' & P & Q).
+- exploit eval_builtin_arg_lessdef; eauto. intros (v1' & P & Q).
destruct IHlist_forall2 as (vl' & U & V).
exists (v1'::vl'); split; constructor; auto.
Qed.