aboutsummaryrefslogtreecommitdiffstats
path: root/cfrontend
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 /cfrontend
parent7a6bb90048db7a254e959b1e3c308bac5fe6c418 (diff)
downloadcompcert-kvx-4d542bc7eafadb16b845cf05d1eb4988eb55ed0f.tar.gz
compcert-kvx-4d542bc7eafadb16b845cf05d1eb4988eb55ed0f.zip
Updated PR by removing whitespaces. Bug 17450.
Diffstat (limited to 'cfrontend')
-rw-r--r--cfrontend/C2C.ml76
-rw-r--r--cfrontend/Cexec.v268
-rw-r--r--cfrontend/Clight.v18
-rw-r--r--cfrontend/ClightBigstep.v62
-rw-r--r--cfrontend/Cminorgen.v8
-rw-r--r--cfrontend/Cminorgenproof.v480
-rw-r--r--cfrontend/Cop.v108
-rw-r--r--cfrontend/Csem.v18
-rw-r--r--cfrontend/Csharpminor.v16
-rw-r--r--cfrontend/Cshmgen.v22
-rw-r--r--cfrontend/Cshmgenproof.v340
-rw-r--r--cfrontend/Cstrategy.v524
-rw-r--r--cfrontend/Csyntax.v10
-rw-r--r--cfrontend/Ctypes.v116
-rw-r--r--cfrontend/Ctyping.v238
-rw-r--r--cfrontend/Initializersproof.v162
-rw-r--r--cfrontend/PrintClight.ml2
-rw-r--r--cfrontend/PrintCsyntax.ml4
-rw-r--r--cfrontend/SimplExpr.v22
-rw-r--r--cfrontend/SimplExprproof.v678
-rw-r--r--cfrontend/SimplExprspec.v164
-rw-r--r--cfrontend/SimplLocals.v2
-rw-r--r--cfrontend/SimplLocalsproof.v722
23 files changed, 2030 insertions, 2030 deletions
diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml
index 9b31dfdb..4835f785 100644
--- a/cfrontend/C2C.ml
+++ b/cfrontend/C2C.ml
@@ -141,8 +141,8 @@ let builtins_generic = {
(* Block copy *)
"__builtin_memcpy_aligned",
(TVoid [],
- [TPtr(TVoid [], []);
- TPtr(TVoid [AConst], []);
+ [TPtr(TVoid [], []);
+ TPtr(TVoid [AConst], []);
TInt(IUInt, []);
TInt(IUInt, [])],
false);
@@ -357,12 +357,12 @@ let make_builtin_memcpy args =
let sz1 =
match Initializers.constval !comp_env sz with
| Errors.OK(Vint n) -> n
- | _ -> error "ill-formed __builtin_memcpy_aligned (3rd argument must be
+ | _ -> error "ill-formed __builtin_memcpy_aligned (3rd argument must be
a constant)"; Integers.Int.zero in
let al1 =
match Initializers.constval !comp_env al with
| Errors.OK(Vint n) -> n
- | _ -> error "ill-formed __builtin_memcpy_aligned (4th argument must be
+ | _ -> error "ill-formed __builtin_memcpy_aligned (4th argument must be
a constant)"; Integers.Int.one in
(* to check: sz1 > 0, al1 divides sz1, al1 = 1|2|4|8 *)
(* Issue #28: must decay array types to pointer types *)
@@ -384,7 +384,7 @@ let va_list_ptr e =
let make_builtin_va_arg_by_val helper ty ty_ret arg =
let ty_fun =
Tfunction(Tcons(Tpointer(Tvoid, noattr), Tnil), ty_ret, cc_default) in
- Ecast
+ Ecast
(Ecall(Evalof(Evar(intern_string helper, ty_fun), ty_fun),
Econs(va_list_ptr arg, Enil),
ty_ret),
@@ -392,13 +392,13 @@ let make_builtin_va_arg_by_val helper ty ty_ret arg =
let make_builtin_va_arg_by_ref helper ty arg =
let ty_fun =
- Tfunction(Tcons(Tpointer(Tvoid, noattr), Tnil),
+ Tfunction(Tcons(Tpointer(Tvoid, noattr), Tnil),
Tpointer(Tvoid, noattr), cc_default) in
let ty_ptr =
Tpointer(ty, noattr) in
let call =
Ecall(Evalof(Evar(intern_string helper, ty_fun), ty_fun),
- Econs(va_list_ptr arg,
+ Econs(va_list_ptr arg,
Econs(Esizeof(ty, Tint(I32, Unsigned, noattr)), Enil)),
Tpointer(Tvoid, noattr)) in
Evalof(Ederef(Ecast(call, ty_ptr), ty), ty)
@@ -406,13 +406,13 @@ let make_builtin_va_arg_by_ref helper ty arg =
let make_builtin_va_arg env ty e =
match ty with
| Tint _ | Tpointer _ ->
- make_builtin_va_arg_by_val
+ make_builtin_va_arg_by_val
"__compcert_va_int32" ty (Tint(I32, Unsigned, noattr)) e
| Tlong _ ->
- make_builtin_va_arg_by_val
+ make_builtin_va_arg_by_val
"__compcert_va_int64" ty (Tlong(Unsigned, noattr)) e
| Tfloat _ ->
- make_builtin_va_arg_by_val
+ make_builtin_va_arg_by_val
"__compcert_va_float64" ty (Tfloat(F64, noattr)) e
| Tstruct _ | Tunion _ ->
make_builtin_va_arg_by_ref
@@ -433,7 +433,7 @@ let rec log2 n = if n = 1 then 0 else 1 + log2 (n lsr 1)
let convertAttr a =
{ attr_volatile = List.mem AVolatile a;
- attr_alignas =
+ attr_alignas =
let n = Cutil.alignas_attribute a in
if n > 0 then Some (N.of_int (log2 n)) else None }
@@ -463,7 +463,7 @@ let convertFkind = function
| C.FFloat -> F32
| C.FDouble -> F64
| C.FLongDouble ->
- if not !Clflags.option_flongdouble then unsupported "'long double' type";
+ if not !Clflags.option_flongdouble then unsupported "'long double' type";
F64
let rec convertTyp env t =
@@ -524,11 +524,11 @@ let convertField env f =
(intern_string f.fld_name, convertTyp env f.fld_typ)
let convertCompositedef env su id attr members =
- let t = match su with
- | C.Struct ->
+ let t = match su with
+ | C.Struct ->
let layout = Cutil.struct_layout env members in
List.iter (fun (a,b) -> Debug.set_member_offset id a b) layout;
- TStruct (id,attr)
+ TStruct (id,attr)
| C.Union -> TUnion (id,attr) in
Debug.set_composite_size id su (Cutil.sizeof env t);
Composite(intern_string id.name,
@@ -763,7 +763,7 @@ let rec convertExpr env e =
EF_debug(P.of_int64 kind, intern_string text,
typlist_of_typelist targs2),
targs2, convertExprList env args2, convertTyp env e.etyp)
-
+
| C.ECall({edesc = C.EVar {name = "__builtin_annot"}}, args) ->
begin match args with
| {edesc = C.EConst(CStr txt)} :: args1 ->
@@ -774,20 +774,20 @@ let rec convertExpr env e =
| _ ->
error "ill-formed __builtin_annot (first argument must be string literal)";
ezero
- end
-
+ end
+
| C.ECall({edesc = C.EVar {name = "__builtin_annot_intval"}}, args) ->
begin match args with
| [ {edesc = C.EConst(CStr txt)}; arg ] ->
let targ = convertTyp env
(Cutil.default_argument_conversion env arg.etyp) in
Ebuiltin(EF_annot_val(coqstring_of_camlstring txt, typ_of_type targ),
- Tcons(targ, Tnil), convertExprList env [arg],
+ Tcons(targ, Tnil), convertExprList env [arg],
convertTyp env e.etyp)
| _ ->
error "ill-formed __builtin_annot_intval (first argument must be string literal)";
ezero
- end
+ end
| C.ECall({edesc = C.EVar {name = "__builtin_memcpy_aligned"}}, args) ->
make_builtin_memcpy (convertExprList env args)
@@ -822,9 +822,9 @@ let rec convertExpr env e =
let sg =
signature_of_type targs tres
{cc_vararg = true; cc_unproto = false; cc_structret = false} in
- Ebuiltin(EF_external(coqstring_of_camlstring "printf", sg),
+ Ebuiltin(EF_external(coqstring_of_camlstring "printf", sg),
targs, convertExprList env args, tres)
-
+
| C.ECall(fn, args) ->
if not (supported_return_type env e.etyp) then
unsupported ("function returning a result of type " ^ string_of_type e.etyp ^ " (consider adding option -fstruct-return)");
@@ -867,17 +867,17 @@ and convertExprList env el =
(* Extended assembly *)
let convertAsm loc env txt outputs inputs clobber =
- let (txt', output', inputs') =
+ let (txt', output', inputs') =
ExtendedAsm.transf_asm loc env txt outputs inputs clobber in
let clobber' =
List.map (fun s -> coqstring_of_camlstring (String.uppercase s)) clobber in
let ty_res =
match output' with None -> TVoid [] | Some e -> e.etyp in
(* Build the Ebuiltin expression *)
- let e =
+ let e =
let tinputs = convertTypArgs env [] inputs' in
let toutput = convertTyp env ty_res in
- Ebuiltin(EF_inline_asm(coqstring_of_camlstring txt',
+ Ebuiltin(EF_inline_asm(coqstring_of_camlstring txt',
signature_of_type tinputs toutput cc_default,
clobber'),
tinputs,
@@ -894,7 +894,7 @@ type switchlabel =
| Case of C.exp
| Default
-type switchbody =
+type switchbody =
| Label of switchlabel
| Stmt of C.stmt
@@ -922,16 +922,16 @@ let rec groupSwitch = function
(Cutil.sseq s.sloc s fst, cases)
(* Test whether the statement contains case and give an *)
-let rec contains_case s =
+let rec contains_case s =
match s.sdesc with
- | C.Sskip
- | C.Sdo _
+ | C.Sskip
+ | C.Sdo _
| C.Sbreak
- | C.Scontinue
+ | C.Scontinue
| C.Sswitch _ (* Stop at a switch *)
- | C.Sgoto _
- | C.Sreturn _
- | C.Sdecl _
+ | C.Sgoto _
+ | C.Sreturn _
+ | C.Sdecl _
| C.Sasm _ -> ()
| C.Sseq (s1,s2)
| C.Sif(_,s1,s2) -> contains_case s1; contains_case s2
@@ -1021,7 +1021,7 @@ and convertSwitch env is_64 = function
match lbl with
| Default ->
None
- | Case e ->
+ | Case e ->
match Ceval.integer_expr env e with
| None -> unsupported "'case' label is not a compile-time integer";
None
@@ -1128,7 +1128,7 @@ let convertInitializer env ty i =
let convertGlobvar loc env (sto, id, ty, optinit) =
let id' = intern_string id.name in
Debug.atom_global_variable id id';
- let ty' = convertTyp env ty in
+ let ty' = convertTyp env ty in
let sz = Ctypes.sizeof !comp_env ty' in
let al = Ctypes.alignof !comp_env ty' in
let attr = Cutil.attributes_of_type env ty in
@@ -1190,7 +1190,7 @@ let rec convertGlobdecls env res gl =
warning ("'#pragma " ^ s ^ "' directive ignored");
convertGlobdecls env res gl'
-(** Convert struct and union declarations.
+(** Convert struct and union declarations.
Result is a list of CompCert C composite declarations. *)
let rec convertCompositedefs env res gl =
@@ -1229,7 +1229,7 @@ let rec translEnv env = function
module IdentSet = Set.Make(struct type t = C.ident let compare = compare end)
let cleanupGlobals p =
-
+
(* First pass: determine what is defined *)
let strong = ref IdentSet.empty (* def functions or variables with inits *)
and weak = ref IdentSet.empty (* variables without inits *)
@@ -1252,7 +1252,7 @@ let cleanupGlobals p =
| _ -> () in
List.iter classify_def p;
- (* Second pass: keep "best" definition for each identifier *)
+ (* Second pass: keep "best" definition for each identifier *)
let rec clean defs accu = function
| [] -> accu
| g :: gl ->
diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v
index 938454c5..7e966ffe 100644
--- a/cfrontend/Cexec.v
+++ b/cfrontend/Cexec.v
@@ -85,7 +85,7 @@ Proof.
intros until ty. destruct a; simpl; congruence.
Qed.
-Local Open Scope option_monad_scope.
+Local Open Scope option_monad_scope.
Fixpoint is_val_list (al: exprlist) : option (list (val * type)) :=
match al with
@@ -110,7 +110,7 @@ Definition eventval_of_val (v: val) (t: typ) : option eventval :=
| Vfloat f, AST.Tfloat => Some (EVfloat f)
| Vsingle f, AST.Tsingle => Some (EVsingle f)
| Vlong n, AST.Tlong => Some (EVlong n)
- | Vptr b ofs, AST.Tint =>
+ | Vptr b ofs, AST.Tint =>
do id <- Genv.invert_symbol ge b;
check (Genv.public_symbol ge id);
Some (EVptr_global id ofs)
@@ -153,7 +153,7 @@ Lemma eventval_of_val_complete:
forall ev t v, eventval_match ge ev t v -> eventval_of_val v t = Some ev.
Proof.
induction 1; simpl; auto.
- rewrite (Genv.find_invert_symbol _ _ H0). simpl in H; rewrite H. auto.
+ rewrite (Genv.find_invert_symbol _ _ H0). simpl in H; rewrite H. auto.
Qed.
Lemma list_eventval_of_val_sound:
@@ -169,7 +169,7 @@ Qed.
Lemma list_eventval_of_val_complete:
forall evl tl vl, eventval_list_match ge evl tl vl -> list_eventval_of_val vl tl = Some evl.
Proof.
- induction 1; simpl. auto.
+ induction 1; simpl. auto.
rewrite (eventval_of_val_complete _ _ _ H). rewrite IHeventval_list_match. auto.
Qed.
@@ -190,7 +190,7 @@ Qed.
(** Volatile memory accesses. *)
-Definition do_volatile_load (w: world) (chunk: memory_chunk) (m: mem) (b: block) (ofs: int)
+Definition do_volatile_load (w: world) (chunk: memory_chunk) (m: mem) (b: block) (ofs: int)
: option (world * trace * val) :=
if Genv.block_is_volatile ge b then
do id <- Genv.invert_symbol ge b;
@@ -230,11 +230,11 @@ Lemma do_volatile_load_sound:
do_volatile_load w chunk m b ofs = Some(w', t, v) ->
volatile_load ge chunk m b ofs t v /\ possible_trace w t w'.
Proof.
- intros until v. unfold do_volatile_load. mydestr.
- destruct p as [ev w'']. mydestr.
- split. constructor; auto. apply Genv.invert_find_symbol; auto.
- apply val_of_eventval_sound; auto.
- econstructor. constructor; eauto. constructor.
+ intros until v. unfold do_volatile_load. mydestr.
+ destruct p as [ev w'']. mydestr.
+ split. constructor; auto. apply Genv.invert_find_symbol; auto.
+ apply val_of_eventval_sound; auto.
+ econstructor. constructor; eauto. constructor.
split. constructor; auto. constructor.
Qed.
@@ -254,10 +254,10 @@ Lemma do_volatile_store_sound:
do_volatile_store w chunk m b ofs v = Some(w', t, m') ->
volatile_store ge chunk m b ofs v t m' /\ possible_trace w t w'.
Proof.
- intros until m'. unfold do_volatile_store. mydestr.
- split. constructor; auto. apply Genv.invert_find_symbol; auto.
- apply eventval_of_val_sound; auto.
- econstructor. constructor; eauto. constructor.
+ intros until m'. unfold do_volatile_store. mydestr.
+ split. constructor; auto. apply Genv.invert_find_symbol; auto.
+ apply eventval_of_val_sound; auto.
+ econstructor. constructor; eauto. constructor.
split. constructor; auto. constructor.
Qed.
@@ -297,7 +297,7 @@ Remark check_assign_copy:
forall (ty: type) (b: block) (ofs: int) (b': block) (ofs': int),
{ assign_copy_ok ty b ofs b' ofs' } + {~ assign_copy_ok ty b ofs b' ofs' }.
Proof with try (right; intuition omega).
- intros. unfold assign_copy_ok.
+ intros. unfold assign_copy_ok.
assert (alignof_blockcopy ge ty > 0) by apply alignof_blockcopy_pos.
destruct (Zdivide_dec (alignof_blockcopy ge ty) (Int.unsigned ofs')); auto...
destruct (Zdivide_dec (alignof_blockcopy ge ty) (Int.unsigned ofs)); auto...
@@ -314,7 +314,7 @@ Proof with try (right; intuition omega).
destruct (zle (Int.unsigned ofs' + sizeof ge ty) (Int.unsigned ofs)); auto.
destruct (zle (Int.unsigned ofs + sizeof ge ty) (Int.unsigned ofs')); auto.
right; intuition omega.
- destruct Y... left; intuition omega.
+ destruct Y... left; intuition omega.
Defined.
Definition do_assign_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: int) (v: val): option (world * trace * mem) :=
@@ -343,8 +343,8 @@ Lemma do_deref_loc_sound:
deref_loc ge ty m b ofs t v /\ possible_trace w t w'.
Proof.
unfold do_deref_loc; intros until v.
- destruct (access_mode ty) eqn:?; mydestr.
- intros. exploit do_volatile_load_sound; eauto. intuition. eapply deref_loc_volatile; eauto.
+ destruct (access_mode ty) eqn:?; mydestr.
+ intros. exploit do_volatile_load_sound; eauto. intuition. eapply deref_loc_volatile; eauto.
split. eapply deref_loc_value; eauto. constructor.
split. eapply deref_loc_reference; eauto. constructor.
split. eapply deref_loc_copy; eauto. constructor.
@@ -368,10 +368,10 @@ Lemma do_assign_loc_sound:
assign_loc ge ty m b ofs v t m' /\ possible_trace w t w'.
Proof.
unfold do_assign_loc; intros until m'.
- destruct (access_mode ty) eqn:?; mydestr.
- intros. exploit do_volatile_store_sound; eauto. intuition. eapply assign_loc_volatile; eauto.
+ destruct (access_mode ty) eqn:?; mydestr.
+ intros. exploit do_volatile_store_sound; eauto. intuition. eapply assign_loc_volatile; eauto.
split. eapply assign_loc_value; eauto. constructor.
- destruct v; mydestr. destruct a as [P [Q R]].
+ destruct v; mydestr. destruct a as [P [Q R]].
split. eapply assign_loc_copy; eauto. constructor.
Qed.
@@ -385,7 +385,7 @@ Proof.
rewrite H1; rewrite H2. apply do_volatile_store_complete; auto.
rewrite H1. destruct (check_assign_copy ty b ofs b' ofs').
inv H0. rewrite H5; rewrite H6; auto.
- elim n. red; tauto.
+ elim n. red; tauto.
Qed.
(** External calls *)
@@ -477,7 +477,7 @@ Remark memcpy_check_args:
forall sz al bdst odst bsrc osrc,
{memcpy_args_ok sz al bdst odst bsrc osrc} + {~memcpy_args_ok sz al bdst odst bsrc osrc}.
Proof with try (right; intuition omega).
- intros.
+ intros.
assert (X: {al = 1 \/ al = 2 \/ al = 4 \/ al = 8} + {~(al = 1 \/ al = 2 \/ al = 4 \/ al = 8)}).
destruct (zeq al 1); auto. destruct (zeq al 2); auto.
destruct (zeq al 4); auto. destruct (zeq al 8); auto...
@@ -486,9 +486,9 @@ Proof with try (right; intuition omega).
destruct (zle 0 sz)...
destruct (Zdivide_dec al sz); auto...
assert(U: forall x, {sz > 0 -> (al | x)} + {~(sz > 0 -> (al | x))}).
- intros. destruct (zeq sz 0).
+ intros. destruct (zeq sz 0).
left; intros; omegaContradiction.
- destruct (Zdivide_dec al x); auto. right; red; intros. elim n0. apply H0. omega.
+ destruct (Zdivide_dec al x); auto. right; red; intros. elim n0. apply H0. omega.
destruct (U osrc); auto...
destruct (U odst); auto...
assert (Y: {bsrc <> bdst \/ osrc = odst \/ osrc + sz <= odst \/ odst + sz <= osrc}
@@ -555,17 +555,17 @@ Proof with try congruence.
intros until m'.
destruct ef; simpl.
(* EF_external *)
- eapply do_external_function_sound; eauto.
+ eapply do_external_function_sound; eauto.
(* EF_builtin *)
- eapply do_external_function_sound; eauto.
+ eapply do_external_function_sound; eauto.
(* EF_vload *)
- unfold do_ef_volatile_load. destruct vargs... destruct v... destruct vargs...
+ unfold do_ef_volatile_load. destruct vargs... destruct v... destruct vargs...
mydestr. destruct p as [[w'' t''] v]; mydestr.
- exploit do_volatile_load_sound; eauto. intuition. econstructor; eauto.
+ exploit do_volatile_load_sound; eauto. intuition. econstructor; eauto.
auto.
(* EF_vstore *)
- unfold do_ef_volatile_store. destruct vargs... destruct v... destruct vargs... destruct vargs...
- mydestr. destruct p as [[w'' t''] m'']. mydestr.
+ unfold do_ef_volatile_store. destruct vargs... destruct v... destruct vargs... destruct vargs...
+ mydestr. destruct p as [[w'' t''] m'']. mydestr.
exploit do_volatile_store_sound; eauto. intuition. econstructor; eauto.
auto.
(* EF_malloc *)
@@ -573,19 +573,19 @@ Proof with try congruence.
destruct (Mem.alloc m (-4) (Int.unsigned i)) as [m1 b] eqn:?. mydestr.
split. econstructor; eauto. constructor.
(* EF_free *)
- unfold do_ef_free. destruct vargs... destruct v... destruct vargs...
- mydestr. destruct v... mydestr.
+ unfold do_ef_free. destruct vargs... destruct v... destruct vargs...
+ mydestr. destruct v... mydestr.
split. econstructor; eauto. omega. constructor.
(* EF_memcpy *)
- unfold do_ef_memcpy. destruct vargs... destruct v... destruct vargs...
- destruct v... destruct vargs... mydestr. red in m0.
+ unfold do_ef_memcpy. destruct vargs... destruct v... destruct vargs...
+ destruct v... destruct vargs... mydestr. red in m0.
split. econstructor; eauto; tauto. constructor.
(* EF_annot *)
- unfold do_ef_annot. mydestr.
+ unfold do_ef_annot. mydestr.
split. constructor. apply list_eventval_of_val_sound; auto.
econstructor. constructor; eauto. constructor.
(* EF_annot_val *)
- unfold do_ef_annot_val. destruct vargs... destruct vargs... mydestr.
+ unfold do_ef_annot_val. destruct vargs... destruct vargs... mydestr.
split. constructor. apply eventval_of_val_sound; auto.
econstructor. constructor; eauto. constructor.
(* EF_inline_asm *)
@@ -611,7 +611,7 @@ Proof.
inv H; unfold do_ef_volatile_store.
exploit do_volatile_store_complete; eauto. intros EQ; rewrite EQ; auto.
(* EF_malloc *)
- inv H; unfold do_ef_malloc.
+ inv H; unfold do_ef_malloc.
inv H0. rewrite H1. rewrite H2. auto.
(* EF_free *)
inv H; unfold do_ef_free.
@@ -619,12 +619,12 @@ Proof.
(* EF_memcpy *)
inv H; unfold do_ef_memcpy.
inv H0. rewrite pred_dec_true. rewrite H7; rewrite H8; auto.
- red. tauto.
+ red. tauto.
(* EF_annot *)
- inv H; unfold do_ef_annot. inv H0. inv H6. inv H4.
+ inv H; unfold do_ef_annot. inv H0. inv H6. inv H4.
rewrite (list_eventval_of_val_complete _ _ _ H1). auto.
(* EF_annot_val *)
- inv H; unfold do_ef_annot_val. inv H0. inv H6. inv H4.
+ inv H; unfold do_ef_annot_val. inv H0. inv H6. inv H4.
rewrite (eventval_of_val_complete _ _ _ H1). auto.
(* EF_inline_asm *)
eapply do_inline_assembly_complete; eauto.
@@ -837,7 +837,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
do w',t, v1 <- do_deref_loc w ty m b ofs;
let op := match id with Incr => Oadd | Decr => Osub end in
let r' :=
- Ecomma (Eassign (Eloc b ofs ty)
+ Ecomma (Eassign (Eloc b ofs ty)
(Ebinop op (Eval v1 ty) (Eval (Vint Int.one) type_int32s) (incrdecr_type ty))
ty)
(Eval v1 ty) ty in
@@ -922,7 +922,7 @@ Inductive imm_safe_t: kind -> expr -> mem -> Prop :=
Remark imm_safe_t_imm_safe:
forall k a m, imm_safe_t k a m -> imm_safe ge e k a m.
Proof.
- induction 1.
+ induction 1.
constructor.
constructor.
eapply imm_safe_lred; eauto.
@@ -975,7 +975,7 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
exists t, exists v1, exists w',
ty = ty1 /\ deref_loc ge ty1 m b ofs t v1 /\ possible_trace w t w'
| Epostincr id (Eloc b ofs ty1) ty =>
- exists t, exists v1, exists w',
+ exists t, exists v1, exists w',
ty = ty1 /\ deref_loc ge ty m b ofs t v1 /\ possible_trace w t w'
| Ecomma (Eval v ty1) r2 ty =>
typeof r2 = ty
@@ -1004,7 +1004,7 @@ Proof.
exists b; auto.
exists b; auto.
exists b; exists ofs; auto.
- exists b; exists ofs; split; auto. exists co, delta; auto.
+ exists b; exists ofs; split; auto. exists co, delta; auto.
exists b; exists ofs; split; auto. exists co; auto.
Qed.
@@ -1072,8 +1072,8 @@ Proof.
intros. elim (H0 a m); auto.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
- red; intros. destruct (C a); auto.
- red; intros. destruct e1; auto. elim (H0 a m); auto.
+ red; intros. destruct (C a); auto.
+ red; intros. destruct e1; auto. elim (H0 a m); auto.
Qed.
Lemma imm_safe_t_inv:
@@ -1086,7 +1086,7 @@ Lemma imm_safe_t_inv:
end.
Proof.
destruct invert_expr_context as [A B].
- intros. inv H.
+ intros. inv H.
auto.
auto.
assert (invert_expr_prop (C l) m).
@@ -1160,7 +1160,7 @@ Proof.
induction rargs; simpl; intros.
inv H. destruct tyargs; simpl in H0; inv H0. constructor.
monadInv. inv H. simpl in H0. destruct p as [v1 t1]. destruct tyargs; try congruence. monadInv.
- inv H0. rewrite (is_val_inv _ _ _ Heqo). constructor. auto. eauto.
+ inv H0. rewrite (is_val_inv _ _ _ Heqo). constructor. auto. eauto.
Qed.
Lemma sem_cast_arguments_complete:
@@ -1170,7 +1170,7 @@ Lemma sem_cast_arguments_complete:
Proof.
induction 1.
exists (@nil (val * type)); auto.
- destruct IHcast_arguments as [vtl [A B]].
+ destruct IHcast_arguments as [vtl [A B]].
exists ((v, ty) :: vtl); simpl. rewrite A; rewrite B; rewrite H. auto.
Qed.
@@ -1179,7 +1179,7 @@ Lemma topred_ok:
reduction_ok k a m rd ->
reducts_ok k a m (topred rd).
Proof.
- intros. unfold topred; split; simpl; intros.
+ intros. unfold topred; split; simpl; intros.
destruct H0; try contradiction. inv H0. exists a; exists k; auto.
congruence.
Qed.
@@ -1199,7 +1199,7 @@ Lemma wrong_kind_ok:
k <> Cstrategy.expr_kind a ->
reducts_ok k a m stuck.
Proof.
- intros. apply stuck_ok. red; intros. exploit Cstrategy.imm_safe_kind; eauto.
+ intros. apply stuck_ok. red; intros. exploit Cstrategy.imm_safe_kind; eauto.
eapply imm_safe_t_imm_safe; eauto.
Qed.
@@ -1212,9 +1212,9 @@ Lemma not_invert_ok:
end ->
reducts_ok k a m stuck.
Proof.
- intros. apply stuck_ok. red; intros.
- exploit imm_safe_t_inv; eauto. destruct a; auto.
-Qed.
+ intros. apply stuck_ok. red; intros.
+ exploit imm_safe_t_inv; eauto. destruct a; auto.
+Qed.
Lemma incontext_ok:
forall k a m C res k' a',
@@ -1272,7 +1272,7 @@ Lemma incontext2_list_ok:
reducts_ok RV a1 m res1 ->
list_reducts_ok a2 m res2 ->
is_val a1 = None \/ is_val_list a2 = None ->
- reducts_ok RV (Ecall a1 a2 ty) m
+ reducts_ok RV (Ecall a1 a2 ty) m
(incontext2 (fun x => Ecall x a2 ty) res1
(fun x => Ecall a1 x ty) res2).
Proof.
@@ -1280,7 +1280,7 @@ Proof.
destruct (in_app_or _ _ _ H4).
exploit list_in_map_inv; eauto. intros [[C' rd'] [P Q]]. inv P.
exploit H; eauto. intros [a'' [k'' [U [V W]]]].
- exists a''; exists k''. split. eauto. rewrite V; auto.
+ exists a''; exists k''. split. eauto. rewrite V; auto.
exploit list_in_map_inv; eauto. intros [[C' rd'] [P Q]]. inv P.
exploit H0; eauto. intros [a'' [k'' [U [V W]]]].
exists a''; exists k''. split. eauto. rewrite V; auto.
@@ -1301,7 +1301,7 @@ Proof.
destruct (in_app_or _ _ _ H3).
exploit list_in_map_inv; eauto. intros [[C' rd'] [P Q]]. inv P.
exploit H; eauto. intros [a'' [k'' [U [V W]]]].
- exists a''; exists k''. split. eauto. rewrite V; auto.
+ exists a''; exists k''. split. eauto. rewrite V; auto.
exploit list_in_map_inv; eauto. intros [[C' rd'] [P Q]]. inv P.
exploit H0; eauto. intros [a'' [k'' [U [V W]]]].
exists a''; exists k''. split. eauto. rewrite V; auto.
@@ -1312,7 +1312,7 @@ Qed.
Lemma is_val_list_all_values:
forall al vtl, is_val_list al = Some vtl -> exprlist_all_values al.
Proof.
- induction al; simpl; intros. auto.
+ induction al; simpl; intros. auto.
destruct (is_val r1) as [[v ty]|] eqn:?; try discriminate.
destruct (is_val_list al) as [vtl'|] eqn:?; try discriminate.
rewrite (is_val_inv _ _ _ Heqo). eauto.
@@ -1344,7 +1344,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
destruct (is_val a) as [[v ty'] | ] eqn:?.
rewrite (is_val_inv _ _ _ Heqo).
destruct v...
- destruct ty'...
+ destruct ty'...
(* top struct *)
destruct (ge.(genv_cenv)!i0) as [co|] eqn:?...
destruct (field_offset ge f (co_members co)) as [delta|] eqn:?...
@@ -1353,7 +1353,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
destruct (ge.(genv_cenv)!i0) as [co|] eqn:?...
apply topred_ok; auto. eapply red_field_union; eauto.
(* in depth *)
- eapply incontext_ok; eauto.
+ eapply incontext_ok; eauto.
(* Evalof *)
destruct (is_loc a) as [[[b ofs] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ Heqo).
(* top *)
@@ -1367,7 +1367,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
(* Ederef *)
destruct (is_val a) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
- destruct v... apply topred_ok; auto. apply red_deref; auto.
+ destruct v... apply topred_ok; auto. apply red_deref; auto.
(* depth *)
eapply incontext_ok; eauto.
(* Eaddrof *)
@@ -1377,31 +1377,31 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
(* depth *)
eapply incontext_ok; eauto.
(* unop *)
- destruct (is_val a) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
+ destruct (is_val a) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
destruct (sem_unary_operation op v ty' m) as [v'|] eqn:?...
- apply topred_ok; auto. split. apply red_unop; auto. exists w; constructor.
+ apply topred_ok; auto. split. apply red_unop; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
(* binop *)
- destruct (is_val a1) as [[v1 ty1] | ] eqn:?.
+ destruct (is_val a1) as [[v1 ty1] | ] eqn:?.
destruct (is_val a2) as [[v2 ty2] | ] eqn:?.
- rewrite (is_val_inv _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
+ rewrite (is_val_inv _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
(* top *)
destruct (sem_binary_operation ge op v1 ty1 v2 ty2 m) as [v|] eqn:?...
apply topred_ok; auto. split. apply red_binop; auto. exists w; constructor.
(* depth *)
- eapply incontext2_ok; eauto.
- eapply incontext2_ok; eauto.
+ eapply incontext2_ok; eauto.
+ eapply incontext2_ok; eauto.
(* cast *)
- destruct (is_val a) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
+ destruct (is_val a) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
destruct (sem_cast v ty' ty) as [v'|] eqn:?...
- apply topred_ok; auto. split. apply red_cast; auto. exists w; constructor.
+ apply topred_ok; auto. split. apply red_cast; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
(* seqand *)
- destruct (is_val a1) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
+ destruct (is_val a1) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
destruct (bool_val v ty' m) as [v'|] eqn:?... destruct v'.
apply topred_ok; auto. split. eapply red_seqand_true; eauto. exists w; constructor.
@@ -1409,7 +1409,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
(* depth *)
eapply incontext_ok; eauto.
(* seqor *)
- destruct (is_val a1) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
+ destruct (is_val a1) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
destruct (bool_val v ty' m) as [v'|] eqn:?... destruct v'.
apply topred_ok; auto. split. eapply red_seqor_true; eauto. exists w; constructor.
@@ -1417,7 +1417,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
(* depth *)
eapply incontext_ok; eauto.
(* condition *)
- destruct (is_val a1) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
+ destruct (is_val a1) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
destruct (bool_val v ty' m) as [v'|] eqn:?...
apply topred_ok; auto. split. eapply red_condition; eauto. exists w; constructor.
@@ -1428,8 +1428,8 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
(* alignof *)
apply topred_ok; auto. split. apply red_alignof. exists w; constructor.
(* assign *)
- destruct (is_loc a1) as [[[b ofs] ty1] | ] eqn:?.
- destruct (is_val a2) as [[v2 ty2] | ] eqn:?.
+ destruct (is_loc a1) as [[[b ofs] ty1] | ] eqn:?.
+ destruct (is_val a2) as [[v2 ty2] | ] eqn:?.
rewrite (is_loc_inv _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
(* top *)
destruct (type_eq ty1 ty)... subst ty1.
@@ -1442,9 +1442,9 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
eapply incontext2_ok; eauto.
eapply incontext2_ok; eauto.
(* assignop *)
- destruct (is_loc a1) as [[[b ofs] ty1] | ] eqn:?.
- destruct (is_val a2) as [[v2 ty2] | ] eqn:?.
- rewrite (is_loc_inv _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
+ destruct (is_loc a1) as [[[b ofs] ty1] | ] eqn:?.
+ destruct (is_val a2) as [[v2 ty2] | ] eqn:?.
+ rewrite (is_loc_inv _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
(* top *)
destruct (type_eq ty1 ty)... subst ty1.
destruct (do_deref_loc w ty m b ofs) as [[[w' t] v] | ] eqn:?.
@@ -1455,7 +1455,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
eapply incontext2_ok; eauto.
eapply incontext2_ok; eauto.
(* postincr *)
- destruct (is_loc a) as [[[b ofs] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ Heqo).
+ destruct (is_loc a) as [[[b ofs] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ Heqo).
(* top *)
destruct (type_eq ty' ty)... subst ty'.
destruct (do_deref_loc w ty m b ofs) as [[[w' t] v] | ] eqn:?.
@@ -1465,22 +1465,22 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
(* depth *)
eapply incontext_ok; eauto.
(* comma *)
- destruct (is_val a1) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
+ destruct (is_val a1) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
destruct (type_eq (typeof a2) ty)... subst ty.
apply topred_ok; auto. split. apply red_comma; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
(* call *)
- destruct (is_val a) as [[vf tyf] | ] eqn:?.
- destruct (is_val_list rargs) as [vtl | ] eqn:?.
+ destruct (is_val a) as [[vf tyf] | ] eqn:?.
+ destruct (is_val_list rargs) as [vtl | ] eqn:?.
rewrite (is_val_inv _ _ _ Heqo). exploit is_val_list_all_values; eauto. intros ALLVAL.
(* top *)
destruct (classify_fun tyf) as [tyargs tyres cconv|] eqn:?...
destruct (Genv.find_funct ge vf) as [fd|] eqn:?...
- destruct (sem_cast_arguments vtl tyargs) as [vargs|] eqn:?...
+ destruct (sem_cast_arguments vtl tyargs) as [vargs|] eqn:?...
destruct (type_eq (type_of_fundef fd) (Tfunction tyargs tyres cconv))...
- apply topred_ok; auto. red. split; auto. eapply red_call; eauto.
+ apply topred_ok; auto. red. split; auto. eapply red_call; eauto.
eapply sem_cast_arguments_sound; eauto.
apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv. congruence.
apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv.
@@ -1491,31 +1491,31 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
eapply incontext2_list_ok; eauto.
eapply incontext2_list_ok; eauto.
(* builtin *)
- destruct (is_val_list rargs) as [vtl | ] eqn:?.
+ destruct (is_val_list rargs) as [vtl | ] eqn:?.
exploit is_val_list_all_values; eauto. intros ALLVAL.
(* top *)
- destruct (sem_cast_arguments vtl tyargs) as [vargs|] eqn:?...
+ destruct (sem_cast_arguments vtl tyargs) as [vargs|] eqn:?...
destruct (do_external ef w vargs m) as [[[[? ?] v] m'] | ] eqn:?...
exploit do_ef_external_sound; eauto. intros [EC PT].
- apply topred_ok; auto. red. split; auto. eapply red_builtin; eauto.
+ apply topred_ok; auto. red. split; auto. eapply red_builtin; eauto.
eapply sem_cast_arguments_sound; eauto.
exists w0; auto.
apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv.
- assert (x = vargs).
+ assert (x = vargs).
exploit sem_cast_arguments_complete; eauto. intros [vtl' [A B]]. congruence.
subst x. exploit do_ef_external_complete; eauto. congruence.
- apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv.
+ apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv.
exploit sem_cast_arguments_complete; eauto. intros [vtl' [A B]]. congruence.
(* depth *)
eapply incontext_list_ok; eauto.
-
+
(* loc *)
split; intros. tauto. simpl; congruence.
(* paren *)
- destruct (is_val a) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
+ destruct (is_val a) as [[v ty'] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
destruct (sem_cast v ty' tycast) as [v'|] eqn:?...
- apply topred_ok; auto. split. apply red_paren; auto. exists w; constructor.
+ apply topred_ok; auto. split. apply red_paren; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
@@ -1529,7 +1529,7 @@ Qed.
Lemma step_exprlist_val_list:
forall m al, is_val_list al <> None -> step_exprlist al m = nil.
Proof.
- induction al; simpl; intros.
+ induction al; simpl; intros.
auto.
destruct (is_val r1) as [[v1 ty1]|] eqn:?; try congruence.
destruct (is_val_list al) eqn:?; try congruence.
@@ -1549,7 +1549,7 @@ Proof.
rewrite H. rewrite dec_eq_true. econstructor; eauto.
(* var global *)
rewrite H; rewrite H0. econstructor; eauto.
-(* deref *)
+(* deref *)
econstructor; eauto.
(* field struct *)
rewrite H, H0; econstructor; eauto.
@@ -1564,7 +1564,7 @@ Lemma rred_topred:
Proof.
induction 1; simpl; intros.
(* valof *)
- rewrite dec_eq_true.
+ rewrite dec_eq_true.
rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0). econstructor; eauto.
(* addrof *)
inv H. econstructor; eauto.
@@ -1591,7 +1591,7 @@ Proof.
econstructor; eauto.
(* assignop *)
rewrite dec_eq_true. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0).
- econstructor; eauto.
+ econstructor; eauto.
(* postincr *)
rewrite dec_eq_true. subst. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H1).
econstructor; eauto.
@@ -1601,7 +1601,7 @@ Proof.
inv H0. rewrite H; econstructor; eauto.
(* builtin *)
exploit sem_cast_arguments_complete; eauto. intros [vtl [A B]].
- exploit do_ef_external_complete; eauto. intros C.
+ exploit do_ef_external_complete; eauto. intros C.
rewrite A. rewrite B. rewrite C. econstructor; eauto.
Qed.
@@ -1624,7 +1624,7 @@ Lemma reducts_incl_trans:
reducts_incl C' res2 res3 ->
reducts_incl (fun x => C'(C x)) res1 res3.
Proof.
- unfold reducts_incl; intros. auto.
+ unfold reducts_incl; intros. auto.
Qed.
Lemma reducts_incl_nil:
@@ -1756,11 +1756,11 @@ Proof.
eapply reducts_incl_trans with (C' := fun x => Ecall x el ty); eauto.
destruct (is_val (C a)) as [[v ty']|] eqn:?; eauto.
(* call right *)
- eapply reducts_incl_trans with (C' := fun x => Ecall e1 x ty). apply step_exprlist_context. auto.
+ eapply reducts_incl_trans with (C' := fun x => Ecall e1 x ty). apply step_exprlist_context. auto.
destruct (is_val e1) as [[v1 ty1]|] eqn:?; eauto.
destruct (is_val_list (C a)) as [vl|] eqn:?; eauto.
(* builtin *)
- eapply reducts_incl_trans with (C' := fun x => Ebuiltin ef tyargs x ty). apply step_exprlist_context. auto.
+ eapply reducts_incl_trans with (C' := fun x => Ebuiltin ef tyargs x ty). apply step_exprlist_context. auto.
destruct (is_val_list (C a)) as [vl|] eqn:?; eauto.
(* comma *)
eapply reducts_incl_trans with (C' := fun x => Ecomma x e2 ty); eauto.
@@ -1785,18 +1785,18 @@ Lemma not_stuckred_imm_safe:
forall m a k,
(forall C, ~In (C, Stuckred) (step_expr k a m)) -> imm_safe_t k a m.
Proof.
- intros. generalize (step_expr_sound a k m). intros [A B].
+ intros. generalize (step_expr_sound a k m). intros [A B].
destruct (step_expr k a m) as [|[C rd] res] eqn:?.
specialize (B (refl_equal _)). destruct k.
destruct a; simpl in B; try congruence. constructor.
destruct a; simpl in B; try congruence. constructor.
assert (NOTSTUCK: rd <> Stuckred).
red; intros. elim (H C); subst rd; auto with coqlib.
- exploit A. eauto with coqlib. intros [a' [k' [P [Q R]]]].
+ exploit A. eauto with coqlib. intros [a' [k' [P [Q R]]]].
destruct k'; destruct rd; simpl in R; intuition.
subst a. eapply imm_safe_t_lred; eauto.
- subst a. destruct H1 as [w' PT]. eapply imm_safe_t_rred; eauto.
- subst. eapply imm_safe_t_callred; eauto.
+ subst a. destruct H1 as [w' PT]. eapply imm_safe_t_rred; eauto.
+ subst. eapply imm_safe_t_callred; eauto.
Qed.
Lemma not_imm_safe_stuck_red:
@@ -1805,14 +1805,14 @@ Lemma not_imm_safe_stuck_red:
~imm_safe_t k a m ->
exists C', In (C', Stuckred) (step_expr RV (C a) m).
Proof.
- intros.
+ intros.
assert (exists C', In (C', Stuckred) (step_expr k a m)).
destruct (classic (exists C', In (C', Stuckred) (step_expr k a m))); auto.
- elim H0. apply not_stuckred_imm_safe. apply not_ex_all_not. auto.
+ elim H0. apply not_stuckred_imm_safe. apply not_ex_all_not. auto.
destruct H1 as [C' IN].
- specialize (step_expr_context _ _ _ H a m). unfold reducts_incl.
+ specialize (step_expr_context _ _ _ H a m). unfold reducts_incl.
intro.
- exists (fun x => (C (C' x))). apply H1; auto.
+ exists (fun x => (C (C' x))). apply H1; auto.
Qed.
(** Connections between [imm_safe_t] and [imm_safe] *)
@@ -1824,12 +1824,12 @@ Lemma imm_safe_imm_safe_t:
exists C, exists a1, exists t, exists a1', exists m',
context RV k C /\ a = C a1 /\ rred ge a1 m t a1' m' /\ forall w', ~possible_trace w t w'.
Proof.
- intros. inv H.
+ intros. inv H.
left. apply imm_safe_t_val.
left. apply imm_safe_t_loc.
left. eapply imm_safe_t_lred; eauto.
destruct (classic (exists w', possible_trace w t w')) as [[w' A] | A].
- left. eapply imm_safe_t_rred; eauto.
+ left. eapply imm_safe_t_rred; eauto.
right. exists C; exists e0; exists t; exists e'; exists m'; intuition. apply A; exists w'; auto.
left. eapply imm_safe_t_callred; eauto.
Qed.
@@ -1847,10 +1847,10 @@ Theorem not_imm_safe_t:
Csem.step ge (ExprState f (C a) k e m) E0 Stuckstate \/ can_crash_world w (ExprState f (C a) k e m).
Proof.
intros. destruct (classic (imm_safe ge e K a m)).
- exploit imm_safe_imm_safe_t; eauto.
+ exploit imm_safe_imm_safe_t; eauto.
intros [A | [C1 [a1 [t [a1' [m' [A [B [D E]]]]]]]]]. contradiction.
- right. red. exists t; econstructor; split; auto.
- left. rewrite B. eapply step_rred with (C := fun x => C(C1 x)). eauto. eauto.
+ right. red. exists t; econstructor; split; auto.
+ left. rewrite B. eapply step_rred with (C := fun x => C(C1 x)). eauto. eauto.
left. left. eapply step_stuck; eauto.
Qed.
@@ -1862,14 +1862,14 @@ Fixpoint do_alloc_variables (e: env) (m: mem) (l: list (ident * type)) {struct l
match l with
| nil => (e,m)
| (id, ty) :: l' =>
- let (m1,b1) := Mem.alloc m 0 (sizeof ge ty) in
+ let (m1,b1) := Mem.alloc m 0 (sizeof ge ty) in
do_alloc_variables (PTree.set id (b1, ty) e) m1 l'
end.
Lemma do_alloc_variables_sound:
forall l e m, alloc_variables ge e m l (fst (do_alloc_variables e m l)) (snd (do_alloc_variables e m l)).
Proof.
- induction l; intros; simpl.
+ induction l; intros; simpl.
constructor.
destruct a as [id ty]. destruct (Mem.alloc m 0 (sizeof ge ty)) as [m1 b1] eqn:?; simpl.
econstructor; eauto.
@@ -1879,12 +1879,12 @@ Lemma do_alloc_variables_complete:
forall e1 m1 l e2 m2, alloc_variables ge e1 m1 l e2 m2 ->
do_alloc_variables e1 m1 l = (e2, m2).
Proof.
- induction 1; simpl.
+ induction 1; simpl.
auto.
- rewrite H; rewrite IHalloc_variables; auto.
+ rewrite H; rewrite IHalloc_variables; auto.
Qed.
-Function sem_bind_parameters (w: world) (e: env) (m: mem) (l: list (ident * type)) (lv: list val)
+Function sem_bind_parameters (w: world) (e: env) (m: mem) (l: list (ident * type)) (lv: list val)
{struct l} : option mem :=
match l, lv with
| nil, nil => Some m
@@ -1900,7 +1900,7 @@ Function sem_bind_parameters (w: world) (e: env) (m: mem) (l: list (ident * type
end.
Lemma sem_bind_parameters_sound : forall w e m l lv m',
- sem_bind_parameters w e m l lv = Some m' ->
+ sem_bind_parameters w e m l lv = Some m' ->
bind_parameters ge e m l lv m'.
Proof.
intros; functional induction (sem_bind_parameters w e m l lv); try discriminate.
@@ -1916,7 +1916,7 @@ Proof.
rewrite H. rewrite dec_eq_true.
assert (possible_trace w E0 w) by constructor.
rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ H0 H2).
- simpl. auto.
+ simpl. auto.
Qed.
Inductive transition : Type := TR (rule: string) (t: trace) (S': state).
@@ -1946,7 +1946,7 @@ Definition do_step (w: world) (s: state) : list transition :=
do b <- bool_val v ty m;
ret "step_ifthenelse_2" (State f (if b then s1 else s2) k e m)
| Kwhile1 x s k =>
- do b <- bool_val v ty m;
+ do b <- bool_val v ty m;
if b
then ret "step_while_true" (State f s (Kwhile2 x s k) e m)
else ret "step_while_false" (State f Sskip k e m)
@@ -2018,7 +2018,7 @@ Definition do_step (w: world) (s: state) : list transition :=
ret "step_return_0" (Returnstate Vundef (call_cont k) m')
| State f (Sreturn (Some x)) k e m =>
ret "step_return_1" (ExprState f x (Kreturn k) e m)
- | State f Sskip ((Kstop | Kcall _ _ _ _ _) as k) e m =>
+ | State f Sskip ((Kstop | Kcall _ _ _ _ _) as k) e m =>
do m' <- Mem.free_list m (blocks_of_env ge e);
ret "step_skip_call" (Returnstate Vundef k m')
@@ -2057,10 +2057,10 @@ Definition do_step (w: world) (s: state) : list transition :=
Ltac myinv :=
match goal with
| [ |- In _ nil -> _ ] => intro X; elim X
- | [ |- In _ (ret _ _) -> _ ] =>
+ | [ |- In _ (ret _ _) -> _ ] =>
intro X; elim X; clear X;
[intro EQ; unfold ret in EQ; inv EQ; myinv | myinv]
- | [ |- In _ (_ :: nil) -> _ ] =>
+ | [ |- In _ (_ :: nil) -> _ ] =>
intro X; elim X; clear X; [intro EQ; inv EQ; myinv | myinv]
| [ |- In _ (match ?x with Some _ => _ | None => _ end) -> _ ] => destruct x eqn:?; myinv
| [ |- In _ (match ?x with false => _ | true => _ end) -> _ ] => destruct x eqn:?; myinv
@@ -2096,7 +2096,7 @@ Proof with try (left; right; econstructor; eauto; fail).
generalize (step_expr_sound e w r RV m). unfold reducts_ok. intros [P Q].
exploit P; eauto. intros [a' [k' [CTX [EQ RD]]]].
unfold expr_final_state in A. simpl in A.
- destruct k'; destruct rd; inv A; simpl in RD; try contradiction.
+ destruct k'; destruct rd; inv A; simpl in RD; try contradiction.
(* lred *)
left; left; apply step_lred; auto.
(* stuck lred *)
@@ -2111,11 +2111,11 @@ Proof with try (left; right; econstructor; eauto; fail).
destruct fd; myinv.
(* internal *)
destruct (do_alloc_variables empty_env m (fn_params f ++ fn_vars f)) as [e m1] eqn:?.
- myinv. left; right; apply step_internal_function with m1. auto.
- change e with (fst (e,m1)). change m1 with (snd (e,m1)) at 2. rewrite <- Heqp.
+ myinv. left; right; apply step_internal_function with m1. auto.
+ change e with (fst (e,m1)). change m1 with (snd (e,m1)) at 2. rewrite <- Heqp.
apply do_alloc_variables_sound. eapply sem_bind_parameters_sound; eauto.
(* external *)
- destruct p as [[[w' tr] v] m']. myinv. left; right; constructor.
+ destruct p as [[[w' tr] v] m']. myinv. left; right; constructor.
eapply do_ef_external_sound; eauto.
(* returnstate *)
destruct k; myinv...
@@ -2126,10 +2126,10 @@ Qed.
Remark estep_not_val:
forall f a k e m t S, estep ge (ExprState f a k e m) t S -> is_val a = None.
Proof.
- intros.
+ intros.
assert (forall b from to C, context from to C -> (from = to /\ C = fun x => x) \/ is_val (C b) = None).
- induction 1; simpl; auto.
- inv H.
+ induction 1; simpl; auto.
+ inv H.
destruct (H0 a0 _ _ _ H9) as [[A B] | A]. subst. inv H8; auto. auto.
destruct (H0 a0 _ _ _ H9) as [[A B] | A]. subst. inv H8; auto. auto.
destruct (H0 a0 _ _ _ H9) as [[A B] | A]. subst. inv H8; auto. auto.
@@ -2141,7 +2141,7 @@ Theorem do_step_complete:
possible_trace w t w' -> Csem.step ge S t S' -> exists rule, In (TR rule t S') (do_step w S).
Proof with (unfold ret; eauto with coqlib).
intros until w'; intros PT H.
- destruct H.
+ destruct H.
(* Expression step *)
inversion H; subst; exploit estep_not_val; eauto; intro NOTVAL.
(* lred *)
@@ -2170,7 +2170,7 @@ Proof with (unfold ret; eauto with coqlib).
change (TR rule E0 (Callstate fd vargs (Kcall f e C ty k) m)) with (expr_final_state f k e (C, Callred rule fd vargs ty m)).
apply in_map.
generalize (step_expr_context e w _ _ _ H1 a m). unfold reducts_incl.
- intro. replace C with (fun x => C x). apply H2.
+ intro. replace C with (fun x => C x). apply H2.
rewrite STEP; unfold topred; auto with coqlib.
apply extensionality; auto.
(* stuck *)
@@ -2179,7 +2179,7 @@ Proof with (unfold ret; eauto with coqlib).
simpl do_step. rewrite NOTVAL.
exists "step_stuck".
change (TR "step_stuck" E0 Stuckstate) with (expr_final_state f k e (C', Stuckred)).
- apply in_map. auto.
+ apply in_map. auto.
(* Statement step *)
inv H; simpl; econstructor...
diff --git a/cfrontend/Clight.v b/cfrontend/Clight.v
index 77511b2c..8722da69 100644
--- a/cfrontend/Clight.v
+++ b/cfrontend/Clight.v
@@ -325,7 +325,7 @@ Fixpoint bind_parameter_temps (formals: list (ident * type)) (args: list val)
| nil, nil => Some le
| (id, t) :: xl, v :: vl => bind_parameter_temps xl vl (PTree.set id v le)
| _, _ => None
- end.
+ end.
(** Return the list of blocks in the codomain of [e], with low and high bounds. *)
@@ -419,7 +419,7 @@ Inductive eval_expr: expr -> val -> Prop :=
| eval_Ealignof: forall ty1 ty,
eval_expr (Ealignof ty1 ty) (Vint (Int.repr (alignof ge ty1)))
| eval_Elvalue: forall a loc ofs v,
- eval_lvalue a loc ofs ->
+ eval_lvalue a loc ofs ->
deref_loc (typeof a) m loc ofs v ->
eval_expr a v
@@ -523,11 +523,11 @@ Inductive state: Type :=
(res: val)
(k: cont)
(m: mem) : state.
-
-(** Find the statement and manufacture the continuation
+
+(** Find the statement and manufacture the continuation
corresponding to a label *)
-Fixpoint find_label (lbl: label) (s: statement) (k: cont)
+Fixpoint find_label (lbl: label) (s: statement) (k: cont)
{struct s}: option (statement * cont) :=
match s with
| Ssequence s1 s2 =>
@@ -552,7 +552,7 @@ Fixpoint find_label (lbl: label) (s: statement) (k: cont)
| _ => None
end
-with find_label_ls (lbl: label) (sl: labeled_statements) (k: cont)
+with find_label_ls (lbl: label) (sl: labeled_statements) (k: cont)
{struct sl}: option (statement * cont) :=
match sl with
| LSnil => None
@@ -646,7 +646,7 @@ Inductive step: state -> trace -> state -> Prop :=
step (State f (Sreturn None) k e le m)
E0 (Returnstate Vundef (call_cont k) m')
| step_return_1: forall f a k e le m v v' m',
- eval_expr e le m a v ->
+ eval_expr e le m a v ->
sem_cast v (typeof a) f.(fn_return) = Some v' ->
Mem.free_list m (blocks_of_env e) = Some m' ->
step (State f (Sreturn (Some a)) k e le m)
@@ -764,10 +764,10 @@ Proof.
intros. subst. inv H0. exists s1; auto.
inversion H; subst; auto.
(* builtin *)
- exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
econstructor; econstructor; eauto.
(* external *)
- exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (Returnstate vres2 k m2). econstructor; eauto.
(* trace length *)
red; simpl; intros. inv H; simpl; try omega.
diff --git a/cfrontend/ClightBigstep.v b/cfrontend/ClightBigstep.v
index ac8931e5..ee653d50 100644
--- a/cfrontend/ClightBigstep.v
+++ b/cfrontend/ClightBigstep.v
@@ -66,9 +66,9 @@ Definition outcome_result_value (out: outcome) (t: type) (v: val) : Prop :=
| Out_return None, Tvoid => v = Vundef
| Out_return (Some (v',t')), ty => ty <> Tvoid /\ sem_cast v' t' t = Some v
| _, _ => False
- end.
+ end.
-(** [exec_stmt ge e m1 s t m2 out] describes the execution of
+(** [exec_stmt ge e m1 s t m2 out] describes the execution of
the statement [s]. [out] is the outcome for this execution.
[m1] is the initial memory state, [m2] the final memory state.
[t] is the trace of input/output events performed during this
@@ -88,7 +88,7 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env ->
| exec_Sset: forall e le m id a v,
eval_expr ge e le m a v ->
exec_stmt e le m (Sset id a)
- E0 (PTree.set id v le) m Out_normal
+ E0 (PTree.set id v le) m Out_normal
| exec_Scall: forall e le m optid a al tyargs tyres cconv vf vargs f t m' vres,
classify_fun (typeof a) = fun_case_f tyargs tyres cconv ->
eval_expr ge e le m a vf ->
@@ -244,7 +244,7 @@ End BIGSTEP.
Inductive bigstep_program_terminates (p: program): trace -> int -> Prop :=
| bigstep_program_terminates_intro: forall b f m0 m1 t r,
- let ge := globalenv p in
+ let ge := globalenv p in
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
@@ -254,7 +254,7 @@ Inductive bigstep_program_terminates (p: program): trace -> int -> Prop :=
Inductive bigstep_program_diverges (p: program): traceinf -> Prop :=
| bigstep_program_diverges_intro: forall b f m0 t,
- let ge := globalenv p in
+ let ge := globalenv p in
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
@@ -282,7 +282,7 @@ Inductive outcome_state_match
outcome_state_match e le m f k Out_continue (State f Scontinue k e le m)
| osm_return_none: forall k',
call_cont k' = call_cont k ->
- outcome_state_match e le m f k
+ outcome_state_match e le m f k
(Out_return None) (State f (Sreturn None) k' e le m)
| osm_return_some: forall a v k',
call_cont k' = call_cont k ->
@@ -322,7 +322,7 @@ Proof.
(* call *)
econstructor; split.
- eapply star_left. econstructor; eauto.
+ eapply star_left. econstructor; eauto.
eapply star_right. apply H5. simpl; auto. econstructor. reflexivity. traceEq.
constructor.
@@ -331,10 +331,10 @@ Proof.
(* sequence 2 *)
destruct (H0 f (Kseq s2 k)) as [S1 [A1 B1]]. inv B1.
- destruct (H2 f k) as [S2 [A2 B2]].
+ destruct (H2 f k) as [S2 [A2 B2]].
econstructor; split.
eapply star_left. econstructor.
- eapply star_trans. eexact A1.
+ eapply star_trans. eexact A1.
eapply star_left. constructor. eexact A2.
reflexivity. reflexivity. traceEq.
auto.
@@ -385,10 +385,10 @@ Proof.
| _ => S1
end).
exists S2; split.
- eapply star_left. eapply step_loop.
+ eapply star_left. eapply step_loop.
eapply star_trans. eexact A1.
unfold S2. inversion H1; subst.
- inv B1. apply star_one. constructor.
+ inv B1. apply star_one. constructor.
apply star_refl.
reflexivity. traceEq.
unfold S2. inversion H1; subst. constructor. inv B1; econstructor; eauto.
@@ -402,14 +402,14 @@ Proof.
| _ => S2
end).
exists S3; split.
- eapply star_left. eapply step_loop.
+ eapply star_left. eapply step_loop.
eapply star_trans. eexact A1.
eapply star_left with (s2 := State f s2 (Kloop2 s1 s2 k) e le1 m1).
inv H1; inv B1; constructor; auto.
eapply star_trans. eexact A2.
unfold S3. inversion H4; subst.
inv B2. apply star_one. constructor. apply star_refl.
- reflexivity. reflexivity. reflexivity. traceEq.
+ reflexivity. reflexivity. reflexivity. traceEq.
unfold S3. inversion H4; subst. constructor. inv B2; econstructor; eauto.
(* loop loop *)
@@ -417,15 +417,15 @@ Proof.
destruct (H3 f (Kloop2 s1 s2 k)) as [S2 [A2 B2]].
destruct (H5 f k) as [S3 [A3 B3]].
exists S3; split.
- eapply star_left. eapply step_loop.
+ eapply star_left. eapply step_loop.
eapply star_trans. eexact A1.
eapply star_left with (s2 := State f s2 (Kloop2 s1 s2 k) e le1 m1).
inv H1; inv B1; constructor; auto.
eapply star_trans. eexact A2.
eapply star_left with (s2 := State f (Sloop s1 s2) k e le2 m2).
- inversion H4; subst; inv B2; constructor; auto.
+ inversion H4; subst; inv B2; constructor; auto.
eexact A3.
- reflexivity. reflexivity. reflexivity. reflexivity. traceEq.
+ reflexivity. reflexivity. reflexivity. reflexivity. traceEq.
auto.
(* switch *)
@@ -438,12 +438,12 @@ Proof.
| _ => S1
end).
exists S2; split.
- eapply star_left. eapply step_switch; eauto.
- eapply star_trans. eexact A1.
- unfold S2; inv B1.
- apply star_one. constructor. auto.
- apply star_one. constructor. auto.
- apply star_one. constructor.
+ eapply star_left. eapply step_switch; eauto.
+ eapply star_trans. eexact A1.
+ unfold S2; inv B1.
+ apply star_one. constructor. auto.
+ apply star_one. constructor. auto.
+ apply star_one. constructor.
apply star_refl.
apply star_refl.
reflexivity. traceEq.
@@ -452,7 +452,7 @@ Proof.
(* call internal *)
destruct (H3 f k) as [S1 [A1 B1]].
eapply star_left. eapply step_internal_function; eauto. econstructor; eauto.
- eapply star_right. eexact A1.
+ eapply star_right. eexact A1.
inv B1; simpl in H4; try contradiction.
(* Out_normal *)
assert (fn_return f = Tvoid /\ vres = Vundef).
@@ -471,7 +471,7 @@ Proof.
reflexivity. traceEq.
(* call external *)
- apply star_one. apply step_external_function; auto.
+ apply star_one. apply step_external_function; auto.
Qed.
Lemma exec_stmt_steps:
@@ -506,7 +506,7 @@ Proof.
(* call *)
eapply forever_N_plus.
- apply plus_one. eapply step_call; eauto.
+ apply plus_one. eapply step_call; eauto.
eapply CIH_FUN. eauto. traceEq.
(* seq 1 *)
@@ -517,25 +517,25 @@ Proof.
destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H0 f (Kseq s2 k)) as [S1 [A1 B1]].
inv B1.
eapply forever_N_plus.
- eapply plus_left. constructor. eapply star_trans. eexact A1.
+ eapply plus_left. constructor. eapply star_trans. eexact A1.
apply star_one. constructor. reflexivity. reflexivity.
apply CIH_STMT; eauto. traceEq.
(* ifthenelse *)
eapply forever_N_plus.
- apply plus_one. eapply step_ifthenelse with (b := b); eauto.
+ apply plus_one. eapply step_ifthenelse with (b := b); eauto.
apply CIH_STMT; eauto. traceEq.
(* loop body 1 *)
eapply forever_N_plus.
- eapply plus_one. constructor.
+ eapply plus_one. constructor.
apply CIH_STMT; eauto. traceEq.
(* loop body 2 *)
destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H0 f (Kloop1 s1 s2 k)) as [S1 [A1 B1]].
eapply forever_N_plus with (s2 := State f s2 (Kloop2 s1 s2 k) e le1 m1).
eapply plus_left. constructor.
eapply star_right. eexact A1.
- inv H1; inv B1; constructor; auto.
+ inv H1; inv B1; constructor; auto.
reflexivity. reflexivity.
apply CIH_STMT; eauto. traceEq.
(* loop loop *)
@@ -571,14 +571,14 @@ Proof.
(* termination *)
inv H. econstructor; econstructor.
split. econstructor; eauto.
- split. eapply eval_funcall_steps. eauto. red; auto.
+ split. eapply eval_funcall_steps. eauto. red; auto.
econstructor.
(* divergence *)
inv H. econstructor.
split. econstructor; eauto.
eapply forever_N_forever with (order := order).
red; intros. constructor; intros. red in H. elim H.
- eapply evalinf_funcall_forever; eauto.
+ eapply evalinf_funcall_forever; eauto.
Qed.
End BIGSTEP_TO_TRANSITIONS.
diff --git a/cfrontend/Cminorgen.v b/cfrontend/Cminorgen.v
index 82509b04..c2b59fbe 100644
--- a/cfrontend/Cminorgen.v
+++ b/cfrontend/Cminorgen.v
@@ -37,7 +37,7 @@ Local Open Scope error_monad_scope.
or assigning such a variable becomes a load or store operation at
that address. Only scalar local variables whose address is never
taken in the Csharpminor code can be mapped to Cminor local
- variable, since the latter do not reside in memory.
+ variable, since the latter do not reside in memory.
Another task performed during the translation to Cminor is to
transform the Clight-like [switch] construct of Csharpminor
@@ -220,7 +220,7 @@ with transl_lblstmt (cenv: compilenv) (xenv: exit_env) (ls: Csharpminor.lbl_stmt
(** * Stack layout *)
-(** Layout of the Cminor stack data block and construction of the
+(** Layout of the Cminor stack data block and construction of the
compilation environment. Every Csharpminor local variable is
allocated a slot in the Cminor stack data. Sufficient padding is
inserted to ensure adequate alignment of addresses. *)
@@ -240,7 +240,7 @@ Definition assign_variable
Definition assign_variables (cenv_stacksize: compilenv * Z) (vars: list (ident * Z)) : compilenv * Z :=
List.fold_left assign_variable vars cenv_stacksize.
-(** Before allocating stack slots, we sort variables by increasing size
+(** Before allocating stack slots, we sort variables by increasing size
so as to minimize padding. *)
Module VarOrder <: TotalLeBool.
@@ -248,7 +248,7 @@ Module VarOrder <: TotalLeBool.
Definition leb (v1 v2: t) : bool := zle (snd v1) (snd v2).
Theorem leb_total: forall v1 v2, leb v1 v2 = true \/ leb v2 v1 = true.
Proof.
- unfold leb; intros.
+ unfold leb; intros.
assert (snd v1 <= snd v2 \/ snd v2 <= snd v1) by omega.
unfold proj_sumbool. destruct H; [left|right]; apply zle_true; auto.
Qed.
diff --git a/cfrontend/Cminorgenproof.v b/cfrontend/Cminorgenproof.v
index dfc69412..7a5d882e 100644
--- a/cfrontend/Cminorgenproof.v
+++ b/cfrontend/Cminorgenproof.v
@@ -70,7 +70,7 @@ Proof (Genv.find_var_info_transf_partial transl_fundef _ TRANSL).
Lemma sig_preserved_body:
forall f tf cenv size,
- transl_funbody cenv size f = OK tf ->
+ transl_funbody cenv size f = OK tf ->
tf.(fn_sig) = Csharpminor.fn_sig f.
Proof.
intros. unfold transl_funbody in H. monadInv H; reflexivity.
@@ -78,13 +78,13 @@ Qed.
Lemma sig_preserved:
forall f tf,
- transl_fundef f = OK tf ->
+ transl_fundef f = OK tf ->
Cminor.funsig tf = Csharpminor.funsig f.
Proof.
intros until tf; destruct f; simpl.
unfold transl_function. destruct (build_compilenv f).
case (zle z Int.max_unsigned); simpl bind; try congruence.
- intros. monadInv H. simpl. eapply sig_preserved_body; eauto.
+ intros. monadInv H. simpl. eapply sig_preserved_body; eauto.
intro. inv H. reflexivity.
Qed.
@@ -92,7 +92,7 @@ Qed.
Lemma load_freelist:
forall fbl chunk m b ofs m',
- (forall b' lo hi, In (b', lo, hi) fbl -> b' <> b) ->
+ (forall b' lo hi, In (b', lo, hi) fbl -> b' <> b) ->
Mem.free_list m fbl = Some m' ->
Mem.load chunk m' b ofs = Mem.load chunk m b ofs.
Proof.
@@ -100,10 +100,10 @@ Proof.
simpl in H0. congruence.
destruct a as [[b' lo] hi].
generalize H0. simpl. case_eq (Mem.free m b' lo hi); try congruence.
- intros m1 FR1 FRL.
+ intros m1 FR1 FRL.
transitivity (Mem.load chunk m1 b ofs).
- eapply IHfbl; eauto. intros. eapply H. eauto with coqlib.
- eapply Mem.load_free; eauto. left. apply sym_not_equal. eapply H. auto with coqlib.
+ eapply IHfbl; eauto. intros. eapply H. eauto with coqlib.
+ eapply Mem.load_free; eauto. left. apply sym_not_equal. eapply H. auto with coqlib.
Qed.
Lemma perm_freelist:
@@ -125,7 +125,7 @@ Lemma nextblock_freelist:
Proof.
induction fbl; intros until m'; simpl.
congruence.
- destruct a as [[b lo] hi].
+ destruct a as [[b lo] hi].
case_eq (Mem.free m b lo hi); intros; try congruence.
transitivity (Mem.nextblock m0). eauto. eapply Mem.nextblock_free; eauto.
Qed.
@@ -141,7 +141,7 @@ Proof.
revert H. destruct a as [[b' lo'] hi'].
caseEq (Mem.free m b' lo' hi'); try congruence.
intros m1 FREE1 FREE2.
- destruct H0. inv H.
+ destruct H0. inv H.
eauto with mem.
red; intros. eapply Mem.perm_free_3; eauto. exploit IHl; eauto.
Qed.
@@ -150,7 +150,7 @@ Lemma nextblock_storev:
forall chunk m addr v m',
Mem.storev chunk m addr v = Some m' -> Mem.nextblock m' = Mem.nextblock m.
Proof.
- unfold Mem.storev; intros. destruct addr; try discriminate.
+ unfold Mem.storev; intros. destruct addr; try discriminate.
eapply Mem.nextblock_store; eauto.
Qed.
@@ -159,7 +159,7 @@ Qed.
(** In C#minor, every variable is stored in a separate memory block.
In the corresponding Cminor code, these variables become sub-blocks
of the stack data block. We capture these changes in memory via a
- memory injection [f]:
+ memory injection [f]:
[f b = Some(b', ofs)] means that C#minor block [b] corresponds
to a sub-block of Cminor block [b] at offset [ofs].
@@ -239,7 +239,7 @@ Record match_env (f: meminj) (cenv: compilenv)
f b = Some(tb, delta) -> Plt b lo -> Plt tb sp
}.
-Ltac geninv x :=
+Ltac geninv x :=
let H := fresh in (generalize x; intro H; inv H).
Lemma match_env_invariant:
@@ -254,7 +254,7 @@ Proof.
(* vars *)
intros. geninv (me_vars0 id); econstructor; eauto.
(* bounded *)
- intros. eauto.
+ intros. eauto.
(* below *)
intros. rewrite H2 in H; eauto.
Qed.
@@ -267,7 +267,7 @@ Remark inject_incr_separated_same:
forall b, Mem.valid_block m1 b -> f2 b = f1 b.
Proof.
intros. case_eq (f1 b).
- intros [b' delta] EQ. apply H; auto.
+ intros [b' delta] EQ. apply H; auto.
intros EQ. case_eq (f2 b).
intros [b'1 delta1] EQ1. exploit H0; eauto. intros [C D]. contradiction.
auto.
@@ -294,7 +294,7 @@ Lemma match_env_external_call:
Proof.
intros. apply match_env_invariant with f1; auto.
intros. eapply inject_incr_separated_same'; eauto.
- intros. eapply inject_incr_separated_same; eauto. red. destruct H. xomega.
+ intros. eapply inject_incr_separated_same; eauto. red. destruct H. xomega.
Qed.
(** [match_env] and allocations *)
@@ -317,7 +317,7 @@ Proof.
(* vars *)
intros. rewrite PTree.gsspec. destruct (peq id0 id).
(* the new var *)
- subst id0. rewrite CENV. constructor. econstructor. eauto.
+ subst id0. rewrite CENV. constructor. econstructor. eauto.
rewrite Int.add_commut; rewrite Int.add_zero; auto.
(* old vars *)
generalize (me_vars0 id0). rewrite PTree.gro; auto. intros M; inv M.
@@ -331,8 +331,8 @@ Proof.
exploit me_bounded0; eauto. rewrite NEXTBLOCK; xomega.
(* inv *)
intros. destruct (eq_block b (Mem.nextblock m1)).
- subst b. rewrite SAME in H; inv H. exists id; exists sz. apply PTree.gss.
- rewrite OTHER in H; auto. exploit me_inv0; eauto.
+ subst b. rewrite SAME in H; inv H. exists id; exists sz. apply PTree.gss.
+ rewrite OTHER in H; auto. exploit me_inv0; eauto.
intros [id1 [sz1 EQ]]. exists id1; exists sz1. rewrite PTree.gso; auto. congruence.
(* incr *)
intros. rewrite OTHER in H. eauto. unfold block in *; xomega.
@@ -351,7 +351,7 @@ Lemma match_bounds_invariant:
PTree.get id e = Some(b, sz) -> Mem.perm m2 b ofs Max p -> Mem.perm m1 b ofs Max p) ->
match_bounds e m2.
Proof.
- intros; red; intros. eapply H; eauto.
+ intros; red; intros. eapply H; eauto.
Qed.
(** ** Permissions on the Cminor stack block *)
@@ -367,7 +367,7 @@ Inductive is_reachable_from_env (f: meminj) (e: Csharpminor.env) (sp: block) (of
is_reachable_from_env f e sp ofs.
Definition padding_freeable (f: meminj) (e: Csharpminor.env) (tm: mem) (sp: block) (sz: Z) : Prop :=
- forall ofs,
+ forall ofs,
0 <= ofs < sz -> Mem.perm tm sp ofs Cur Freeable \/ is_reachable_from_env f e sp ofs.
Lemma padding_freeable_invariant:
@@ -382,7 +382,7 @@ Proof.
exploit H; eauto. intros [A | A].
left; auto.
right. inv A. exploit me_bounded; eauto. intros [D E].
- econstructor; eauto. rewrite H2; auto.
+ econstructor; eauto. rewrite H2; auto.
Qed.
(** Decidability of the [is_reachable_from_env] predicate. *)
@@ -390,7 +390,7 @@ Qed.
Lemma is_reachable_from_env_dec:
forall f e sp ofs, is_reachable_from_env f e sp ofs \/ ~is_reachable_from_env f e sp ofs.
Proof.
- intros.
+ intros.
set (pred := fun id_b_sz : ident * (block * Z) =>
match id_b_sz with
| (id, (b, sz)) =>
@@ -404,22 +404,22 @@ Proof.
end).
destruct (List.existsb pred (PTree.elements e)) eqn:?.
(* yes *)
- rewrite List.existsb_exists in Heqb.
+ rewrite List.existsb_exists in Heqb.
destruct Heqb as [[id [b sz]] [A B]].
simpl in B. destruct (f b) as [[sp' delta] |] eqn:?; try discriminate.
destruct (eq_block sp sp'); try discriminate.
destruct (andb_prop _ _ B).
left. apply is_reachable_intro with id b sz delta.
- apply PTree.elements_complete; auto.
+ apply PTree.elements_complete; auto.
congruence.
split; eapply proj_sumbool_true; eauto.
(* no *)
- right; red; intro NE; inv NE.
+ right; red; intro NE; inv NE.
assert (existsb pred (PTree.elements e) = true).
rewrite List.existsb_exists. exists (id, (b, sz)); split.
- apply PTree.elements_correct; auto.
+ apply PTree.elements_correct; auto.
simpl. rewrite H0. rewrite dec_eq_true.
- unfold proj_sumbool. destruct H1. rewrite zle_true; auto. rewrite zlt_true; auto.
+ unfold proj_sumbool. destruct H1. rewrite zle_true; auto. rewrite zlt_true; auto.
congruence.
Qed.
@@ -443,8 +443,8 @@ Remark inj_preserves_globals:
Proof.
intros. inv H.
split. intros. apply DOMAIN. eapply SYMBOLS. eauto.
- split. intros. apply DOMAIN. eapply VARINFOS. eauto.
- intros. symmetry. eapply IMAGE; eauto.
+ split. intros. apply DOMAIN. eapply VARINFOS. eauto.
+ intros. symmetry. eapply IMAGE; eauto.
Qed.
(** * Invariant on abstract call stacks *)
@@ -481,7 +481,7 @@ Inductive match_callstack (f: meminj) (m: mem) (tm: mem):
forall hi bound tbound,
match_globalenvs f hi ->
Ple hi bound -> Ple hi tbound ->
- match_callstack f m tm nil bound tbound
+ match_callstack f m tm nil bound tbound
| mcs_cons:
forall cenv tf e le te sp lo hi cs bound tbound
(BOUND: Ple hi bound)
@@ -524,16 +524,16 @@ Proof.
assert (Ple lo hi) by (eapply me_low_high; eauto).
econstructor; eauto.
eapply match_temps_invariant; eauto.
- eapply match_env_invariant; eauto.
+ eapply match_env_invariant; eauto.
intros. apply H3. xomega.
eapply match_bounds_invariant; eauto.
- intros. eapply H1; eauto.
- exploit me_bounded; eauto. xomega.
- eapply padding_freeable_invariant; eauto.
- intros. apply H3. xomega.
- eapply IHmatch_callstack; eauto.
- intros. eapply H1; eauto. xomega.
- intros. eapply H2; eauto. xomega.
+ intros. eapply H1; eauto.
+ exploit me_bounded; eauto. xomega.
+ eapply padding_freeable_invariant; eauto.
+ intros. apply H3. xomega.
+ eapply IHmatch_callstack; eauto.
+ intros. eapply H1; eauto. xomega.
+ intros. eapply H2; eauto. xomega.
intros. eapply H3; eauto. xomega.
intros. eapply H4; eauto. xomega.
Qed.
@@ -545,7 +545,7 @@ Lemma match_callstack_incr_bound:
match_callstack f m tm cs bound' tbound'.
Proof.
intros. inv H.
- econstructor; eauto. xomega. xomega.
+ econstructor; eauto. xomega. xomega.
constructor; auto. xomega. xomega.
Qed.
@@ -558,7 +558,7 @@ Lemma match_callstack_set_temp:
match_callstack f m tm (Frame cenv tf e (PTree.set id v le) (PTree.set id tv te) sp lo hi :: cs) bound tbound.
Proof.
intros. inv H0. constructor; auto.
- eapply match_temps_assign; eauto.
+ eapply match_temps_assign; eauto.
Qed.
(** Preservation of [match_callstack] by freeing all blocks allocated
@@ -569,7 +569,7 @@ Lemma in_blocks_of_env:
forall e id b sz,
e!id = Some(b, sz) -> In (b, 0, sz) (blocks_of_env e).
Proof.
- unfold blocks_of_env; intros.
+ unfold blocks_of_env; intros.
change (b, 0, sz) with (block_of_binding (id, (b, sz))).
apply List.in_map. apply PTree.elements_correct. auto.
Qed.
@@ -579,9 +579,9 @@ Lemma in_blocks_of_env_inv:
In (b, lo, hi) (blocks_of_env e) ->
exists id, e!id = Some(b, hi) /\ lo = 0.
Proof.
- unfold blocks_of_env; intros.
+ unfold blocks_of_env; intros.
exploit list_in_map_inv; eauto. intros [[id [b' sz]] [A B]].
- unfold block_of_binding in A. inv A.
+ unfold block_of_binding in A. inv A.
exists id; intuition. apply PTree.elements_complete. auto.
Qed.
@@ -601,10 +601,10 @@ Proof.
red; intros.
exploit PERM; eauto. intros [A | A].
auto.
- inv A. assert (Mem.range_perm m b 0 sz Cur Freeable).
+ inv A. assert (Mem.range_perm m b 0 sz Cur Freeable).
eapply free_list_freeable; eauto. eapply in_blocks_of_env; eauto.
- replace ofs with ((ofs - delta) + delta) by omega.
- eapply Mem.perm_inject; eauto. apply H3. omega.
+ replace ofs with ((ofs - delta) + delta) by omega.
+ eapply Mem.perm_inject; eauto. apply H3. omega.
destruct X as [tm' FREE].
exploit nextblock_freelist; eauto. intro NEXT.
exploit Mem.nextblock_free; eauto. intro NEXT'.
@@ -615,10 +615,10 @@ Proof.
intros. eapply perm_freelist; eauto.
intros. eapply Mem.perm_free_1; eauto. left; unfold block; xomega. xomega. xomega.
eapply Mem.free_inject; eauto.
- intros. exploit me_inv0; eauto. intros [id [sz A]].
+ intros. exploit me_inv0; eauto. intros [id [sz A]].
exists 0; exists sz; split.
eapply in_blocks_of_env; eauto.
- eapply BOUND0; eauto. eapply Mem.perm_max. eauto.
+ eapply BOUND0; eauto. eapply Mem.perm_max. eauto.
Qed.
(** Preservation of [match_callstack] by external calls. *)
@@ -635,33 +635,33 @@ Lemma match_callstack_external_call:
Ple bound (Mem.nextblock m1) -> Ple tbound (Mem.nextblock m1') ->
match_callstack f2 m2 m2' cs bound tbound.
Proof.
- intros until m2'.
+ intros until m2'.
intros UNMAPPED OUTOFREACH INCR SEPARATED MAXPERMS.
induction 1; intros.
(* base case *)
apply mcs_nil with hi; auto.
inv H. constructor; auto.
- intros. case_eq (f1 b1).
- intros [b2' delta'] EQ. rewrite (INCR _ _ _ EQ) in H. inv H. eauto.
- intro EQ. exploit SEPARATED; eauto. intros [A B]. elim B. red. xomega.
+ intros. case_eq (f1 b1).
+ intros [b2' delta'] EQ. rewrite (INCR _ _ _ EQ) in H. inv H. eauto.
+ intro EQ. exploit SEPARATED; eauto. intros [A B]. elim B. red. xomega.
(* inductive case *)
- constructor. auto. auto.
+ constructor. auto. auto.
eapply match_temps_invariant; eauto.
- eapply match_env_invariant; eauto.
- red in SEPARATED. intros. destruct (f1 b) as [[b' delta']|] eqn:?.
+ eapply match_env_invariant; eauto.
+ red in SEPARATED. intros. destruct (f1 b) as [[b' delta']|] eqn:?.
exploit INCR; eauto. congruence.
- exploit SEPARATED; eauto. intros [A B]. elim B. red. xomega.
+ exploit SEPARATED; eauto. intros [A B]. elim B. red. xomega.
intros. assert (Ple lo hi) by (eapply me_low_high; eauto).
- destruct (f1 b) as [[b' delta']|] eqn:?.
- apply INCR; auto.
+ destruct (f1 b) as [[b' delta']|] eqn:?.
+ apply INCR; auto.
destruct (f2 b) as [[b' delta']|] eqn:?; auto.
exploit SEPARATED; eauto. intros [A B]. elim A. red. xomega.
- eapply match_bounds_invariant; eauto.
- intros. eapply MAXPERMS; eauto. red. exploit me_bounded; eauto. xomega.
+ eapply match_bounds_invariant; eauto.
+ intros. eapply MAXPERMS; eauto. red. exploit me_bounded; eauto. xomega.
(* padding-freeable *)
red; intros.
destruct (is_reachable_from_env_dec f1 e sp ofs).
- inv H3. right. apply is_reachable_intro with id b sz delta; auto.
+ inv H3. right. apply is_reachable_intro with id b sz delta; auto.
exploit PERM; eauto. intros [A|A]; try contradiction.
left. eapply Mem.perm_unchanged_on; eauto.
red; intros; red; intros. elim H3.
@@ -698,10 +698,10 @@ Proof.
xomega.
rewrite PTree.gempty in H4; discriminate.
eelim Mem.fresh_block_alloc; eauto. eapply Mem.valid_block_inject_2; eauto.
- rewrite RES. change (Mem.valid_block tm tb). eapply Mem.valid_block_inject_2; eauto.
+ rewrite RES. change (Mem.valid_block tm tb). eapply Mem.valid_block_inject_2; eauto.
red; intros. rewrite PTree.gempty in H4. discriminate.
- red; intros. left. eapply Mem.perm_alloc_2; eauto.
- eapply match_callstack_invariant with (tm1 := tm); eauto.
+ red; intros. left. eapply Mem.perm_alloc_2; eauto.
+ eapply match_callstack_invariant with (tm1 := tm); eauto.
rewrite RES; auto.
intros. eapply Mem.perm_alloc_1; eauto.
Qed.
@@ -721,10 +721,10 @@ Lemma match_callstack_alloc_left:
(Frame cenv tf (PTree.set id (b, sz) e) le te sp lo (Mem.nextblock m2) :: cs)
(Mem.nextblock m2) (Mem.nextblock tm).
Proof.
- intros. inv H.
+ intros. inv H.
exploit Mem.nextblock_alloc; eauto. intros NEXTBLOCK.
exploit Mem.alloc_result; eauto. intros RES.
- assert (LO: Ple lo (Mem.nextblock m1)) by (eapply me_low_high; eauto).
+ assert (LO: Ple lo (Mem.nextblock m1)) by (eapply me_low_high; eauto).
constructor.
xomega.
auto.
@@ -732,17 +732,17 @@ Proof.
eapply match_env_alloc; eauto.
red; intros. rewrite PTree.gsspec in H. destruct (peq id0 id).
inversion H. subst b0 sz0 id0. eapply Mem.perm_alloc_3; eauto.
- eapply BOUND0; eauto. eapply Mem.perm_alloc_4; eauto.
+ eapply BOUND0; eauto. eapply Mem.perm_alloc_4; eauto.
exploit me_bounded; eauto. unfold block in *; xomega.
- red; intros. exploit PERM; eauto. intros [A|A]. auto. right.
+ red; intros. exploit PERM; eauto. intros [A|A]. auto. right.
inv A. apply is_reachable_intro with id0 b0 sz0 delta; auto.
rewrite PTree.gso. auto. congruence.
- eapply match_callstack_invariant with (m1 := m1); eauto.
+ eapply match_callstack_invariant with (m1 := m1); eauto.
intros. eapply Mem.perm_alloc_4; eauto.
unfold block in *; xomega.
intros. apply H4. unfold block in *; xomega.
- intros. destruct (eq_block b0 b).
- subst b0. rewrite H3 in H. inv H. xomegaContradiction.
+ intros. destruct (eq_block b0 b).
+ subst b0. rewrite H3 in H. inv H. xomegaContradiction.
rewrite H4 in H; auto.
Qed.
@@ -761,8 +761,8 @@ Remark cenv_remove_gso:
Proof.
induction vars; simpl; intros.
auto.
- rewrite PTree.gro. apply IHvars. intuition. intuition.
-Qed.
+ rewrite PTree.gro. apply IHvars. intuition. intuition.
+Qed.
Remark cenv_remove_gss:
forall id vars cenv,
@@ -778,8 +778,8 @@ Qed.
Definition cenv_compat (cenv: compilenv) (vars: list (ident * Z)) (tsz: Z) : Prop :=
forall id sz,
In (id, sz) vars ->
- exists ofs,
- PTree.get id cenv = Some ofs
+ exists ofs,
+ PTree.get id cenv = Some ofs
/\ Mem.inj_offset_aligned ofs sz
/\ 0 <= ofs
/\ ofs + Zmax 0 sz <= tsz.
@@ -794,7 +794,7 @@ Definition cenv_separated (cenv: compilenv) (vars: list (ident * Z)) : Prop :=
Definition cenv_mem_separated (cenv: compilenv) (vars: list (ident * Z)) (f: meminj) (sp: block) (m: mem) : Prop :=
forall id sz ofs b delta ofs' k p,
In (id, sz) vars -> PTree.get id cenv = Some ofs ->
- f b = Some (sp, delta) ->
+ f b = Some (sp, delta) ->
Mem.perm m b ofs' k p ->
ofs <= ofs' + delta < sz + ofs -> False.
@@ -825,36 +825,36 @@ Proof.
intros until cs; intros VALID REPRES STKSIZE STKPERMS.
induction 1; intros f1 NOREPET COMPAT SEP1 SEP2 UNBOUND MCS MINJ.
(* base case *)
- simpl in MCS. exists f1; auto.
+ simpl in MCS. exists f1; auto.
(* inductive case *)
simpl in NOREPET. inv NOREPET.
(* exploit Mem.alloc_result; eauto. intros RES.
exploit Mem.nextblock_alloc; eauto. intros NB.*)
exploit (COMPAT id sz). auto with coqlib. intros [ofs [CENV [ALIGNED [LOB HIB]]]].
- exploit Mem.alloc_left_mapped_inject.
+ exploit Mem.alloc_left_mapped_inject.
eexact MINJ.
eexact H.
eexact VALID.
- instantiate (1 := ofs). zify. omega.
- intros. exploit STKSIZE; eauto. omega.
+ instantiate (1 := ofs). zify. omega.
+ intros. exploit STKSIZE; eauto. omega.
intros. apply STKPERMS. zify. omega.
replace (sz - 0) with sz by omega. auto.
- intros. eapply SEP2. eauto with coqlib. eexact CENV. eauto. eauto. omega.
+ intros. eapply SEP2. eauto with coqlib. eexact CENV. eauto. eauto. omega.
intros [f2 [A [B [C D]]]].
exploit (IHalloc_variables f2); eauto.
red; intros. eapply COMPAT. auto with coqlib.
red; intros. eapply SEP1; eauto with coqlib.
red; intros. exploit Mem.perm_alloc_inv; eauto. destruct (eq_block b b1); intros P.
- subst b. rewrite C in H5; inv H5.
- exploit SEP1. eapply in_eq. eapply in_cons; eauto. eauto. eauto.
- red; intros; subst id0. elim H3. change id with (fst (id, sz0)). apply in_map; auto.
+ subst b. rewrite C in H5; inv H5.
+ exploit SEP1. eapply in_eq. eapply in_cons; eauto. eauto. eauto.
+ red; intros; subst id0. elim H3. change id with (fst (id, sz0)). apply in_map; auto.
omega.
- eapply SEP2. apply in_cons; eauto. eauto.
- rewrite D in H5; eauto. eauto. auto.
- intros. rewrite PTree.gso. eapply UNBOUND; eauto with coqlib.
+ eapply SEP2. apply in_cons; eauto. eauto.
+ rewrite D in H5; eauto. eauto. auto.
+ intros. rewrite PTree.gso. eapply UNBOUND; eauto with coqlib.
red; intros; subst id0. elim H3. change id with (fst (id, sz0)). apply in_map; auto.
- eapply match_callstack_alloc_left; eauto.
- rewrite cenv_remove_gso; auto.
+ eapply match_callstack_alloc_left; eauto.
+ rewrite cenv_remove_gso; auto.
apply UNBOUND with sz; auto with coqlib.
Qed.
@@ -881,14 +881,14 @@ Proof.
intros. eapply Mem.perm_alloc_3; eauto.
intros. apply Mem.perm_implies with Freeable; auto with mem. eapply Mem.perm_alloc_2; eauto.
instantiate (1 := f1). red; intros. eelim Mem.fresh_block_alloc; eauto.
- eapply Mem.valid_block_inject_2; eauto.
+ eapply Mem.valid_block_inject_2; eauto.
intros. apply PTree.gempty.
- eapply match_callstack_alloc_right; eauto.
+ eapply match_callstack_alloc_right; eauto.
intros. destruct (In_dec peq id (map fst vars)).
apply cenv_remove_gss; auto.
rewrite cenv_remove_gso; auto.
- destruct (cenv!id) as [ofs|] eqn:?; auto. elim n; eauto.
- eapply Mem.alloc_right_inject; eauto.
+ destruct (cenv!id) as [ofs|] eqn:?; auto. elim n; eauto.
+ eapply Mem.alloc_right_inject; eauto.
Qed.
(** Properties of the compilation environment produced by [build_compilenv] *)
@@ -896,8 +896,8 @@ Qed.
Remark block_alignment_pos:
forall sz, block_alignment sz > 0.
Proof.
- unfold block_alignment; intros.
- destruct (zlt sz 2). omega.
+ unfold block_alignment; intros.
+ destruct (zlt sz 2). omega.
destruct (zlt sz 4). omega.
destruct (zlt sz 8); omega.
Qed.
@@ -906,7 +906,7 @@ Remark assign_variable_incr:
forall id sz cenv stksz cenv' stksz',
assign_variable (cenv, stksz) (id, sz) = (cenv', stksz') -> stksz <= stksz'.
Proof.
- simpl; intros. inv H.
+ simpl; intros. inv H.
generalize (align_le stksz (block_alignment sz) (block_alignment_pos sz)).
assert (0 <= Zmax 0 sz). apply Zmax_bound_l. omega.
omega.
@@ -920,7 +920,7 @@ Proof.
simpl; intros. inv H. omega.
Opaque assign_variable.
destruct a as [id s]. simpl. intros.
- destruct (assign_variable (cenv, sz) (id, s)) as [cenv1 sz1] eqn:?.
+ destruct (assign_variable (cenv, sz) (id, s)) as [cenv1 sz1] eqn:?.
apply Zle_trans with sz1. eapply assign_variable_incr; eauto. eauto.
Transparent assign_variable.
Qed.
@@ -951,9 +951,9 @@ Remark inj_offset_aligned_block':
forall stacksize sz,
Mem.inj_offset_aligned (align stacksize (block_alignment sz)) (Zmax 0 sz).
Proof.
- intros.
+ intros.
replace (block_alignment sz) with (block_alignment (Zmax 0 sz)).
- apply inj_offset_aligned_block.
+ apply inj_offset_aligned_block.
rewrite Zmax_spec. destruct (zlt sz 0); auto.
transitivity 1. reflexivity. unfold block_alignment. rewrite zlt_true. auto. omega.
Qed.
@@ -974,31 +974,31 @@ Proof.
assert (EITHER: forall id' sz',
In (id', sz') (vars ++ (id, sz) :: nil) ->
In (id', sz') vars /\ id' <> id \/ (id', sz') = (id, sz)).
- intros. rewrite in_app in H. destruct H.
- left; split; auto. red; intros; subst id'. elim NOREPET.
+ intros. rewrite in_app in H. destruct H.
+ left; split; auto. red; intros; subst id'. elim NOREPET.
change id with (fst (id, sz')). apply in_map; auto.
simpl in H. destruct H. auto. contradiction.
split; red; intros.
apply EITHER in H. destruct H as [[P Q] | P].
- exploit COMPAT; eauto. intros [ofs [A [B [C D]]]].
+ exploit COMPAT; eauto. intros [ofs [A [B [C D]]]].
exists ofs.
split. rewrite PTree.gso; auto.
- split. auto. split. auto. zify; omega.
+ split. auto. split. auto. zify; omega.
inv P. exists (align sz1 (block_alignment sz)).
split. apply PTree.gss.
split. apply inj_offset_aligned_block.
- split. omega.
+ split. omega.
omega.
apply EITHER in H; apply EITHER in H0.
destruct H as [[P Q] | P]; destruct H0 as [[R S] | R].
- rewrite PTree.gso in *; auto. eapply SEP; eauto.
+ rewrite PTree.gso in *; auto. eapply SEP; eauto.
inv R. rewrite PTree.gso in H1; auto. rewrite PTree.gss in H2; inv H2.
- exploit COMPAT; eauto. intros [ofs [A [B [C D]]]].
- assert (ofs = ofs1) by congruence. subst ofs.
- left. zify; omega.
+ exploit COMPAT; eauto. intros [ofs [A [B [C D]]]].
+ assert (ofs = ofs1) by congruence. subst ofs.
+ left. zify; omega.
inv P. rewrite PTree.gso in H2; auto. rewrite PTree.gss in H1; inv H1.
- exploit COMPAT; eauto. intros [ofs [A [B [C D]]]].
- assert (ofs = ofs2) by congruence. subst ofs.
+ exploit COMPAT; eauto. intros [ofs [A [B [C D]]]].
+ assert (ofs = ofs2) by congruence. subst ofs.
right. zify; omega.
congruence.
Qed.
@@ -1026,13 +1026,13 @@ Proof.
exploit IHvars'.
eauto.
instantiate (1 := vars ++ ((id, sz) :: nil)).
- rewrite list_norepet_app. split. auto.
- split. rewrite map_app. apply list_norepet_append_commut. simpl. constructor; auto.
- rewrite map_app. simpl. red; intros. rewrite in_app in H4. destruct H4.
+ rewrite list_norepet_app. split. auto.
+ split. rewrite map_app. apply list_norepet_append_commut. simpl. constructor; auto.
+ rewrite map_app. simpl. red; intros. rewrite in_app in H4. destruct H4.
eauto. simpl in H4. destruct H4. subst y. red; intros; subst x. tauto. tauto.
generalize (assign_variable_incr _ _ _ _ _ _ Heqp). omega.
auto. auto.
- rewrite app_ass. auto.
+ rewrite app_ass. auto.
Qed.
Remark permutation_norepet:
@@ -1051,7 +1051,7 @@ Lemma build_compilenv_sound:
list_norepet (map fst (Csharpminor.fn_vars f)) ->
cenv_compat cenv (Csharpminor.fn_vars f) sz /\ cenv_separated cenv (Csharpminor.fn_vars f).
Proof.
- unfold build_compilenv; intros.
+ unfold build_compilenv; intros.
set (vars1 := Csharpminor.fn_vars f) in *.
generalize (VarSort.Permuted_sort vars1). intros P.
set (vars2 := VarSort.sort vars1) in *.
@@ -1061,11 +1061,11 @@ Proof.
eexact H.
simpl. rewrite app_nil_r. apply permutation_norepet with (map fst vars1); auto.
apply Permutation_map. auto.
- omega.
+ omega.
red; intros. contradiction.
red; intros. contradiction.
destruct H1 as [A B]. split.
- red; intros. apply A. apply Permutation_in with vars1; auto.
+ red; intros. apply A. apply Permutation_in with vars1; auto.
red; intros. eapply B; eauto; apply Permutation_in with vars1; auto.
Qed.
@@ -1077,7 +1077,7 @@ Proof.
induction vars; simpl; intros.
auto.
exploit IHvars; eauto. unfold assign_variable. destruct a as [id1 sz1].
- destruct cesz as [cenv stksz]. simpl.
+ destruct cesz as [cenv stksz]. simpl.
rewrite PTree.gsspec. destruct (peq id id1). auto. tauto.
Qed.
@@ -1086,12 +1086,12 @@ Lemma build_compilenv_domain:
build_compilenv f = (cenv, sz) ->
cenv!id = Some ofs -> In id (map fst (Csharpminor.fn_vars f)).
Proof.
- unfold build_compilenv; intros.
+ unfold build_compilenv; intros.
set (vars1 := Csharpminor.fn_vars f) in *.
generalize (VarSort.Permuted_sort vars1). intros P.
set (vars2 := VarSort.sort vars1) in *.
generalize (assign_variables_domain id vars2 (PTree.empty Z, 0)).
- rewrite H. simpl. intros. destruct H1. congruence.
+ rewrite H. simpl. intros. destruct H1. congruence.
rewrite PTree.gempty in H1. congruence.
apply Permutation_in with (map fst vars2); auto.
apply Permutation_map. apply Permutation_sym; auto.
@@ -1106,7 +1106,7 @@ Proof.
rewrite PTree.gempty in H. congruence.
rewrite PTree.gsspec in H. destruct (peq id a).
split. auto. congruence.
- exploit IHtemps; eauto. tauto.
+ exploit IHtemps; eauto. tauto.
Qed.
Fixpoint set_params' (vl: list val) (il: list ident) (te: Cminor.env) : Cminor.env :=
@@ -1125,10 +1125,10 @@ Lemma bind_parameters_agree_rec:
Proof.
Opaque PTree.set.
induction vars; simpl; intros.
- destruct vals; try discriminate. inv H. auto.
+ destruct vals; try discriminate. inv H. auto.
destruct vals; try discriminate. inv H0.
simpl. eapply IHvars; eauto.
- red; intros. rewrite PTree.gsspec in *. destruct (peq id a).
+ red; intros. rewrite PTree.gsspec in *. destruct (peq id a).
inv H0. exists v'; auto.
apply H1; auto.
Qed.
@@ -1136,7 +1136,7 @@ Qed.
Lemma set_params'_outside:
forall id il vl te, ~In id il -> (set_params' vl il te)!id = te!id.
Proof.
- induction il; simpl; intros. auto.
+ induction il; simpl; intros. auto.
destruct vl; rewrite IHil.
apply PTree.gso. intuition. intuition.
apply PTree.gso. intuition. intuition.
@@ -1161,17 +1161,17 @@ Lemma set_params_set_params':
Proof.
induction il; simpl; intros.
auto.
- inv H. destruct vl.
- rewrite PTree.gsspec. destruct (peq id a).
+ inv H. destruct vl.
+ rewrite PTree.gsspec. destruct (peq id a).
subst a. rewrite set_params'_outside; auto. rewrite PTree.gss; auto.
rewrite IHil; auto.
destruct (List.in_dec peq id il). apply set_params'_inside; auto.
- repeat rewrite set_params'_outside; auto. rewrite PTree.gso; auto.
- rewrite PTree.gsspec. destruct (peq id a).
+ repeat rewrite set_params'_outside; auto. rewrite PTree.gso; auto.
+ rewrite PTree.gsspec. destruct (peq id a).
subst a. rewrite set_params'_outside; auto. rewrite PTree.gss; auto.
rewrite IHil; auto.
destruct (List.in_dec peq id il). apply set_params'_inside; auto.
- repeat rewrite set_params'_outside; auto. rewrite PTree.gso; auto.
+ repeat rewrite set_params'_outside; auto. rewrite PTree.gso; auto.
Qed.
Lemma set_locals_outside:
@@ -1180,7 +1180,7 @@ Lemma set_locals_outside:
Proof.
induction il; simpl; intros.
auto.
- rewrite PTree.gso. apply IHil. tauto. intuition.
+ rewrite PTree.gso. apply IHil. tauto. intuition.
Qed.
Lemma set_locals_inside:
@@ -1189,8 +1189,8 @@ Lemma set_locals_inside:
Proof.
induction il; simpl; intros.
contradiction.
- destruct H. subst a. apply PTree.gss.
- rewrite PTree.gsspec. destruct (peq id a). auto. auto.
+ destruct H. subst a. apply PTree.gss.
+ rewrite PTree.gsspec. destruct (peq id a). auto. auto.
Qed.
Lemma set_locals_set_params':
@@ -1203,11 +1203,11 @@ Proof.
intros. destruct (in_dec peq id vars).
assert (~In id params). apply list_disjoint_notin with vars; auto. apply list_disjoint_sym; auto.
rewrite set_locals_inside; auto. rewrite set_params'_outside; auto. rewrite set_locals_inside; auto.
- rewrite set_locals_outside; auto. rewrite set_params_set_params'; auto.
- destruct (in_dec peq id params).
+ rewrite set_locals_outside; auto. rewrite set_params_set_params'; auto.
+ destruct (in_dec peq id params).
apply set_params'_inside; auto.
- repeat rewrite set_params'_outside; auto.
- rewrite set_locals_outside; auto.
+ repeat rewrite set_params'_outside; auto.
+ rewrite set_locals_outside; auto.
Qed.
Lemma bind_parameters_agree:
@@ -1253,7 +1253,7 @@ Proof.
exploit build_compilenv_sound; eauto. intros [C1 C2].
eapply match_callstack_alloc_variables; eauto.
intros. eapply build_compilenv_domain; eauto.
- eapply bind_parameters_agree; eauto.
+ eapply bind_parameters_agree; eauto.
Qed.
(** * Compatibility of evaluation functions with respect to memory injections. *)
@@ -1348,18 +1348,18 @@ Proof.
inv H; inv H0; inv H1; TrivialExists.
apply Int.sub_add_l.
simpl. destruct (eq_block b1 b0); auto.
- subst b1. rewrite H in H0; inv H0.
+ subst b1. rewrite H in H0; inv H0.
rewrite dec_eq_true. rewrite Int.sub_shifted. auto.
inv H; inv H0; inv H1; TrivialExists.
- inv H0; try discriminate; inv H1; try discriminate. simpl in *.
+ inv H0; try discriminate; inv H1; try discriminate. simpl in *.
destruct (Int.eq i0 Int.zero
|| Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H; TrivialExists.
- inv H0; try discriminate; inv H1; try discriminate. simpl in *.
+ inv H0; try discriminate; inv H1; try discriminate. simpl in *.
destruct (Int.eq i0 Int.zero); inv H. TrivialExists.
- inv H0; try discriminate; inv H1; try discriminate. simpl in *.
+ inv H0; try discriminate; inv H1; try discriminate. simpl in *.
destruct (Int.eq i0 Int.zero
|| Int.eq i (Int.repr Int.min_signed) && Int.eq i0 Int.mone); inv H; TrivialExists.
- inv H0; try discriminate; inv H1; try discriminate. simpl in *.
+ inv H0; try discriminate; inv H1; try discriminate. simpl in *.
destruct (Int.eq i0 Int.zero); inv H. TrivialExists.
inv H; inv H0; inv H1; TrivialExists.
inv H; inv H0; inv H1; TrivialExists.
@@ -1378,15 +1378,15 @@ Proof.
inv H; inv H0; inv H1; TrivialExists.
inv H; inv H0; inv H1; TrivialExists.
inv H; inv H0; inv H1; TrivialExists.
- inv H0; try discriminate; inv H1; try discriminate. simpl in *.
+ inv H0; try discriminate; inv H1; try discriminate. simpl in *.
destruct (Int64.eq i0 Int64.zero
|| Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H; TrivialExists.
- inv H0; try discriminate; inv H1; try discriminate. simpl in *.
+ inv H0; try discriminate; inv H1; try discriminate. simpl in *.
destruct (Int64.eq i0 Int64.zero); inv H. TrivialExists.
- inv H0; try discriminate; inv H1; try discriminate. simpl in *.
+ inv H0; try discriminate; inv H1; try discriminate. simpl in *.
destruct (Int64.eq i0 Int64.zero
|| Int64.eq i (Int64.repr Int64.min_signed) && Int64.eq i0 Int64.mone); inv H; TrivialExists.
- inv H0; try discriminate; inv H1; try discriminate. simpl in *.
+ inv H0; try discriminate; inv H1; try discriminate. simpl in *.
destruct (Int64.eq i0 Int64.zero); inv H. TrivialExists.
inv H; inv H0; inv H1; TrivialExists.
inv H; inv H0; inv H1; TrivialExists.
@@ -1396,8 +1396,8 @@ Proof.
inv H; inv H0; inv H1; TrivialExists. simpl. destruct (Int.ltu i0 Int64.iwordsize'); auto.
inv H; inv H0; inv H1; TrivialExists. apply val_inject_val_of_optbool.
(* cmpu *)
- inv H. econstructor; split; eauto.
- unfold Val.cmpu.
+ inv H. econstructor; split; eauto.
+ unfold Val.cmpu.
destruct (Val.cmpu_bool (Mem.valid_pointer m) c v1 v2) as [b|] eqn:E.
replace (Val.cmpu_bool (Mem.valid_pointer tm) c tv1 tv2) with (Some b).
destruct b; simpl; constructor.
@@ -1437,7 +1437,7 @@ Proof.
inv H1; inv H0; try congruence.
(* local *)
exists (Vptr sp (Int.repr ofs)); split.
- constructor. simpl. rewrite Int.add_zero_l; auto.
+ constructor. simpl. rewrite Int.add_zero_l; auto.
congruence.
(* global *)
exploit match_callstack_match_globalenvs; eauto. intros [bnd MG]. inv MG.
@@ -1486,7 +1486,7 @@ Lemma transl_constant_correct:
eval_constant tge sp (transl_constant cst) = Some tv
/\ Val.inject f v tv.
Proof.
- destruct cst; simpl; intros; inv H.
+ destruct cst; simpl; intros; inv H.
exists (Vint i); auto.
exists (Vfloat f0); auto.
exists (Vsingle f0); auto.
@@ -1509,17 +1509,17 @@ Lemma transl_expr_correct:
Proof.
induction 3; intros; simpl in TR; try (monadInv TR).
(* Etempvar *)
- inv MATCH. exploit MTMP; eauto. intros [tv [A B]].
+ inv MATCH. exploit MTMP; eauto. intros [tv [A B]].
exists tv; split. constructor; auto. auto.
(* Eaddrof *)
eapply var_addr_correct; eauto.
(* Econst *)
exploit transl_constant_correct; eauto. intros [tv [A B]].
- exists tv; split; eauto. constructor; eauto.
+ exists tv; split; eauto. constructor; eauto.
(* Eunop *)
exploit IHeval_expr; eauto. intros [tv1 [EVAL1 INJ1]].
exploit eval_unop_compat; eauto. intros [tv [EVAL INJ]].
- exists tv; split; auto. econstructor; eauto.
+ exists tv; split; auto. econstructor; eauto.
(* Ebinop *)
exploit IHeval_expr1; eauto. intros [tv1 [EVAL1 INJ1]].
exploit IHeval_expr2; eauto. intros [tv2 [EVAL2 INJ2]].
@@ -1672,17 +1672,17 @@ Inductive lbl_stmt_tail: lbl_stmt -> nat -> lbl_stmt -> Prop :=
Lemma switch_table_default:
forall sl base,
- exists n,
+ exists n,
lbl_stmt_tail sl n (select_switch_default sl)
/\ snd (switch_table sl base) = (n + base)%nat.
Proof.
induction sl; simpl; intros.
- exists O; split. constructor. omega.
-- destruct o.
- + destruct (IHsl (S base)) as (n & P & Q). exists (S n); split.
- constructor; auto.
+- destruct o.
+ + destruct (IHsl (S base)) as (n & P & Q). exists (S n); split.
+ constructor; auto.
destruct (switch_table sl (S base)) as [tbl dfl]; simpl in *. omega.
- + exists O; split. constructor.
+ + exists O; split. constructor.
destruct (switch_table sl (S base)) as [tbl dfl]; simpl in *. auto.
Qed.
@@ -1699,17 +1699,17 @@ Lemma switch_table_case:
Proof.
induction sl; simpl; intros.
- auto.
-- destruct (switch_table sl (S base)) as [tbl1 dfl1] eqn:ST.
+- destruct (switch_table sl (S base)) as [tbl1 dfl1] eqn:ST.
destruct o; simpl.
rewrite dec_eq_sym. destruct (zeq i z).
exists O; split; auto. constructor.
specialize (IHsl (S base) dfl). rewrite ST in IHsl. simpl in *.
destruct (select_switch_case i sl).
- destruct IHsl as (x & P & Q). exists (S x); split. constructor; auto. omega.
+ destruct IHsl as (x & P & Q). exists (S x); split. constructor; auto. omega.
auto.
specialize (IHsl (S base) dfl). rewrite ST in IHsl. simpl in *.
destruct (select_switch_case i sl).
- destruct IHsl as (x & P & Q). exists (S x); split. constructor; auto. omega.
+ destruct IHsl as (x & P & Q). exists (S x); split. constructor; auto. omega.
auto.
Qed.
@@ -1720,7 +1720,7 @@ Lemma switch_table_select:
(select_switch i sl).
Proof.
unfold select_switch; intros.
- generalize (switch_table_case i sl O (snd (switch_table sl O))).
+ generalize (switch_table_case i sl O (snd (switch_table sl O))).
destruct (select_switch_case i sl) as [sl'|].
intros (n & P & Q). replace (n + O)%nat with n in Q by omega. congruence.
intros E; rewrite E.
@@ -1744,15 +1744,15 @@ Lemma switch_descent:
/\ (forall f sp e m,
plus step tge (State f s k sp e m) E0 (State f body k' sp e m)).
Proof.
- induction ls; intros.
+ induction ls; intros.
- monadInv H. econstructor; split.
econstructor.
- intros. eapply plus_two. constructor. constructor. auto.
-- monadInv H. exploit IHls; eauto. intros [k' [A B]].
+ intros. eapply plus_two. constructor. constructor. auto.
+- monadInv H. exploit IHls; eauto. intros [k' [A B]].
econstructor; split.
econstructor; eauto.
- intros. eapply plus_star_trans. eauto.
- eapply star_left. constructor. apply star_one. constructor.
+ intros. eapply plus_star_trans. eauto.
+ eapply star_left. constructor. apply star_one. constructor.
reflexivity. traceEq.
Qed.
@@ -1766,12 +1766,12 @@ Lemma switch_ascent:
E0 (State f (Sexit O) k2 sp e m)
/\ transl_lblstmt_cont cenv xenv ls' k k2.
Proof.
- induction 1; intros.
+ induction 1; intros.
- exists k1; split; auto. apply star_refl.
-- inv H0. exploit IHlbl_stmt_tail; eauto. intros (k2 & P & Q).
+- inv H0. exploit IHlbl_stmt_tail; eauto. intros (k2 & P & Q).
exists k2; split; auto.
eapply star_left. constructor. eapply star_left. constructor. eexact P.
- eauto. auto.
+ eauto. auto.
Qed.
Lemma switch_match_cont:
@@ -1782,7 +1782,7 @@ Lemma switch_match_cont:
Proof.
induction ls; intros; simpl.
inv H0. apply match_Kblock2. econstructor; eauto.
- inv H0. apply match_Kblock2. eapply match_Kseq2. auto. eauto.
+ inv H0. apply match_Kblock2. eapply match_Kseq2. auto. eauto.
Qed.
Lemma switch_match_states:
@@ -1799,9 +1799,9 @@ Lemma switch_match_states:
plus step tge (State tfn (Sexit O) tk' (Vptr sp Int.zero) te tm) E0 S
/\ match_states (Csharpminor.State fn (seq_of_lbl_stmt ls) k e le m) S.
Proof.
- intros. inv TK.
-- econstructor; split. eapply plus_two. constructor. constructor. auto.
- eapply match_state; eauto.
+ intros. inv TK.
+- econstructor; split. eapply plus_two. constructor. constructor. auto.
+ eapply match_state; eauto.
- econstructor; split. eapply plus_left. constructor. apply star_one. constructor. auto.
simpl. eapply match_state_seq; eauto. simpl. eapply switch_match_cont; eauto.
Qed.
@@ -1812,9 +1812,9 @@ Lemma transl_lblstmt_suffix:
forall body ts, transl_lblstmt cenv (switch_env ls xenv) ls body = OK ts ->
exists body', exists ts', transl_lblstmt cenv (switch_env ls' xenv) ls' body' = OK ts'.
Proof.
- induction 1; intros.
+ induction 1; intros.
- exists body, ts; auto.
-- monadInv H0. eauto.
+- monadInv H0. eauto.
Qed.
(** Commutation between [find_label] and compilation *)
@@ -1834,7 +1834,7 @@ Lemma transl_lblstmt_find_label_context:
Proof.
induction ls; intros.
- monadInv H. inv H0. simpl. rewrite H1. auto.
-- monadInv H. inv H0. simpl in H6. eapply IHls; eauto.
+- monadInv H. inv H0. simpl in H6. eapply IHls; eauto.
replace x with ts0 by congruence. simpl. rewrite H1. auto.
Qed.
@@ -1868,25 +1868,25 @@ with transl_lblstmt_find_label:
Proof.
intros. destruct s; try (monadInv H); simpl; auto.
(* seq *)
- exploit (transl_find_label s1). eauto. eapply match_Kseq. eexact EQ1. eauto.
+ exploit (transl_find_label s1). eauto. eapply match_Kseq. eexact EQ1. eauto.
destruct (Csharpminor.find_label lbl s1 (Csharpminor.Kseq s2 k)) as [[s' k'] | ].
- intros [ts' [tk' [xenv' [A [B C]]]]].
+ intros [ts' [tk' [xenv' [A [B C]]]]].
exists ts'; exists tk'; exists xenv'. intuition. rewrite A; auto.
- intro. rewrite H. apply transl_find_label with xenv; auto.
+ intro. rewrite H. apply transl_find_label with xenv; auto.
(* ifthenelse *)
- exploit (transl_find_label s1). eauto. eauto.
+ exploit (transl_find_label s1). eauto. eauto.
destruct (Csharpminor.find_label lbl s1 k) as [[s' k'] | ].
- intros [ts' [tk' [xenv' [A [B C]]]]].
+ intros [ts' [tk' [xenv' [A [B C]]]]].
exists ts'; exists tk'; exists xenv'. intuition. rewrite A; auto.
- intro. rewrite H. apply transl_find_label with xenv; auto.
+ intro. rewrite H. apply transl_find_label with xenv; auto.
(* loop *)
apply transl_find_label with xenv. auto. econstructor; eauto. simpl. rewrite EQ; auto.
(* block *)
- apply transl_find_label with (true :: xenv). auto. constructor; auto.
+ apply transl_find_label with (true :: xenv). auto. constructor; auto.
(* switch *)
- simpl in H. destruct (switch_table l O) as [tbl dfl]. monadInv H.
+ simpl in H. destruct (switch_table l O) as [tbl dfl]. monadInv H.
exploit switch_descent; eauto. intros [k' [A B]].
- eapply transl_lblstmt_find_label. eauto. eauto. eauto. reflexivity.
+ eapply transl_lblstmt_find_label. eauto. eauto. eauto. reflexivity.
(* return *)
destruct o; monadInv H; auto.
(* label *)
@@ -1899,14 +1899,14 @@ Proof.
inv H1. rewrite H2. auto.
(* cons *)
inv H1. simpl in H7.
- exploit (transl_find_label s). eauto. eapply switch_match_cont; eauto.
+ exploit (transl_find_label s). eauto. eapply switch_match_cont; eauto.
destruct (Csharpminor.find_label lbl s (Csharpminor.Kseq (seq_of_lbl_stmt ls) k)) as [[s' k''] | ].
intros [ts' [tk' [xenv' [A [B C]]]]].
- exists ts'; exists tk'; exists xenv'; intuition.
- eapply transl_lblstmt_find_label_context; eauto.
+ exists ts'; exists tk'; exists xenv'; intuition.
+ eapply transl_lblstmt_find_label_context; eauto.
+ simpl. replace x with ts0 by congruence. rewrite H2. auto.
+ intro. eapply transl_lblstmt_find_label. eauto. auto. eauto.
simpl. replace x with ts0 by congruence. rewrite H2. auto.
- intro. eapply transl_lblstmt_find_label. eauto. auto. eauto.
- simpl. replace x with ts0 by congruence. rewrite H2. auto.
Qed.
End FIND_LABEL.
@@ -1921,7 +1921,7 @@ Lemma transl_find_label_body:
/\ transl_stmt cenv xenv' s' = OK ts'
/\ match_cont k' tk' cenv xenv' cs.
Proof.
- intros. monadInv H. simpl.
+ intros. monadInv H. simpl.
exploit transl_find_label. eexact EQ. eapply match_call_cont. eexact H0.
instantiate (1 := lbl). rewrite H1. auto.
Qed.
@@ -1951,12 +1951,12 @@ Proof.
(* skip seq *)
monadInv TR. left.
dependent induction MK.
- econstructor; split.
+ econstructor; split.
apply plus_one. constructor.
econstructor; eauto.
econstructor; split.
apply plus_one. constructor.
- eapply match_state_seq; eauto.
+ eapply match_state_seq; eauto.
exploit IHMK; eauto. intros [T2 [A B]].
exists T2; split. eapply plus_left. constructor. apply plus_star; eauto. traceEq.
auto.
@@ -1971,7 +1971,7 @@ Proof.
auto.
(* skip call *)
monadInv TR. left.
- exploit match_is_call_cont; eauto. intros [tk' [A [B C]]].
+ exploit match_is_call_cont; eauto. intros [tk' [A [B C]]].
exploit match_callstack_freelist; eauto. intros [tm' [P [Q R]]].
econstructor; split.
eapply plus_right. eexact A. apply step_skip_call. auto. eauto. traceEq.
@@ -1981,25 +1981,25 @@ Proof.
monadInv TR.
exploit transl_expr_correct; eauto. intros [tv [EVAL VINJ]].
left; econstructor; split.
- apply plus_one. econstructor; eauto.
- econstructor; eauto.
- eapply match_callstack_set_temp; eauto.
+ apply plus_one. econstructor; eauto.
+ econstructor; eauto.
+ eapply match_callstack_set_temp; eauto.
(* store *)
monadInv TR.
- exploit transl_expr_correct. eauto. eauto. eexact H. eauto.
+ exploit transl_expr_correct. eauto. eauto. eexact H. eauto.
intros [tv1 [EVAL1 VINJ1]].
- exploit transl_expr_correct. eauto. eauto. eexact H0. eauto.
+ exploit transl_expr_correct. eauto. eauto. eexact H0. eauto.
intros [tv2 [EVAL2 VINJ2]].
- exploit Mem.storev_mapped_inject; eauto. intros [tm' [STORE' MINJ']].
+ exploit Mem.storev_mapped_inject; eauto. intros [tm' [STORE' MINJ']].
left; econstructor; split.
- apply plus_one. econstructor; eauto.
+ apply plus_one. econstructor; eauto.
econstructor; eauto.
inv VINJ1; simpl in H1; try discriminate. unfold Mem.storev in STORE'.
rewrite (Mem.nextblock_store _ _ _ _ _ _ H1).
rewrite (Mem.nextblock_store _ _ _ _ _ _ STORE').
eapply match_callstack_invariant with f0 m tm; eauto.
- intros. eapply Mem.perm_store_2; eauto.
+ intros. eapply Mem.perm_store_2; eauto.
intros. eapply Mem.perm_store_1; eauto.
(* call *)
@@ -2024,11 +2024,11 @@ Proof.
exploit transl_exprlist_correct; eauto.
intros [tvargs [EVAL2 VINJ2]].
exploit match_callstack_match_globalenvs; eauto. intros [hi' MG].
- exploit external_call_mem_inject; eauto.
+ exploit external_call_mem_inject; eauto.
eapply inj_preserves_globals; eauto.
intros [f' [vres' [tm' [EC [VINJ [MINJ' [UNMAPPED [OUTOFREACH [INCR SEPARATED]]]]]]]]].
left; econstructor; split.
- apply plus_one. econstructor. eauto.
+ apply plus_one. econstructor. eauto.
eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. eexact varinfo_preserved.
assert (MCS': match_callstack f' m' tm'
@@ -2037,23 +2037,23 @@ Proof.
apply match_callstack_incr_bound with (Mem.nextblock m) (Mem.nextblock tm).
eapply match_callstack_external_call; eauto.
intros. eapply external_call_max_perm; eauto.
- xomega. xomega.
+ xomega. xomega.
eapply external_call_nextblock; eauto.
eapply external_call_nextblock; eauto.
econstructor; eauto.
Opaque PTree.set.
- unfold set_optvar. destruct optid; simpl.
- eapply match_callstack_set_temp; eauto.
+ unfold set_optvar. destruct optid; simpl.
+ eapply match_callstack_set_temp; eauto.
auto.
(* seq *)
- monadInv TR.
+ monadInv TR.
left; econstructor; split.
- apply plus_one. constructor.
+ apply plus_one. constructor.
econstructor; eauto.
econstructor; eauto.
(* seq 2 *)
- right. split. auto. split. auto. econstructor; eauto.
+ right. split. auto. split. auto. econstructor; eauto.
(* ifthenelse *)
monadInv TR.
@@ -2065,21 +2065,21 @@ Opaque PTree.set.
(* loop *)
monadInv TR.
left; econstructor; split.
- apply plus_one. constructor.
+ apply plus_one. constructor.
econstructor; eauto.
- econstructor; eauto. simpl. rewrite EQ; auto.
+ econstructor; eauto. simpl. rewrite EQ; auto.
(* block *)
monadInv TR.
left; econstructor; split.
- apply plus_one. constructor.
+ apply plus_one. constructor.
econstructor; eauto.
econstructor; eauto.
(* exit seq *)
monadInv TR. left.
dependent induction MK.
- econstructor; split.
+ econstructor; split.
apply plus_one. constructor.
econstructor; eauto. simpl. auto.
exploit IHMK; eauto. intros [T2 [A B]].
@@ -2092,20 +2092,20 @@ Opaque PTree.set.
monadInv TR. left.
dependent induction MK.
econstructor; split.
- simpl. apply plus_one. constructor.
+ simpl. apply plus_one. constructor.
econstructor; eauto.
exploit IHMK; eauto. intros [T2 [A B]].
- exists T2; split; auto. simpl.
+ exists T2; split; auto. simpl.
eapply plus_left. constructor. apply plus_star; eauto. traceEq.
(* exit block n+1 *)
monadInv TR. left.
dependent induction MK.
econstructor; split.
- simpl. apply plus_one. constructor.
- econstructor; eauto. auto.
+ simpl. apply plus_one. constructor.
+ econstructor; eauto. auto.
exploit IHMK; eauto. intros [T2 [A B]].
- exists T2; split; auto. simpl.
+ exists T2; split; auto. simpl.
eapply plus_left. constructor. apply plus_star; eauto. traceEq.
(* switch *)
@@ -2120,9 +2120,9 @@ Opaque PTree.set.
simpl. intros [body' [ts' E]].
exploit switch_match_states; eauto. intros [T2 [F G]].
left; exists T2; split.
- eapply plus_star_trans. eapply B.
- eapply star_left. econstructor; eauto.
- eapply star_trans. eexact C.
+ eapply plus_star_trans. eapply B.
+ eapply star_left. econstructor; eauto.
+ eapply star_trans. eexact C.
apply plus_star. eexact F.
reflexivity. reflexivity. traceEq.
auto.
@@ -2136,11 +2136,11 @@ Opaque PTree.set.
simpl; auto.
(* return some *)
- monadInv TR. left.
+ monadInv TR. left.
exploit transl_expr_correct; eauto. intros [tv [EVAL VINJ]].
exploit match_callstack_freelist; eauto. intros [tm' [A [B C]]].
econstructor; split.
- apply plus_one. eapply step_return_1. eauto. eauto.
+ apply plus_one. eapply step_return_1. eauto. eauto.
econstructor; eauto. eapply match_call_cont; eauto.
(* label *)
@@ -2151,10 +2151,10 @@ Opaque PTree.set.
(* goto *)
monadInv TR.
- exploit transl_find_label_body; eauto.
+ exploit transl_find_label_body; eauto.
intros [ts' [tk' [xenv' [A [B C]]]]].
left; econstructor; split.
- apply plus_one. apply step_goto. eexact A.
+ apply plus_one. apply step_goto. eexact A.
econstructor; eauto.
(* internal call *)
@@ -2163,7 +2163,7 @@ Opaque PTree.set.
destruct (zle sz Int.max_unsigned); try congruence.
intro TRBODY.
generalize TRBODY; intro TMP. monadInv TMP.
- set (tf := mkfunction (Csharpminor.fn_sig f)
+ set (tf := mkfunction (Csharpminor.fn_sig f)
(Csharpminor.fn_params f)
(Csharpminor.fn_temps f)
sz
@@ -2172,14 +2172,14 @@ Opaque PTree.set.
exploit match_callstack_function_entry; eauto. simpl; eauto. simpl; auto.
intros [f2 [MCS2 MINJ2]].
left; econstructor; split.
- apply plus_one. constructor; simpl; eauto.
+ apply plus_one. constructor; simpl; eauto.
econstructor. eexact TRBODY. eauto. eexact MINJ2. eexact MCS2.
inv MK; simpl in ISCC; contradiction || econstructor; eauto.
(* external call *)
- monadInv TR.
+ monadInv TR.
exploit match_callstack_match_globalenvs; eauto. intros [hi MG].
- exploit external_call_mem_inject; eauto.
+ exploit external_call_mem_inject; eauto.
eapply inj_preserves_globals; eauto.
intros [f' [vres' [tm' [EC [VINJ [MINJ' [UNMAPPED [OUTOFREACH [INCR SEPARATED]]]]]]]]].
left; econstructor; split.
@@ -2197,9 +2197,9 @@ Opaque PTree.set.
(* return *)
inv MK. simpl.
left; econstructor; split.
- apply plus_one. econstructor; eauto.
+ apply plus_one. econstructor; eauto.
unfold set_optvar. destruct optid; simpl; econstructor; eauto.
- eapply match_callstack_set_temp; eauto.
+ eapply match_callstack_set_temp; eauto.
Qed.
Lemma match_globalenvs_init:
@@ -2208,12 +2208,12 @@ Lemma match_globalenvs_init:
match_globalenvs (Mem.flat_inj (Mem.nextblock m)) (Mem.nextblock m).
Proof.
intros. constructor.
- intros. unfold Mem.flat_inj. apply pred_dec_true; auto.
- intros. unfold Mem.flat_inj in H0.
+ intros. unfold Mem.flat_inj. apply pred_dec_true; auto.
+ intros. unfold Mem.flat_inj in H0.
destruct (plt b1 (Mem.nextblock m)); congruence.
intros. eapply Genv.find_symbol_not_fresh; eauto.
intros. eapply Genv.find_funct_ptr_not_fresh; eauto.
- intros. eapply Genv.find_var_info_not_fresh; eauto.
+ intros. eapply Genv.find_var_info_not_fresh; eauto.
Qed.
Lemma transl_initial_states:
@@ -2224,19 +2224,19 @@ Proof.
exploit function_ptr_translated; eauto. intros [tf [FIND TR]].
econstructor; split.
econstructor.
- apply (Genv.init_mem_transf_partial _ _ TRANSL). eauto.
+ apply (Genv.init_mem_transf_partial _ _ TRANSL). eauto.
simpl. fold tge. rewrite symbols_preserved.
replace (prog_main tprog) with (prog_main prog). eexact H0.
symmetry. unfold transl_program in TRANSL.
eapply transform_partial_program_main; eauto.
- eexact FIND.
+ eexact FIND.
rewrite <- H2. apply sig_preserved; auto.
eapply match_callstate with (f := Mem.flat_inj (Mem.nextblock m0)) (cs := @nil frame) (cenv := PTree.empty Z).
auto.
eapply Genv.initmem_inject; eauto.
apply mcs_nil with (Mem.nextblock m0). apply match_globalenvs_init; auto. xomega. xomega.
constructor. red; auto.
- constructor.
+ constructor.
Qed.
Lemma transl_final_states:
@@ -2253,7 +2253,7 @@ Proof.
eexact public_preserved.
eexact transl_initial_states.
eexact transl_final_states.
- eexact transl_step_correct.
+ eexact transl_step_correct.
Qed.
End TRANSLATION.
diff --git a/cfrontend/Cop.v b/cfrontend/Cop.v
index 948ccaca..b4784028 100644
--- a/cfrontend/Cop.v
+++ b/cfrontend/Cop.v
@@ -65,7 +65,7 @@ Inductive incr_or_decr : Type := Incr | Decr.
The [sem_*] functions below compute the result of an operator
application. Since operators are overloaded, the result depends
both on the static types of the arguments and on their run-time values.
- The corresponding [classify_*] function is first called on the
+ The corresponding [classify_*] function is first called on the
types of the arguments to resolve static overloading. It is then
followed by a case analysis on the values of the arguments. *)
@@ -139,7 +139,7 @@ Definition cast_int_int (sz: intsize) (sg: signedness) (i: int) : int :=
| I8, Signed => Int.sign_ext 8 i
| I8, Unsigned => Int.zero_ext 8 i
| I16, Signed => Int.sign_ext 16 i
- | I16, Unsigned => Int.zero_ext 16 i
+ | I16, Unsigned => Int.zero_ext 16 i
| I32, _ => i
| IBool, _ => if Int.eq i Int.zero then Int.zero else Int.one
end.
@@ -343,8 +343,8 @@ Definition sem_cast (v: val) (t1 t2: type) : option val :=
end.
(** The following describes types that can be interpreted as a boolean:
- integers, floats, pointers. It is used for the semantics of
- the [!] and [?] operators, as well as the [if], [while],
+ integers, floats, pointers. It is used for the semantics of
+ the [!] and [?] operators, as well as the [if], [while],
and [for] statements. *)
Inductive classify_bool_cases : Type :=
@@ -638,33 +638,33 @@ Definition classify_add (ty1: type) (ty2: type) :=
end.
Definition sem_add (cenv: composite_env) (v1:val) (t1:type) (v2: val) (t2:type) : option val :=
- match classify_add t1 t2 with
+ match classify_add t1 t2 with
| add_case_pi ty => (**r pointer plus integer *)
match v1,v2 with
- | Vptr b1 ofs1, Vint n2 =>
+ | Vptr b1 ofs1, Vint n2 =>
Some (Vptr b1 (Int.add ofs1 (Int.mul (Int.repr (sizeof cenv ty)) n2)))
| _, _ => None
- end
+ end
| add_case_ip ty => (**r integer plus pointer *)
match v1,v2 with
- | Vint n1, Vptr b2 ofs2 =>
+ | Vint n1, Vptr b2 ofs2 =>
Some (Vptr b2 (Int.add ofs2 (Int.mul (Int.repr (sizeof cenv ty)) n1)))
| _, _ => None
- end
+ end
| add_case_pl ty => (**r pointer plus long *)
match v1,v2 with
- | Vptr b1 ofs1, Vlong n2 =>
+ | Vptr b1 ofs1, Vlong n2 =>
let n2 := Int.repr (Int64.unsigned n2) in
Some (Vptr b1 (Int.add ofs1 (Int.mul (Int.repr (sizeof cenv ty)) n2)))
| _, _ => None
- end
+ end
| add_case_lp ty => (**r long plus pointer *)
match v1,v2 with
- | Vlong n1, Vptr b2 ofs2 =>
+ | Vlong n1, Vptr b2 ofs2 =>
let n1 := Int.repr (Int64.unsigned n1) in
Some (Vptr b2 (Int.add ofs2 (Int.mul (Int.repr (sizeof cenv ty)) n1)))
| _, _ => None
- end
+ end
| add_default =>
sem_binarith
(fun sg n1 n2 => Some(Vint(Int.add n1 n2)))
@@ -694,13 +694,13 @@ Definition sem_sub (cenv: composite_env) (v1:val) (t1:type) (v2: val) (t2:type)
match classify_sub t1 t2 with
| sub_case_pi ty => (**r pointer minus integer *)
match v1,v2 with
- | Vptr b1 ofs1, Vint n2 =>
+ | Vptr b1 ofs1, Vint n2 =>
Some (Vptr b1 (Int.sub ofs1 (Int.mul (Int.repr (sizeof cenv ty)) n2)))
| _, _ => None
end
| sub_case_pl ty => (**r pointer minus long *)
match v1,v2 with
- | Vptr b1 ofs1, Vlong n2 =>
+ | Vptr b1 ofs1, Vlong n2 =>
let n2 := Int.repr (Int64.unsigned n2) in
Some (Vptr b1 (Int.sub ofs1 (Int.mul (Int.repr (sizeof cenv ty)) n2)))
| _, _ => None
@@ -724,7 +724,7 @@ Definition sem_sub (cenv: composite_env) (v1:val) (t1:type) (v2: val) (t2:type)
(fun n1 n2 => Some(Vsingle(Float32.sub n1 n2)))
v1 t1 v2 t2
end.
-
+
(** *** Multiplication, division, modulus *)
Definition sem_mul (v1:val) (t1:type) (v2: val) (t2:type) : option val :=
@@ -842,28 +842,28 @@ Definition sem_shift
match classify_shift t1 t2 with
| shift_case_ii sg =>
match v1, v2 with
- | Vint n1, Vint n2 =>
+ | Vint n1, Vint n2 =>
if Int.ltu n2 Int.iwordsize
then Some(Vint(sem_int sg n1 n2)) else None
| _, _ => None
end
| shift_case_il sg =>
match v1, v2 with
- | Vint n1, Vlong n2 =>
+ | Vint n1, Vlong n2 =>
if Int64.ltu n2 (Int64.repr 32)
then Some(Vint(sem_int sg n1 (Int64.loword n2))) else None
| _, _ => None
end
| shift_case_li sg =>
match v1, v2 with
- | Vlong n1, Vint n2 =>
+ | Vlong n1, Vint n2 =>
if Int.ltu n2 Int64.iwordsize'
then Some(Vlong(sem_long sg n1 (Int64.repr (Int.unsigned n2)))) else None
| _, _ => None
end
| shift_case_ll sg =>
match v1, v2 with
- | Vlong n1, Vlong n2 =>
+ | Vlong n1, Vlong n2 =>
if Int64.ltu n2 Int64.iwordsize
then Some(Vlong(sem_long sg n1 n2)) else None
| _, _ => None
@@ -892,7 +892,7 @@ Inductive classify_cmp_cases : Type :=
| cmp_default. (**r numerical, numerical *)
Definition classify_cmp (ty1: type) (ty2: type) :=
- match typeconv ty1, typeconv ty2 with
+ match typeconv ty1, typeconv ty2 with
| Tpointer _ _ , Tpointer _ _ => cmp_case_pp
| Tpointer _ _ , Tint _ _ _ => cmp_case_pp
| Tint _ _ _, Tpointer _ _ => cmp_case_pp
@@ -909,14 +909,14 @@ Definition sem_cmp (c:comparison)
option_map Val.of_bool (Val.cmpu_bool (Mem.valid_pointer m) c v1 v2)
| cmp_case_pl =>
match v2 with
- | Vlong n2 =>
+ | Vlong n2 =>
let n2 := Int.repr (Int64.unsigned n2) in
option_map Val.of_bool (Val.cmpu_bool (Mem.valid_pointer m) c v1 (Vint n2))
| _ => None
end
| cmp_case_lp =>
match v1 with
- | Vlong n1 =>
+ | Vlong n1 =>
let n1 := Int.repr (Int64.unsigned n1) in
option_map Val.of_bool (Val.cmpu_bool (Mem.valid_pointer m) c (Vint n1) v2)
| _ => None
@@ -941,7 +941,7 @@ Inductive classify_fun_cases : Type :=
| fun_default.
Definition classify_fun (ty: type) :=
- match ty with
+ match ty with
| Tfunction args res cc => fun_case_f args res cc
| Tpointer (Tfunction args res cc) _ => fun_case_f args res cc
| _ => fun_default
@@ -989,15 +989,15 @@ Definition sem_binary_operation
(m: mem): option val :=
match op with
| Oadd => sem_add cenv v1 t1 v2 t2
- | Osub => sem_sub cenv v1 t1 v2 t2
+ | Osub => sem_sub cenv v1 t1 v2 t2
| Omul => sem_mul v1 t1 v2 t2
| Omod => sem_mod v1 t1 v2 t2
- | Odiv => sem_div v1 t1 v2 t2
+ | Odiv => sem_div v1 t1 v2 t2
| Oand => sem_and v1 t1 v2 t2
| Oor => sem_or v1 t1 v2 t2
| Oxor => sem_xor v1 t1 v2 t2
| Oshl => sem_shl v1 t1 v2 t2
- | Oshr => sem_shr v1 t1 v2 t2
+ | Oshr => sem_shr v1 t1 v2 t2
| Oeq => sem_cmp Ceq v1 t1 v2 t2 m
| One => sem_cmp Cne v1 t1 v2 t2 m
| Olt => sem_cmp Clt v1 t1 v2 t2 m
@@ -1063,7 +1063,7 @@ Remark val_inject_vfalse: forall f, Val.inject f Vfalse Vfalse.
Proof. unfold Vfalse; auto. Qed.
Remark val_inject_of_bool: forall f b, Val.inject f (Val.of_bool b) (Val.of_bool b).
-Proof. intros. unfold Val.of_bool. destruct b; [apply val_inject_vtrue|apply val_inject_vfalse].
+Proof. intros. unfold Val.of_bool. destruct b; [apply val_inject_vtrue|apply val_inject_vfalse].
Qed.
Hint Resolve val_inject_vtrue val_inject_vfalse val_inject_of_bool.
@@ -1082,7 +1082,7 @@ Lemma sem_cast_inject:
Proof.
unfold sem_cast; intros; destruct (classify_cast ty1 ty);
inv H0; inv H; TrivialInject.
-- econstructor; eauto.
+- econstructor; eauto.
- destruct (cast_float_int si2 f0); inv H1; TrivialInject.
- destruct (cast_single_int si2 f0); inv H1; TrivialInject.
- destruct (cast_float_long si2 f0); inv H1; TrivialInject.
@@ -1102,7 +1102,7 @@ Proof.
(* notbool *)
unfold sem_notbool in *; destruct (classify_bool ty); inv H0; inv H; TrivialInject.
destruct (Mem.weak_valid_pointer m b1 (Int.unsigned ofs1)) eqn:VP; inv H2.
- erewrite weak_valid_pointer_inj by eauto. TrivialInject.
+ erewrite weak_valid_pointer_inj by eauto. TrivialInject.
(* notint *)
unfold sem_notint in *; destruct (classify_notint ty); inv H0; inv H; TrivialInject.
(* neg *)
@@ -1127,7 +1127,7 @@ Remark sem_binarith_inject:
(forall n1 n2, optval_self_injects (sem_single n1 n2)) ->
exists v', sem_binarith sem_int sem_long sem_float sem_single v1' t1 v2' t2 = Some v' /\ Val.inject f v v'.
Proof.
- intros.
+ intros.
assert (SELF: forall ov v, ov = Some v -> optval_self_injects ov -> Val.inject f v v).
{
intros. subst ov; simpl in H7. destruct v0; contradiction || constructor.
@@ -1169,22 +1169,22 @@ Proof.
- (* pointer - pointer *)
destruct (Val.cmpu_bool (Mem.valid_pointer m) cmp v1 v2) as [b|] eqn:E; simpl in H; inv H.
replace (Val.cmpu_bool (Mem.valid_pointer m') cmp tv1 tv2) with (Some b).
- simpl. TrivialInject.
- symmetry. eapply Val.cmpu_bool_inject; eauto.
+ simpl. TrivialInject.
+ symmetry. eapply Val.cmpu_bool_inject; eauto.
- (* pointer - long *)
- destruct v2; try discriminate. inv H1.
+ destruct v2; try discriminate. inv H1.
set (v2 := Vint (Int.repr (Int64.unsigned i))) in *.
destruct (Val.cmpu_bool (Mem.valid_pointer m) cmp v1 v2) as [b|] eqn:E; simpl in H; inv H.
replace (Val.cmpu_bool (Mem.valid_pointer m') cmp tv1 v2) with (Some b).
- simpl. TrivialInject.
- symmetry. eapply Val.cmpu_bool_inject with (v2 := v2); eauto. constructor.
+ simpl. TrivialInject.
+ symmetry. eapply Val.cmpu_bool_inject with (v2 := v2); eauto. constructor.
- (* long - pointer *)
- destruct v1; try discriminate. inv H0.
+ destruct v1; try discriminate. inv H0.
set (v1 := Vint (Int.repr (Int64.unsigned i))) in *.
destruct (Val.cmpu_bool (Mem.valid_pointer m) cmp v1 v2) as [b|] eqn:E; simpl in H; inv H.
replace (Val.cmpu_bool (Mem.valid_pointer m') cmp v1 tv2) with (Some b).
- simpl. TrivialInject.
- symmetry. eapply Val.cmpu_bool_inject with (v1 := v1); eauto. constructor.
+ simpl. TrivialInject.
+ symmetry. eapply Val.cmpu_bool_inject with (v1 := v1); eauto. constructor.
- (* numerical - numerical *)
assert (SELF: forall b, optval_self_injects (Some (Val.of_bool b))).
{
@@ -1202,13 +1202,13 @@ Proof.
unfold sem_binary_operation; intros; destruct op.
- (* add *)
unfold sem_add in *; destruct (classify_add ty1 ty2).
- + inv H0; inv H1; inv H. TrivialInject.
+ + inv H0; inv H1; inv H. TrivialInject.
econstructor. eauto. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut.
- + inv H0; inv H1; inv H. TrivialInject.
+ + inv H0; inv H1; inv H. TrivialInject.
econstructor. eauto. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut.
- + inv H0; inv H1; inv H. TrivialInject.
+ + inv H0; inv H1; inv H. TrivialInject.
econstructor. eauto. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut.
- + inv H0; inv H1; inv H. TrivialInject.
+ + inv H0; inv H1; inv H. TrivialInject.
econstructor. eauto. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut.
+ eapply sem_binarith_inject; eauto; intros; exact I.
- (* sub *)
@@ -1216,8 +1216,8 @@ Proof.
+ inv H0; inv H1; inv H. TrivialInject.
econstructor. eauto. rewrite Int.sub_add_l. auto.
+ inv H0; inv H1; inv H. TrivialInject.
- destruct (eq_block b1 b0); try discriminate. subst b1.
- rewrite H0 in H2; inv H2. rewrite dec_eq_true.
+ destruct (eq_block b1 b0); try discriminate. subst b1.
+ rewrite H0 in H2; inv H2. rewrite dec_eq_true.
destruct (zlt 0 (sizeof cenv ty) && zle (sizeof cenv ty) Int.max_signed); inv H3.
rewrite Int.sub_shifted. TrivialInject.
+ inv H0; inv H1; inv H. TrivialInject.
@@ -1274,7 +1274,7 @@ Lemma bool_val_inj:
Val.inject f v tv ->
bool_val tv ty m' = Some b.
Proof.
- unfold bool_val; intros.
+ unfold bool_val; intros.
destruct (classify_bool ty); inv H0; try congruence.
destruct (Mem.weak_valid_pointer m b1 (Int.unsigned ofs1)) eqn:VP; inv H.
erewrite weak_valid_pointer_inj by eauto. auto.
@@ -1289,7 +1289,7 @@ Lemma sem_unary_operation_inject:
Mem.inject f m m' ->
exists tv, sem_unary_operation op tv1 ty1 m' = Some tv /\ Val.inject f v tv.
Proof.
- intros. eapply sem_unary_operation_inj; eauto.
+ intros. eapply sem_unary_operation_inj; eauto.
intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
Qed.
@@ -1300,7 +1300,7 @@ Lemma sem_binary_operation_inject:
Mem.inject f m m' ->
exists tv, sem_binary_operation cenv op tv1 ty1 tv2 ty2 m' = Some tv /\ Val.inject f v tv.
Proof.
- intros. eapply sem_binary_operation_inj; eauto.
+ intros. eapply sem_binary_operation_inj; eauto.
intros; eapply Mem.valid_pointer_inject_val; eauto.
intros; eapply Mem.weak_valid_pointer_inject_val; eauto.
intros; eapply Mem.weak_valid_pointer_inject_no_overflow; eauto.
@@ -1359,9 +1359,9 @@ Proof.
destruct f; auto.
destruct (Float32.cmp Ceq f0 Float32.zero); auto.
destruct f; auto.
- destruct (Int.eq i Int.zero); auto.
- destruct (Int.eq i Int.zero); auto.
- destruct (Int.eq i Int.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
Qed.
(** Relation between Boolean value and Boolean negation. *)
@@ -1371,9 +1371,9 @@ Lemma notbool_bool_val:
sem_notbool v t m =
match bool_val v t m with None => None | Some b => Some(Val.of_bool (negb b)) end.
Proof.
- intros. unfold sem_notbool, bool_val.
+ intros. unfold sem_notbool, bool_val.
destruct (classify_bool t); auto; destruct v; auto; rewrite ? negb_involutive; auto.
- destruct (Mem.weak_valid_pointer m b (Int.unsigned i)); auto.
+ destruct (Mem.weak_valid_pointer m b (Int.unsigned i)); auto.
Qed.
(** Relation with the arithmetic conversions of ISO C99, section 6.3.1 *)
@@ -1558,4 +1558,4 @@ End ArithConv.
-
+
diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v
index 3e9017c9..539b6826 100644
--- a/cfrontend/Csem.v
+++ b/cfrontend/Csem.v
@@ -292,7 +292,7 @@ Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop :=
deref_loc ty m b ofs t v1 ->
op = match id with Incr => Oadd | Decr => Osub end ->
rred (Epostincr id (Eloc b ofs ty) ty) m
- t (Ecomma (Eassign (Eloc b ofs ty)
+ t (Ecomma (Eassign (Eloc b ofs ty)
(Ebinop op (Eval v1 ty)
(Eval (Vint Int.one) type_int32s)
(incrdecr_type ty))
@@ -408,7 +408,7 @@ with contextlist: kind -> (expr -> exprlist) -> Prop :=
This expression is never stuck because the evaluation of [f()] can make
infinitely many transitions. Yet it contains a subexpression [10 / x]
that can go wrong if [x = 0], and the compiler may choose to evaluate
- [10 / x] first, before calling [f()].
+ [10 / x] first, before calling [f()].
Therefore, we must make sure that not only an expression cannot get stuck,
but none of its subexpressions can either. We say that a subexpression
@@ -437,10 +437,10 @@ Inductive imm_safe: kind -> expr -> mem -> Prop :=
is immediately stuck. *)
(*
Definition not_stuck (e: expr) (m: mem) : Prop :=
- forall k C e' ,
+ forall k C e' ,
context k RV C -> e = C e' -> not_imm_stuck k e' m.
*)
-End EXPR.
+End EXPR.
(** ** Transition semantics. *)
@@ -527,11 +527,11 @@ Inductive state: Type :=
(k: cont)
(m: mem) : state
| Stuckstate. (**r undefined behavior occurred *)
-
-(** Find the statement and manufacture the continuation
+
+(** Find the statement and manufacture the continuation
corresponding to a label. *)
-Fixpoint find_label (lbl: label) (s: statement) (k: cont)
+Fixpoint find_label (lbl: label) (s: statement) (k: cont)
{struct s}: option (statement * cont) :=
match s with
| Ssequence s1 s2 =>
@@ -564,7 +564,7 @@ Fixpoint find_label (lbl: label) (s: statement) (k: cont)
| _ => None
end
-with find_label_ls (lbl: label) (sl: labeled_statements) (k: cont)
+with find_label_ls (lbl: label) (sl: labeled_statements) (k: cont)
{struct sl}: option (statement * cont) :=
match sl with
| LSnil => None
@@ -793,7 +793,7 @@ Definition semantics (p: program) :=
(** This semantics has the single-event property. *)
-Lemma semantics_single_events:
+Lemma semantics_single_events:
forall p, single_events (semantics p).
Proof.
unfold semantics; intros; red; simpl; intros.
diff --git a/cfrontend/Csharpminor.v b/cfrontend/Csharpminor.v
index c34b10a4..e42091af 100644
--- a/cfrontend/Csharpminor.v
+++ b/cfrontend/Csharpminor.v
@@ -104,7 +104,7 @@ Definition funsig (fd: fundef) :=
(** Three evaluation environments are involved:
- [genv]: global environments, map symbols and functions to memory blocks,
and maps symbols to variable informations (type [var_kind])
-- [env]: local environments, map local variables
+- [env]: local environments, map local variables
to pairs (memory block, size)
- [temp_env]: local environments, map temporary variables to
their current values.
@@ -133,7 +133,7 @@ Fixpoint bind_parameters (formals: list ident) (args: list val)
| nil, nil => Some le
| id :: xl, v :: vl => bind_parameters xl vl (PTree.set id v le)
| _, _ => None
- end.
+ end.
(** Continuations *)
@@ -211,10 +211,10 @@ Fixpoint seq_of_lbl_stmt (sl: lbl_stmt) : stmt :=
| LScons c s sl' => Sseq s (seq_of_lbl_stmt sl')
end.
-(** Find the statement and manufacture the continuation
+(** Find the statement and manufacture the continuation
corresponding to a label *)
-Fixpoint find_label (lbl: label) (s: stmt) (k: cont)
+Fixpoint find_label (lbl: label) (s: stmt) (k: cont)
{struct s}: option (stmt * cont) :=
match s with
| Sseq s1 s2 =>
@@ -238,7 +238,7 @@ Fixpoint find_label (lbl: label) (s: stmt) (k: cont)
| _ => None
end
-with find_label_ls (lbl: label) (sl: lbl_stmt) (k: cont)
+with find_label_ls (lbl: label) (sl: lbl_stmt) (k: cont)
{struct sl}: option (stmt * cont) :=
match sl with
| LSnil => None
@@ -290,8 +290,8 @@ Section RELSEM.
Variable ge: genv.
-(* Evaluation of the address of a variable:
- [eval_var_addr prg ge e id b] states that variable [id]
+(* Evaluation of the address of a variable:
+ [eval_var_addr prg ge e id b] states that variable [id]
in environment [e] evaluates to block [b]. *)
Inductive eval_var_addr: env -> ident -> block -> Prop :=
@@ -460,7 +460,7 @@ Inductive step: state -> trace -> state -> Prop :=
| step_external_function: forall ef vargs k m t vres m',
external_call ef ge vargs m t vres m' ->
step (Callstate (External ef) vargs k m)
- t (Returnstate vres k m')
+ t (Returnstate vres k m')
| step_return: forall v optid f e le k m,
step (Returnstate v (Kcall optid f e le k) m)
diff --git a/cfrontend/Cshmgen.v b/cfrontend/Cshmgen.v
index a80f4c15..825a563c 100644
--- a/cfrontend/Cshmgen.v
+++ b/cfrontend/Cshmgen.v
@@ -130,10 +130,10 @@ Definition make_cmp_ne_zero (e: expr) :=
Definition make_cast_int (e: expr) (sz: intsize) (si: signedness) :=
match sz, si with
- | I8, Signed => Eunop Ocast8signed e
- | I8, Unsigned => Eunop Ocast8unsigned e
- | I16, Signed => Eunop Ocast16signed e
- | I16, Unsigned => Eunop Ocast16unsigned e
+ | I8, Signed => Eunop Ocast8signed e
+ | I8, Unsigned => Eunop Ocast8unsigned e
+ | I16, Signed => Eunop Ocast16signed e
+ | I16, Unsigned => Eunop Ocast16unsigned e
| I32, _ => e
| IBool, _ => make_cmp_ne_zero e
end.
@@ -356,7 +356,7 @@ Definition make_memcpy (ce: composite_env) (dst src: expr) (ty: type) :=
(** [make_store addr ty rhs] stores the value of the
Csharpminor expression [rhs] into the memory location denoted by the
- Csharpminor expression [addr].
+ Csharpminor expression [addr].
[ty] is the type of the memory location. *)
Definition make_store (ce: composite_env) (addr: expr) (ty: type) (rhs: expr) :=
@@ -376,7 +376,7 @@ Definition transl_unop (op: Cop.unary_operation) (a: expr) (ta: type) : res expr
| Cop.Oabsfloat => make_absfloat a ta
end.
-Definition transl_binop (ce: composite_env)
+Definition transl_binop (ce: composite_env)
(op: Cop.binary_operation)
(a: expr) (ta: type)
(b: expr) (tb: type) : res expr :=
@@ -473,10 +473,10 @@ with transl_lvalue (ce: composite_env) (a: Clight.expr) {struct a} : res expr :=
OK (Eaddrof id)
| Clight.Ederef b _ =>
transl_expr ce b
- | Clight.Efield b i ty =>
+ | Clight.Efield b i ty =>
do tb <- transl_expr ce b;
make_field_access ce (typeof b) i tb
- | _ =>
+ | _ =>
Error(msg "Cshmgen.transl_lvalue")
end.
@@ -492,7 +492,7 @@ Fixpoint transl_arglist (ce: composite_env) (al: list Clight.expr) (tyl: typelis
| a1 :: a2, Tcons ty1 ty2 =>
do ta1 <- transl_expr ce a1;
do ta1' <- make_cast (typeof a1) ty1 ta1;
- do ta2 <- transl_arglist ce a2 ty2;
+ do ta2 <- transl_arglist ce a2 ty2;
OK (ta1' :: ta2)
| a1 :: a2, Tnil =>
(* Tolerance for calls to K&R or variadic functions *)
@@ -630,7 +630,7 @@ Definition signature_of_function (f: Clight.function) :=
Definition transl_function (ce: composite_env) (f: Clight.function) : res function :=
do tbody <- transl_statement ce f.(Clight.fn_return) 1%nat 0%nat (Clight.fn_body f);
- OK (mkfunction
+ OK (mkfunction
(signature_of_function f)
(map fst (Clight.fn_params f))
(map (transl_var ce) (Clight.fn_vars f))
@@ -639,7 +639,7 @@ Definition transl_function (ce: composite_env) (f: Clight.function) : res functi
Definition transl_fundef (ce: composite_env) (f: Clight.fundef) : res fundef :=
match f with
- | Clight.Internal g =>
+ | Clight.Internal g =>
do tg <- transl_function ce g; OK(AST.Internal tg)
| Clight.External ef args res cconv =>
if signature_eq (ef_sig ef) (signature_of_type args res cconv)
diff --git a/cfrontend/Cshmgenproof.v b/cfrontend/Cshmgenproof.v
index c69d0c0a..e25e21c9 100644
--- a/cfrontend/Cshmgenproof.v
+++ b/cfrontend/Cshmgenproof.v
@@ -45,8 +45,8 @@ Lemma transl_fundef_sig1:
classify_fun (type_of_fundef f) = fun_case_f args res cc ->
funsig tf = signature_of_type args res cc.
Proof.
- intros. destruct f; simpl in *.
- monadInv H. monadInv EQ. simpl. inversion H0.
+ intros. destruct f; simpl in *.
+ monadInv H. monadInv EQ. simpl. inversion H0.
unfold signature_of_function, signature_of_type.
f_equal. apply transl_params_types.
destruct (signature_eq (ef_sig e) (signature_of_type t t0 c)); inv H.
@@ -81,9 +81,9 @@ Proof.
(* deref *)
monadInv TR. exists x; auto.
(* field struct *)
- monadInv TR. exists x0; split; auto. simpl; rewrite EQ; auto.
+ monadInv TR. exists x0; split; auto. simpl; rewrite EQ; auto.
(* field union *)
- monadInv TR. exists x0; split; auto. simpl; rewrite EQ; auto.
+ monadInv TR. exists x0; split; auto. simpl; rewrite EQ; auto.
Qed.
(** Properties of labeled statements *)
@@ -98,7 +98,7 @@ Proof.
transl_lbl_stmt ce tyret nbrk ncnt sl = OK tsl ->
transl_lbl_stmt ce tyret nbrk ncnt (Clight.select_switch_default sl) = OK (select_switch_default tsl)).
{
- induction sl; simpl; intros.
+ induction sl; simpl; intros.
inv H; auto.
monadInv H. simpl. destruct o; eauto. simpl; rewrite EQ; simpl; rewrite EQ1; auto.
}
@@ -114,16 +114,16 @@ Proof.
end).
{
induction sl; simpl; intros.
- inv H; auto.
+ inv H; auto.
monadInv H; simpl. destruct o. destruct (zeq z n).
econstructor; split; eauto. simpl; rewrite EQ; simpl; rewrite EQ1; auto.
apply IHsl; auto.
apply IHsl; auto.
}
- intros. specialize (CASE _ _ H). unfold Clight.select_switch, select_switch.
- destruct (Clight.select_switch_case n sl) as [sl'|].
+ intros. specialize (CASE _ _ H). unfold Clight.select_switch, select_switch.
+ destruct (Clight.select_switch_case n sl) as [sl'|].
destruct CASE as [tsl' [P Q]]. rewrite P, Q. auto.
- rewrite CASE. auto.
+ rewrite CASE. auto.
Qed.
Lemma transl_lbl_stmt_2:
@@ -132,7 +132,7 @@ Lemma transl_lbl_stmt_2:
transl_statement ce tyret nbrk ncnt (seq_of_labeled_statement sl) = OK (seq_of_lbl_stmt tsl).
Proof.
induction sl; intros.
- monadInv H. auto.
+ monadInv H. auto.
monadInv H. simpl. rewrite EQ; simpl. rewrite (IHsl _ EQ1). simpl. auto.
Qed.
@@ -147,28 +147,28 @@ Lemma make_intconst_correct:
forall n e le m,
eval_expr ge e le m (make_intconst n) (Vint n).
Proof.
- intros. unfold make_intconst. econstructor. reflexivity.
+ intros. unfold make_intconst. econstructor. reflexivity.
Qed.
Lemma make_floatconst_correct:
forall n e le m,
eval_expr ge e le m (make_floatconst n) (Vfloat n).
Proof.
- intros. unfold make_floatconst. econstructor. reflexivity.
+ intros. unfold make_floatconst. econstructor. reflexivity.
Qed.
Lemma make_singleconst_correct:
forall n e le m,
eval_expr ge e le m (make_singleconst n) (Vsingle n).
Proof.
- intros. unfold make_singleconst. econstructor. reflexivity.
+ intros. unfold make_singleconst. econstructor. reflexivity.
Qed.
Lemma make_longconst_correct:
forall n e le m,
eval_expr ge e le m (make_longconst n) (Vlong n).
Proof.
- intros. unfold make_floatconst. econstructor. reflexivity.
+ intros. unfold make_floatconst. econstructor. reflexivity.
Qed.
Lemma make_singleoffloat_correct:
@@ -193,7 +193,7 @@ Lemma make_floatofint_correct:
eval_expr ge e le m (make_floatofint a sg) (Vfloat(cast_int_float sg n)).
Proof.
intros. unfold make_floatofint, cast_int_float.
- destruct sg; econstructor; eauto.
+ destruct sg; econstructor; eauto.
Qed.
Hint Resolve make_intconst_correct make_floatconst_correct make_longconst_correct
@@ -207,33 +207,33 @@ Lemma make_cmp_ne_zero_correct:
eval_expr ge e le m a (Vint n) ->
eval_expr ge e le m (make_cmp_ne_zero a) (Vint (if Int.eq n Int.zero then Int.zero else Int.one)).
Proof.
- intros.
+ intros.
assert (DEFAULT: eval_expr ge e le m (Ebinop (Ocmp Cne) a (make_intconst Int.zero))
(Vint (if Int.eq n Int.zero then Int.zero else Int.one))).
- econstructor; eauto with cshm. simpl. unfold Val.cmp, Val.cmp_bool.
- unfold Int.cmp. destruct (Int.eq n Int.zero); auto.
+ econstructor; eauto with cshm. simpl. unfold Val.cmp, Val.cmp_bool.
+ unfold Int.cmp. destruct (Int.eq n Int.zero); auto.
assert (CMP: forall ob,
Val.of_optbool ob = Vint n ->
n = (if Int.eq n Int.zero then Int.zero else Int.one)).
- intros. destruct ob; simpl in H0; inv H0. destruct b; inv H2.
+ intros. destruct ob; simpl in H0; inv H0. destruct b; inv H2.
rewrite Int.eq_false. auto. apply Int.one_not_zero.
rewrite Int.eq_true. auto.
- destruct a; simpl; auto. destruct b; auto.
- inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ destruct a; simpl; auto. destruct b; auto.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
simpl in H6. inv H6. unfold Val.cmp in H0. eauto.
- inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
simpl in H6. inv H6. unfold Val.cmp in H0. eauto.
- inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
simpl in H6. inv H6. unfold Val.cmp in H0. eauto.
- inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
simpl in H6. unfold Val.cmpfs in H6.
- destruct (Val.cmpfs_bool c v1 v2) as [[]|]; inv H6; reflexivity.
- inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ destruct (Val.cmpfs_bool c v1 v2) as [[]|]; inv H6; reflexivity.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
simpl in H6. unfold Val.cmpl in H6.
- destruct (Val.cmpl_bool c v1 v2) as [[]|]; inv H6; reflexivity.
- inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ destruct (Val.cmpl_bool c v1 v2) as [[]|]; inv H6; reflexivity.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
simpl in H6. unfold Val.cmplu in H6.
- destruct (Val.cmplu_bool c v1 v2) as [[]|]; inv H6; reflexivity.
+ destruct (Val.cmplu_bool c v1 v2) as [[]|]; inv H6; reflexivity.
Qed.
Lemma make_cast_int_correct:
@@ -241,7 +241,7 @@ Lemma make_cast_int_correct:
eval_expr ge e le m a (Vint n) ->
eval_expr ge e le m (make_cast_int a sz si) (Vint (cast_int_int sz si n)).
Proof.
- intros. unfold make_cast_int, cast_int_int.
+ intros. unfold make_cast_int, cast_int_int.
destruct sz.
destruct si; eauto with cshm.
destruct si; eauto with cshm.
@@ -261,15 +261,15 @@ Proof.
intros. unfold make_cast, sem_cast in *;
destruct (classify_cast ty1 ty2); inv H; destruct v; inv H1; eauto with cshm.
(* single -> int *)
- unfold make_singleofint, cast_int_float. destruct si1; eauto with cshm.
+ unfold make_singleofint, cast_int_float. destruct si1; eauto with cshm.
(* float -> int *)
destruct (cast_float_int si2 f) as [i|] eqn:E; inv H2.
- apply make_cast_int_correct.
+ apply make_cast_int_correct.
unfold cast_float_int in E. unfold make_intoffloat.
destruct si2; econstructor; eauto; simpl; rewrite E; auto.
(* single -> int *)
destruct (cast_single_int si2 f) as [i|] eqn:E; inv H2.
- apply make_cast_int_correct.
+ apply make_cast_int_correct.
unfold cast_single_int in E. unfold make_intofsingle.
destruct si2; econstructor; eauto with cshm; simpl; rewrite E; auto.
(* long -> int *)
@@ -316,36 +316,36 @@ Lemma make_boolean_correct:
eval_expr ge e le m (make_boolean a ty) vb
/\ Val.bool_of_val vb b.
Proof.
- intros. unfold make_boolean. unfold bool_val in H0.
+ intros. unfold make_boolean. unfold bool_val in H0.
destruct (classify_bool ty); destruct v; inv H0.
(* int *)
- econstructor; split. apply make_cmp_ne_zero_correct with (n := i); auto.
- destruct (Int.eq i Int.zero); simpl; constructor.
+ econstructor; split. apply make_cmp_ne_zero_correct with (n := i); auto.
+ destruct (Int.eq i Int.zero); simpl; constructor.
(* float *)
- econstructor; split. econstructor; eauto with cshm. simpl. eauto.
- unfold Val.cmpf, Val.cmpf_bool. simpl. rewrite <- Float.cmp_ne_eq.
- destruct (Float.cmp Cne f Float.zero); constructor.
+ econstructor; split. econstructor; eauto with cshm. simpl. eauto.
+ unfold Val.cmpf, Val.cmpf_bool. simpl. rewrite <- Float.cmp_ne_eq.
+ destruct (Float.cmp Cne f Float.zero); constructor.
(* single *)
- econstructor; split. econstructor; eauto with cshm. simpl. eauto.
- unfold Val.cmpfs, Val.cmpfs_bool. simpl. rewrite <- Float32.cmp_ne_eq.
- destruct (Float32.cmp Cne f Float32.zero); constructor.
+ econstructor; split. econstructor; eauto with cshm. simpl. eauto.
+ unfold Val.cmpfs, Val.cmpfs_bool. simpl. rewrite <- Float32.cmp_ne_eq.
+ destruct (Float32.cmp Cne f Float32.zero); constructor.
(* pointer *)
- econstructor; split. econstructor; eauto with cshm. simpl. eauto.
+ econstructor; split. econstructor; eauto with cshm. simpl. eauto.
unfold Val.cmpu, Val.cmpu_bool. simpl.
destruct (Int.eq i Int.zero); simpl; constructor.
econstructor; split. econstructor; eauto with cshm. simpl. eauto.
destruct (Mem.weak_valid_pointer m b0 (Int.unsigned i)) eqn:V; inv H2.
unfold Val.cmpu, Val.cmpu_bool. simpl.
- unfold Mem.weak_valid_pointer in V; rewrite V. constructor.
+ unfold Mem.weak_valid_pointer in V; rewrite V. constructor.
(* long *)
- econstructor; split. econstructor; eauto with cshm. simpl. unfold Val.cmpl. simpl. eauto.
- destruct (Int64.eq i Int64.zero); simpl; constructor.
+ econstructor; split. econstructor; eauto with cshm. simpl. unfold Val.cmpl. simpl. eauto.
+ destruct (Int64.eq i Int64.zero); simpl; constructor.
Qed.
Lemma make_neg_correct:
forall a tya c va v e le m,
sem_neg va tya = Some v ->
- make_neg a tya = OK c ->
+ make_neg a tya = OK c ->
eval_expr ge e le m a va ->
eval_expr ge e le m c v.
Proof.
@@ -356,21 +356,21 @@ Qed.
Lemma make_absfloat_correct:
forall a tya c va v e le m,
sem_absfloat va tya = Some v ->
- make_absfloat a tya = OK c ->
+ make_absfloat a tya = OK c ->
eval_expr ge e le m a va ->
eval_expr ge e le m c v.
Proof.
unfold sem_absfloat, make_absfloat; intros until m; intros SEM MAKE EV1;
destruct (classify_neg tya); inv MAKE; destruct va; inv SEM; eauto with cshm.
unfold make_floatoflong, cast_long_float. destruct s.
- econstructor. econstructor; simpl; eauto. simpl; eauto. simpl; eauto.
- econstructor. econstructor; simpl; eauto. simpl; eauto. simpl; eauto.
+ econstructor. econstructor; simpl; eauto. simpl; eauto. simpl; eauto.
+ econstructor. econstructor; simpl; eauto. simpl; eauto. simpl; eauto.
Qed.
Lemma make_notbool_correct:
forall a tya c va v e le m,
sem_notbool va tya m = Some v ->
- make_notbool a tya = OK c ->
+ make_notbool a tya = OK c ->
eval_expr ge e le m a va ->
eval_expr ge e le m c v.
Proof.
@@ -378,13 +378,13 @@ Proof.
destruct (classify_bool tya); inv MAKE; destruct va; inv SEM; eauto with cshm.
destruct (Mem.weak_valid_pointer m b (Int.unsigned i)) eqn:V; inv H0.
econstructor; eauto with cshm. simpl. unfold Val.cmpu, Val.cmpu_bool.
- unfold Mem.weak_valid_pointer in V; rewrite V. auto.
+ unfold Mem.weak_valid_pointer in V; rewrite V. auto.
Qed.
Lemma make_notint_correct:
forall a tya c va v e le m,
sem_notint va tya = Some v ->
- make_notint a tya = OK c ->
+ make_notint a tya = OK c ->
eval_expr ge e le m a va ->
eval_expr ge e le m c v.
Proof.
@@ -397,7 +397,7 @@ Definition binary_constructor_correct
(sem: val -> type -> val -> type -> option val): Prop :=
forall a tya b tyb c va vb v e le m,
sem va tya vb tyb = Some v ->
- make a tya b tyb = OK c ->
+ make a tya b tyb = OK c ->
eval_expr ge e le m a va ->
eval_expr ge e le m b vb ->
eval_expr ge e le m c v.
@@ -438,9 +438,9 @@ Proof.
exploit make_cast_correct. eexact EQ. eauto. eauto. intros EV1'.
exploit make_cast_correct. eexact EQ1. eauto. eauto. intros EV2'.
destruct cls; inv EQ2; destruct va'; try discriminate; destruct vb'; try discriminate.
-- destruct s; inv H0; econstructor; eauto with cshm.
+- destruct s; inv H0; econstructor; eauto with cshm.
rewrite iop_ok; auto. rewrite iopu_ok; auto.
-- destruct s; inv H0; econstructor; eauto with cshm.
+- destruct s; inv H0; econstructor; eauto with cshm.
rewrite lop_ok; auto. rewrite lopu_ok; auto.
- erewrite <- fop_ok in SEM; eauto with cshm.
- erewrite <- sop_ok in SEM; eauto with cshm.
@@ -461,9 +461,9 @@ Proof.
exploit make_cast_correct. eexact EQ. eauto. eauto. intros EV1'.
exploit make_cast_correct. eexact EQ1. eauto. eauto. intros EV2'.
destruct cls; inv EQ2; destruct va'; try discriminate; destruct vb'; try discriminate.
-- destruct s; inv H0; econstructor; eauto with cshm.
+- destruct s; inv H0; econstructor; eauto with cshm.
rewrite iop_ok; auto. rewrite iopu_ok; auto.
-- destruct s; inv H0; econstructor; eauto with cshm.
+- destruct s; inv H0; econstructor; eauto with cshm.
rewrite lop_ok; auto. rewrite lopu_ok; auto.
Qed.
@@ -492,17 +492,17 @@ Proof.
- destruct va; try discriminate; destruct vb; inv SEM.
destruct (eq_block b0 b1); try discriminate.
set (sz := sizeof ce ty) in *.
- destruct (zlt 0 sz); try discriminate.
+ destruct (zlt 0 sz); try discriminate.
destruct (zle sz Int.max_signed); simpl in H0; inv H0.
- econstructor; eauto with cshm.
- rewrite dec_eq_true; simpl.
- assert (E: Int.signed (Int.repr sz) = sz).
+ econstructor; eauto with cshm.
+ rewrite dec_eq_true; simpl.
+ assert (E: Int.signed (Int.repr sz) = sz).
{ apply Int.signed_repr. generalize Int.min_signed_neg; omega. }
predSpec Int.eq Int.eq_spec (Int.repr sz) Int.zero.
rewrite H in E; rewrite Int.signed_zero in E; omegaContradiction.
predSpec Int.eq Int.eq_spec (Int.repr sz) Int.mone.
rewrite H0 in E; rewrite Int.signed_mone in E; omegaContradiction.
- rewrite andb_false_r; auto.
+ rewrite andb_false_r; auto.
- destruct va; try discriminate; destruct vb; inv SEM; eauto with cshm.
- eapply make_binarith_correct; eauto; intros; auto.
Qed.
@@ -546,10 +546,10 @@ Remark small_shift_amount_1:
Int.ltu (Int64.loword i) Int64.iwordsize' = true
/\ Int64.unsigned i = Int.unsigned (Int64.loword i).
Proof.
- intros. apply Int64.ltu_inv in H. comput (Int64.unsigned Int64.iwordsize).
+ intros. apply Int64.ltu_inv in H. comput (Int64.unsigned Int64.iwordsize).
assert (Int64.unsigned i = Int.unsigned (Int64.loword i)).
{
- unfold Int64.loword. rewrite Int.unsigned_repr; auto.
+ unfold Int64.loword. rewrite Int.unsigned_repr; auto.
comput Int.max_unsigned; omega.
}
split; auto. unfold Int.ltu. apply zlt_true. rewrite <- H0. tauto.
@@ -563,7 +563,7 @@ Proof.
intros. apply Int64.ltu_inv in H. comput (Int64.unsigned (Int64.repr 32)).
assert (Int64.unsigned i = Int.unsigned (Int64.loword i)).
{
- unfold Int64.loword. rewrite Int.unsigned_repr; auto.
+ unfold Int64.loword. rewrite Int.unsigned_repr; auto.
comput Int.max_unsigned; omega.
}
unfold Int.ltu. apply zlt_true. rewrite <- H0. tauto.
@@ -574,7 +574,7 @@ Lemma small_shift_amount_3:
Int.ltu i Int64.iwordsize' = true ->
Int64.unsigned (Int64.repr (Int.unsigned i)) = Int.unsigned i.
Proof.
- intros. apply Int.ltu_inv in H. comput (Int.unsigned Int64.iwordsize').
+ intros. apply Int.ltu_inv in H. comput (Int.unsigned Int64.iwordsize').
apply Int64.unsigned_repr. comput Int64.max_unsigned; omega.
Qed.
@@ -588,12 +588,12 @@ Proof.
econstructor; eauto. simpl; rewrite E; auto.
- destruct (Int64.ltu i0 Int64.iwordsize) eqn:E; inv SEM.
exploit small_shift_amount_1; eauto. intros [A B].
- econstructor; eauto with cshm. simpl. rewrite A.
+ econstructor; eauto with cshm. simpl. rewrite A.
f_equal; f_equal. unfold Int64.shl', Int64.shl. rewrite B; auto.
- destruct (Int64.ltu i0 (Int64.repr 32)) eqn:E; inv SEM.
- econstructor; eauto with cshm. simpl. rewrite small_shift_amount_2; auto.
-- destruct (Int.ltu i0 Int64.iwordsize') eqn:E; inv SEM.
- econstructor; eauto with cshm. simpl. rewrite E.
+ econstructor; eauto with cshm. simpl. rewrite small_shift_amount_2; auto.
+- destruct (Int.ltu i0 Int64.iwordsize') eqn:E; inv SEM.
+ econstructor; eauto with cshm. simpl. rewrite E.
unfold Int64.shl', Int64.shl. rewrite small_shift_amount_3; auto.
Qed.
@@ -612,9 +612,9 @@ Proof.
unfold Int64.shr', Int64.shr; rewrite B; auto.
unfold Int64.shru', Int64.shru; rewrite B; auto.
- destruct (Int64.ltu i0 (Int64.repr 32)) eqn:E; inv SEM.
- destruct s; inv H0; econstructor; eauto with cshm; simpl; rewrite small_shift_amount_2; auto.
+ destruct s; inv H0; econstructor; eauto with cshm; simpl; rewrite small_shift_amount_2; auto.
- destruct (Int.ltu i0 Int64.iwordsize') eqn:E; inv SEM.
- destruct s; inv H0; econstructor; eauto with cshm; simpl; rewrite E.
+ destruct s; inv H0; econstructor; eauto with cshm; simpl; rewrite E.
unfold Int64.shr', Int64.shr; rewrite small_shift_amount_3; auto.
unfold Int64.shru', Int64.shru; rewrite small_shift_amount_3; auto.
Qed.
@@ -622,7 +622,7 @@ Qed.
Lemma make_cmp_correct:
forall cmp a tya b tyb c va vb v e le m,
sem_cmp cmp va tya vb tyb m = Some v ->
- make_cmp cmp a tya b tyb = OK c ->
+ make_cmp cmp a tya b tyb = OK c ->
eval_expr ge e le m a va ->
eval_expr ge e le m b vb ->
eval_expr ge e le m c v.
@@ -632,38 +632,38 @@ Proof.
- inv MAKE. destruct (Val.cmpu_bool (Mem.valid_pointer m) cmp va vb) as [bv|] eqn:E;
simpl in SEM; inv SEM.
econstructor; eauto. simpl. unfold Val.cmpu. rewrite E. auto.
-- inv MAKE. destruct vb; try discriminate.
+- inv MAKE. destruct vb; try discriminate.
set (vb := Vint (Int.repr (Int64.unsigned i))) in *.
destruct (Val.cmpu_bool (Mem.valid_pointer m) cmp va vb) as [bv|] eqn:E;
simpl in SEM; inv SEM.
- econstructor; eauto with cshm. simpl. change (Vint (Int64.loword i)) with vb.
+ econstructor; eauto with cshm. simpl. change (Vint (Int64.loword i)) with vb.
unfold Val.cmpu. rewrite E. auto.
-- inv MAKE. destruct va; try discriminate.
+- inv MAKE. destruct va; try discriminate.
set (va := Vint (Int.repr (Int64.unsigned i))) in *.
destruct (Val.cmpu_bool (Mem.valid_pointer m) cmp va vb) as [bv|] eqn:E;
simpl in SEM; inv SEM.
- econstructor; eauto with cshm. simpl. change (Vint (Int64.loword i)) with va.
+ econstructor; eauto with cshm. simpl. change (Vint (Int64.loword i)) with va.
unfold Val.cmpu. rewrite E. auto.
- eapply make_binarith_correct; eauto; intros; auto.
Qed.
Lemma transl_unop_correct:
- forall op a tya c va v e le m,
+ forall op a tya c va v e le m,
transl_unop op a tya = OK c ->
sem_unary_operation op va tya m = Some v ->
eval_expr ge e le m a va ->
eval_expr ge e le m c v.
Proof.
intros. destruct op; simpl in *.
- eapply make_notbool_correct; eauto.
- eapply make_notint_correct; eauto.
+ eapply make_notbool_correct; eauto.
+ eapply make_notint_correct; eauto.
eapply make_neg_correct; eauto.
eapply make_absfloat_correct; eauto.
Qed.
Lemma transl_binop_correct:
forall op a tya b tyb c va vb v e le m,
- transl_binop ce op a tya b tyb = OK c ->
+ transl_binop ce op a tya b tyb = OK c ->
sem_binary_operation ce op va tya vb tyb m = Some v ->
eval_expr ge e le m a va ->
eval_expr ge e le m b vb ->
@@ -686,7 +686,7 @@ Proof.
eapply make_cmp_correct; eauto.
eapply make_cmp_correct; eauto.
eapply make_cmp_correct; eauto.
-Qed.
+Qed.
Lemma make_load_correct:
forall addr ty code b ofs v e le m,
@@ -696,7 +696,7 @@ Lemma make_load_correct:
eval_expr ge e le m code v.
Proof.
unfold make_load; intros until m; intros MKLOAD EVEXP DEREF.
- inv DEREF.
+ inv DEREF.
(* scalar *)
rewrite H in MKLOAD. inv MKLOAD. apply eval_Eload with (Vptr b ofs); auto.
(* by reference *)
@@ -713,16 +713,16 @@ Lemma make_memcpy_correct:
access_mode ty = By_copy ->
step ge (State f (make_memcpy ce dst src ty) k e le m) E0 (State f Sskip k e le m').
Proof.
- intros. inv H1; try congruence.
- unfold make_memcpy. change le with (set_optvar None Vundef le) at 2.
+ intros. inv H1; try congruence.
+ unfold make_memcpy. change le with (set_optvar None Vundef le) at 2.
econstructor.
- econstructor. eauto. econstructor. eauto. constructor.
- econstructor; eauto.
+ econstructor. eauto. econstructor. eauto. constructor.
+ econstructor; eauto.
apply alignof_blockcopy_1248.
apply sizeof_pos.
apply sizeof_alignof_blockcopy_compat.
Qed.
-
+
Lemma make_store_correct:
forall addr ty rhs code e le m b ofs v m' f k,
make_store ce addr ty rhs = OK code ->
@@ -735,10 +735,10 @@ Proof.
inversion ASSIGN; subst.
(* nonvolatile scalar *)
rewrite H in MKSTORE; inv MKSTORE.
- econstructor; eauto.
+ econstructor; eauto.
(* by copy *)
- rewrite H in MKSTORE; inv MKSTORE.
- eapply make_memcpy_correct; eauto.
+ rewrite H in MKSTORE; inv MKSTORE.
+ eapply make_memcpy_correct; eauto.
Qed.
End CONSTRUCTORS.
@@ -813,19 +813,19 @@ Proof.
match x, y with
| (b1, ty), (b2, sz) => b2 = b1 /\ sz = sizeof ge ty
end).
- assert (list_forall2
+ assert (list_forall2
(fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y))
(PTree.elements e) (PTree.elements te)).
apply PTree.elements_canonical_order.
intros id [b ty] GET. exists (b, sizeof ge ty); split. eapply me_local; eauto. red; auto.
intros id [b sz] GET. exploit me_local_inv; eauto. intros [ty EQ].
- exploit me_local; eauto. intros EQ1.
+ exploit me_local; eauto. intros EQ1.
exists (b, ty); split. auto. red; split; congruence.
unfold blocks_of_env, Clight.blocks_of_env.
- generalize H0. induction 1. auto.
+ generalize H0. induction 1. auto.
simpl. f_equal; auto.
- unfold block_of_binding, Clight.block_of_binding.
+ unfold block_of_binding, Clight.block_of_binding.
destruct a1 as [id1 [blk1 ty1]]. destruct b1 as [id2 [blk2 sz2]].
simpl in *. destruct H1 as [A [B C]]. congruence.
Qed.
@@ -867,19 +867,19 @@ Proof.
constructor.
(* me_local *)
intros until ty0. repeat rewrite PTree.gsspec.
- destruct (peq id0 id); intros. congruence. eapply me_local; eauto.
+ destruct (peq id0 id); intros. congruence. eapply me_local; eauto.
(* me_local_inv *)
- intros until sz. repeat rewrite PTree.gsspec.
- destruct (peq id0 id); intros. exists ty; congruence. eapply me_local_inv; eauto.
+ intros until sz. repeat rewrite PTree.gsspec.
+ destruct (peq id0 id); intros. exists ty; congruence. eapply me_local_inv; eauto.
intros [te2 [ALLOC MENV]].
exists te2; split. econstructor; eauto. auto.
-Qed.
+Qed.
Lemma create_undef_temps_match:
forall temps,
create_undef_temps (map fst temps) = Clight.create_undef_temps temps.
Proof.
- induction temps; simpl. auto.
+ induction temps; simpl. auto.
destruct a as [id ty]. simpl. decEq. auto.
Qed.
@@ -889,8 +889,8 @@ Lemma bind_parameter_temps_match:
bind_parameters (map fst vars) vals le1 = Some le2.
Proof.
induction vars; simpl; intros.
- destruct vals; inv H. auto.
- destruct a as [id ty]. destruct vals; try discriminate. auto.
+ destruct vals; inv H. auto.
+ destruct a as [id ty]. destruct vals; try discriminate. auto.
Qed.
(** * Proof of semantic preservation *)
@@ -909,9 +909,9 @@ Qed.
>>
Left: evaluation of r-value expression [a] in Clight.
Right: evaluation of its translation [ta] in Csharpminor.
- Top (precondition): matching between environments [e], [te],
+ Top (precondition): matching between environments [e], [te],
plus well-typedness of expression [a].
- Bottom (postcondition): the result values [v]
+ Bottom (postcondition): the result values [v]
are identical in both evaluations.
We state these diagrams as the following properties, parameterized
@@ -947,7 +947,7 @@ Proof.
(* temp var *)
constructor; auto.
(* addrof *)
- simpl in TR. auto.
+ simpl in TR. auto.
(* unop *)
eapply transl_unop_correct; eauto.
(* binop *)
@@ -960,21 +960,21 @@ Proof.
apply make_intconst_correct.
(* rvalue out of lvalue *)
exploit transl_expr_lvalue; eauto. intros [tb [TRLVAL MKLOAD]].
- eapply make_load_correct; eauto.
+ eapply make_load_correct; eauto.
(* var local *)
exploit (me_local _ _ MENV); eauto. intros EQ.
econstructor. eapply eval_var_addr_local. eauto.
(* var global *)
- econstructor. eapply eval_var_addr_global.
+ econstructor. eapply eval_var_addr_global.
eapply match_env_globals; eauto.
rewrite symbols_preserved. auto.
(* deref *)
- simpl in TR. eauto.
+ simpl in TR. eauto.
(* field struct *)
change (prog_comp_env prog) with (genv_cenv ge) in EQ0.
unfold make_field_access in EQ0; rewrite H1, H2 in EQ0; monadInv EQ0.
eapply eval_Ebinop; eauto.
- apply make_intconst_correct.
+ apply make_intconst_correct.
simpl. congruence.
(* field union *)
unfold make_field_access in EQ0; rewrite H1 in EQ0; monadInv EQ0.
@@ -1003,8 +1003,8 @@ Lemma transl_arglist_correct:
Proof.
induction 1; intros.
monadInv H. constructor.
- monadInv H2. constructor.
- eapply make_cast_correct; eauto. eapply transl_expr_correct; eauto. auto.
+ monadInv H2. constructor.
+ eapply make_cast_correct; eauto. eapply transl_expr_correct; eauto. auto.
Qed.
Lemma typlist_of_arglist_eq:
@@ -1026,9 +1026,9 @@ End EXPR.
<<
I
S1 ------- R1
- | |
+ | |
t | + | t
- v v
+ v v
S2 ------- R2
I I
>>
@@ -1047,8 +1047,8 @@ Lemma match_transl_step:
match_transl (Sblock ts) tk ts' tk' ->
star step tge (State f ts' tk' te le m) E0 (State f ts (Kblock tk) te le m).
Proof.
- intros. inv H.
- apply star_one. constructor.
+ intros. inv H.
+ apply star_one. constructor.
apply star_refl.
Qed.
@@ -1084,7 +1084,7 @@ Inductive match_cont: type -> nat -> nat -> Clight.cont -> Csharpminor.cont -> P
transl_function ge f = OK tf ->
match_env e te ->
match_cont (Clight.fn_return f) nbrk' ncnt' k tk ->
- match_cont tyret nbrk ncnt
+ match_cont tyret nbrk ncnt
(Clight.Kcall id f e le k)
(Kcall id tf te le tk).
@@ -1107,7 +1107,7 @@ Inductive match_states: Clight.state -> Csharpminor.state -> Prop :=
match_states (Clight.Callstate fd args k m)
(Callstate tfd args tk m)
| match_returnstate:
- forall res k m tk
+ forall res k m tk
(MK: match_cont Tvoid 0%nat 0%nat k tk),
match_states (Clight.Returnstate res k m)
(Returnstate res tk m).
@@ -1119,7 +1119,7 @@ Remark match_states_skip:
match_cont (Clight.fn_return f) nbrk ncnt k tk ->
match_states (Clight.State f Clight.Sskip k e le m) (State tf Sskip tk te le m).
Proof.
- intros. econstructor; eauto. simpl; reflexivity. constructor.
+ intros. econstructor; eauto. simpl; reflexivity. constructor.
Qed.
(** Commutation between label resolution and compilation *)
@@ -1168,13 +1168,13 @@ Proof.
(* builtin *)
auto.
(* seq *)
- exploit (transl_find_label s0 nbrk ncnt (Clight.Kseq s1 k)); eauto. constructor; eauto.
+ exploit (transl_find_label s0 nbrk ncnt (Clight.Kseq s1 k)); eauto. constructor; eauto.
destruct (Clight.find_label lbl s0 (Clight.Kseq s1 k)) as [[s' k'] | ].
intros [ts' [tk' [nbrk' [ncnt' [A [B C]]]]]].
rewrite A. exists ts'; exists tk'; exists nbrk'; exists ncnt'; auto.
intro. rewrite H. eapply transl_find_label; eauto.
(* ifthenelse *)
- exploit (transl_find_label s0); eauto.
+ exploit (transl_find_label s0); eauto.
destruct (Clight.find_label lbl s0 k) as [[s' k'] | ].
intros [ts' [tk' [nbrk' [ncnt' [A [B C]]]]]].
rewrite A. exists ts'; exists tk'; exists nbrk'; exists ncnt'; auto.
@@ -1185,20 +1185,20 @@ Proof.
intros [ts' [tk' [nbrk' [ncnt' [A [B C]]]]]].
rewrite A. exists ts'; exists tk'; exists nbrk'; exists ncnt'; auto.
intro. rewrite H.
- eapply transl_find_label; eauto. econstructor; eauto.
+ eapply transl_find_label; eauto. econstructor; eauto.
(* break *)
auto.
(* continue *)
auto.
(* return *)
- simpl in TR. destruct o; monadInv TR. auto. auto.
+ simpl in TR. destruct o; monadInv TR. auto. auto.
(* switch *)
assert (exists b, ts = Sblock (Sswitch b x x0)).
{ destruct (classify_switch (typeof e)); inv EQ2; econstructor; eauto. }
destruct H as [b EQ3]; rewrite EQ3; simpl.
- eapply transl_find_label_ls with (k := Clight.Kswitch k); eauto. econstructor; eauto.
+ eapply transl_find_label_ls with (k := Clight.Kswitch k); eauto. econstructor; eauto.
(* label *)
- destruct (ident_eq lbl l).
+ destruct (ident_eq lbl l).
exists x; exists tk; exists nbrk; exists ncnt; auto.
eapply transl_find_label; eauto.
(* goto *)
@@ -1208,7 +1208,7 @@ Proof.
(* nil *)
auto.
(* cons *)
- exploit (transl_find_label s nbrk ncnt (Clight.Kseq (seq_of_labeled_statement l) k)); eauto.
+ exploit (transl_find_label s nbrk ncnt (Clight.Kseq (seq_of_labeled_statement l) k)); eauto.
econstructor; eauto. apply transl_lbl_stmt_2; eauto.
destruct (Clight.find_label lbl s (Clight.Kseq (seq_of_labeled_statement l) k)) as [[s' k'] | ].
intros [ts' [tk' [nbrk' [ncnt' [A [B C]]]]]].
@@ -1228,7 +1228,7 @@ Lemma match_cont_call_cont:
Proof.
induction 1; simpl; auto.
constructor.
- econstructor; eauto.
+ econstructor; eauto.
Qed.
Lemma match_cont_is_call_cont:
@@ -1254,41 +1254,41 @@ Proof.
(* assign *)
monadInv TR.
assert (SAME: ts' = ts /\ tk' = tk).
- inversion MTR. auto.
+ inversion MTR. auto.
subst ts. unfold make_store, make_memcpy in EQ3. destruct (access_mode (typeof a1)); congruence.
destruct SAME; subst ts' tk'.
econstructor; split.
apply plus_one. eapply make_store_correct; eauto.
eapply transl_lvalue_correct; eauto. eapply make_cast_correct; eauto.
- eapply transl_expr_correct; eauto.
+ eapply transl_expr_correct; eauto.
eapply match_states_skip; eauto.
(* set *)
monadInv TR. inv MTR. econstructor; split.
- apply plus_one. econstructor. eapply transl_expr_correct; eauto.
+ apply plus_one. econstructor. eapply transl_expr_correct; eauto.
eapply match_states_skip; eauto.
(* call *)
revert TR. simpl. case_eq (classify_fun (typeof a)); try congruence.
- intros targs tres cc CF TR. monadInv TR. inv MTR.
+ intros targs tres cc CF TR. monadInv TR. inv MTR.
exploit functions_translated; eauto. intros [tfd [FIND TFD]].
rewrite H in CF. simpl in CF. inv CF.
econstructor; split.
- apply plus_one. econstructor; eauto.
+ apply plus_one. econstructor; eauto.
exploit transl_expr_correct; eauto.
exploit transl_arglist_correct; eauto.
- erewrite typlist_of_arglist_eq by eauto.
+ erewrite typlist_of_arglist_eq by eauto.
eapply transl_fundef_sig1; eauto.
rewrite H3. auto.
- econstructor; eauto.
+ econstructor; eauto.
econstructor; eauto.
simpl. auto.
(* builtin *)
- monadInv TR. inv MTR.
+ monadInv TR. inv MTR.
econstructor; split.
- apply plus_one. econstructor.
- eapply transl_arglist_correct; eauto.
+ apply plus_one. econstructor.
+ eapply transl_arglist_correct; eauto.
eapply external_call_symbols_preserved_gen with (ge1 := ge).
exact symbols_preserved. exact public_preserved. exact block_is_volatile_preserved. eauto.
eapply match_states_skip; eauto.
@@ -1296,31 +1296,31 @@ Proof.
(* seq *)
monadInv TR. inv MTR.
econstructor; split.
- apply plus_one. constructor.
- econstructor; eauto. constructor.
+ apply plus_one. constructor.
+ econstructor; eauto. constructor.
econstructor; eauto.
(* skip seq *)
monadInv TR. inv MTR. inv MK.
econstructor; split.
- apply plus_one. apply step_skip_seq.
+ apply plus_one. apply step_skip_seq.
econstructor; eauto. constructor.
(* continue seq *)
monadInv TR. inv MTR. inv MK.
econstructor; split.
- apply plus_one. constructor.
+ apply plus_one. constructor.
econstructor; eauto. simpl. reflexivity. constructor.
(* break seq *)
monadInv TR. inv MTR. inv MK.
econstructor; split.
- apply plus_one. constructor.
+ apply plus_one. constructor.
econstructor; eauto. simpl. reflexivity. constructor.
(* ifthenelse *)
monadInv TR. inv MTR.
- exploit make_boolean_correct; eauto.
+ exploit make_boolean_correct; eauto.
exploit transl_expr_correct; eauto.
intros [v [A B]].
econstructor; split.
@@ -1330,12 +1330,12 @@ Proof.
(* loop *)
monadInv TR.
econstructor; split.
- eapply star_plus_trans. eapply match_transl_step; eauto.
- eapply plus_left. constructor.
+ eapply star_plus_trans. eapply match_transl_step; eauto.
+ eapply plus_left. constructor.
eapply star_left. constructor.
apply star_one. constructor.
reflexivity. reflexivity. traceEq.
- econstructor; eauto. constructor. econstructor; eauto.
+ econstructor; eauto. constructor. econstructor; eauto.
(* skip-or-continue loop *)
assert ((ts' = Sskip \/ ts' = Sexit ncnt) /\ tk' = tk).
@@ -1345,7 +1345,7 @@ Proof.
eapply plus_left.
destruct H0; subst ts'. 2:constructor. constructor.
apply star_one. constructor. traceEq.
- econstructor; eauto. constructor. econstructor; eauto.
+ econstructor; eauto. constructor. econstructor; eauto.
(* break loop1 *)
monadInv TR. inv MTR. inv MK.
@@ -1362,9 +1362,9 @@ Proof.
econstructor; split.
apply plus_one. constructor.
econstructor; eauto.
-Local Opaque ge.
- simpl. rewrite H5; simpl. rewrite H7; simpl. eauto.
- constructor.
+Local Opaque ge.
+ simpl. rewrite H5; simpl. rewrite H7; simpl. eauto.
+ constructor.
(* break loop2 *)
monadInv TR. inv MTR. inv MK.
@@ -1375,21 +1375,21 @@ Local Opaque ge.
eapply match_states_skip; eauto.
(* return none *)
- monadInv TR. inv MTR.
+ monadInv TR. inv MTR.
econstructor; split.
apply plus_one. constructor.
- eapply match_env_free_blocks; eauto.
+ eapply match_env_free_blocks; eauto.
econstructor; eauto.
- eapply match_cont_call_cont. eauto.
+ eapply match_cont_call_cont. eauto.
(* return some *)
- monadInv TR. inv MTR.
+ monadInv TR. inv MTR.
econstructor; split.
apply plus_one. constructor.
eapply make_cast_correct; eauto. eapply transl_expr_correct; eauto.
eapply match_env_free_blocks; eauto.
econstructor; eauto.
- eapply match_cont_call_cont. eauto.
+ eapply match_cont_call_cont. eauto.
(* skip call *)
monadInv TR. inv MTR.
@@ -1405,13 +1405,13 @@ Local Opaque ge.
{ unfold sem_switch_arg in H0.
destruct (classify_switch (typeof a)); inv EQ2; econstructor; split; eauto;
destruct v; inv H0; constructor. }
- destruct E as (b & A & B). subst ts.
+ destruct E as (b & A & B). subst ts.
exploit transl_expr_correct; eauto. intro EV.
econstructor; split.
eapply star_plus_trans. eapply match_transl_step; eauto.
- apply plus_one. econstructor; eauto. traceEq.
+ apply plus_one. econstructor; eauto. traceEq.
econstructor; eauto.
- apply transl_lbl_stmt_2. apply transl_lbl_stmt_1. eauto.
+ apply transl_lbl_stmt_2. apply transl_lbl_stmt_1. eauto.
constructor.
econstructor. eauto.
@@ -1427,29 +1427,29 @@ Local Opaque ge.
(* continue switch *)
monadInv TR. inv MTR. inv MK.
econstructor; split.
- apply plus_one. constructor.
+ apply plus_one. constructor.
econstructor; eauto. simpl. reflexivity. constructor.
(* label *)
- monadInv TR. inv MTR.
+ monadInv TR. inv MTR.
econstructor; split.
- apply plus_one. constructor.
+ apply plus_one. constructor.
econstructor; eauto. constructor.
(* goto *)
monadInv TR. inv MTR.
generalize TRF. unfold transl_function. intro TRF'. monadInv TRF'.
exploit (transl_find_label lbl). eexact EQ. eapply match_cont_call_cont. eauto.
- rewrite H.
+ rewrite H.
intros [ts' [tk'' [nbrk' [ncnt' [A [B C]]]]]].
econstructor; split.
- apply plus_one. constructor. simpl. eexact A.
+ apply plus_one. constructor. simpl. eexact A.
econstructor; eauto. constructor.
(* internal function *)
inv H. monadInv TR. monadInv EQ.
exploit match_cont_is_call_cont; eauto. intros [A B].
- exploit match_env_alloc_variables; eauto.
+ exploit match_env_alloc_variables; eauto.
apply match_env_empty.
intros [te1 [C D]].
econstructor; split.
@@ -1464,17 +1464,17 @@ Local Opaque ge.
constructor.
(* external function *)
- simpl in TR.
+ simpl in TR.
destruct (signature_eq (ef_sig ef) (signature_of_type targs tres cconv)); inv TR.
exploit match_cont_is_call_cont; eauto. intros [A B].
econstructor; split.
- apply plus_one. constructor. eauto.
+ apply plus_one. constructor. eauto.
eapply external_call_symbols_preserved_gen with (ge1 := ge).
exact symbols_preserved. exact public_preserved. exact block_is_volatile_preserved. eauto.
econstructor; eauto.
(* returnstate *)
- inv MK.
+ inv MK.
econstructor; split.
apply plus_one. constructor.
econstructor; eauto. simpl; reflexivity. constructor.
@@ -1492,10 +1492,10 @@ Proof.
change (prog_main prog) with (AST.prog_main (program_of_program prog)).
eapply transform_partial_program2_main; eauto.
assert (funsig tf = signature_of_type Tnil type_int32s cc_default).
- eapply transl_fundef_sig2; eauto.
+ eapply transl_fundef_sig2; eauto.
econstructor; split.
- econstructor; eauto. eapply Genv.init_mem_transf_partial2; eauto.
- econstructor; eauto. constructor; auto. exact I.
+ econstructor; eauto. eapply Genv.init_mem_transf_partial2; eauto.
+ econstructor; eauto. constructor; auto. exact I.
Qed.
Lemma transl_final_states:
diff --git a/cfrontend/Cstrategy.v b/cfrontend/Cstrategy.v
index b082ea56..b3cbacca 100644
--- a/cfrontend/Cstrategy.v
+++ b/cfrontend/Cstrategy.v
@@ -40,7 +40,7 @@ Variable ge: genv.
(** We now formalize a particular strategy for reducing expressions which
is the one implemented by the CompCert compiler. It evaluates effectful
- subexpressions first, in leftmost-innermost order, then finishes
+ subexpressions first, in leftmost-innermost order, then finishes
with the evaluation of the remaining simple expression. *)
(** Simple expressions are defined as follows. *)
@@ -99,7 +99,7 @@ Inductive eval_simple_lvalue: expr -> block -> int -> Prop :=
eval_simple_lvalue (Ederef r ty) b ofs
| esl_field_struct: forall r f ty b ofs id co a delta,
eval_simple_rvalue r (Vptr b ofs) ->
- typeof r = Tstruct id a ->
+ typeof r = Tstruct id a ->
ge.(genv_cenv)!id = Some co ->
field_offset ge f (co_members co) = OK delta ->
eval_simple_lvalue (Efield r f ty) b (Int.add ofs (Int.repr delta))
@@ -402,7 +402,7 @@ Hint Resolve context_compose contextlist_compose.
(** * Safe executions. *)
-(** A state is safe according to the nondeterministic semantics
+(** A state is safe according to the nondeterministic semantics
if it cannot get stuck by doing silent transitions only. *)
Definition safe (s: Csem.state) : Prop :=
@@ -413,8 +413,8 @@ Lemma safe_steps:
forall s s',
safe s -> star Csem.step ge s E0 s' -> safe s'.
Proof.
- intros; red; intros.
- eapply H. eapply star_trans; eauto.
+ intros; red; intros.
+ eapply H. eapply star_trans; eauto.
Qed.
Lemma star_safe:
@@ -433,7 +433,7 @@ Proof.
intros. eapply star_plus_trans; eauto. apply H1. eapply safe_steps; eauto. auto.
Qed.
-Require Import Classical.
+Require Import Classical.
Lemma safe_imm_safe:
forall f C a k e m K,
@@ -442,10 +442,10 @@ Lemma safe_imm_safe:
imm_safe ge e K a m.
Proof.
intros. destruct (classic (imm_safe ge e K a m)); auto.
- destruct (H Stuckstate).
+ destruct (H Stuckstate).
apply star_one. left. econstructor; eauto.
- destruct H2 as [r F]. inv F.
- destruct H2 as [t [s' S]]. inv S. inv H2. inv H2.
+ destruct H2 as [r F]. inv F.
+ destruct H2 as [t [s' S]]. inv S. inv H2. inv H2.
Qed.
(** Safe expressions are well-formed with respect to l-values and r-values. *)
@@ -649,11 +649,11 @@ Proof.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
destruct e1; auto. intros. elim (H0 a m); auto.
- intros. elim (H0 a m); auto.
+ intros. elim (H0 a m); auto.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
- red; intros. destruct (C a); auto.
- red; intros. destruct e1; auto. elim (H0 a m); auto.
+ red; intros. destruct (C a); auto.
+ red; intros. destruct e1; auto. elim (H0 a m); auto.
Qed.
Lemma imm_safe_inv:
@@ -666,7 +666,7 @@ Lemma imm_safe_inv:
end.
Proof.
destruct invert_expr_context as [A B].
- intros. inv H.
+ intros. inv H.
auto.
auto.
assert (invert_expr_prop (C e0) m).
@@ -733,7 +733,7 @@ Ltac FinishL := apply star_one; left; apply step_lred; eauto; simpl; try (econst
Steps H2 (fun x => C(Ebinop op (Eval v1 (typeof r1)) x ty)).
FinishR.
(* cast *)
- Steps H0 (fun x => C(Ecast x ty)). FinishR.
+ Steps H0 (fun x => C(Ecast x ty)). FinishR.
(* sizeof *)
FinishR.
(* alignof *)
@@ -741,11 +741,11 @@ Ltac FinishL := apply star_one; left; apply step_lred; eauto; simpl; try (econst
(* loc *)
apply star_refl.
(* var local *)
- FinishL.
+ FinishL.
(* var global *)
FinishL. apply red_var_global; auto.
(* deref *)
- Steps H0 (fun x => C(Ederef x ty)). FinishL.
+ Steps H0 (fun x => C(Ederef x ty)). FinishL.
(* field struct *)
Steps H0 (fun x => C(Efield x f0 ty)). rewrite H1 in *. FinishL.
(* field union *)
@@ -796,13 +796,13 @@ Ltac StepL REC C' a :=
let b := fresh "b" in let ofs := fresh "ofs" in
let E := fresh "E" in let S := fresh "SAFE" in
exploit (REC LV C'); eauto; intros [b [ofs E]];
- assert (S: safe (ExprState f (C' (Eloc b ofs (typeof a))) k e m)) by
+ assert (S: safe (ExprState f (C' (Eloc b ofs (typeof a))) k e m)) by
(eapply (eval_simple_lvalue_safe C'); eauto);
simpl in S.
Ltac StepR REC C' a :=
let v := fresh "v" in let E := fresh "E" in let S := fresh "SAFE" in
exploit (REC RV C'); eauto; intros [v E];
- assert (S: safe (ExprState f (C' (Eval v (typeof a))) k e m)) by
+ assert (S: safe (ExprState f (C' (Eval v (typeof a))) k e m)) by
(eapply (eval_simple_rvalue_safe C'); eauto);
simpl in S.
@@ -857,22 +857,22 @@ Ltac StepR REC C' a :=
Qed.
Lemma simple_can_eval_rval:
- forall r C,
+ forall r C,
simple r = true -> context RV RV C -> safe (ExprState f (C r) k e m) ->
exists v, eval_simple_rvalue e m r v
/\ safe (ExprState f (C (Eval v (typeof r))) k e m).
Proof.
- intros. exploit (simple_can_eval r RV); eauto. intros [v A].
+ intros. exploit (simple_can_eval r RV); eauto. intros [v A].
exists v; split; auto. eapply eval_simple_rvalue_safe; eauto.
Qed.
Lemma simple_can_eval_lval:
- forall l C,
+ forall l C,
simple l = true -> context LV RV C -> safe (ExprState f (C l) k e m) ->
exists b, exists ofs, eval_simple_lvalue e m l b ofs
/\ safe (ExprState f (C (Eloc b ofs (typeof l))) k e m).
Proof.
- intros. exploit (simple_can_eval l LV); eauto. intros [b [ofs A]].
+ intros. exploit (simple_can_eval l LV); eauto. intros [b [ofs A]].
exists b; exists ofs; split; auto. eapply eval_simple_lvalue_safe; eauto.
Qed.
@@ -892,7 +892,7 @@ Inductive eval_simple_list': exprlist -> list val -> Prop :=
Lemma eval_simple_list_implies:
forall rl tyl vl,
- eval_simple_list e m rl tyl vl ->
+ eval_simple_list e m rl tyl vl ->
exists vl', cast_arguments (rval_list vl' rl) tyl vl /\ eval_simple_list' rl vl'.
Proof.
induction 1.
@@ -908,9 +908,9 @@ Lemma can_eval_simple_list:
cast_arguments (rval_list vl rl) tyl vl' ->
eval_simple_list e m rl tyl vl'.
Proof.
- induction 1; simpl; intros.
+ induction 1; simpl; intros.
inv H. constructor.
- inv H1. econstructor; eauto.
+ inv H1. econstructor; eauto.
Qed.
Fixpoint exprlist_app (rl1 rl2: exprlist) : exprlist :=
@@ -924,7 +924,7 @@ Lemma exprlist_app_assoc:
exprlist_app (exprlist_app rl1 rl2) rl3 =
exprlist_app rl1 (exprlist_app rl2 rl3).
Proof.
- induction rl1; auto. simpl. congruence.
+ induction rl1; auto. simpl. congruence.
Qed.
Inductive contextlist' : (exprlist -> expr) -> Prop :=
@@ -939,9 +939,9 @@ Lemma exprlist_app_context:
forall rl1 rl2,
contextlist RV (fun x => exprlist_app rl1 (Econs x rl2)).
Proof.
- induction rl1; simpl; intros.
+ induction rl1; simpl; intros.
apply ctx_list_head. constructor.
- apply ctx_list_tail. auto.
+ apply ctx_list_tail. auto.
Qed.
Lemma contextlist'_head:
@@ -949,14 +949,14 @@ Lemma contextlist'_head:
contextlist' C ->
context RV RV (fun x => C (Econs x rl)).
Proof.
- intros. inv H.
+ intros. inv H.
set (C' := fun x => Ecall r1 (exprlist_app rl0 (Econs x rl)) ty).
assert (context RV RV C'). constructor. apply exprlist_app_context.
- change (context RV RV (fun x => C0 (C' x))).
+ change (context RV RV (fun x => C0 (C' x))).
eapply context_compose; eauto.
set (C' := fun x => Ebuiltin ef tyargs (exprlist_app rl0 (Econs x rl)) ty).
assert (context RV RV C'). constructor. apply exprlist_app_context.
- change (context RV RV (fun x => C0 (C' x))).
+ change (context RV RV (fun x => C0 (C' x))).
eapply context_compose; eauto.
Qed.
@@ -965,14 +965,14 @@ Lemma contextlist'_tail:
contextlist' C ->
contextlist' (fun x => C (Econs r1 x)).
Proof.
- intros. inv H.
+ intros. inv H.
replace (fun x => C0 (Ecall r0 (exprlist_app rl0 (Econs r1 x)) ty))
with (fun x => C0 (Ecall r0 (exprlist_app (exprlist_app rl0 (Econs r1 Enil)) x) ty)).
- constructor. auto.
+ constructor. auto.
apply extensionality; intros. f_equal. f_equal. apply exprlist_app_assoc.
replace (fun x => C0 (Ebuiltin ef tyargs (exprlist_app rl0 (Econs r1 x)) ty))
with (fun x => C0 (Ebuiltin ef tyargs (exprlist_app (exprlist_app rl0 (Econs r1 Enil)) x) ty)).
- constructor. auto.
+ constructor. auto.
apply extensionality; intros. f_equal. f_equal. apply exprlist_app_assoc.
Qed.
@@ -984,7 +984,7 @@ Lemma eval_simple_list_steps:
star Csem.step ge (ExprState f (C rl) k e m)
E0 (ExprState f (C (rval_list vl rl)) k e m).
Proof.
- induction 1; intros.
+ induction 1; intros.
(* nil *)
apply star_refl.
(* cons *)
@@ -1003,7 +1003,7 @@ Lemma simple_list_can_eval:
Proof.
induction rl; intros.
econstructor; constructor.
- simpl in H. destruct (andb_prop _ _ H).
+ simpl in H. destruct (andb_prop _ _ H).
exploit (simple_can_eval r1 RV (fun x => C(Econs x rl))); eauto.
intros [v1 EV1].
exploit (IHrl (fun x => C(Econs (Eval v1 (typeof r1)) x))); eauto.
@@ -1015,7 +1015,7 @@ Qed.
Lemma rval_list_all_values:
forall vl rl, exprlist_all_values (rval_list vl rl).
Proof.
- induction vl; simpl; intros. auto.
+ induction vl; simpl; intros. auto.
destruct rl; simpl; auto.
Qed.
@@ -1105,13 +1105,13 @@ Ltac Base :=
(* comma *)
Kind. Rec H RV C (fun x => Ecomma x r2 ty). Base.
(* call *)
- Kind. Rec H RV C (fun x => Ecall x rargs ty).
+ Kind. Rec H RV C (fun x => Ecall x rargs ty).
destruct (H0 (fun x => C (Ecall r1 x ty))) as [A | [C' [a' [D [A B]]]]].
eapply contextlist'_call with (C := C) (rl0 := Enil). auto. auto.
Base.
right; exists (fun x => Ecall r1 (C' x) ty); exists a'. rewrite D; simpl; auto.
(* builtin *)
- Kind.
+ Kind.
destruct (H (fun x => C (Ebuiltin ef tyargs x ty))) as [A | [C' [a' [D [A B]]]]].
eapply contextlist'_builtin with (C := C) (rl0 := Enil). auto. auto.
Base.
@@ -1123,7 +1123,7 @@ Ltac Base :=
eapply contextlist'_head; eauto. auto.
destruct (H0 (fun x => C (Econs r1 x))) as [A' | [C' [a' [A' [B D]]]]].
eapply contextlist'_tail; eauto. auto.
- rewrite A; rewrite A'; auto.
+ rewrite A; rewrite A'; auto.
right; exists (fun x => Econs r1 (C' x)); exists a'. rewrite A'; eauto.
right; exists (fun x => Econs (C' x) rl); exists a'. rewrite A; eauto.
Qed.
@@ -1145,43 +1145,43 @@ Lemma estep_simulation:
forall S t S',
estep S t S' -> plus Csem.step ge S t S'.
Proof.
- intros. inv H.
+ intros. inv H.
(* simple *)
exploit eval_simple_rvalue_steps; eauto. simpl; intros STEPS.
- exploit star_inv; eauto. intros [[EQ1 EQ2] | A]; eauto.
+ exploit star_inv; eauto. intros [[EQ1 EQ2] | A]; eauto.
inversion EQ1. rewrite <- H2 in H1; contradiction.
(* valof volatile *)
- eapply plus_right.
+ eapply plus_right.
eapply eval_simple_lvalue_steps with (C := fun x => C(Evalof x (typeof l))); eauto.
left. apply step_rred; eauto. econstructor; eauto. auto.
(* seqand true *)
eapply plus_right.
- eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqand x r2 ty)); eauto.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqand x r2 ty)); eauto.
left. apply step_rred; eauto. apply red_seqand_true; auto. traceEq.
(* seqand false *)
eapply plus_right.
- eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqand x r2 ty)); eauto.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqand x r2 ty)); eauto.
left. apply step_rred; eauto. apply red_seqand_false; auto. traceEq.
(* seqor true *)
eapply plus_right.
- eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqor x r2 ty)); eauto.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqor x r2 ty)); eauto.
left. apply step_rred; eauto. apply red_seqor_true; auto. traceEq.
(* seqor false *)
eapply plus_right.
- eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqor x r2 ty)); eauto.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqor x r2 ty)); eauto.
left. apply step_rred; eauto. apply red_seqor_false; auto. traceEq.
(* condition *)
eapply plus_right.
eapply eval_simple_rvalue_steps with (C := fun x => C(Econdition x r2 r3 ty)); eauto.
- left; apply step_rred; eauto. constructor; auto. auto.
+ left; apply step_rred; eauto. constructor; auto. auto.
(* assign *)
eapply star_plus_trans.
eapply eval_simple_lvalue_steps with (C := fun x => C(Eassign x r (typeof l))); eauto.
eapply plus_right.
eapply eval_simple_rvalue_steps with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto.
left; apply step_rred; eauto. econstructor; eauto.
- reflexivity. auto.
-(* assignop *)
+ reflexivity. auto.
+(* assignop *)
eapply star_plus_trans.
eapply eval_simple_lvalue_steps with (C := fun x => C(Eassignop op x r tyres (typeof l))); eauto.
eapply star_plus_trans.
@@ -1189,9 +1189,9 @@ Proof.
eapply plus_left.
left; apply step_rred; auto. econstructor; eauto.
eapply star_left.
- left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto. econstructor; eauto.
- apply star_one.
- left; apply step_rred; auto. econstructor; eauto.
+ left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto. econstructor; eauto.
+ apply star_one.
+ left; apply step_rred; auto. econstructor; eauto.
reflexivity. reflexivity. reflexivity. traceEq.
(* assignop stuck *)
eapply star_plus_trans.
@@ -1202,8 +1202,8 @@ Proof.
left; apply step_rred; auto. econstructor; eauto.
destruct (sem_binary_operation ge op v1 (typeof l) v2 (typeof r) m) as [v3|] eqn:?.
eapply star_left.
- left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto. econstructor; eauto.
- apply star_one.
+ left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto. econstructor; eauto.
+ apply star_one.
left; eapply step_stuck; eauto.
red; intros. exploit imm_safe_inv; eauto. simpl. intros [v4' [m' [t' [A [B D]]]]].
rewrite B in H4. eelim H4; eauto.
@@ -1220,12 +1220,12 @@ Proof.
left; apply step_rred; auto. econstructor; eauto.
eapply star_left.
left; apply step_rred with (C := fun x => C (Ecomma (Eassign (Eloc b ofs (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto.
- econstructor. instantiate (1 := v2). destruct id; assumption.
+ econstructor. instantiate (1 := v2). destruct id; assumption.
eapply star_left.
left; apply step_rred with (C := fun x => C (Ecomma x (Eval v1 (typeof l)) (typeof l))); eauto.
econstructor; eauto.
apply star_one.
- left; apply step_rred; auto. econstructor; eauto.
+ left; apply step_rred; auto. econstructor; eauto.
reflexivity. reflexivity. reflexivity. traceEq.
(* postincr stuck *)
eapply star_plus_trans.
@@ -1248,7 +1248,7 @@ Proof.
apply star_one.
left; eapply step_stuck with (C := fun x => C (Ecomma (Eassign (Eloc b ofs (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto.
red; intros. exploit imm_safe_inv; eauto. simpl. intros [v2 A]. congruence.
- reflexivity.
+ reflexivity.
traceEq.
(* comma *)
eapply plus_right.
@@ -1260,7 +1260,7 @@ Proof.
left; apply step_rred; eauto. econstructor; eauto. auto.
(* call *)
exploit eval_simple_list_implies; eauto. intros [vl' [A B]].
- eapply star_plus_trans.
+ eapply star_plus_trans.
eapply eval_simple_rvalue_steps with (C := fun x => C(Ecall x rargs ty)); eauto.
eapply plus_right.
eapply eval_simple_list_steps with (C := fun x => C(Ecall (Eval vf (typeof rf)) x ty)); eauto.
@@ -1273,7 +1273,7 @@ Proof.
eapply eval_simple_list_steps with (C := fun x => C(Ebuiltin ef tyargs x ty)); eauto.
eapply contextlist'_builtin with (rl0 := Enil); auto.
left; apply Csem.step_rred; eauto. econstructor; eauto.
- traceEq.
+ traceEq.
Qed.
Lemma can_estep:
@@ -1285,34 +1285,34 @@ Proof.
intros. destruct (decompose_topexpr f k e m a H) as [A | [C [b [P [Q R]]]]].
(* simple expr *)
exploit (simple_can_eval f k e m a RV (fun x => x)); auto. intros [v P].
- econstructor; econstructor; eapply step_expr; eauto.
+ econstructor; econstructor; eapply step_expr; eauto.
(* side effect *)
- clear H0. subst a. red in Q. destruct b; try contradiction.
+ clear H0. subst a. red in Q. destruct b; try contradiction.
(* valof volatile *)
destruct Q.
exploit (simple_can_eval_lval f k e m b (fun x => C(Evalof x ty))); eauto.
intros [b1 [ofs [E1 S1]]].
exploit safe_inv. eexact S1. eauto. simpl. intros [A [t [v B]]].
- econstructor; econstructor; eapply step_rvalof_volatile; eauto. congruence.
+ econstructor; econstructor; eapply step_rvalof_volatile; eauto. congruence.
(* seqand *)
exploit (simple_can_eval_rval f k e m b1 (fun x => C(Eseqand x b2 ty))); eauto.
intros [v1 [E1 S1]].
exploit safe_inv. eexact S1. eauto. simpl. intros [b BV].
destruct b.
- econstructor; econstructor; eapply step_seqand_true; eauto.
- econstructor; econstructor; eapply step_seqand_false; eauto.
+ econstructor; econstructor; eapply step_seqand_true; eauto.
+ econstructor; econstructor; eapply step_seqand_false; eauto.
(* seqor *)
exploit (simple_can_eval_rval f k e m b1 (fun x => C(Eseqor x b2 ty))); eauto.
intros [v1 [E1 S1]].
exploit safe_inv. eexact S1. eauto. simpl. intros [b BV].
destruct b.
- econstructor; econstructor; eapply step_seqor_true; eauto.
- econstructor; econstructor; eapply step_seqor_false; eauto.
+ econstructor; econstructor; eapply step_seqor_true; eauto.
+ econstructor; econstructor; eapply step_seqor_false; eauto.
(* condition *)
exploit (simple_can_eval_rval f k e m b1 (fun x => C(Econdition x b2 b3 ty))); eauto.
intros [v1 [E1 S1]].
exploit safe_inv. eexact S1. eauto. simpl. intros [b BV].
- econstructor; econstructor. eapply step_condition; eauto.
+ econstructor; econstructor. eapply step_condition; eauto.
(* assign *)
destruct Q.
exploit (simple_can_eval_lval f k e m b1 (fun x => C(Eassign x b2 ty))); eauto.
@@ -1320,7 +1320,7 @@ Proof.
exploit (simple_can_eval_rval f k e m b2 (fun x => C(Eassign (Eloc b ofs (typeof b1)) x ty))); eauto.
intros [v [E2 S2]].
exploit safe_inv. eexact S2. eauto. simpl. intros [v' [m' [t [A [B D]]]]].
- econstructor; econstructor; eapply step_assign; eauto.
+ econstructor; econstructor; eapply step_assign; eauto.
(* assignop *)
destruct Q.
exploit (simple_can_eval_lval f k e m b1 (fun x => C(Eassignop op x b2 tyres ty))); eauto.
@@ -1333,11 +1333,11 @@ Proof.
destruct (classic (exists t2, exists m', assign_loc ge (typeof b1) m b ofs v4 t2 m')).
destruct H2 as [t2 [m' D]].
econstructor; econstructor; eapply step_assignop; eauto.
- econstructor; econstructor; eapply step_assignop_stuck; eauto.
+ econstructor; econstructor; eapply step_assignop_stuck; eauto.
rewrite Heqo. rewrite Heqo0. intros; red; intros. elim H2. exists t2; exists m'; auto.
- econstructor; econstructor; eapply step_assignop_stuck; eauto.
+ econstructor; econstructor; eapply step_assignop_stuck; eauto.
rewrite Heqo. rewrite Heqo0. auto.
- econstructor; econstructor; eapply step_assignop_stuck; eauto.
+ econstructor; econstructor; eapply step_assignop_stuck; eauto.
rewrite Heqo. auto.
(* postincr *)
exploit (simple_can_eval_lval f k e m b (fun x => C(Epostincr id x ty))); eauto.
@@ -1348,11 +1348,11 @@ Proof.
destruct (classic (exists t2, exists m', assign_loc ge ty m b1 ofs v3 t2 m')).
destruct H0 as [t2 [m' D]].
econstructor; econstructor; eapply step_postincr; eauto.
- econstructor; econstructor; eapply step_postincr_stuck; eauto.
+ econstructor; econstructor; eapply step_postincr_stuck; eauto.
rewrite Heqo. rewrite Heqo0. intros; red; intros. elim H0. exists t2; exists m'; congruence.
- econstructor; econstructor; eapply step_postincr_stuck; eauto.
+ econstructor; econstructor; eapply step_postincr_stuck; eauto.
rewrite Heqo. rewrite Heqo0. auto.
- econstructor; econstructor; eapply step_postincr_stuck; eauto.
+ econstructor; econstructor; eapply step_postincr_stuck; eauto.
rewrite Heqo. auto.
(* comma *)
exploit (simple_can_eval_rval f k e m b1 (fun x => C(Ecomma x b2 ty))); eauto.
@@ -1370,9 +1370,9 @@ Proof.
exploit safe_inv. 2: eapply leftcontext_context; eexact R.
eapply safe_steps. eexact S1.
apply (eval_simple_list_steps f k e m rargs vl E2 C'); auto.
- simpl. intros X. exploit X. eapply rval_list_all_values.
+ simpl. intros X. exploit X. eapply rval_list_all_values.
intros [tyargs [tyres [cconv [fd [vargs [P [Q [U V]]]]]]]].
- econstructor; econstructor; eapply step_call; eauto. eapply can_eval_simple_list; eauto.
+ econstructor; econstructor; eapply step_call; eauto. eapply can_eval_simple_list; eauto.
(* builtin *)
pose (C' := fun x => C(Ebuiltin ef tyargs x ty)).
assert (contextlist' C'). unfold C'; eapply contextlist'_builtin with (rl0 := Enil); auto.
@@ -1384,7 +1384,7 @@ Proof.
simpl. intros X. exploit X. eapply rval_list_all_values.
intros [vargs [t [vres [m' [U V]]]]].
econstructor; econstructor; eapply step_builtin; eauto.
- eapply can_eval_simple_list; eauto.
+ eapply can_eval_simple_list; eauto.
(* paren *)
exploit (simple_can_eval_rval f k e m b (fun x => C(Eparen x tycast ty))); eauto.
intros [v1 [E1 S1]].
@@ -1415,9 +1415,9 @@ Proof.
assert (exists t, exists S', estep S t S').
inv H0.
(* lred *)
- eapply can_estep; eauto. inv H2; auto.
+ eapply can_estep; eauto. inv H2; auto.
(* rred *)
- eapply can_estep; eauto. inv H2; auto. inv H1; auto.
+ eapply can_estep; eauto. inv H2; auto. inv H1; auto.
(* callred *)
eapply can_estep; eauto. inv H2; auto. inv H1; auto.
(* stuck *)
@@ -1444,7 +1444,7 @@ Remark deref_loc_trace:
deref_loc ge ty m b ofs t v ->
match t with nil => True | ev :: nil => True | _ => False end.
Proof.
- intros. inv H; simpl; auto. inv H2; simpl; auto.
+ intros. inv H; simpl; auto. inv H2; simpl; auto.
Qed.
Remark deref_loc_receptive:
@@ -1455,7 +1455,7 @@ Remark deref_loc_receptive:
Proof.
intros.
assert (t1 = nil). exploit deref_loc_trace; eauto. destruct t1; simpl; tauto.
- inv H. exploit volatile_load_receptive; eauto. intros [v' A].
+ inv H. exploit volatile_load_receptive; eauto. intros [v' A].
split; auto; exists v'; econstructor; eauto.
Qed.
@@ -1464,7 +1464,7 @@ Remark assign_loc_trace:
assign_loc ge ty m b ofs v t m' ->
match t with nil => True | ev :: nil => output_event ev | _ => False end.
Proof.
- intros. inv H; simpl; auto. inv H2; simpl; auto.
+ intros. inv H; simpl; auto. inv H2; simpl; auto.
Qed.
Remark assign_loc_receptive:
@@ -1475,10 +1475,10 @@ Remark assign_loc_receptive:
Proof.
intros.
assert (t1 = nil). exploit assign_loc_trace; eauto. destruct t1; simpl; tauto.
- inv H. eapply volatile_store_receptive; eauto.
+ inv H. eapply volatile_store_receptive; eauto.
Qed.
-Lemma semantics_strongly_receptive:
+Lemma semantics_strongly_receptive:
forall p, strongly_receptive (semantics p).
Proof.
intros. constructor; simpl; intros.
@@ -1487,78 +1487,78 @@ Proof.
inversion H; subst.
inv H1.
(* valof volatile *)
- exploit deref_loc_receptive; eauto. intros [A [v' B]].
- econstructor; econstructor. left; eapply step_rvalof_volatile; eauto.
+ exploit deref_loc_receptive; eauto. intros [A [v' B]].
+ econstructor; econstructor. left; eapply step_rvalof_volatile; eauto.
(* assign *)
exploit assign_loc_receptive; eauto. intro EQ; rewrite EQ in H.
econstructor; econstructor; eauto.
(* assignop *)
- destruct t0 as [ | ev0 t0]; simpl in H10.
+ destruct t0 as [ | ev0 t0]; simpl in H10.
subst t2. exploit assign_loc_receptive; eauto. intros EQ; rewrite EQ in H.
econstructor; econstructor; eauto.
inv H10. exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t0.
destruct (sem_binary_operation ge op v1' (typeof l) v2 (typeof r) m) as [v3'|] eqn:?.
destruct (sem_cast v3' tyres (typeof l)) as [v4'|] eqn:?.
destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v4' t2' m'')).
- destruct H1 as [t2' [m'' P]].
- econstructor; econstructor. left; eapply step_assignop with (v1 := v1'); eauto. simpl; reflexivity.
- econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
+ destruct H1 as [t2' [m'' P]].
+ econstructor; econstructor. left; eapply step_assignop with (v1 := v1'); eauto. simpl; reflexivity.
+ econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t0; exists m'0; auto.
- econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
+ econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0; auto.
- econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
+ econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
rewrite Heqo; auto.
(* assignop stuck *)
exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t1.
destruct (sem_binary_operation ge op v1' (typeof l) v2 (typeof r) m) as [v3'|] eqn:?.
destruct (sem_cast v3' tyres (typeof l)) as [v4'|] eqn:?.
destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v4' t2' m'')).
- destruct H1 as [t2' [m'' P]].
- econstructor; econstructor. left; eapply step_assignop with (v1 := v1'); eauto. simpl; reflexivity.
- econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
+ destruct H1 as [t2' [m'' P]].
+ econstructor; econstructor. left; eapply step_assignop with (v1 := v1'); eauto. simpl; reflexivity.
+ econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t2; exists m'; auto.
- econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
+ econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0; auto.
- econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
+ econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
rewrite Heqo; auto.
(* postincr *)
- destruct t0 as [ | ev0 t0]; simpl in H9.
+ destruct t0 as [ | ev0 t0]; simpl in H9.
subst t2. exploit assign_loc_receptive; eauto. intros EQ; rewrite EQ in H.
econstructor; econstructor; eauto.
inv H9. exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t0.
destruct (sem_incrdecr ge id v1' (typeof l)) as [v2'|] eqn:?.
destruct (sem_cast v2' (incrdecr_type (typeof l)) (typeof l)) as [v3'|] eqn:?.
destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v3' t2' m'')).
- destruct H1 as [t2' [m'' P]].
- econstructor; econstructor. left; eapply step_postincr with (v1 := v1'); eauto. simpl; reflexivity.
- econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
+ destruct H1 as [t2' [m'' P]].
+ econstructor; econstructor. left; eapply step_postincr with (v1 := v1'); eauto. simpl; reflexivity.
+ econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t0; exists m'0; auto.
- econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
+ econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0; auto.
- econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
+ econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; auto.
(* postincr stuck *)
exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t1.
destruct (sem_incrdecr ge id v1' (typeof l)) as [v2'|] eqn:?.
destruct (sem_cast v2' (incrdecr_type (typeof l)) (typeof l)) as [v3'|] eqn:?.
destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v3' t2' m'')).
- destruct H1 as [t2' [m'' P]].
- econstructor; econstructor. left; eapply step_postincr with (v1 := v1'); eauto. simpl; reflexivity.
- econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
+ destruct H1 as [t2' [m'' P]].
+ econstructor; econstructor. left; eapply step_postincr with (v1 := v1'); eauto. simpl; reflexivity.
+ econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t2; exists m'; auto.
- econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
+ econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0; auto.
- econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
+ econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; auto.
(* builtin *)
- exploit external_call_trace_length; eauto. destruct t1; simpl; intros.
- exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exploit external_call_trace_length; eauto. destruct t1; simpl; intros.
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
econstructor; econstructor. left; eapply step_builtin; eauto.
omegaContradiction.
(* external calls *)
inv H1.
- exploit external_call_trace_length; eauto. destruct t1; simpl; intros.
- exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exploit external_call_trace_length; eauto. destruct t1; simpl; intros.
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (Returnstate vres2 k m2); exists E0; right; econstructor; eauto.
omegaContradiction.
(* well-behaved traces *)
@@ -1569,15 +1569,15 @@ Proof.
exploit assign_loc_trace; eauto. destruct t; auto. destruct t; simpl; tauto.
(* assignop *)
exploit deref_loc_trace; eauto. exploit assign_loc_trace; eauto.
- destruct t1. destruct t2. simpl; auto. destruct t2; simpl; tauto.
- destruct t1. destruct t2. simpl; auto. destruct t2; simpl; tauto.
+ destruct t1. destruct t2. simpl; auto. destruct t2; simpl; tauto.
+ destruct t1. destruct t2. simpl; auto. destruct t2; simpl; tauto.
tauto.
(* assignop stuck *)
exploit deref_loc_trace; eauto. destruct t; auto. destruct t; tauto.
(* postincr *)
exploit deref_loc_trace; eauto. exploit assign_loc_trace; eauto.
- destruct t1. destruct t2. simpl; auto. destruct t2; simpl; tauto.
- destruct t1. destruct t2. simpl; auto. destruct t2; simpl; tauto.
+ destruct t1. destruct t2. simpl; auto. destruct t2; simpl; tauto.
+ destruct t1. destruct t2. simpl; auto. destruct t2; simpl; tauto.
tauto.
(* postincr stuck *)
exploit deref_loc_trace; eauto. destruct t; auto. destruct t; tauto.
@@ -1603,9 +1603,9 @@ Proof.
(* initial states match *)
intros. exists s2; auto.
(* final states match *)
- intros. subst s2. auto.
+ intros. subst s2. auto.
(* progress *)
- intros. subst s2. apply progress. auto.
+ intros. subst s2. apply progress. auto.
(* simulation *)
intros. subst s1. exists s2'; split; auto. apply step_simulation; auto.
Qed.
@@ -1647,7 +1647,7 @@ Definition outcome_result_value (out: outcome) (t: type) (v: val) : Prop :=
| Out_return None, Tvoid => v = Vundef
| Out_return (Some (v', ty')), ty => ty <> Tvoid /\ sem_cast v' ty' ty = Some v
| _, _ => False
- end.
+ end.
(** [eval_expression ge e m1 a t m2 a'] describes the evaluation of the
complex expression e. [v] is the resulting value, [m2] the final
@@ -1775,7 +1775,7 @@ with eval_exprlist: env -> mem -> exprlist -> trace -> mem -> exprlist -> Prop :
eval_expr e m RV a1 t1 m1 a1' -> eval_exprlist e m1 al t2 m2 al' ->
eval_exprlist e m (Econs a1 al) (t1**t2) m2 (Econs a1' al')
-(** [exec_stmt ge e m1 s t m2 out] describes the execution of
+(** [exec_stmt ge e m1 s t m2 out] describes the execution of
the statement [s]. [out] is the outcome for this execution.
[m1] is the initial memory state, [m2] the final memory state.
[t] is the trace of input/output events performed during this
@@ -1856,12 +1856,12 @@ with exec_stmt: env -> mem -> statement -> trace -> mem -> outcome -> Prop :=
eval_expression e m1 a t2 m2 v ->
bool_val v (typeof a) m2 = Some true ->
exec_stmt e m2 (Sdowhile a s) t3 m3 out ->
- exec_stmt e m (Sdowhile a s)
+ exec_stmt e m (Sdowhile a s)
(t1 ** t2 ** t3) m3 out
| exec_Sfor_start: forall e m s a1 a2 a3 out m1 m2 t1 t2,
exec_stmt e m a1 t1 m1 Out_normal ->
exec_stmt e m1 (Sfor Sskip a2 a3 s) t2 m2 out ->
- exec_stmt e m (Sfor a1 a2 a3 s)
+ exec_stmt e m (Sfor a1 a2 a3 s)
(t1 ** t2) m2 out
| exec_Sfor_false: forall e m s a2 a3 t m' v,
eval_expression e m a2 t m' v ->
@@ -1998,7 +1998,7 @@ CoInductive evalinf_expr: env -> mem -> kind -> expr -> traceinf -> Prop :=
evalinf_expr e m RV a1 t1 ->
evalinf_expr e m RV (Ecall a1 a2 ty) t1
| evalinf_call_right: forall e m a1 t1 m1 a1' a2 t2 ty,
- eval_expr e m RV a1 t1 m1 a1' ->
+ eval_expr e m RV a1 t1 m1 a1' ->
evalinf_exprlist e m1 a2 t2 ->
evalinf_expr e m RV (Ecall a1 a2 ty) (t1 *** t2)
| evalinf_call: forall e m rf rargs ty t1 m1 rf' t2 m2 rargs' vf vargs
@@ -2020,7 +2020,7 @@ with evalinf_exprlist: env -> mem -> exprlist -> traceinf -> Prop :=
eval_expr e m RV a1 t1 m1 a1' -> evalinf_exprlist e m1 al t2 ->
evalinf_exprlist e m (Econs a1 al) (t1***t2)
-(** [execinf_stmt ge e m1 s t] describes the diverging execution of
+(** [execinf_stmt ge e m1 s t] describes the diverging execution of
the statement [s]. *)
with execinf_stmt: env -> mem -> statement -> traceinf -> Prop :=
@@ -2137,7 +2137,7 @@ Inductive outcome_state_match
outcome_state_match e m f k Out_continue (State f Scontinue k e m)
| osm_return_none: forall k',
call_cont k' = call_cont k ->
- outcome_state_match e m f k
+ outcome_state_match e m f k
(Out_return None) (State f (Sreturn None) k' e m)
| osm_return_some: forall v ty k',
call_cont k' = call_cont k ->
@@ -2168,16 +2168,16 @@ Lemma exprlist_app_leftcontext:
forall rl1 rl2,
simplelist rl1 = true -> leftcontextlist RV (fun x => exprlist_app rl1 (Econs x rl2)).
Proof.
- induction rl1; simpl; intros.
+ induction rl1; simpl; intros.
apply lctx_list_head. constructor.
- destruct (andb_prop _ _ H). apply lctx_list_tail. auto. auto.
+ destruct (andb_prop _ _ H). apply lctx_list_tail. auto. auto.
Qed.
Lemma exprlist_app_simple:
forall rl1 rl2,
simplelist (exprlist_app rl1 rl2) = simplelist rl1 && simplelist rl2.
Proof.
- induction rl1; intros; simpl. auto. rewrite IHrl1. apply andb_assoc.
+ induction rl1; intros; simpl. auto. rewrite IHrl1. apply andb_assoc.
Qed.
Lemma bigstep_to_steps:
@@ -2212,9 +2212,9 @@ Proof.
exploit (H0 (fun x => x) f k). constructor. intros [A [B C]].
assert (match a' with Eval _ _ => False | _ => True end ->
star step ge (ExprState f a k e m) t (ExprState f (Eval v (typeof a)) k e m')).
- intro. eapply star_right. eauto. left. eapply step_expr; eauto. traceEq.
+ intro. eapply star_right. eauto. left. eapply step_expr; eauto. traceEq.
destruct a'; auto.
- simpl in B. rewrite B in C. inv H1. auto.
+ simpl in B. rewrite B in C. inv H1. auto.
(* val *)
simpl; intuition. apply star_refl.
@@ -2223,7 +2223,7 @@ Proof.
(* field *)
exploit (H0 (fun x => C(Efield x f ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
- simpl; intuition; eauto.
+ simpl; intuition; eauto.
(* valof *)
exploit (H1 (fun x => C(Evalof x ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
@@ -2232,8 +2232,8 @@ Proof.
exploit (H1 (fun x => C(Evalof x ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
simpl; intuition.
- eapply star_right. eexact D.
- left. eapply step_rvalof_volatile; eauto. rewrite H4; eauto. congruence. congruence.
+ eapply star_right. eexact D.
+ left. eapply step_rvalof_volatile; eauto. rewrite H4; eauto. congruence. congruence.
traceEq.
(* deref *)
exploit (H0 (fun x => C(Ederef x ty))).
@@ -2252,7 +2252,7 @@ Proof.
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
exploit (H2 (fun x => C(Ebinop op a1' x ty))).
eapply leftcontext_compose; eauto. repeat constructor. auto. intros [E [F G]].
- simpl; intuition. eapply star_trans; eauto.
+ simpl; intuition. eapply star_trans; eauto.
(* cast *)
exploit (H0 (fun x => C(Ecast x ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
@@ -2262,15 +2262,15 @@ Proof.
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
exploit (H4 (fun x => C(Eparen x type_bool ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [E [F G]].
- simpl; intuition. eapply star_trans. eexact D.
- eapply star_left. left; eapply step_seqand_true; eauto. rewrite B; auto.
+ simpl; intuition. eapply star_trans. eexact D.
+ eapply star_left. left; eapply step_seqand_true; eauto. rewrite B; auto.
eapply star_right. eexact G.
left; eapply step_paren; eauto. rewrite F; eauto.
- eauto. eauto. traceEq.
+ eauto. eauto. traceEq.
(* seqand false *)
exploit (H0 (fun x => C(Eseqand x a2 ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
- simpl; intuition. eapply star_right. eexact D.
+ simpl; intuition. eapply star_right. eexact D.
left; eapply step_seqand_false; eauto. rewrite B; auto.
traceEq.
(* seqor false *)
@@ -2278,15 +2278,15 @@ Proof.
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
exploit (H4 (fun x => C(Eparen x type_bool ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [E [F G]].
- simpl; intuition. eapply star_trans. eexact D.
- eapply star_left. left; eapply step_seqor_false; eauto. rewrite B; auto.
+ simpl; intuition. eapply star_trans. eexact D.
+ eapply star_left. left; eapply step_seqor_false; eauto. rewrite B; auto.
eapply star_right. eexact G.
left; eapply step_paren; eauto. rewrite F; eauto.
- eauto. eauto. traceEq.
+ eauto. eauto. traceEq.
(* seqor true *)
exploit (H0 (fun x => C(Eseqor x a2 ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
- simpl; intuition. eapply star_right. eexact D.
+ simpl; intuition. eapply star_right. eexact D.
left; eapply step_seqor_true; eauto. rewrite B; auto.
traceEq.
(* condition *)
@@ -2294,10 +2294,10 @@ Proof.
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
exploit (H4 (fun x => C(Eparen x ty ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [E [F G]].
- simpl. split; auto. split; auto.
+ simpl. split; auto. split; auto.
eapply star_trans. eexact D.
- eapply star_left. left; eapply step_condition; eauto. rewrite B; eauto.
- eapply star_right. eexact G. left; eapply step_paren; eauto. congruence.
+ eapply star_left. left; eapply step_condition; eauto. rewrite B; eauto.
+ eapply star_right. eexact G. left; eapply step_paren; eauto. congruence.
reflexivity. reflexivity. traceEq.
(* sizeof *)
simpl; intuition. apply star_refl.
@@ -2309,7 +2309,7 @@ Proof.
exploit (H2 (fun x => C(Eassign l' x ty))).
eapply leftcontext_compose; eauto. repeat constructor. auto. intros [E [F G]].
simpl; intuition.
- eapply star_trans. eexact D.
+ eapply star_trans. eexact D.
eapply star_right. eexact G.
left. eapply step_assign; eauto. congruence. rewrite B; eauto. congruence.
reflexivity. traceEq.
@@ -2319,7 +2319,7 @@ Proof.
exploit (H2 (fun x => C(Eassignop op l' x tyres ty))).
eapply leftcontext_compose; eauto. repeat constructor. auto. intros [E [F G]].
simpl; intuition.
- eapply star_trans. eexact D.
+ eapply star_trans. eexact D.
eapply star_right. eexact G.
left. eapply step_assignop; eauto.
rewrite B; eauto. rewrite B; rewrite F; eauto. congruence. rewrite B; eauto. congruence.
@@ -2328,8 +2328,8 @@ Proof.
exploit (H0 (fun x => C(Epostincr id x ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
simpl; intuition.
- eapply star_right. eexact D.
- left. eapply step_postincr; eauto. congruence.
+ eapply star_right. eexact D.
+ left. eapply step_postincr; eauto. congruence.
traceEq.
(* comma *)
exploit (H0 (fun x => C(Ecomma x r2 ty))).
@@ -2337,19 +2337,19 @@ Proof.
exploit (H3 C). auto. intros [E [F G]].
simpl; intuition. congruence.
eapply star_trans. eexact D.
- eapply star_left. left; eapply step_comma; eauto.
+ eapply star_left. left; eapply step_comma; eauto.
eexact G.
reflexivity. traceEq.
(* call *)
exploit (H0 (fun x => C(Ecall x rargs ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
exploit (H2 rf' Enil ty C); eauto. intros [E F].
- simpl; intuition.
+ simpl; intuition.
eapply star_trans. eexact D.
- eapply star_trans. eexact F.
- eapply star_left. left; eapply step_call; eauto. congruence.
- eapply star_right. eapply H9. red; auto.
- right; constructor.
+ eapply star_trans. eexact F.
+ eapply star_left. left; eapply step_call; eauto. congruence.
+ eapply star_right. eapply H9. red; auto.
+ right; constructor.
reflexivity. reflexivity. reflexivity. traceEq.
(* nil *)
simpl; intuition. apply star_refl.
@@ -2358,18 +2358,18 @@ Proof.
eapply leftcontext_compose; eauto. repeat constructor. auto.
apply exprlist_app_leftcontext; auto. intros [A [B D]].
exploit (H2 a0 (exprlist_app al2 (Econs a1' Enil))); eauto.
- rewrite exprlist_app_simple. simpl. rewrite H5; rewrite A; auto.
- repeat rewrite exprlist_app_assoc. simpl.
+ rewrite exprlist_app_simple. simpl. rewrite H5; rewrite A; auto.
+ repeat rewrite exprlist_app_assoc. simpl.
intros [E F].
- simpl; intuition.
+ simpl; intuition.
eapply star_trans; eauto.
(* skip *)
econstructor; split. apply star_refl. constructor.
(* do *)
- econstructor; split.
- eapply star_left. right; constructor.
+ econstructor; split.
+ eapply star_left. right; constructor.
eapply star_right. apply H0. right; constructor.
reflexivity. traceEq.
constructor.
@@ -2379,7 +2379,7 @@ Proof.
destruct (H2 f k) as [S2 [A2 B2]]; auto.
econstructor; split.
eapply star_left. right; econstructor.
- eapply star_trans. eexact A1.
+ eapply star_trans. eexact A1.
eapply star_left. right; constructor. eexact A2.
reflexivity. reflexivity. traceEq.
auto.
@@ -2420,7 +2420,7 @@ Proof.
econstructor; split.
eapply star_left. right; apply step_return_1.
eapply H0. traceEq.
- econstructor; eauto.
+ econstructor; eauto.
(* break *)
econstructor; split. apply star_refl. constructor.
@@ -2431,7 +2431,7 @@ Proof.
(* while false *)
econstructor; split.
eapply star_left. right; apply step_while.
- eapply star_right. apply H0. right; eapply step_while_false; eauto.
+ eapply star_right. apply H0. right; eapply step_while_false; eauto.
reflexivity. traceEq.
constructor.
@@ -2448,7 +2448,7 @@ Proof.
eapply star_left. right; eapply step_while_true; eauto.
eapply star_trans. eexact A1.
unfold S2. inversion H4; subst.
- inv B1. apply star_one. right; constructor.
+ inv B1. apply star_one. right; constructor.
apply star_refl.
reflexivity. reflexivity. reflexivity. traceEq.
unfold S2. inversion H4; subst. constructor. inv B1; econstructor; eauto.
@@ -2474,9 +2474,9 @@ Proof.
eapply star_trans. eexact A1.
eapply star_left.
inv H1; inv B1; right; eapply step_skip_or_continue_dowhile; eauto.
- eapply star_right. apply H3.
- right; eapply step_dowhile_false; eauto.
- reflexivity. reflexivity. reflexivity. traceEq.
+ eapply star_right. apply H3.
+ right; eapply step_dowhile_false; eauto.
+ reflexivity. reflexivity. reflexivity. traceEq.
constructor.
(* dowhile stop *)
@@ -2487,7 +2487,7 @@ Proof.
| _ => S1
end).
exists S2; split.
- eapply star_left. right; apply step_dowhile.
+ eapply star_left. right; apply step_dowhile.
eapply star_trans. eexact A1.
unfold S2. inversion H1; subst.
inv B1. apply star_one. right; constructor.
@@ -2510,13 +2510,13 @@ Proof.
auto.
(* for start *)
- assert (a1 = Sskip \/ a1 <> Sskip). destruct a1; auto; right; congruence.
+ assert (a1 = Sskip \/ a1 <> Sskip). destruct a1; auto; right; congruence.
destruct H3.
subst a1. inv H. apply H2; auto.
destruct (H0 f (Kseq (Sfor Sskip a2 a3 s) k)) as [S1 [A1 B1]]; auto. inv B1.
destruct (H2 f k) as [S2 [A2 B2]]; auto.
exists S2; split.
- eapply star_left. right; apply step_for_start; auto.
+ eapply star_left. right; apply step_for_start; auto.
eapply star_trans. eexact A1.
eapply star_left. right; constructor. eexact A2.
reflexivity. reflexivity. traceEq.
@@ -2524,8 +2524,8 @@ Proof.
(* for false *)
econstructor; split.
- eapply star_left. right; apply step_for.
- eapply star_right. apply H0. right; eapply step_for_false; eauto.
+ eapply star_left. right; apply step_for.
+ eapply star_right. apply H0. right; eapply step_for_false; eauto.
reflexivity. traceEq.
constructor.
@@ -2537,12 +2537,12 @@ Proof.
| _ => S1
end).
exists S2; split.
- eapply star_left. right; apply step_for.
- eapply star_trans. apply H0.
+ eapply star_left. right; apply step_for.
+ eapply star_trans. apply H0.
eapply star_left. right; eapply step_for_true; eauto.
- eapply star_trans. eexact A1.
+ eapply star_trans. eexact A1.
unfold S2. inversion H4; subst.
- inv B1. apply star_one. right; constructor.
+ inv B1. apply star_one. right; constructor.
apply star_refl.
reflexivity. reflexivity. reflexivity. traceEq.
unfold S2. inversion H4; subst. constructor. inv B1; econstructor; eauto.
@@ -2552,15 +2552,15 @@ Proof.
destruct (H6 f (Kfor4 a2 a3 s k)) as [S2 [A2 B2]]; auto. inv B2.
destruct (H8 f k) as [S3 [A3 B3]]; auto.
exists S3; split.
- eapply star_left. right; apply step_for.
- eapply star_trans. apply H0.
+ eapply star_left. right; apply step_for.
+ eapply star_trans. apply H0.
eapply star_left. right; eapply step_for_true; eauto.
eapply star_trans. eexact A1.
eapply star_trans with (s2 := State f a3 (Kfor4 a2 a3 s k) e m2).
inv H4; inv B1.
- apply star_one. right; constructor; auto.
- apply star_one. right; constructor; auto.
- eapply star_trans. eexact A2.
+ apply star_one. right; constructor; auto.
+ apply star_one. right; constructor; auto.
+ eapply star_trans. eexact A2.
eapply star_left. right; constructor.
eexact A3.
reflexivity. reflexivity. reflexivity. reflexivity.
@@ -2578,13 +2578,13 @@ Proof.
end).
exists S2; split.
eapply star_left. right; eapply step_switch.
- eapply star_trans. apply H0.
- eapply star_left. right; eapply step_expr_switch. eauto.
- eapply star_trans. eexact A1.
+ eapply star_trans. apply H0.
+ eapply star_left. right; eapply step_expr_switch. eauto.
+ eapply star_trans. eexact A1.
unfold S2; inv B1.
- apply star_one. right; constructor. auto.
- apply star_one. right; constructor. auto.
- apply star_one. right; constructor.
+ apply star_one. right; constructor. auto.
+ apply star_one. right; constructor. auto.
+ apply star_one. right; constructor.
apply star_refl.
apply star_refl.
reflexivity. reflexivity. reflexivity. traceEq.
@@ -2593,7 +2593,7 @@ Proof.
(* call internal *)
destruct (H3 f k) as [S1 [A1 B1]].
eapply star_left. right; eapply step_internal_function; eauto.
- eapply star_right. eexact A1.
+ eapply star_right. eexact A1.
inv B1; simpl in H4; try contradiction.
(* Out_normal *)
assert (fn_return f = Tvoid /\ vres = Vundef).
@@ -2611,7 +2611,7 @@ Proof.
reflexivity. traceEq.
(* call external *)
- apply star_one. right; apply step_external_function; auto.
+ apply star_one. right; apply step_external_function; auto.
Qed.
Lemma eval_expression_to_steps:
@@ -2713,7 +2713,7 @@ Lemma evalinf_funcall_steps:
Proof.
cofix COF.
- assert (COS:
+ assert (COS:
forall e m s t f k,
execinf_stmt e m s t ->
forever_N step lt ge O (State f s k e m) t).
@@ -2735,180 +2735,180 @@ Proof.
cofix COEL.
intros. inv H.
(* cons left *)
- eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
+ eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
eapply COE with (C := fun x => C(Ecall a1 (exprlist_app al (Econs x al0)) ty)).
- eauto. eapply leftcontext_compose; eauto. constructor. auto.
+ eauto. eapply leftcontext_compose; eauto. constructor. auto.
apply exprlist_app_leftcontext; auto. traceEq.
-(* cons right *)
- destruct (eval_expr_to_steps _ _ _ _ _ _ _ H3
+(* cons right *)
+ destruct (eval_expr_to_steps _ _ _ _ _ _ _ H3
(fun x => C(Ecall a1 (exprlist_app al (Econs x al0)) ty)) f k)
- as [P [Q R]].
- eapply leftcontext_compose; eauto. repeat constructor. auto.
- apply exprlist_app_leftcontext; auto.
+ as [P [Q R]].
+ eapply leftcontext_compose; eauto. repeat constructor. auto.
+ apply exprlist_app_leftcontext; auto.
eapply forever_N_star with (a2 := (esizelist al0)).
eexact R. simpl; omega.
change (Econs a1' al0) with (exprlist_app (Econs a1' Enil) al0).
- rewrite <- exprlist_app_assoc.
- eapply COEL. eauto. auto. auto.
+ rewrite <- exprlist_app_assoc.
+ eapply COEL. eauto. auto. auto.
rewrite exprlist_app_simple. simpl. rewrite H2; rewrite P; auto.
auto.
intros. inv H.
(* field *)
eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Efield x f0 ty)). eauto.
+ eapply COE with (C := fun x => C(Efield x f0 ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* valof *)
eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Evalof x ty)). eauto.
+ eapply COE with (C := fun x => C(Evalof x ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* deref *)
eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Ederef x ty)). eauto.
+ eapply COE with (C := fun x => C(Ederef x ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* addrof *)
eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Eaddrof x ty)). eauto.
+ eapply COE with (C := fun x => C(Eaddrof x ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* unop *)
eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Eunop op x ty)). eauto.
+ eapply COE with (C := fun x => C(Eunop op x ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* binop left *)
eapply forever_N_star with (a2 := (esize a1)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Ebinop op x a2 ty)). eauto.
+ eapply COE with (C := fun x => C(Ebinop op x a2 ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* binop right *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Ebinop op x a2 ty)) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
eapply forever_N_star with (a2 := (esize a2)). eexact R. simpl; omega.
- eapply COE with (C := fun x => C(Ebinop op a1' x ty)). eauto.
+ eapply COE with (C := fun x => C(Ebinop op a1' x ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. auto. traceEq.
(* cast *)
eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Ecast x ty)). eauto.
+ eapply COE with (C := fun x => C(Ecast x ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* seqand left *)
eapply forever_N_star with (a2 := (esize a1)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Eseqand x a2 ty)). eauto.
+ eapply COE with (C := fun x => C(Eseqand x a2 ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* seqand 2 *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Eseqand x a2 ty)) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
- eapply forever_N_plus. eapply plus_right. eexact R.
+ eapply forever_N_plus. eapply plus_right. eexact R.
left; eapply step_seqand_true; eauto. rewrite Q; eauto.
- reflexivity.
+ reflexivity.
eapply COE with (C := fun x => (C (Eparen x type_bool ty))). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* seqor left *)
eapply forever_N_star with (a2 := (esize a1)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Eseqor x a2 ty)). eauto.
+ eapply COE with (C := fun x => C(Eseqor x a2 ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* seqor 2 *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Eseqor x a2 ty)) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
- eapply forever_N_plus. eapply plus_right. eexact R.
+ eapply forever_N_plus. eapply plus_right. eexact R.
left; eapply step_seqor_false; eauto. rewrite Q; eauto.
- reflexivity.
+ reflexivity.
eapply COE with (C := fun x => (C (Eparen x type_bool ty))). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* condition top *)
eapply forever_N_star with (a2 := (esize a1)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Econdition x a2 a3 ty)). eauto.
+ eapply COE with (C := fun x => C(Econdition x a2 a3 ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* condition *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Econdition x a2 a3 ty)) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
- eapply forever_N_plus. eapply plus_right. eexact R.
+ eapply forever_N_plus. eapply plus_right. eexact R.
left; eapply step_condition; eauto. rewrite Q; eauto.
- reflexivity.
+ reflexivity.
eapply COE with (C := fun x => (C (Eparen x ty ty))). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* assign left *)
eapply forever_N_star with (a2 := (esize a1)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Eassign x a2 ty)). eauto.
+ eapply COE with (C := fun x => C(Eassign x a2 ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* assign right *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Eassign x a2 ty)) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
eapply forever_N_star with (a2 := (esize a2)). eexact R. simpl; omega.
- eapply COE with (C := fun x => C(Eassign a1' x ty)). eauto.
+ eapply COE with (C := fun x => C(Eassign a1' x ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. auto. traceEq.
(* assignop left *)
eapply forever_N_star with (a2 := (esize a1)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Eassignop op x a2 tyres ty)). eauto.
+ eapply COE with (C := fun x => C(Eassignop op x a2 tyres ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* assignop right *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Eassignop op x a2 tyres ty)) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
eapply forever_N_star with (a2 := (esize a2)). eexact R. simpl; omega.
- eapply COE with (C := fun x => C(Eassignop op a1' x tyres ty)). eauto.
+ eapply COE with (C := fun x => C(Eassignop op a1' x tyres ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. auto. traceEq.
(* postincr *)
eapply forever_N_star with (a2 := (esize a0)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Epostincr id x ty)). eauto.
+ eapply COE with (C := fun x => C(Epostincr id x ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* comma left *)
eapply forever_N_star with (a2 := (esize a1)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Ecomma x a2 ty)). eauto.
+ eapply COE with (C := fun x => C(Ecomma x a2 ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* comma right *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Ecomma x a2 (typeof a2))) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
- eapply forever_N_plus. eapply plus_right. eexact R.
- left; eapply step_comma; eauto. reflexivity.
+ eapply forever_N_plus. eapply plus_right. eexact R.
+ left; eapply step_comma; eauto. reflexivity.
eapply COE with (C := C); eauto. traceEq.
(* call left *)
eapply forever_N_star with (a2 := (esize a1)). apply star_refl. simpl; omega.
- eapply COE with (C := fun x => C(Ecall x a2 ty)). eauto.
+ eapply COE with (C := fun x => C(Ecall x a2 ty)). eauto.
eapply leftcontext_compose; eauto. repeat constructor. traceEq.
(* call right *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Ecall x a2 ty)) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
eapply forever_N_star with (a2 := (esizelist a2)). eexact R. simpl; omega.
eapply COEL with (al := Enil). eauto. auto. auto. auto. traceEq.
(* call *)
destruct (eval_expr_to_steps _ _ _ _ _ _ _ H1 (fun x => C(Ecall x rargs ty)) f k)
- as [P [Q R]].
+ as [P [Q R]].
eapply leftcontext_compose; eauto. repeat constructor.
destruct (eval_exprlist_to_steps _ _ _ _ _ _ H2 rf' Enil ty C f k)
as [S T]. auto. auto. simpl; auto.
eapply forever_N_plus. eapply plus_right.
eapply star_trans. eexact R. eexact T. reflexivity.
- simpl. left; eapply step_call; eauto. congruence. reflexivity.
- apply COF. eauto. traceEq.
+ simpl. left; eapply step_call; eauto. congruence. reflexivity.
+ apply COF. eauto. traceEq.
(* statements *)
intros. inv H.
(* do *)
- eapply forever_N_plus. apply plus_one; right; constructor.
+ eapply forever_N_plus. apply plus_one; right; constructor.
eapply COE with (C := fun x => x); eauto. constructor. traceEq.
(* seq 1 *)
- eapply forever_N_plus. apply plus_one; right; constructor.
+ eapply forever_N_plus. apply plus_one; right; constructor.
eapply COS; eauto. traceEq.
(* seq 2 *)
destruct (exec_stmt_to_steps _ _ _ _ _ _ H0 f (Kseq s2 k)) as [S1 [A1 B1]]; auto. inv B1.
- eapply forever_N_plus.
- eapply plus_left. right; constructor.
+ eapply forever_N_plus.
+ eapply plus_left. right; constructor.
eapply star_right. eauto. right; constructor.
- reflexivity. reflexivity.
+ reflexivity. reflexivity.
eapply COS; eauto. traceEq.
(* if test *)
- eapply forever_N_plus. apply plus_one; right; constructor.
+ eapply forever_N_plus. apply plus_one; right; constructor.
eapply COE with (C := fun x => x); eauto. constructor. traceEq.
(* if true/false *)
eapply forever_N_plus.
eapply plus_left. right; constructor.
eapply star_right. eapply eval_expression_to_steps; eauto.
- right. eapply step_ifthenelse_2 with (b := b). auto.
+ right. eapply step_ifthenelse_2 with (b := b). auto.
reflexivity. reflexivity.
eapply COS; eauto. traceEq.
(* return some *)
@@ -2919,7 +2919,7 @@ Proof.
eapply COE with (C := fun x => x); eauto. constructor. traceEq.
(* while body *)
eapply forever_N_plus.
- eapply plus_left. right; constructor.
+ eapply plus_left. right; constructor.
eapply star_right. eapply eval_expression_to_steps; eauto.
right; apply step_while_true; auto.
reflexivity. reflexivity.
@@ -2927,10 +2927,10 @@ Proof.
(* while loop *)
destruct (exec_stmt_to_steps _ _ _ _ _ _ H2 f (Kwhile2 a s0 k)) as [S1 [A1 B1]]; auto.
eapply forever_N_plus.
- eapply plus_left. right; constructor.
+ eapply plus_left. right; constructor.
eapply star_trans. eapply eval_expression_to_steps; eauto.
eapply star_left. right; apply step_while_true; auto.
- eapply star_trans. eexact A1.
+ eapply star_trans. eexact A1.
inv H3; inv B1; apply star_one; right; apply step_skip_or_continue_while; auto.
reflexivity. reflexivity. reflexivity. reflexivity.
eapply COS; eauto. traceEq.
@@ -2940,7 +2940,7 @@ Proof.
(* dowhile test *)
destruct (exec_stmt_to_steps _ _ _ _ _ _ H0 f (Kdowhile1 a s0 k)) as [S1 [A1 B1]]; auto.
eapply forever_N_plus.
- eapply plus_left. right; constructor.
+ eapply plus_left. right; constructor.
eapply star_trans. eexact A1.
eapply star_one. right. inv H1; inv B1; apply step_skip_or_continue_dowhile; auto.
reflexivity. reflexivity.
@@ -2948,7 +2948,7 @@ Proof.
(* dowhile loop *)
destruct (exec_stmt_to_steps _ _ _ _ _ _ H0 f (Kdowhile1 a s0 k)) as [S1 [A1 B1]]; auto.
eapply forever_N_plus.
- eapply plus_left. right; constructor.
+ eapply plus_left. right; constructor.
eapply star_trans. eexact A1.
eapply star_left. right. inv H1; inv B1; apply step_skip_or_continue_dowhile; auto.
eapply star_right. eapply eval_expression_to_steps; eauto.
@@ -2962,9 +2962,9 @@ Proof.
(* for start 2 *)
destruct (exec_stmt_to_steps _ _ _ _ _ _ H0 f (Kseq (Sfor Sskip a2 a3 s0) k)) as [S1 [A1 B1]]; auto. inv B1.
eapply forever_N_plus.
- eapply plus_left. right; constructor. auto.
+ eapply plus_left. right; constructor. auto.
eapply star_trans. eexact A1.
- apply star_one. right; constructor.
+ apply star_one. right; constructor.
reflexivity. reflexivity.
eapply COS; eauto. traceEq.
(* for test *)
@@ -2983,7 +2983,7 @@ Proof.
eapply plus_left. right; apply step_for.
eapply star_trans. eapply eval_expression_to_steps; eauto.
eapply star_left. right; apply step_for_true; auto.
- eapply star_trans. eexact A1.
+ eapply star_trans. eexact A1.
inv H3; inv B1; apply star_one; right; apply step_skip_or_continue_for3; auto.
reflexivity. reflexivity. reflexivity. reflexivity.
eapply COS; eauto. traceEq.
@@ -2997,7 +2997,7 @@ Proof.
eapply star_trans. eexact A1.
eapply star_left.
inv H3; inv B1; right; apply step_skip_or_continue_for3; auto.
- eapply star_right. eexact A2.
+ eapply star_right. eexact A2.
right; constructor.
reflexivity. reflexivity. reflexivity. reflexivity. reflexivity. reflexivity.
eapply COS; eauto. traceEq.
@@ -3006,15 +3006,15 @@ Proof.
eapply COE with (C := fun x => x); eauto. constructor. traceEq.
(* switch body *)
eapply forever_N_plus.
- eapply plus_left. right; constructor.
- eapply star_right. eapply eval_expression_to_steps; eauto.
- right; constructor. eauto.
- reflexivity. reflexivity.
- eapply COS; eauto. traceEq.
+ eapply plus_left. right; constructor.
+ eapply star_right. eapply eval_expression_to_steps; eauto.
+ right; constructor. eauto.
+ reflexivity. reflexivity.
+ eapply COS; eauto. traceEq.
(* funcalls *)
intros. inv H.
- eapply forever_N_plus. apply plus_one. right; econstructor; eauto.
+ eapply forever_N_plus. apply plus_one. right; econstructor; eauto.
eapply COS; eauto. traceEq.
Qed.
@@ -3024,7 +3024,7 @@ End BIGSTEP.
Inductive bigstep_program_terminates (p: program): trace -> int -> Prop :=
| bigstep_program_terminates_intro: forall b f m0 m1 t r,
- let ge := globalenv p in
+ let ge := globalenv p in
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
@@ -3034,7 +3034,7 @@ Inductive bigstep_program_terminates (p: program): trace -> int -> Prop :=
Inductive bigstep_program_diverges (p: program): traceinf -> Prop :=
| bigstep_program_diverges_intro: forall b f m0 t,
- let ge := globalenv p in
+ let ge := globalenv p in
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
@@ -3050,7 +3050,7 @@ Theorem bigstep_semantics_sound:
Proof.
intros; constructor; intros.
(* termination *)
- inv H. econstructor; econstructor.
+ inv H. econstructor; econstructor.
split. econstructor; eauto.
split. apply eval_funcall_to_steps. eauto. red; auto.
econstructor.
diff --git a/cfrontend/Csyntax.v b/cfrontend/Csyntax.v
index db059791..89d0b2bf 100644
--- a/cfrontend/Csyntax.v
+++ b/cfrontend/Csyntax.v
@@ -74,18 +74,18 @@ ranged over by [l] and [r], respectively, in the grammar above.
L-values are those expressions that can occur to the left of an assignment.
They denote memory locations. (Indeed, the reduction semantics for
expression reduces them to [Eloc b ofs] expressions.) L-values are
-variables ([Evar]), pointer dereferences ([Ederef]), field accesses ([Efield]).
+variables ([Evar]), pointer dereferences ([Ederef]), field accesses ([Efield]).
R-values are all other expressions. They denote values, and the reduction
-semantics reduces them to [Eval v] expressions.
+semantics reduces them to [Eval v] expressions.
A l-value can be used in a r-value context, but this use must be marked
-explicitly with the [Evalof] operator, which is not materialized in the
+explicitly with the [Evalof] operator, which is not materialized in the
concrete syntax of C but denotes a read from the location corresponding to
the l-value [l] argument of [Evalof l].
The grammar above contains some forms that cannot appear in source terms
but appear during reduction. These forms are:
-- [Eval v] where [v] is a pointer or [Vundef]. ([Eval] of an integer or
+- [Eval v] where [v] is a pointer or [Vundef]. ([Eval] of an integer or
float value can occur in a source term and represents a numeric literal.)
- [Eloc b ofs], which appears during reduction of l-values.
- [Eparen r tycast ty], which appears during reduction of conditionals
@@ -102,7 +102,7 @@ Definition Eindex (r1 r2: expr) (ty: type) :=
[l += 1] and [l -= 1], respectively. *)
Definition Epreincr (id: incr_or_decr) (l: expr) (ty: type) :=
- Eassignop (match id with Incr => Oadd | Decr => Osub end)
+ Eassignop (match id with Incr => Oadd | Decr => Osub end)
l (Eval (Vint Int.one) type_int32s) (typeconv ty) ty.
(** Extract the type part of a type-annotated expression. *)
diff --git a/cfrontend/Ctypes.v b/cfrontend/Ctypes.v
index 1f55da7f..78345b42 100644
--- a/cfrontend/Ctypes.v
+++ b/cfrontend/Ctypes.v
@@ -24,7 +24,7 @@ Require Archi.
(** * Syntax of types *)
(** Compcert C types are similar to those of C. They include numeric types,
- pointers, arrays, function types, and composite types (struct and
+ pointers, arrays, function types, and composite types (struct and
union). Numeric types (integers and floats) fully specify the
bit size of the type. An integer type is a pair of a signed/unsigned
flag and a bit size: 8, 16, or 32 bits, or the special [IBool] size
@@ -140,7 +140,7 @@ Definition attr_union (a1 a2: attr) : attr :=
| None, al => al
| al, None => al
| Some n1, Some n2 => Some (N.max n1 n2)
- end
+ end
|}.
Definition merge_attributes (ty: type) (a: attr) : type :=
@@ -184,7 +184,7 @@ Definition type_int32s := Tint I32 Signed noattr.
Definition type_bool := Tint IBool Signed noattr.
(** The usual unary conversion. Promotes small integer types to [signed int32]
- and degrades array types and function types to pointer types.
+ and degrades array types and function types to pointer types.
Attributes are erased. *)
Definition typeconv (ty: type) : type :=
@@ -272,7 +272,7 @@ Remark align_attr_two_p:
(exists n, al = two_power_nat n) ->
(exists n, align_attr a al = two_power_nat n).
Proof.
- intros. unfold align_attr. destruct (attr_alignas a).
+ intros. unfold align_attr. destruct (attr_alignas a).
exists (N.to_nat n). rewrite two_power_nat_two_p. rewrite N_nat_Z. auto.
auto.
Qed.
@@ -309,7 +309,7 @@ Qed.
(** In the ISO C standard, size is defined only for complete
types. However, it is convenient that [sizeof] is a total
function. For [void] and function types, we follow GCC and define
- their size to be 1. For undefined structures and unions, the size is
+ their size to be 1. For undefined structures and unions, the size is
arbitrarily taken to be 0.
*)
@@ -355,16 +355,16 @@ Fixpoint naturally_aligned (t: type) : Prop :=
Lemma sizeof_alignof_compat:
forall env t, naturally_aligned t -> (alignof env t | sizeof env t).
Proof.
- induction t; intros [A B]; unfold alignof, align_attr; rewrite A; simpl.
+ induction t; intros [A B]; unfold alignof, align_attr; rewrite A; simpl.
- apply Zdivide_refl.
- destruct i; apply Zdivide_refl.
-- exists (8 / Archi.align_int64); reflexivity.
-- destruct f. apply Zdivide_refl. exists (8 / Archi.align_float64); reflexivity.
+- exists (8 / Archi.align_int64); reflexivity.
+- destruct f. apply Zdivide_refl. exists (8 / Archi.align_float64); reflexivity.
- apply Zdivide_refl.
-- apply Z.divide_mul_l; auto.
+- apply Z.divide_mul_l; auto.
- apply Zdivide_refl.
-- destruct (env!i). apply co_sizeof_alignof. apply Zdivide_0.
-- destruct (env!i). apply co_sizeof_alignof. apply Zdivide_0.
+- destruct (env!i). apply co_sizeof_alignof. apply Zdivide_0.
+- destruct (env!i). apply co_sizeof_alignof. apply Zdivide_0.
Qed.
(** ** Size and alignment for composite definitions *)
@@ -399,7 +399,7 @@ Fixpoint sizeof_union (env: composite_env) (m: members) : Z :=
Lemma alignof_composite_two_p:
forall env m, exists n, alignof_composite env m = two_power_nat n.
Proof.
- induction m as [|[id t]]; simpl.
+ induction m as [|[id t]]; simpl.
- exists 0%nat; auto.
- apply Z.max_case; auto. apply alignof_two_p.
Qed.
@@ -408,8 +408,8 @@ Lemma alignof_composite_pos:
forall env m a, align_attr a (alignof_composite env m) > 0.
Proof.
intros.
- exploit align_attr_two_p. apply (alignof_composite_two_p env m).
- instantiate (1 := a). intros [n EQ].
+ exploit align_attr_two_p. apply (alignof_composite_two_p env m).
+ instantiate (1 := a). intros [n EQ].
rewrite EQ; apply two_power_nat_pos.
Qed.
@@ -418,10 +418,10 @@ Lemma sizeof_struct_incr:
Proof.
induction m as [|[id t]]; simpl; intros.
- omega.
-- apply Zle_trans with (align cur (alignof env t)).
+- apply Zle_trans with (align cur (alignof env t)).
apply align_le. apply alignof_pos.
apply Zle_trans with (align cur (alignof env t) + sizeof env t).
- generalize (sizeof_pos env t); omega.
+ generalize (sizeof_pos env t); omega.
apply IHm.
Qed.
@@ -457,7 +457,7 @@ Fixpoint field_type (id: ident) (fld: members) {struct fld} : res type :=
| (id', t) :: fld' => if ident_eq id id' then OK t else field_type id fld'
end.
-(** Some sanity checks about field offsets. First, field offsets are
+(** Some sanity checks about field offsets. First, field offsets are
within the range of acceptable offsets. *)
Remark field_offset_rec_in_range:
@@ -470,9 +470,9 @@ Proof.
- destruct (ident_eq id i); intros.
inv H. inv H0. split.
apply align_le. apply alignof_pos. apply sizeof_struct_incr.
- exploit IHfld; eauto. intros [A B]. split; auto.
+ exploit IHfld; eauto. intros [A B]. split; auto.
eapply Zle_trans; eauto. apply Zle_trans with (align pos (alignof env t)).
- apply align_le. apply alignof_pos. generalize (sizeof_pos env t). omega.
+ apply align_le. apply alignof_pos. generalize (sizeof_pos env t). omega.
Qed.
Lemma field_offset_in_range:
@@ -496,11 +496,11 @@ Proof.
induction fld as [|[i t]]; simpl; intros.
- discriminate.
- destruct (ident_eq id1 i); destruct (ident_eq id2 i).
-+ congruence.
++ congruence.
+ subst i. inv H; inv H0.
- exploit field_offset_rec_in_range. eexact H1. eauto. tauto.
+ exploit field_offset_rec_in_range. eexact H1. eauto. tauto.
+ subst i. inv H1; inv H2.
- exploit field_offset_rec_in_range. eexact H. eauto. tauto.
+ exploit field_offset_rec_in_range. eexact H. eauto. tauto.
+ eapply IHfld; eauto.
Qed.
@@ -512,10 +512,10 @@ Lemma field_offset_prefix:
field_offset env id fld1 = OK ofs ->
field_offset env id (fld1 ++ fld2) = OK ofs.
Proof.
- intros until fld1. unfold field_offset. generalize 0 as pos.
+ intros until fld1. unfold field_offset. generalize 0 as pos.
induction fld1 as [|[i t]]; simpl; intros.
- discriminate.
-- destruct (ident_eq id i); auto.
+- destruct (ident_eq id i); auto.
Qed.
(** Fourth, the position of each field respects its alignment. *)
@@ -525,7 +525,7 @@ Lemma field_offset_aligned:
field_offset env id fld = OK ofs -> field_type id fld = OK ty ->
(alignof env ty | ofs).
Proof.
- intros until ty. unfold field_offset. generalize 0 as pos. revert fld.
+ intros until ty. unfold field_offset. generalize 0 as pos. revert fld.
induction fld as [|[i t]]; simpl; intros.
- discriminate.
- destruct (ident_eq id i).
@@ -613,18 +613,18 @@ Proof.
assert (X: forall co, let a := Zmin 8 (co_alignof co) in
a = 1 \/ a = 2 \/ a = 4 \/ a = 8).
{
- intros. destruct (co_alignof_two_p co) as [n EQ]. unfold a; rewrite EQ.
+ intros. destruct (co_alignof_two_p co) as [n EQ]. unfold a; rewrite EQ.
destruct n; auto.
destruct n; auto.
destruct n; auto.
right; right; right. apply Z.min_l.
- rewrite two_power_nat_two_p. rewrite ! inj_S.
+ rewrite two_power_nat_two_p. rewrite ! inj_S.
change 8 with (two_p 3). apply two_p_monotone. omega.
}
induction ty; simpl; auto.
destruct i; auto.
destruct f; auto.
- destruct (env!i); auto.
+ destruct (env!i); auto.
destruct (env!i); auto.
Qed.
@@ -644,9 +644,9 @@ Proof.
destruct n. apply Zdivide_refl.
destruct n. apply Zdivide_refl.
destruct n. apply Zdivide_refl.
- apply Z.min_case.
- exists (two_p (Z.of_nat n)).
- change 8 with (two_p 3).
+ apply Z.min_case.
+ exists (two_p (Z.of_nat n)).
+ change 8 with (two_p 3).
rewrite <- two_p_is_exp by omega.
rewrite two_power_nat_two_p. rewrite !inj_S. f_equal. omega.
apply Zdivide_refl.
@@ -659,8 +659,8 @@ Proof.
apply Zdivide_refl.
apply Z.divide_mul_l. auto.
apply Zdivide_refl.
- destruct (env!i). apply X. apply Zdivide_0.
- destruct (env!i). apply X. apply Zdivide_0.
+ destruct (env!i). apply X. apply Zdivide_0.
+ destruct (env!i). apply X. apply Zdivide_0.
Qed.
(** Type ranks *)
@@ -761,7 +761,7 @@ Qed.
The size and alignment of the composite are determined at this time.
The alignment takes into account the [__Alignas] attributes
associated with the definition. The size is rounded up to a multiple
- of the alignment.
+ of the alignment.
The conversion fails if a type of a member is not complete. This rules
out incorrect recursive definitions such as
@@ -800,7 +800,7 @@ Program Definition composite_of_def
co_sizeof_alignof := _ |}
end.
Next Obligation.
- apply Zle_ge. eapply Zle_trans. eapply sizeof_composite_pos.
+ apply Zle_ge. eapply Zle_trans. eapply sizeof_composite_pos.
apply align_le; apply alignof_composite_pos.
Defined.
Next Obligation.
@@ -842,9 +842,9 @@ Lemma alignof_stable:
forall t, complete_type env t = true -> alignof env' t = alignof env t.
Proof.
induction t; simpl; intros; f_equal; auto.
- destruct (env!i) as [co|] eqn:E; try discriminate.
+ destruct (env!i) as [co|] eqn:E; try discriminate.
erewrite extends by eauto. auto.
- destruct (env!i) as [co|] eqn:E; try discriminate.
+ destruct (env!i) as [co|] eqn:E; try discriminate.
erewrite extends by eauto. auto.
Qed.
@@ -853,9 +853,9 @@ Lemma sizeof_stable:
Proof.
induction t; simpl; intros; auto.
rewrite IHt by auto. auto.
- destruct (env!i) as [co|] eqn:E; try discriminate.
+ destruct (env!i) as [co|] eqn:E; try discriminate.
erewrite extends by eauto. auto.
- destruct (env!i) as [co|] eqn:E; try discriminate.
+ destruct (env!i) as [co|] eqn:E; try discriminate.
erewrite extends by eauto. auto.
Qed.
@@ -863,9 +863,9 @@ Lemma complete_type_stable:
forall t, complete_type env t = true -> complete_type env' t = true.
Proof.
induction t; simpl; intros; auto.
- destruct (env!i) as [co|] eqn:E; try discriminate.
+ destruct (env!i) as [co|] eqn:E; try discriminate.
erewrite extends by eauto. auto.
- destruct (env!i) as [co|] eqn:E; try discriminate.
+ destruct (env!i) as [co|] eqn:E; try discriminate.
erewrite extends by eauto. auto.
Qed.
@@ -873,16 +873,16 @@ Lemma rank_type_stable:
forall t, complete_type env t = true -> rank_type env' t = rank_type env t.
Proof.
induction t; simpl; intros; auto.
- destruct (env!i) as [co|] eqn:E; try discriminate.
+ destruct (env!i) as [co|] eqn:E; try discriminate.
erewrite extends by eauto. auto.
- destruct (env!i) as [co|] eqn:E; try discriminate.
+ destruct (env!i) as [co|] eqn:E; try discriminate.
erewrite extends by eauto. auto.
Qed.
Lemma alignof_composite_stable:
forall m, complete_members env m = true -> alignof_composite env' m = alignof_composite env m.
Proof.
- induction m as [|[id t]]; simpl; intros.
+ induction m as [|[id t]]; simpl; intros.
auto.
InvBooleans. rewrite alignof_stable by auto. rewrite IHm by auto. auto.
Qed.
@@ -892,7 +892,7 @@ Lemma sizeof_struct_stable:
Proof.
induction m as [|[id t]]; simpl; intros.
auto.
- InvBooleans. rewrite alignof_stable by auto. rewrite sizeof_stable by auto.
+ InvBooleans. rewrite alignof_stable by auto. rewrite sizeof_stable by auto.
rewrite IHm by auto. auto.
Qed.
@@ -907,8 +907,8 @@ Qed.
Lemma sizeof_composite_stable:
forall su m, complete_members env m = true -> sizeof_composite env' su m = sizeof_composite env su m.
Proof.
- intros. destruct su; simpl.
- apply sizeof_struct_stable; auto.
+ intros. destruct su; simpl.
+ apply sizeof_struct_stable; auto.
apply sizeof_union_stable; auto.
Qed.
@@ -937,7 +937,7 @@ Lemma add_composite_definitions_incr:
Proof.
induction defs; simpl; intros.
- inv H; auto.
-- destruct a; monadInv H.
+- destruct a; monadInv H.
eapply IHdefs; eauto. rewrite PTree.gso; auto.
red; intros; subst id0. unfold composite_of_def in EQ. rewrite H0 in EQ; discriminate.
Qed.
@@ -974,22 +974,22 @@ Proof.
eapply IHdefs; eauto.
set (env1 := PTree.set id x env0) in *.
unfold composite_of_def in EQ.
- destruct (env0!id) eqn:E; try discriminate.
+ destruct (env0!id) eqn:E; try discriminate.
destruct (complete_members env0 m) eqn:C; inversion EQ; clear EQ.
assert (forall id1 co1, env0!id1 = Some co1 -> env1!id1 = Some co1).
{ intros. unfold env1. rewrite PTree.gso; auto. congruence. }
red; intros. unfold env1 in H2; rewrite PTree.gsspec in H2; destruct (peq id0 id).
-+ subst id0. inversion H2; clear H2. subst co.
++ subst id0. inversion H2; clear H2. subst co.
(*
assert (A: alignof_composite env1 m = alignof_composite env0 m)
by (apply alignof_composite_stable; assumption).
*)
rewrite <- H1; constructor; simpl.
-* eapply complete_members_stable; eauto.
-* f_equal. symmetry. apply alignof_composite_stable; auto.
-* f_equal. symmetry. apply sizeof_composite_stable; auto.
+* eapply complete_members_stable; eauto.
+* f_equal. symmetry. apply alignof_composite_stable; auto.
+* f_equal. symmetry. apply sizeof_composite_stable; auto.
* symmetry. apply rank_members_stable; auto.
-+ exploit H0; eauto. intros [P Q R S].
++ exploit H0; eauto. intros [P Q R S].
constructor; intros.
* eapply complete_members_stable; eauto.
* rewrite Q. f_equal. symmetry. apply alignof_composite_stable; auto.
@@ -1005,13 +1005,13 @@ Theorem build_composite_env_charact:
In (Composite id su m a) defs ->
exists co, env!id = Some co /\ co_members co = m /\ co_attr co = a /\ co_su co = su.
Proof.
- intros until defs. unfold build_composite_env. generalize (PTree.empty composite) as env0.
+ intros until defs. unfold build_composite_env. generalize (PTree.empty composite) as env0.
revert defs. induction defs as [|d1 defs]; simpl; intros.
- contradiction.
- destruct d1; monadInv H.
- destruct H0; [idtac|eapply IHdefs;eauto]. inv H.
- unfold composite_of_def in EQ.
- destruct (env0!id) eqn:E; try discriminate.
+ destruct H0; [idtac|eapply IHdefs;eauto]. inv H.
+ unfold composite_of_def in EQ.
+ destruct (env0!id) eqn:E; try discriminate.
destruct (complete_members env0 m) eqn:C; simplify_eq EQ. clear EQ; intros EQ.
exists x.
split. eapply add_composite_definitions_incr; eauto. apply PTree.gss.
diff --git a/cfrontend/Ctyping.v b/cfrontend/Ctyping.v
index cde9ad11..aa320f20 100644
--- a/cfrontend/Ctyping.v
+++ b/cfrontend/Ctyping.v
@@ -180,7 +180,7 @@ Definition floatsize_eq: forall (x y: floatsize), {x=y} + {x<>y}.
Proof. decide equality. Defined.
Definition callconv_combine (cc1 cc2: calling_convention) : res calling_convention :=
- if bool_eq cc1.(cc_vararg) cc2.(cc_vararg) then
+ if bool_eq cc1.(cc_vararg) cc2.(cc_vararg) then
OK {| cc_vararg := cc1.(cc_vararg);
cc_unproto := cc1.(cc_unproto) && cc2.(cc_unproto);
cc_structret := cc1.(cc_structret) |}
@@ -356,11 +356,11 @@ Inductive wt_rvalue : expr -> Prop :=
wt_rvalue r -> wt_cast (typeof r) ty ->
wt_rvalue (Ecast r ty)
| wt_Eseqand: forall r1 r2,
- wt_rvalue r1 -> wt_rvalue r2 ->
+ wt_rvalue r1 -> wt_rvalue r2 ->
wt_bool (typeof r1) -> wt_bool (typeof r2) ->
wt_rvalue (Eseqand r1 r2 (Tint I32 Signed noattr))
| wt_Eseqor: forall r1 r2,
- wt_rvalue r1 -> wt_rvalue r2 ->
+ wt_rvalue r1 -> wt_rvalue r2 ->
wt_bool (typeof r1) -> wt_bool (typeof r2) ->
wt_rvalue (Eseqor r1 r2 (Tint I32 Signed noattr))
| wt_Econdition: forall r1 r2 r3 ty,
@@ -873,7 +873,7 @@ Fixpoint retype_stmt (ce: composite_env) (e: typenv) (rt: type) (s: statement) :
| Sreturn None =>
OK (Sreturn None)
| Sreturn (Some a) =>
- do a' <- retype_expr ce e a;
+ do a' <- retype_expr ce e a;
sreturn rt a'
| Sswitch a sl =>
do a' <- retype_expr ce e a;
@@ -924,14 +924,14 @@ Definition typecheck_program (p: program) : res program :=
Lemma check_cast_sound:
forall t1 t2 x, check_cast t1 t2 = OK x -> wt_cast t1 t2.
Proof.
- unfold check_cast, wt_cast; intros.
+ unfold check_cast, wt_cast; intros.
destruct (classify_cast t1 t2); congruence.
Qed.
Lemma check_bool_sound:
forall t x, check_bool t = OK x -> wt_bool t.
Proof.
- unfold check_bool, wt_bool; intros.
+ unfold check_bool, wt_bool; intros.
destruct (classify_bool t); congruence.
Qed.
@@ -940,7 +940,7 @@ Hint Resolve check_cast_sound check_bool_sound: ty.
Lemma check_arguments_sound:
forall el tl x, check_arguments el tl = OK x -> wt_arguments el tl.
Proof.
- intros el tl; revert tl el.
+ intros el tl; revert tl el.
induction tl; destruct el; simpl; intros; try discriminate.
constructor.
destruct strict eqn:S; try discriminate. constructor; auto.
@@ -970,13 +970,13 @@ Qed.
Lemma typeconv_cast:
forall t1 t2, wt_cast (typeconv t1) t2 -> wt_cast t1 t2.
Proof.
- unfold typeconv, wt_cast; intros. destruct t1; auto.
+ unfold typeconv, wt_cast; intros. destruct t1; auto.
assert (classify_cast (Tint I32 Signed a) t2 <> cast_case_default ->
classify_cast (Tint i s a) t2 <> cast_case_default).
{
unfold classify_cast. destruct t2; try congruence. destruct f; congruence.
}
- destruct i; auto.
+ destruct i; auto.
Qed.
Lemma type_combine_cast:
@@ -1006,20 +1006,20 @@ Lemma type_conditional_cast:
forall t1 t2 t,
type_conditional t1 t2 = OK t -> wt_cast t1 t /\ wt_cast t2 t.
Proof.
- intros.
+ intros.
assert (A: forall x, match typeconv x with Tarray _ _ _ => False | Tfunction _ _ _ => False | _ => True end).
{ destruct x; simpl; auto. destruct i; auto. }
assert (D: type_combine (typeconv t1) (typeconv t2) = OK t -> wt_cast t1 t /\ wt_cast t2 t).
- { intros. apply type_combine_cast in H0. destruct H0; split; apply typeconv_cast; auto.
+ { intros. apply type_combine_cast in H0. destruct H0; split; apply typeconv_cast; auto.
apply A. apply A. }
clear A. unfold type_conditional in H.
destruct (typeconv t1) eqn:T1; try discriminate;
destruct (typeconv t2) eqn:T2; inv H; eauto using D, binarith_type_cast.
-- split; apply typeconv_cast; unfold wt_cast.
+- split; apply typeconv_cast; unfold wt_cast.
rewrite T1; simpl; congruence. rewrite T2; simpl; congruence.
-- split; apply typeconv_cast; unfold wt_cast.
+- split; apply typeconv_cast; unfold wt_cast.
rewrite T1; simpl; congruence. rewrite T2; simpl; congruence.
-- split; apply typeconv_cast; unfold wt_cast.
+- split; apply typeconv_cast; unfold wt_cast.
rewrite T1; simpl; congruence. rewrite T2; simpl; congruence.
Qed.
@@ -1053,13 +1053,13 @@ Qed.
Lemma ederef_sound:
forall r a, ederef r = OK a -> wt_expr ce e r -> wt_expr ce e a.
Proof.
- intros. monadInv H. eauto with ty.
+ intros. monadInv H. eauto with ty.
Qed.
Lemma efield_sound:
forall r f a, efield ce r f = OK a -> wt_expr ce e r -> wt_expr ce e a.
Proof.
- intros. monadInv H.
+ intros. monadInv H.
destruct (typeof r) eqn:TR; try discriminate;
destruct (ce!i) as [co|] eqn:CE; monadInv EQ0; eauto with ty.
Qed.
@@ -1085,13 +1085,13 @@ Qed.
Lemma econst_float_sound:
forall n, wt_expr ce e (econst_float n).
Proof.
- unfold econst_float; auto with ty.
+ unfold econst_float; auto with ty.
Qed.
Lemma econst_single_sound:
forall n, wt_expr ce e (econst_single n).
Proof.
- unfold econst_single; auto with ty.
+ unfold econst_single; auto with ty.
Qed.
Lemma evalof_sound:
@@ -1140,14 +1140,14 @@ Lemma econdition_sound:
forall r1 r2 r3 a, econdition r1 r2 r3 = OK a ->
wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e r3 -> wt_expr ce e a.
Proof.
- intros. monadInv H. apply type_conditional_cast in EQ3. destruct EQ3. eauto 10 with ty.
+ intros. monadInv H. apply type_conditional_cast in EQ3. destruct EQ3. eauto 10 with ty.
Qed.
Lemma econdition'_sound:
forall r1 r2 r3 ty a, econdition' r1 r2 r3 ty = OK a ->
wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e r3 -> wt_expr ce e a.
Proof.
- intros. monadInv H. eauto 10 with ty.
+ intros. monadInv H. eauto 10 with ty.
Qed.
Lemma esizeof_sound:
@@ -1189,16 +1189,16 @@ Qed.
Lemma ecall_sound:
forall fn args a, ecall fn args = OK a -> wt_expr ce e fn -> wt_exprlist ce e args -> wt_expr ce e a.
Proof.
- intros. monadInv H.
- destruct (classify_fun (typeof fn)) eqn:CF; monadInv EQ2.
- econstructor; eauto with ty. eapply check_arguments_sound; eauto.
+ intros. monadInv H.
+ destruct (classify_fun (typeof fn)) eqn:CF; monadInv EQ2.
+ econstructor; eauto with ty. eapply check_arguments_sound; eauto.
Qed.
Lemma ebuiltin_sound:
forall ef tyargs args tyres a,
ebuiltin ef tyargs args tyres = OK a -> wt_exprlist ce e args -> wt_expr ce e a.
Proof.
- intros. monadInv H.
+ intros. monadInv H.
destruct (type_eq tyres Tvoid); simpl in EQ2; try discriminate.
destruct (opt_typ_eq (sig_res (ef_sig ef)) None); inv EQ2.
econstructor; eauto. eapply check_arguments_sound; eauto.
@@ -1242,14 +1242,14 @@ Qed.
Lemma sreturn_sound:
forall a s, sreturn rt a = OK s -> wt_expr ce e a -> wt_stmt ce e rt s.
Proof.
- intros. monadInv H; eauto with ty.
+ intros. monadInv H; eauto with ty.
Qed.
Lemma sswitch_sound:
forall a sl s, sswitch a sl = OK s ->
wt_expr ce e a -> wt_lblstmts ce e rt sl -> wt_stmt ce e rt s.
Proof.
- intros. monadInv H. destruct (typeof a) eqn:TA; inv EQ0.
+ intros. monadInv H. destruct (typeof a) eqn:TA; inv EQ0.
eauto with ty.
eapply wt_Sswitch with (sz := I32); eauto with ty.
Qed.
@@ -1262,11 +1262,11 @@ Proof.
- destruct a; simpl; intros a' RT; try (monadInv RT).
+ destruct v; try discriminate.
destruct ty; inv RT. apply econst_int_sound. apply econst_ptr_int_sound.
- destruct ty; inv RT. apply econst_long_sound.
+ destruct ty; inv RT. apply econst_long_sound.
inv RT. apply econst_float_sound.
inv RT. apply econst_single_sound.
+ eapply evar_sound; eauto.
-+ eapply efield_sound; eauto.
++ eapply efield_sound; eauto.
+ eapply evalof_sound; eauto.
+ eapply ederef_sound; eauto.
+ eapply eaddrof_sound; eauto.
@@ -1282,7 +1282,7 @@ Proof.
+ eapply eassignop_sound; eauto.
+ eapply epostincrdecr_sound; eauto.
+ eapply ecomma_sound; eauto.
-+ eapply ecall_sound; eauto.
++ eapply ecall_sound; eauto.
+ eapply ebuiltin_sound; eauto.
- destruct al; simpl; intros al' RT; monadInv RT.
+ constructor.
@@ -1297,7 +1297,7 @@ Proof.
- destruct s; simpl; intros s' RT; try (monadInv RT).
+ constructor.
+ eapply sdo_sound; eauto using retype_expr_sound.
-+ constructor; eauto.
++ constructor; eauto.
+ eapply sifthenelse_sound; eauto using retype_expr_sound.
+ eapply swhile_sound; eauto using retype_expr_sound.
+ eapply sdowhile_sound; eauto using retype_expr_sound.
@@ -1306,9 +1306,9 @@ Proof.
+ constructor.
+ destruct o; monadInv RT. eapply sreturn_sound; eauto using retype_expr_sound. constructor.
+ eapply sswitch_sound; eauto using retype_expr_sound.
-+ constructor; eauto.
++ constructor; eauto.
+ constructor.
-- destruct sl; simpl; intros sl' RT; monadInv RT.
+- destruct sl; simpl; intros sl' RT; monadInv RT.
+ constructor.
+ constructor; eauto.
Qed.
@@ -1318,13 +1318,13 @@ End SOUNDNESS_CONSTRUCTORS.
Lemma retype_function_sound:
forall ce e f f', retype_function ce e f = OK f' -> wt_function ce e f'.
Proof.
- intros. monadInv H. constructor; simpl. eapply retype_stmt_sound; eauto.
+ intros. monadInv H. constructor; simpl. eapply retype_stmt_sound; eauto.
Qed.
Theorem typecheck_program_sound:
forall p p', typecheck_program p = OK p' -> wt_program p'.
Proof.
- unfold typecheck_program; intros. monadInv H.
+ unfold typecheck_program; intros. monadInv H.
rename x into defs.
constructor; simpl.
set (ce := prog_comp_env p) in *.
@@ -1335,22 +1335,22 @@ Proof.
{
revert EQ; generalize (prog_defs p) defs.
induction l as [ | [id gd] l ]; intros l'; simpl; intros.
- inv EQ. constructor.
+ inv EQ. constructor.
destruct gd as [f | v].
destruct (retype_fundef ce e f) as [tf|msg] eqn:R; monadInv EQ.
- constructor; auto. constructor; auto.
+ constructor; auto. constructor; auto.
monadInv EQ. constructor; auto. destruct v; constructor; auto. }
assert (ENVS: e' = e).
- { unfold e, e'. revert MATCH; generalize (prog_defs p) defs (PTree.empty type).
+ { unfold e, e'. revert MATCH; generalize (prog_defs p) defs (PTree.empty type).
induction l as [ | [id gd] l ]; intros l' t M; inv M.
- auto. inv H1; simpl; auto. replace (type_of_fundef f2) with (type_of_fundef f1); auto.
- unfold retype_fundef in H4. destruct f1; monadInv H4; auto. monadInv EQ0; auto.
+ auto. inv H1; simpl; auto. replace (type_of_fundef f2) with (type_of_fundef f1); auto.
+ unfold retype_fundef in H4. destruct f1; monadInv H4; auto. monadInv EQ0; auto.
}
rewrite ENVS.
intros id f. revert MATCH; generalize (prog_defs p) defs. induction 1; simpl; intros.
contradiction.
- destruct H0; auto. subst b1; inv H. destruct f1; simpl in H2.
- monadInv H2. eapply retype_function_sound; eauto. congruence.
+ destruct H0; auto. subst b1; inv H. destruct f1; simpl in H2.
+ monadInv H2. eapply retype_function_sound; eauto. congruence.
Qed.
(** * Subject reduction *)
@@ -1360,7 +1360,7 @@ Qed.
Lemma pres_cast_int_int:
forall sz sg n, wt_int (cast_int_int sz sg n) sz sg.
Proof.
- intros. unfold cast_int_int. destruct sz; simpl.
+ intros. unfold cast_int_int. destruct sz; simpl.
- destruct sg. apply Int.sign_ext_idem; omega. apply Int.zero_ext_idem; omega.
- destruct sg. apply Int.sign_ext_idem; omega. apply Int.zero_ext_idem; omega.
- auto.
@@ -1389,7 +1389,7 @@ Proof.
Qed.
Lemma pres_sem_binarith:
- forall
+ forall
(sem_int: signedness -> int -> int -> option val)
(sem_long: signedness -> int64 -> int64 -> option val)
(sem_float: float -> float -> option val)
@@ -1411,19 +1411,19 @@ Proof with (try discriminate).
set (ty' := Cop.binarith_type (classify_binarith ty1 ty2)) in *.
destruct (sem_cast v1 ty1 ty') as [v1'|] eqn:CAST1...
destruct (sem_cast v2 ty2 ty') as [v2'|] eqn:CAST2...
- DestructCases.
-- specialize (H s i i0). rewrite H3 in H.
+ DestructCases.
+- specialize (H s i i0). rewrite H3 in H.
destruct v; auto with ty; contradiction.
-- specialize (H0 s i i0). rewrite H3 in H0.
+- specialize (H0 s i i0). rewrite H3 in H0.
destruct v; auto with ty; contradiction.
-- specialize (H1 f f0). rewrite H3 in H1.
+- specialize (H1 f f0). rewrite H3 in H1.
destruct v; auto with ty; contradiction.
-- specialize (H2 f f0). rewrite H3 in H2.
+- specialize (H2 f f0). rewrite H3 in H2.
destruct v; auto with ty; contradiction.
Qed.
Lemma pres_sem_binarith_int:
- forall
+ forall
(sem_int: signedness -> int -> int -> option val)
(sem_long: signedness -> int64 -> int64 -> option val)
v1 ty1 v2 ty2 v ty msg,
@@ -1435,9 +1435,9 @@ Lemma pres_sem_binarith_int:
binarith_int_type ty1 ty2 msg = OK ty ->
wt_val v ty.
Proof.
- intros. eapply pres_sem_binarith with (msg := msg); eauto.
+ intros. eapply pres_sem_binarith with (msg := msg); eauto.
simpl; auto. simpl; auto.
- unfold binarith_int_type, binarith_type in *.
+ unfold binarith_int_type, binarith_type in *.
destruct (classify_binarith ty1 ty2); congruence.
Qed.
@@ -1463,13 +1463,13 @@ Proof with (try discriminate).
}
assert (Y: forall ob, option_map Val.of_bool ob = Some v -> wt_val v (Tint I32 Signed noattr)).
{
- intros ob EQ. destruct ob as [b|]; inv EQ. eauto.
+ intros ob EQ. destruct ob as [b|]; inv EQ. eauto.
}
destruct (classify_cmp ty1 ty2).
- inv H; eauto.
- DestructCases; eauto.
- DestructCases; eauto.
-- unfold sem_binarith in H0.
+- unfold sem_binarith in H0.
set (ty' := Cop.binarith_type (classify_binarith ty1 ty2)) in *.
destruct (sem_cast v1 ty1 ty') as [v1'|]...
destruct (sem_cast v2 ty2 ty') as [v2'|]...
@@ -1510,16 +1510,16 @@ Proof.
- (* xor *)
unfold sem_xor in SEM. eapply pres_sem_binarith_int; eauto; intros; exact I.
- (* shl *)
- unfold sem_shl in SEM. eapply pres_sem_shift; eauto.
+ unfold sem_shl in SEM. eapply pres_sem_shift; eauto.
- (* shr *)
- unfold sem_shr in SEM. eapply pres_sem_shift; eauto.
+ unfold sem_shr in SEM. eapply pres_sem_shift; eauto.
- (* comparisons *)
- eapply pres_sem_cmp; eauto.
-- eapply pres_sem_cmp; eauto.
-- eapply pres_sem_cmp; eauto.
-- eapply pres_sem_cmp; eauto.
-- eapply pres_sem_cmp; eauto.
-- eapply pres_sem_cmp; eauto.
+ eapply pres_sem_cmp; eauto.
+- eapply pres_sem_cmp; eauto.
+- eapply pres_sem_cmp; eauto.
+- eapply pres_sem_cmp; eauto.
+- eapply pres_sem_cmp; eauto.
+- eapply pres_sem_cmp; eauto.
Qed.
Lemma pres_sem_unop:
@@ -1529,7 +1529,7 @@ Lemma pres_sem_unop:
wt_val v1 ty1 ->
wt_val v ty.
Proof.
- intros until v; intros TY SEM WT1.
+ intros until v; intros TY SEM WT1.
destruct op; simpl in TY; simpl in SEM.
- (* notbool *)
unfold sem_notbool in SEM; DestructCases.
@@ -1542,10 +1542,10 @@ Proof.
- (* notint *)
unfold sem_notint in SEM; DestructCases; auto with ty.
- (* neg *)
- unfold sem_neg in SEM; DestructCases; auto with ty.
+ unfold sem_neg in SEM; DestructCases; auto with ty.
- (* absfloat *)
unfold sem_absfloat in SEM; DestructCases; auto with ty.
-Qed.
+Qed.
Lemma wt_load_result:
forall ty chunk v,
@@ -1569,10 +1569,10 @@ Lemma wt_decode_val:
access_mode ty = By_value chunk ->
wt_val (decode_val chunk vl) ty.
Proof.
- intros until vl; intros ACC.
+ intros until vl; intros ACC.
destruct ty; simpl in ACC; try discriminate.
- destruct i; [destruct s|destruct s|idtac|idtac]; inv ACC; unfold decode_val;
- destruct (proj_bytes vl); auto with ty.
+ destruct (proj_bytes vl); auto with ty.
constructor; red. apply Int.sign_ext_idem; omega.
constructor; red. apply Int.zero_ext_idem; omega.
constructor; red. apply Int.sign_ext_idem; omega.
@@ -1643,7 +1643,7 @@ Lemma type_add_int32s:
type_binop Oadd ty1 type_int32s = OK ty2 ->
ty2 = incrdecr_type ty1.
Proof.
- simpl; intros. unfold classify_add in H; destruct ty1; simpl in H;
+ simpl; intros. unfold classify_add in H; destruct ty1; simpl in H;
try (eapply binarith_type_int32s; eauto; fail).
destruct i; eapply binarith_type_int32s; eauto.
inv H; auto.
@@ -1656,7 +1656,7 @@ Lemma type_sub_int32s:
type_binop Osub ty1 type_int32s = OK ty2 ->
ty2 = incrdecr_type ty1.
Proof.
- simpl; intros. unfold classify_sub in H; destruct ty1; simpl in H;
+ simpl; intros. unfold classify_sub in H; destruct ty1; simpl in H;
try (eapply binarith_type_int32s; eauto; fail).
destruct i; eapply binarith_type_int32s; eauto.
inv H; auto.
@@ -1669,37 +1669,37 @@ Lemma wt_rred:
rred ge a m t a' m' -> wt_rvalue ge tenv a -> wt_rvalue ge tenv a'.
Proof.
induction 1; intros WT; inversion WT.
-- (* valof *) simpl in *. constructor. eapply wt_deref_loc; eauto.
+- (* valof *) simpl in *. constructor. eapply wt_deref_loc; eauto.
- (* addrof *) constructor; auto with ty.
-- (* unop *) simpl in H4. inv H2. constructor. eapply pres_sem_unop; eauto.
-- (* binop *)
+- (* unop *) simpl in H4. inv H2. constructor. eapply pres_sem_unop; eauto.
+- (* binop *)
simpl in H6. inv H3. inv H5. constructor. eapply pres_sem_binop; eauto.
- (* cast *) inv H2. constructor. eapply pres_sem_cast; eauto.
- (* sequand true *) subst. constructor. auto. apply wt_bool_cast; auto.
- red; intros. inv H0; auto with ty.
+ red; intros. inv H0; auto with ty.
- (* sequand false *) constructor. auto with ty.
- (* seqor true *) constructor. auto with ty.
- (* seqor false *) subst. constructor. auto. apply wt_bool_cast; auto.
- red; intros. inv H0; auto with ty.
-- (* condition *) constructor. destruct b; auto. destruct b; auto. red; auto.
+ red; intros. inv H0; auto with ty.
+- (* condition *) constructor. destruct b; auto. destruct b; auto. red; auto.
- (* sizeof *) constructor; auto with ty.
- (* alignof *) constructor; auto with ty.
- (* assign *) inversion H5. constructor. eapply pres_sem_cast; eauto.
-- (* assignop *) subst tyres l r. constructor. auto.
+- (* assignop *) subst tyres l r. constructor. auto.
constructor. constructor. eapply wt_deref_loc; eauto.
- auto. auto. auto.
-- (* postincr *) simpl in *. subst id0 l.
+ auto. auto. auto.
+- (* postincr *) simpl in *. subst id0 l.
exploit wt_deref_loc; eauto. intros WTV1.
- constructor.
- constructor. auto. rewrite <- H0 in H5. constructor.
+ constructor.
+ constructor. auto. rewrite <- H0 in H5. constructor.
constructor; auto. constructor. constructor. auto with ty.
- subst op. destruct id.
+ subst op. destruct id.
erewrite <- type_add_int32s by eauto. auto.
erewrite <- type_sub_int32s by eauto. auto.
simpl; auto.
constructor; auto.
- (* comma *) auto.
-- (* paren *) inv H3. constructor. apply H5. eapply pres_sem_cast; eauto.
+- (* paren *) inv H3. constructor. apply H5. eapply pres_sem_cast; eauto.
- (* builtin *) subst. auto with ty.
Qed.
@@ -1707,7 +1707,7 @@ Lemma wt_lred:
forall tenv ge e a m a' m',
lred ge e a m a' m' -> wt_lvalue ge tenv a -> wt_lvalue ge tenv a'.
Proof.
- induction 1; intros WT; constructor.
+ induction 1; intros WT; constructor.
Qed.
Lemma rred_same_type:
@@ -1733,7 +1733,7 @@ Hypothesis SAMETY: typeof a' = typeof a.
Lemma wt_subexpr:
forall from to C,
- context from to C ->
+ context from to C ->
wt_expr_kind cenv tenv to (C a) ->
wt_expr_kind cenv tenv from a
with wt_subexprlist:
@@ -1756,10 +1756,10 @@ Lemma wt_arguments_context:
forall k C, contextlist k C ->
forall tyl, wt_arguments (C a) tyl -> wt_arguments (C a') tyl.
Proof.
- induction 1; intros.
+ induction 1; intros.
- inv H0. constructor; auto. rewrite (typeof_context _ _ _ H); auto.
constructor; auto.
-- inv H0. constructor; auto. constructor; auto.
+- inv H0. constructor; auto. constructor; auto.
Qed.
Lemma wt_context:
@@ -1777,11 +1777,11 @@ with wt_contextlist:
Proof.
- induction 1; intros WT BASE;
auto;
- inv WT;
+ inv WT;
try (pose (EQTY := typeof_context _ _ _ H); rewrite <- ? EQTY; econstructor;
try (apply IHcontext; assumption); rewrite ? EQTY; eauto).
-* red. econstructor; eauto. eapply wt_arguments_context; eauto.
-* red. econstructor; eauto. eapply wt_arguments_context; eauto.
+* red. econstructor; eauto. eapply wt_arguments_context; eauto.
+* red. econstructor; eauto. eapply wt_arguments_context; eauto.
- induction 1; intros WT BASE.
* inv WT. constructor. apply (wt_context _ _ _ H); auto. auto.
* inv WT. constructor; auto.
@@ -1798,9 +1798,9 @@ Proof.
unfold select_switch; intros.
assert (A: wt_lblstmts ce e rt (select_switch_default sl)).
{
- revert sl H. induction 1; simpl; intros.
+ revert sl H. induction 1; simpl; intros.
constructor.
- destruct case. auto. constructor; auto.
+ destruct case. auto. constructor; auto.
}
assert (B: forall sl', select_switch_case n sl = Some sl' -> wt_lblstmts ce e rt sl').
{
@@ -1812,7 +1812,7 @@ Proof.
Qed.
Lemma wt_seq_of_ls:
- forall ce e rt sl,
+ forall ce e rt sl,
wt_lblstmts ce e rt sl -> wt_stmt ce e rt (seq_of_labeled_statement sl).
Proof.
induction 1; simpl.
@@ -1830,7 +1830,7 @@ Let ge := globalenv prog.
Let gtenv := bind_globdef (PTree.empty _) prog.(prog_defs).
Hypothesis WT_EXTERNAL:
- forall id ef args res cc vargs m t vres m',
+ forall id ef args res cc vargs m t vres m',
In (id, Gfun (External ef args res cc)) prog.(prog_defs) ->
external_call ef ge vargs m t vres m' ->
wt_val vres res.
@@ -1916,7 +1916,7 @@ Qed.
Lemma wt_call_cont_stmt_cont:
forall te f k, wt_call_cont k f.(fn_return) -> wt_stmt_cont te f k.
Proof.
- intros. inversion H; subst. constructor. constructor; auto.
+ intros. inversion H; subst. constructor. constructor; auto.
Qed.
Lemma call_cont_wt:
@@ -1984,7 +1984,7 @@ Scheme wt_stmt_ind2 := Minimality for wt_stmt Sort Prop
Lemma wt_find_label:
forall lbl e f s, wt_stmt ge e f.(fn_return) s ->
- forall k s' k',
+ forall k s' k',
find_label lbl s k = Some (s', k') ->
wt_stmt_cont e f k ->
wt_stmt ge e f.(fn_return) s' /\ wt_stmt_cont e f k'.
@@ -1998,10 +1998,10 @@ Proof.
wt_stmt ge e f.(fn_return) s' /\ wt_stmt_cont e f k');
simpl; intros; try discriminate.
+ destruct (find_label lbl s1 (Kseq s2 k)) as [[sx kx] | ] eqn:F.
- inv H3. eauto with ty.
+ inv H3. eauto with ty.
eauto with ty.
+ destruct (find_label lbl s1 k) as [[sx kx] | ] eqn:F.
- inv H5. eauto with ty.
+ inv H5. eauto with ty.
eauto with ty.
+ eauto with ty.
+ eauto with ty.
@@ -2011,13 +2011,13 @@ Proof.
inv H7. eauto with ty.
eauto with ty.
+ eauto with ty.
- + destruct (ident_eq lbl lbl0).
- inv H1. auto.
+ + destruct (ident_eq lbl lbl0).
+ inv H1. auto.
eauto.
+ destruct (find_label lbl s (Kseq (seq_of_labeled_statement ls) k)) as [[sx kx] | ] eqn:F.
- inv H4. eapply H0; eauto. constructor. auto. apply wt_seq_of_ls; auto.
+ inv H4. eapply H0; eauto. constructor. auto. apply wt_seq_of_ls; auto.
eauto.
- + assumption.
+ + assumption.
Qed.
End WT_FIND_LABEL.
@@ -2031,21 +2031,21 @@ Proof.
econstructor; eauto. change (wt_expr_kind ge te RV (C a')).
eapply wt_context with (a := a); eauto.
eapply lred_same_type; eauto.
- eapply wt_lred; eauto. change (wt_expr_kind ge te LV a). eapply wt_subexpr; eauto.
+ eapply wt_lred; eauto. change (wt_expr_kind ge te LV a). eapply wt_subexpr; eauto.
- (* rred *)
econstructor; eauto. change (wt_expr_kind ge te RV (C a')).
eapply wt_context with (a := a); eauto.
eapply rred_same_type; eauto.
- eapply wt_rred; eauto. change (wt_expr_kind ge te RV a). eapply wt_subexpr; eauto.
+ eapply wt_rred; eauto. change (wt_expr_kind ge te RV a). eapply wt_subexpr; eauto.
- (* call *)
- assert (A: wt_expr_kind ge te RV a) by (eapply wt_subexpr; eauto).
+ assert (A: wt_expr_kind ge te RV a) by (eapply wt_subexpr; eauto).
simpl in A. inv H. inv A. simpl in H9; rewrite H4 in H9; inv H9.
assert (fundef_return fd = ty).
{ destruct fd; simpl in *.
unfold type_of_function in H3. congruence.
congruence. }
econstructor.
- rewrite H. econstructor; eauto.
+ rewrite H. econstructor; eauto.
intros. change (wt_expr_kind ge te RV (C (Eval v ty))).
eapply wt_context with (a := Ecall (Eval vf tyf) el ty); eauto.
red; constructor; auto.
@@ -2062,7 +2062,7 @@ Proof.
- inv WTS; eauto with ty.
- inv WTK; eauto with ty.
- inv WTS; eauto with ty.
-- inv WTK; eauto with ty.
+- inv WTK; eauto with ty.
- inv WTK; eauto with ty.
- inv WTK; eauto with ty.
- inv WTS; eauto with ty.
@@ -2084,35 +2084,35 @@ Proof.
- inv WTK; eauto with ty.
- inv WTK; eauto with ty.
- inv WTK; eauto with ty.
-- econstructor. eapply call_cont_wt; eauto. constructor.
-- inv WTS. eauto with ty.
-- inv WTK. econstructor. eapply call_cont_wt; eauto.
- inv WTE. eapply pres_sem_cast; eauto.
-- econstructor. eapply is_wt_call_cont; eauto. constructor.
-- inv WTS; eauto with ty.
-- inv WTK. econstructor; eauto with ty.
- apply wt_seq_of_ls. apply wt_select_switch; auto.
+- econstructor. eapply call_cont_wt; eauto. constructor.
+- inv WTS. eauto with ty.
+- inv WTK. econstructor. eapply call_cont_wt; eauto.
+ inv WTE. eapply pres_sem_cast; eauto.
+- econstructor. eapply is_wt_call_cont; eauto. constructor.
+- inv WTS; eauto with ty.
+- inv WTK. econstructor; eauto with ty.
+ apply wt_seq_of_ls. apply wt_select_switch; auto.
- inv WTK; eauto with ty.
- inv WTK; eauto with ty.
- inv WTS; eauto with ty.
-- exploit wt_find_label. eexact WTB. eauto. eapply call_cont_wt'; eauto.
+- exploit wt_find_label. eexact WTB. eauto. eapply call_cont_wt'; eauto.
intros [A B]. eauto with ty.
-- simpl in WTFD; inv WTFD. econstructor; eauto. apply wt_call_cont_stmt_cont; auto.
-- exploit (Genv.find_funct_inversion prog); eauto. intros (id & A).
- econstructor; eauto.
+- simpl in WTFD; inv WTFD. econstructor; eauto. apply wt_call_cont_stmt_cont; auto.
+- exploit (Genv.find_funct_inversion prog); eauto. intros (id & A).
+ econstructor; eauto.
- inv WTK. eauto with ty.
Qed.
Theorem preservation:
forall S t S', step ge S t S' -> wt_state S -> wt_state S'.
Proof.
- intros. destruct H. eapply preservation_estep; eauto. eapply preservation_sstep; eauto.
+ intros. destruct H. eapply preservation_estep; eauto. eapply preservation_sstep; eauto.
Qed.
Theorem wt_initial_state:
forall S, initial_state prog S -> wt_state S.
Proof.
- intros. inv H. econstructor. constructor.
+ intros. inv H. econstructor. constructor.
apply Genv.find_funct_ptr_prop with (p := prog) (b := b); auto.
intros. inv WTPROG. destruct f0; simpl; auto. apply H4 with id; auto.
instantiate (1 := (Vptr b Int.zero)). rewrite Genv.find_funct_find_funct_ptr. auto.
diff --git a/cfrontend/Initializersproof.v b/cfrontend/Initializersproof.v
index 790877bd..3a7b5593 100644
--- a/cfrontend/Initializersproof.v
+++ b/cfrontend/Initializersproof.v
@@ -143,7 +143,7 @@ with eval_simple_rvalue: expr -> val -> Prop :=
eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) m = Some true ->
eval_simple_rvalue (Eseqor r1 r2 ty) (Vint Int.one)
| esr_condition: forall r1 r2 r3 ty v v1 b v',
- eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) m = Some b ->
+ eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) m = Some b ->
eval_simple_rvalue (if b then r2 else r3) v' ->
sem_cast v' (typeof (if b then r2 else r3)) ty = Some v ->
eval_simple_rvalue (Econdition r1 r2 r3 ty) v
@@ -190,7 +190,7 @@ Qed.
Lemma rred_simple:
forall r m t r' m', rred ge r m t r' m' -> simple r -> simple r'.
Proof.
- induction 1; simpl; intuition. destruct b; auto.
+ induction 1; simpl; intuition. destruct b; auto.
Qed.
Lemma rred_compat:
@@ -199,20 +199,20 @@ Lemma rred_compat:
m = m' /\ compat_eval RV e r r' m.
Proof.
intros until m'; intros RED SIMP. inv RED; simpl in SIMP; try contradiction; split; auto; split; auto; intros vx EV.
- inv EV. econstructor. constructor. auto. auto.
+ inv EV. econstructor. constructor. auto. auto.
inv EV. econstructor. constructor.
- inv EV. econstructor; eauto. constructor.
+ inv EV. econstructor; eauto. constructor.
inv EV. econstructor; eauto. constructor. constructor.
inv EV. econstructor; eauto. constructor.
inv EV. eapply esr_seqand_true; eauto. constructor.
inv EV. eapply esr_seqand_false; eauto. constructor.
inv EV. eapply esr_seqor_true; eauto. constructor.
inv EV. eapply esr_seqor_false; eauto. constructor.
- inv EV. eapply esr_condition; eauto. constructor.
+ inv EV. eapply esr_condition; eauto. constructor.
inv EV. constructor.
inv EV. constructor.
econstructor; eauto. constructor.
- inv EV. econstructor. constructor. auto.
+ inv EV. econstructor. constructor. auto.
Qed.
Lemma compat_eval_context:
@@ -225,19 +225,19 @@ Proof.
try (generalize (IHcontext CE); intros [TY EV]; red; split; simpl; auto; intros).
inv H0. constructor; auto.
inv H0.
- eapply esl_field_struct; eauto. rewrite TY; eauto.
+ eapply esl_field_struct; eauto. rewrite TY; eauto.
eapply esl_field_union; eauto. rewrite TY; eauto.
inv H0. econstructor. eauto. auto. auto.
- inv H0. econstructor; eauto.
+ inv H0. econstructor; eauto.
inv H0. econstructor; eauto. congruence.
inv H0. econstructor; eauto. congruence.
inv H0. econstructor; eauto. congruence.
inv H0. econstructor; eauto. congruence.
- inv H0.
- eapply esr_seqand_true; eauto. rewrite TY; auto.
+ inv H0.
+ eapply esr_seqand_true; eauto. rewrite TY; auto.
eapply esr_seqand_false; eauto. rewrite TY; auto.
- inv H0.
- eapply esr_seqor_false; eauto. rewrite TY; auto.
+ inv H0.
+ eapply esr_seqor_false; eauto. rewrite TY; auto.
eapply esr_seqor_true; eauto. rewrite TY; auto.
inv H0. eapply esr_condition; eauto. congruence.
inv H0.
@@ -249,19 +249,19 @@ Proof.
red; split; intros. auto. inv H0.
red; split; intros. auto. inv H0.
inv H0. econstructor; eauto.
- inv H0. econstructor; eauto. congruence.
+ inv H0. econstructor; eauto. congruence.
Qed.
Lemma simple_context_1:
forall a from to C, context from to C -> simple (C a) -> simple a.
Proof.
- induction 1; simpl; tauto.
+ induction 1; simpl; tauto.
Qed.
Lemma simple_context_2:
forall a a', simple a' -> forall from to C, context from to C -> simple (C a) -> simple (C a').
Proof.
- induction 2; simpl; try tauto.
+ induction 2; simpl; try tauto.
Qed.
Lemma compat_eval_steps_aux f r e m r' m' s2 :
@@ -296,16 +296,16 @@ Qed.
Lemma compat_eval_steps:
forall f r e m r' m',
star step ge (ExprState f r Kstop e m) E0 (ExprState f r' Kstop e m') ->
- simple r ->
+ simple r ->
m' = m /\ compat_eval RV e r r' m.
Proof.
- intros.
+ intros.
remember (ExprState f r Kstop e m) as S1.
remember E0 as t.
remember (ExprState f r' Kstop e m') as S2.
revert S1 t S2 H r m r' m' HeqS1 Heqt HeqS2 H0.
induction 1; intros; subst.
- (* base case *)
+ (* base case *)
inv HeqS2. split. auto. red; auto.
(* inductive case *)
destruct (app_eq_nil t1 t2); auto. subst. inv H.
@@ -313,7 +313,7 @@ Proof.
exploit compat_eval_steps_aux; eauto.
intros [r1 [A [B C]]]. subst s2.
exploit IHstar; eauto. intros [D E].
- split. auto. destruct B; destruct E. split. congruence. auto.
+ split. auto. destruct B; destruct E. split. congruence. auto.
(* statement steps *)
inv H1.
Qed.
@@ -324,7 +324,7 @@ Theorem eval_simple_steps:
simple r ->
m' = m /\ ty = typeof r /\ eval_simple_rvalue e m r v.
Proof.
- intros. exploit compat_eval_steps; eauto. intros [A [B C]].
+ intros. exploit compat_eval_steps; eauto. intros [A [B C]].
intuition. apply C. constructor.
Qed.
@@ -344,7 +344,7 @@ Lemma mem_empty_not_valid_pointer:
forall b ofs, Mem.valid_pointer Mem.empty b ofs = false.
Proof.
intros. unfold Mem.valid_pointer. destruct (Mem.perm_dec Mem.empty b ofs Cur Nonempty); auto.
- eelim Mem.perm_empty; eauto.
+ eelim Mem.perm_empty; eauto.
Qed.
Lemma mem_empty_not_weak_valid_pointer:
@@ -362,7 +362,7 @@ Lemma sem_cast_match:
Val.inject inj v2' v2.
Proof.
intros. unfold do_cast in H0. destruct (sem_cast v1' ty1 ty2) as [v2''|] eqn:E; inv H0.
- exploit sem_cast_inject. eexact E. eauto.
+ exploit sem_cast_inject. eexact E. eauto.
intros [v' [A B]]. congruence.
Qed.
@@ -395,14 +395,14 @@ Proof.
(* val *)
destruct v; monadInv CV; constructor.
(* rval *)
- inv H1; rewrite H2 in CV; try congruence. eauto. eauto.
+ inv H1; rewrite H2 in CV; try congruence. eauto. eauto.
(* addrof *)
eauto.
(* unop *)
destruct (sem_unary_operation op x (typeof r1) Mem.empty) as [v1'|] eqn:E; inv EQ0.
exploit (sem_unary_operation_inj inj Mem.empty m).
intros. rewrite mem_empty_not_weak_valid_pointer in H2; discriminate.
- eexact E. eauto.
+ eexact E. eauto.
intros [v' [A B]]. congruence.
(* binop *)
destruct (sem_binary_operation ge op x (typeof r1) x0 (typeof r2) Mem.empty) as [v1'|] eqn:E; inv EQ2.
@@ -411,34 +411,34 @@ Proof.
intros. rewrite mem_empty_not_weak_valid_pointer in H3; discriminate.
intros. rewrite mem_empty_not_weak_valid_pointer in H3; discriminate.
intros. rewrite mem_empty_not_valid_pointer in H3; discriminate.
- eauto. eauto. eauto.
+ eauto. eauto. eauto.
intros [v' [A B]]. congruence.
(* cast *)
- eapply sem_cast_match; eauto.
+ eapply sem_cast_match; eauto.
(* sizeof *)
constructor.
(* alignof *)
constructor.
(* seqand *)
- destruct (bool_val x (typeof r1) Mem.empty) as [b|] eqn:E; inv EQ2.
+ destruct (bool_val x (typeof r1) Mem.empty) as [b|] eqn:E; inv EQ2.
exploit bool_val_match. eexact E. eauto. instantiate (1 := m). intros E'.
assert (b = true) by congruence. subst b.
eapply sem_cast_match; eauto.
- destruct (bool_val x (typeof r1) Mem.empty) as [b|] eqn:E; inv EQ2.
+ destruct (bool_val x (typeof r1) Mem.empty) as [b|] eqn:E; inv EQ2.
exploit bool_val_match. eexact E. eauto. instantiate (1 := m). intros E'.
assert (b = false) by congruence. subst b. inv H2. auto.
(* seqor *)
- destruct (bool_val x (typeof r1) Mem.empty) as [b|] eqn:E; inv EQ2.
+ destruct (bool_val x (typeof r1) Mem.empty) as [b|] eqn:E; inv EQ2.
exploit bool_val_match. eexact E. eauto. instantiate (1 := m). intros E'.
- assert (b = false) by congruence. subst b.
+ assert (b = false) by congruence. subst b.
eapply sem_cast_match; eauto.
- destruct (bool_val x (typeof r1) Mem.empty) as [b|] eqn:E; inv EQ2.
+ destruct (bool_val x (typeof r1) Mem.empty) as [b|] eqn:E; inv EQ2.
exploit bool_val_match. eexact E. eauto. instantiate (1 := m). intros E'.
assert (b = true) by congruence. subst b. inv H2. auto.
(* conditional *)
destruct (bool_val x (typeof r1) Mem.empty) as [b'|] eqn:E; inv EQ3.
exploit bool_val_match. eexact E. eauto. instantiate (1 := m). intros E'.
- assert (b' = b) by congruence. subst b'.
+ assert (b' = b) by congruence. subst b'.
destruct b; eapply sem_cast_match; eauto.
(* comma *)
auto.
@@ -450,14 +450,14 @@ Proof.
(* var local *)
unfold empty_env in H. rewrite PTree.gempty in H. congruence.
(* var_global *)
- econstructor. unfold inj. rewrite H0. eauto. auto.
+ econstructor. unfold inj. rewrite H0. eauto. auto.
(* deref *)
eauto.
(* field struct *)
- rewrite H0 in CV. monadInv CV. unfold lookup_composite in EQ; rewrite H1 in EQ; monadInv EQ.
+ rewrite H0 in CV. monadInv CV. unfold lookup_composite in EQ; rewrite H1 in EQ; monadInv EQ.
exploit constval_rvalue; eauto. intro MV. inv MV.
- simpl. replace x0 with delta by congruence. econstructor; eauto.
- rewrite ! Int.add_assoc. f_equal. apply Int.add_commut.
+ simpl. replace x0 with delta by congruence. econstructor; eauto.
+ rewrite ! Int.add_assoc. f_equal. apply Int.add_commut.
simpl. auto.
(* field union *)
rewrite H0 in CV. eauto.
@@ -481,7 +481,7 @@ Theorem constval_steps:
constval ge r = OK v ->
m' = m /\ ty = typeof r /\ Val.inject inj v v'.
Proof.
- intros. exploit eval_simple_steps; eauto. eapply constval_simple; eauto.
+ intros. exploit eval_simple_steps; eauto. eapply constval_simple; eauto.
intros [A [B C]]. intuition. eapply constval_rvalue; eauto.
Qed.
@@ -498,33 +498,33 @@ Theorem transl_init_single_steps:
Mem.store chunk m' b ofs v = Some m'' ->
Genv.store_init_data ge m b ofs data = Some m''.
Proof.
- intros. monadInv H.
+ intros. monadInv H.
exploit constval_steps; eauto. intros [A [B C]]. subst m' ty1.
exploit sem_cast_match; eauto. intros D.
- unfold Genv.store_init_data.
- inv D.
+ unfold Genv.store_init_data.
+ inv D.
(* int *)
- destruct ty; try discriminate.
+ destruct ty; try discriminate.
destruct i0; inv EQ2.
destruct s; simpl in H2; inv H2. rewrite <- Mem.store_signed_unsigned_8; auto. auto.
destruct s; simpl in H2; inv H2. rewrite <- Mem.store_signed_unsigned_16; auto. auto.
simpl in H2; inv H2. assumption.
- simpl in H2; inv H2. assumption.
+ simpl in H2; inv H2. assumption.
inv EQ2. simpl in H2; inv H2. assumption.
(* long *)
destruct ty; inv EQ2. simpl in H2; inv H2. assumption.
(* float *)
- destruct ty; try discriminate.
+ destruct ty; try discriminate.
destruct f1; inv EQ2; simpl in H2; inv H2; assumption.
(* single *)
- destruct ty; try discriminate.
+ destruct ty; try discriminate.
destruct f1; inv EQ2; simpl in H2; inv H2; assumption.
(* pointer *)
unfold inj in H.
assert (data = Init_addrof b1 ofs1 /\ chunk = Mint32).
destruct ty; inv EQ2; inv H2.
destruct i; inv H5. intuition congruence. auto.
- destruct H4; subst. destruct (Genv.find_symbol ge b1); inv H.
+ destruct H4; subst. destruct (Genv.find_symbol ge b1); inv H.
rewrite Int.add_zero in H3. auto.
(* undef *)
discriminate.
@@ -537,9 +537,9 @@ Lemma transl_init_single_size:
transl_init_single ge ty a = OK data ->
Genv.init_data_size data = sizeof ge ty.
Proof.
- intros. monadInv H. destruct x0.
+ intros. monadInv H. destruct x0.
- monadInv EQ2.
-- destruct ty; try discriminate.
+- destruct ty; try discriminate.
destruct i0; inv EQ2; auto.
inv EQ2; auto.
- destruct ty; inv EQ2; auto.
@@ -557,15 +557,15 @@ Notation idlsize := Genv.init_data_list_size.
Remark padding_size:
forall frm to, frm <= to -> idlsize (padding frm to) = to - frm.
Proof.
- unfold padding; intros. destruct (zlt frm to).
- simpl. xomega.
+ unfold padding; intros. destruct (zlt frm to).
+ simpl. xomega.
simpl. omega.
Qed.
Remark idlsize_app:
forall d1 d2, idlsize (d1 ++ d2) = idlsize d1 + idlsize d2.
Proof.
- induction d1; simpl; intros.
+ induction d1; simpl; intros.
auto.
rewrite IHd1. omega.
Qed.
@@ -573,11 +573,11 @@ Qed.
Remark union_field_size:
forall f ty fl, field_type f fl = OK ty -> sizeof ge ty <= sizeof_union ge fl.
Proof.
- induction fl as [|[i t]]; simpl; intros.
+ induction fl as [|[i t]]; simpl; intros.
- inv H.
-- destruct (ident_eq f i).
- + inv H. xomega.
- + specialize (IHfl H). xomega.
+- destruct (ident_eq f i).
+ + inv H. xomega.
+ + specialize (IHfl H). xomega.
Qed.
Hypothesis ce_consistent: composite_env_consistent ge.
@@ -599,9 +599,9 @@ with transl_init_list_size:
idlsize data + pos = sizeof ge ty).
Proof.
-- induction i; intros.
+- induction i; intros.
+ (* single *)
- monadInv H. simpl. erewrite transl_init_single_size by eauto. omega.
+ monadInv H. simpl. erewrite transl_init_single_size by eauto. omega.
+ (* array *)
simpl in H. destruct ty; try discriminate.
simpl. eapply (proj1 (transl_init_list_size il)); eauto.
@@ -611,19 +611,19 @@ Proof.
replace (idlsize data) with (idlsize data + 0) by omega.
eapply (proj2 (transl_init_list_size il)). eauto.
unfold lookup_composite in EQ. simpl. destruct (ge.(genv_cenv)!i) as [co|] eqn:?; inv EQ.
- erewrite co_consistent_sizeof by (eapply ce_consistent; eauto).
+ erewrite co_consistent_sizeof by (eapply ce_consistent; eauto).
unfold sizeof_composite. rewrite Heqs. apply align_le.
destruct (co_alignof_two_p x) as [n EQ]. rewrite EQ. apply two_power_nat_pos.
+ (* union *)
simpl in H. destruct ty; try discriminate.
monadInv H. destruct (co_su x) eqn:?; try discriminate.
- monadInv EQ0.
+ monadInv EQ0.
rewrite idlsize_app. rewrite (IHi _ _ EQ0).
unfold lookup_composite in EQ. simpl. destruct (ge.(genv_cenv)!i0) as [co|] eqn:?; inv EQ.
rewrite padding_size. omega.
- apply Zle_trans with (sizeof_union ge (co_members x)).
+ apply Zle_trans with (sizeof_union ge (co_members x)).
eapply union_field_size; eauto.
- erewrite co_consistent_sizeof by (eapply ce_consistent; eauto).
+ erewrite co_consistent_sizeof by (eapply ce_consistent; eauto).
unfold sizeof_composite. rewrite Heqs. apply align_le.
destruct (co_alignof_two_p x) as [n EQ]. rewrite EQ. apply two_power_nat_pos.
@@ -631,9 +631,9 @@ Proof.
+ (* base cases *)
simpl. intuition auto.
* (* arrays *)
- destruct (zeq sz 0). inv H. simpl; ring.
- destruct (zle 0 sz); inv H. simpl.
- rewrite Z.mul_comm.
+ destruct (zeq sz 0). inv H. simpl; ring.
+ destruct (zle 0 sz); inv H. simpl.
+ rewrite Z.mul_comm.
assert (0 <= sizeof ge ty * sz).
{ apply Zmult_gt_0_le_0_compat. omega. generalize (sizeof_pos ge ty); omega. }
zify; omega.
@@ -645,17 +645,17 @@ Proof.
* (* arrays *)
intros. monadInv H.
rewrite idlsize_app.
- rewrite (transl_init_size _ _ _ EQ).
+ rewrite (transl_init_size _ _ _ EQ).
rewrite (A _ _ _ EQ1).
ring.
* (* structs *)
intros. simpl in H. destruct fl as [|[i1 t1]]; monadInv H.
- rewrite ! idlsize_app.
- simpl in H0.
+ rewrite ! idlsize_app.
+ simpl in H0.
rewrite padding_size.
- rewrite (transl_init_size _ _ _ EQ).
+ rewrite (transl_init_size _ _ _ EQ).
rewrite <- (B _ _ _ _ EQ1). omega.
- auto. apply align_le. apply alignof_pos.
+ auto. apply align_le. apply alignof_pos.
Qed.
(** A semantics for general initializers *)
@@ -671,7 +671,7 @@ Fixpoint fields_of_struct (fl: members) (pos: Z) : list (Z * type) :=
Inductive exec_init: mem -> block -> Z -> type -> initializer -> mem -> Prop :=
| exec_init_single: forall m b ofs ty a v1 ty1 chunk m' v m'',
- star step ge (ExprState dummy_function a Kstop empty_env m)
+ star step ge (ExprState dummy_function a Kstop empty_env m)
E0 (ExprState dummy_function (Eval v1 ty1) Kstop empty_env m') ->
sem_cast v1 ty1 ty = Some v ->
access_mode ty = By_value chunk ->
@@ -713,7 +713,7 @@ Scheme exec_init_ind3 := Minimality for exec_init Sort Prop
Combined Scheme exec_init_scheme from exec_init_ind3, exec_init_array_ind3, exec_init_list_ind3.
Remark exec_init_array_length:
- forall m b ofs ty sz il m',
+ forall m b ofs ty sz il m',
exec_init_array m b ofs ty sz il m' -> sz >= 0.
Proof.
induction 1; omega.
@@ -725,7 +725,7 @@ Lemma store_init_data_list_app:
Genv.store_init_data_list ge m' b (ofs + idlsize data1) data2 = Some m'' ->
Genv.store_init_data_list ge m b ofs (data1 ++ data2) = Some m''.
Proof.
- induction data1; simpl; intros.
+ induction data1; simpl; intros.
inv H. rewrite Zplus_0_r in H0. auto.
destruct (Genv.store_init_data ge m b ofs a); try discriminate.
rewrite Zplus_assoc in H0. eauto.
@@ -735,7 +735,7 @@ Remark store_init_data_list_padding:
forall frm to b ofs m,
Genv.store_init_data_list ge m b ofs (padding frm to) = Some m.
Proof.
- intros. unfold padding. destruct (zlt frm to); auto.
+ intros. unfold padding. destruct (zlt frm to); auto.
Qed.
Lemma transl_init_sound_gen:
@@ -757,33 +757,33 @@ Local Opaque sizeof.
monadInv H3. simpl. erewrite transl_init_single_steps by eauto. auto.
- (* array *)
replace (Z.max 0 sz) with sz in H1. eauto.
- assert (sz >= 0) by (eapply exec_init_array_length; eauto). xomega.
+ assert (sz >= 0) by (eapply exec_init_array_length; eauto). xomega.
- (* struct *)
- unfold lookup_composite in H3. rewrite H in H3. simpl in H3. rewrite H0 in H3.
+ unfold lookup_composite in H3. rewrite H in H3. simpl in H3. rewrite H0 in H3.
replace ofs with (ofs + 0) by omega. eauto.
- (* union *)
- unfold lookup_composite in H4. rewrite H in H4. simpl in H4. rewrite H0 in H4.
- monadInv H4. assert (x = ty) by congruence. subst x.
+ unfold lookup_composite in H4. rewrite H in H4. simpl in H4. rewrite H0 in H4.
+ monadInv H4. assert (x = ty) by congruence. subst x.
eapply store_init_data_list_app. eauto.
- apply store_init_data_list_padding.
+ apply store_init_data_list_padding.
- (* array, empty *)
destruct (zeq sz 0).
inv H0. auto.
rewrite zle_true in H0 by omega. inv H0. auto.
- (* array, nonempty *)
- monadInv H3.
+ monadInv H3.
eapply store_init_data_list_app.
eauto.
rewrite (transl_init_size _ _ _ EQ). eauto.
- (* struct, empty *)
- destruct fl as [|[i t]]; simpl in H0; inv H0.
+ destruct fl as [|[i t]]; simpl in H0; inv H0.
apply store_init_data_list_padding.
- (* struct, nonempty *)
destruct fl as [|[i t]]; simpl in H4; monadInv H4.
simpl in H3; inv H3.
eapply store_init_data_list_app. apply store_init_data_list_padding.
- rewrite padding_size.
+ rewrite padding_size.
replace (ofs + pos0 + (align pos0 (alignof ge t) - pos0))
with (ofs + align pos0 (alignof ge t)) by omega.
eapply store_init_data_list_app.
@@ -801,7 +801,7 @@ Theorem transl_init_sound:
transl_init (prog_comp_env p) ty i = OK data ->
Genv.store_init_data_list (globalenv p) m b 0 data = Some m'.
Proof.
- intros.
+ intros.
set (ge := globalenv p) in *.
change (prog_comp_env p) with (genv_cenv ge) in H0.
destruct (transl_init_sound_gen ge) as (A & B & C).
diff --git a/cfrontend/PrintClight.ml b/cfrontend/PrintClight.ml
index f1c3ef18..ed19e178 100644
--- a/cfrontend/PrintClight.ml
+++ b/cfrontend/PrintClight.ml
@@ -67,7 +67,7 @@ let rec expr p (prec, e) =
if assoc = LtoR
then (prec', prec' + 1)
else (prec' + 1, prec') in
- if prec' < prec
+ if prec' < prec
then fprintf p "@[<hov 2>("
else fprintf p "@[<hov 2>";
begin match e with
diff --git a/cfrontend/PrintCsyntax.ml b/cfrontend/PrintCsyntax.ml
index ce912a8c..4f2a8d0c 100644
--- a/cfrontend/PrintCsyntax.ml
+++ b/cfrontend/PrintCsyntax.ml
@@ -80,7 +80,7 @@ let attributes a =
sprintf " _Alignas(%Ld)%s" (Int64.shift_left 1L (N.to_int l)) s1
let attributes_space a =
- let s = attributes a in
+ let s = attributes a in
if String.length s = 0 then s else s ^ " "
let name_optid id =
@@ -202,7 +202,7 @@ let rec expr p (prec, e) =
if assoc = LtoR
then (prec', prec' + 1)
else (prec' + 1, prec') in
- if prec' < prec
+ if prec' < prec
then fprintf p "@[<hov 2>("
else fprintf p "@[<hov 2>";
begin match e with
diff --git a/cfrontend/SimplExpr.v b/cfrontend/SimplExpr.v
index 097dc589..4fe8105d 100644
--- a/cfrontend/SimplExpr.v
+++ b/cfrontend/SimplExpr.v
@@ -10,7 +10,7 @@
(* *)
(* *********************************************************************)
-(** Translation from Compcert C to Clight.
+(** Translation from Compcert C to Clight.
Side effects are pulled out of Compcert C expressions. *)
Require Import Coqlib.
@@ -82,11 +82,11 @@ Local Open Scope gensym_monad_scope.
Parameter first_unused_ident: unit -> ident.
-Definition initial_generator (x: unit) : generator :=
+Definition initial_generator (x: unit) : generator :=
mkgenerator (first_unused_ident x) nil.
Definition gensym (ty: type): mon ident :=
- fun (g: generator) =>
+ fun (g: generator) =>
Res (gen_next g)
(mkgenerator (Psucc (gen_next g)) ((gen_next g, ty) :: gen_trail g))
(Ple_succ (gen_next g)).
@@ -119,7 +119,7 @@ Function eval_simpl_expr (a: expr) : option val :=
| Econst_float n _ => Some(Vfloat n)
| Econst_single n _ => Some(Vsingle n)
| Econst_long n _ => Some(Vlong n)
- | Ecast b ty =>
+ | Ecast b ty =>
match eval_simpl_expr b with
| None => None
| Some v => sem_cast v (typeof b) ty
@@ -149,7 +149,7 @@ Definition transl_incrdecr (id: incr_or_decr) (a: expr) (ty: type) : expr :=
to dereference a l-value [l] and store its result in temporary variable [id]. *)
Definition chunk_for_volatile_type (ty: type) : option memory_chunk :=
- if type_is_volatile ty
+ if type_is_volatile ty
then match access_mode ty with By_value chunk => Some chunk | _ => None end
else None.
@@ -201,7 +201,7 @@ Inductive set_destination : Type :=
| SDbase (tycast ty: type) (tmp: ident)
| SDcons (tycast ty: type) (tmp: ident) (sd: set_destination).
-Inductive destination : Type :=
+Inductive destination : Type :=
| For_val
| For_effects
| For_set (sd: set_destination).
@@ -277,7 +277,7 @@ Fixpoint transl_expr (dst: destination) (a: Csyntax.expr) : mon (list statement
| For_val =>
do t <- gensym ty;
do (sl2, a2) <- transl_expr (For_set (sd_seqbool_val t ty)) r2;
- ret (sl1 ++
+ ret (sl1 ++
makeif a1 (makeseq sl2) (Sset t (Econst_int Int.zero ty)) :: nil,
Etempvar t ty)
| For_effects =>
@@ -285,7 +285,7 @@ Fixpoint transl_expr (dst: destination) (a: Csyntax.expr) : mon (list statement
ret (sl1 ++ makeif a1 (makeseq sl2) Sskip :: nil, dummy_expr)
| For_set sd =>
do (sl2, a2) <- transl_expr (For_set (sd_seqbool_set ty sd)) r2;
- ret (sl1 ++
+ ret (sl1 ++
makeif a1 (makeseq sl2) (makeseq (do_set sd (Econst_int Int.zero ty))) :: nil,
dummy_expr)
end
@@ -295,7 +295,7 @@ Fixpoint transl_expr (dst: destination) (a: Csyntax.expr) : mon (list statement
| For_val =>
do t <- gensym ty;
do (sl2, a2) <- transl_expr (For_set (sd_seqbool_val t ty)) r2;
- ret (sl1 ++
+ ret (sl1 ++
makeif a1 (Sset t (Econst_int Int.one ty)) (makeseq sl2) :: nil,
Etempvar t ty)
| For_effects =>
@@ -303,7 +303,7 @@ Fixpoint transl_expr (dst: destination) (a: Csyntax.expr) : mon (list statement
ret (sl1 ++ makeif a1 Sskip (makeseq sl2) :: nil, dummy_expr)
| For_set sd =>
do (sl2, a2) <- transl_expr (For_set (sd_seqbool_set ty sd)) r2;
- ret (sl1 ++
+ ret (sl1 ++
makeif a1 (makeseq (do_set sd (Econst_int Int.one ty))) (makeseq sl2) :: nil,
dummy_expr)
end
@@ -336,7 +336,7 @@ Fixpoint transl_expr (dst: destination) (a: Csyntax.expr) : mon (list statement
match dst with
| For_val | For_set _ =>
do t <- gensym ty2;
- ret (finish dst
+ ret (finish dst
(sl1 ++ sl2 ++ Sset t a2 :: make_assign a1 (Etempvar t ty2) :: nil)
(Ecast (Etempvar t ty2) ty1))
| For_effects =>
diff --git a/cfrontend/SimplExprproof.v b/cfrontend/SimplExprproof.v
index 7ef1cbe2..8f06e777 100644
--- a/cfrontend/SimplExprproof.v
+++ b/cfrontend/SimplExprproof.v
@@ -45,20 +45,20 @@ Let tge := Clight.globalenv tprog.
Lemma comp_env_preserved:
Clight.genv_cenv tge = Csem.genv_cenv ge.
Proof.
- monadInv TRANSL. unfold tge; rewrite <- H0; auto.
+ monadInv TRANSL. unfold tge; rewrite <- H0; auto.
Qed.
Lemma symbols_preserved:
forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s.
Proof.
- intros. eapply Genv.find_symbol_match. eapply transl_program_spec; eauto.
+ intros. eapply Genv.find_symbol_match. eapply transl_program_spec; eauto.
simpl. tauto.
Qed.
Lemma public_preserved:
forall (s: ident), Genv.public_symbol tge s = Genv.public_symbol ge s.
Proof.
- intros. eapply Genv.public_symbol_match. eapply transl_program_spec; eauto.
+ intros. eapply Genv.public_symbol_match. eapply transl_program_spec; eauto.
simpl. tauto.
Qed.
@@ -87,11 +87,11 @@ Qed.
Lemma varinfo_preserved:
forall b, Genv.find_var_info tge b = Genv.find_var_info ge b.
Proof.
- intros. destruct (Genv.find_var_info ge b) as [v|] eqn:V.
-- exploit Genv.find_var_info_match. eapply transl_program_spec; eauto. eassumption.
+ intros. destruct (Genv.find_var_info ge b) as [v|] eqn:V.
+- exploit Genv.find_var_info_match. eapply transl_program_spec; eauto. eassumption.
intros [tv [A B]]. inv B. assumption.
- destruct (Genv.find_var_info tge b) as [v'|] eqn:V'; auto.
- exploit Genv.find_var_info_rev_match. eapply transl_program_spec; eauto. eassumption.
+ exploit Genv.find_var_info_rev_match. eapply transl_program_spec; eauto. eassumption.
simpl. destruct (plt b (Genv.genv_next (Genv.globalenv prog))); try tauto.
intros [v [A B]]. inv B. change (Genv.globalenv prog) with (Csem.genv_genv ge) in A. congruence.
Qed.
@@ -115,7 +115,7 @@ Lemma function_return_preserved:
forall f tf, tr_function f tf ->
fn_return tf = Csyntax.fn_return f.
Proof.
- intros. inv H; auto.
+ intros. inv H; auto.
Qed.
(** Translation of simple expressions. *)
@@ -133,7 +133,7 @@ Proof.
rewrite H0; auto. simpl; auto.
destruct H1; congruence.
destruct (andb_prop _ _ H6). inv H1.
- rewrite H0; eauto. simpl; auto.
+ rewrite H0; eauto. simpl; auto.
unfold chunk_for_volatile_type in H9.
destruct (type_is_volatile (Csyntax.typeof e1)); simpl in H8; congruence.
rewrite H0; auto. simpl; auto.
@@ -163,7 +163,7 @@ Remark deref_loc_translated:
| Some chunk => volatile_load tge chunk m b ofs t v
end.
Proof.
- intros. unfold chunk_for_volatile_type. inv H.
+ intros. unfold chunk_for_volatile_type. inv H.
(* By_value, not volatile *)
rewrite H1. split; auto. eapply deref_loc_value; eauto.
(* By_value, volatile *)
@@ -183,14 +183,14 @@ Remark assign_loc_translated:
| Some chunk => volatile_store tge chunk m b ofs v t m'
end.
Proof.
- intros. unfold chunk_for_volatile_type. inv H.
+ intros. unfold chunk_for_volatile_type. inv H.
(* By_value, not volatile *)
rewrite H1. split; auto. eapply assign_loc_value; eauto.
(* By_value, volatile *)
rewrite H0; rewrite H1. eapply volatile_store_preserved with (ge1 := ge); auto.
exact symbols_preserved. exact public_preserved. exact block_is_volatile_preserved.
(* By copy *)
- rewrite H0. rewrite <- comp_env_preserved in *.
+ rewrite H0. rewrite <- comp_env_preserved in *.
destruct (type_is_volatile ty); split; auto; eapply assign_loc_copy; eauto.
Qed.
@@ -233,12 +233,12 @@ Opaque makeif.
rewrite <- B.
exploit deref_loc_translated; eauto. unfold chunk_for_volatile_type; rewrite H2. tauto.
destruct dst; auto.
- econstructor. split. simpl; eauto. auto.
+ econstructor. split. simpl; eauto. auto.
(* addrof *)
exploit H0; eauto. intros [A [B C]].
subst sl1; simpl.
assert (eval_expr tge e le m (Eaddrof a1 ty) (Vptr b ofs)). econstructor; eauto.
- destruct dst; auto. simpl; econstructor; eauto.
+ destruct dst; auto. simpl; econstructor; eauto.
(* unop *)
exploit H0; eauto. intros [A [B C]].
subst sl1; simpl.
@@ -256,19 +256,19 @@ Opaque makeif.
assert (eval_expr tge e le m (Ecast a1 ty) v). econstructor; eauto. congruence.
destruct dst; auto. simpl; econstructor; eauto.
(* sizeof *)
- rewrite <- comp_env_preserved.
+ rewrite <- comp_env_preserved.
destruct dst.
split; auto. split; auto. constructor.
auto.
exists (Esizeof ty1 ty). split. auto. split. auto. constructor.
(* alignof *)
- rewrite <- comp_env_preserved.
+ rewrite <- comp_env_preserved.
destruct dst.
split; auto. split; auto. constructor.
auto.
exists (Ealignof ty1 ty). split. auto. split. auto. constructor.
(* var local *)
- split; auto. split; auto. apply eval_Evar_local; auto.
+ split; auto. split; auto. apply eval_Evar_local; auto.
(* var global *)
split; auto. split; auto. apply eval_Evar_global; auto.
rewrite symbols_preserved; auto.
@@ -302,7 +302,7 @@ Proof.
intros e m. exact (proj1 (tr_simple e m)).
Qed.
-Lemma tr_simple_lvalue:
+Lemma tr_simple_lvalue:
forall e m l b ofs,
eval_simple_lvalue ge e m l b ofs ->
forall le sl a tmps,
@@ -319,7 +319,7 @@ Lemma tr_simple_exprlist:
eval_simple_list ge e m rl tyl vl ->
sl = nil /\ eval_exprlist tge e le m al tyl vl.
Proof.
- induction 1; intros.
+ induction 1; intros.
inv H. split. auto. constructor.
inv H4.
exploit tr_simple_rvalue; eauto. intros [A [B C]].
@@ -334,7 +334,7 @@ Lemma typeof_context:
forall e1 e2, Csyntax.typeof e1 = Csyntax.typeof e2 ->
Csyntax.typeof (C e1) = Csyntax.typeof (C e2).
Proof.
- induction 1; intros; auto.
+ induction 1; intros; auto.
Qed.
Scheme leftcontext_ind2 := Minimality for leftcontext Sort Prop
@@ -395,131 +395,131 @@ Ltac UNCHANGED :=
(* base *)
TR. rewrite <- app_nil_end; auto. red; auto.
- intros. rewrite <- app_nil_end; auto.
+ intros. rewrite <- app_nil_end; auto.
(* deref *)
- inv H1.
+ inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1; rewrite app_ass; eauto. auto.
- intros. rewrite <- app_ass. econstructor; eauto.
+ TR. subst sl1; rewrite app_ass; eauto. auto.
+ intros. rewrite <- app_ass. econstructor; eauto.
(* field *)
inv H1.
exploit H0. eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1; rewrite app_ass; eauto. auto.
+ TR. subst sl1; rewrite app_ass; eauto. auto.
intros. rewrite <- app_ass. econstructor; eauto.
(* rvalof *)
inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1; rewrite app_ass; eauto. red; eauto.
- intros. rewrite <- app_ass; econstructor; eauto.
+ TR. subst sl1; rewrite app_ass; eauto. red; eauto.
+ intros. rewrite <- app_ass; econstructor; eauto.
exploit typeof_context; eauto. congruence.
(* addrof *)
- inv H1.
+ inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1; rewrite app_ass; eauto. auto.
+ TR. subst sl1; rewrite app_ass; eauto. auto.
intros. rewrite <- app_ass. econstructor; eauto.
-(* unop *)
- inv H1.
+(* unop *)
+ inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1; rewrite app_ass; eauto. auto.
- intros. rewrite <- app_ass. econstructor; eauto.
+ TR. subst sl1; rewrite app_ass; eauto. auto.
+ intros. rewrite <- app_ass. econstructor; eauto.
(* binop left *)
inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1. rewrite app_ass. eauto.
- red; auto.
+ TR. subst sl1. rewrite app_ass. eauto.
+ red; auto.
intros. rewrite <- app_ass. econstructor; eauto.
- eapply tr_expr_invariant; eauto. UNCHANGED.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
(* binop right *)
- inv H2.
+ inv H2.
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl2. rewrite app_ass. eauto.
- red; auto.
+ TR. subst sl2. rewrite app_ass. eauto.
+ red; auto.
intros. rewrite <- app_ass. change (sl3 ++ sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor; eauto.
- eapply tr_expr_invariant; eauto. UNCHANGED.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
(* cast *)
inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1; rewrite app_ass; eauto. auto.
+ TR. subst sl1; rewrite app_ass; eauto. auto.
intros. rewrite <- app_ass. econstructor; eauto.
(* seqand *)
inv H1.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. apply S; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto.
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. apply S; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto.
(* for set *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. apply S; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto.
(* seqor *)
inv H1.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. apply S; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto.
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. apply S; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto.
(* for set *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. apply S; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto.
(* condition *)
inv H1.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. apply S; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto. auto. auto.
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. eapply tr_condition_effects. apply S; auto.
+ intros. rewrite <- app_ass. eapply tr_condition_effects. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto. auto.
(* for set *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR.
- rewrite Q. rewrite app_ass. eauto.
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. eapply tr_condition_set. apply S; auto.
+ intros. rewrite <- app_ass. eapply tr_condition_set. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto. auto. auto.
@@ -527,120 +527,120 @@ Ltac UNCHANGED :=
inv H1.
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1. rewrite app_ass. eauto.
- red; auto.
+ TR. subst sl1. rewrite app_ass. eauto.
+ red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1. rewrite app_ass. eauto.
- red; auto.
+ TR. subst sl1. rewrite app_ass. eauto.
+ red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto. auto. auto.
eapply typeof_context; eauto.
- auto.
+ auto.
(* assign right *)
- inv H2.
+ inv H2.
(* for effects *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl2. rewrite app_ass. eauto.
- red; auto.
- intros. rewrite <- app_ass. change (sl3 ++ sl2') with (nil ++ (sl3 ++ sl2')). rewrite app_ass.
- econstructor.
+ TR. subst sl2. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. change (sl3 ++ sl2') with (nil ++ (sl3 ++ sl2')). rewrite app_ass.
+ econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
apply S; auto. auto. auto. auto.
(* for val *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl2. rewrite app_ass. eauto.
- red; auto.
- intros. rewrite <- app_ass. change (sl3 ++ sl2') with (nil ++ (sl3 ++ sl2')). rewrite app_ass.
- econstructor.
+ TR. subst sl2. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. change (sl3 ++ sl2') with (nil ++ (sl3 ++ sl2')). rewrite app_ass.
+ econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
- apply S; auto. auto. auto. auto. auto. auto. auto. auto.
+ apply S; auto. auto. auto. auto. auto. auto. auto. auto.
eapply typeof_context; eauto.
(* assignop left *)
inv H1.
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1. rewrite app_ass. eauto.
- red; auto.
+ TR. subst sl1. rewrite app_ass. eauto.
+ red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
- eapply tr_expr_invariant; eauto. UNCHANGED.
- symmetry; eapply typeof_context; eauto. eauto.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
+ symmetry; eapply typeof_context; eauto. eauto.
auto. auto. auto. auto. auto. auto.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1. rewrite app_ass. eauto.
- red; auto.
+ TR. subst sl1. rewrite app_ass. eauto.
+ red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
- eapply tr_expr_invariant; eauto. UNCHANGED.
- eauto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
+ eauto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto.
eapply typeof_context; eauto.
(* assignop right *)
inv H2.
(* for effects *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl2. rewrite app_ass. eauto.
+ TR. subst sl2. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. change (sl0 ++ sl2') with (nil ++ sl0 ++ sl2'). rewrite app_ass. econstructor.
+ intros. rewrite <- app_ass. change (sl0 ++ sl2') with (nil ++ sl0 ++ sl2'). rewrite app_ass. econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
- apply S; auto. auto. eauto. auto. auto. auto. auto. auto. auto.
+ apply S; auto. auto. eauto. auto. auto. auto. auto. auto. auto.
(* for val *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl2. rewrite app_ass. eauto.
- red; auto.
- intros. rewrite <- app_ass. change (sl0 ++ sl2') with (nil ++ sl0 ++ sl2'). rewrite app_ass. econstructor.
+ TR. subst sl2. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. change (sl0 ++ sl2') with (nil ++ sl0 ++ sl2'). rewrite app_ass. econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
- apply S; auto. eauto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto.
+ apply S; auto. eauto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto.
(* postincr *)
inv H1.
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. rewrite Q; rewrite app_ass; eauto. red; auto.
- intros. rewrite <- app_ass. econstructor; eauto.
- symmetry; eapply typeof_context; eauto.
+ TR. rewrite Q; rewrite app_ass; eauto. red; auto.
+ intros. rewrite <- app_ass. econstructor; eauto.
+ symmetry; eapply typeof_context; eauto.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. rewrite Q; rewrite app_ass; eauto. red; auto.
- intros. rewrite <- app_ass. econstructor; eauto.
+ TR. rewrite Q; rewrite app_ass; eauto. red; auto.
+ intros. rewrite <- app_ass. econstructor; eauto.
eapply typeof_context; eauto.
(* call left *)
inv H1.
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. rewrite Q; rewrite app_ass; eauto. red; auto.
+ TR. rewrite Q; rewrite app_ass; eauto. red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_exprlist_invariant; eauto. UNCHANGED.
auto. auto. auto.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. rewrite Q; rewrite app_ass; eauto. red; auto.
+ TR. rewrite Q; rewrite app_ass; eauto. red; auto.
intros. rewrite <- app_ass. econstructor. auto. apply S; auto.
eapply tr_exprlist_invariant; eauto. UNCHANGED.
- auto. auto. auto. auto.
+ auto. auto. auto. auto.
(* call right *)
inv H2.
(* for effects *)
- assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
+ assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto.
(*destruct dst'; constructor||contradiction.*)
- red; auto.
+ red; auto.
intros. rewrite <- app_ass. change (sl3++sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
apply S; auto. auto. auto. auto.
(* for val *)
- assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
+ assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto.
(*destruct dst'; constructor||contradiction.*)
- red; auto.
+ red; auto.
intros. rewrite <- app_ass. change (sl3++sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor.
auto. eapply tr_expr_invariant; eauto. UNCHANGED.
apply S; auto.
@@ -650,27 +650,27 @@ Ltac UNCHANGED :=
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto.
- red; auto.
+ red; auto.
intros. rewrite <- app_ass. change (sl3++sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor.
apply S; auto. auto.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto.
- red; auto.
+ red; auto.
intros. rewrite <- app_ass. change (sl3++sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor.
auto. apply S; auto. auto. auto.
(* comma *)
inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. rewrite Q; rewrite app_ass; eauto. red; auto.
+ TR. rewrite Q; rewrite app_ass; eauto. red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto.
(* paren *)
- inv H1.
+ inv H1.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. rewrite Q. eauto. red; auto.
+ TR. rewrite Q. eauto. red; auto.
intros. econstructor; eauto.
(* for effects *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
@@ -683,17 +683,17 @@ Ltac UNCHANGED :=
(* cons left *)
inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl1. rewrite app_ass. eauto.
- red; auto.
+ TR. subst sl1. rewrite app_ass. eauto.
+ red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_exprlist_invariant; eauto. UNCHANGED.
auto. auto. auto.
(* cons right *)
inv H2.
- assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
+ assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
- TR. subst sl2. eauto.
- red; auto.
+ TR. subst sl2. eauto.
+ red; auto.
intros. change sl3 with (nil ++ sl3). rewrite app_ass. econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
apply S; auto.
@@ -747,7 +747,7 @@ Proof.
exists dst'; exists sl1; exists sl2; exists a'; exists tmp'.
split. apply tr_top_base; auto.
split. auto. split. auto.
- intros. apply tr_top_base. apply S; auto.
+ intros. apply tr_top_base. apply S; auto.
Qed.
(** Semantics of smart constructors *)
@@ -756,18 +756,18 @@ Lemma eval_simpl_expr_sound:
forall e le m a v, eval_expr tge e le m a v ->
match eval_simpl_expr a with Some v' => v' = v | None => True end.
Proof.
- induction 1; simpl; auto.
- destruct (eval_simpl_expr a); auto. subst. rewrite H0. auto.
+ induction 1; simpl; auto.
+ destruct (eval_simpl_expr a); auto. subst. rewrite H0. auto.
inv H; simpl; auto.
Qed.
Lemma static_bool_val_sound:
forall v t m b, bool_val v t Mem.empty = Some b -> bool_val v t m = Some b.
Proof.
- intros until b; unfold bool_val. destruct (classify_bool t); destruct v; auto.
- intros E. unfold Mem.weak_valid_pointer, Mem.valid_pointer, proj_sumbool in E.
- rewrite ! pred_dec_false in E by (apply Mem.perm_empty). discriminate.
-Qed.
+ intros until b; unfold bool_val. destruct (classify_bool t); destruct v; auto.
+ intros E. unfold Mem.weak_valid_pointer, Mem.valid_pointer, proj_sumbool in E.
+ rewrite ! pred_dec_false in E by (apply Mem.perm_empty). discriminate.
+Qed.
Lemma step_makeif:
forall f a s1 s2 k e le m v1 b,
@@ -778,13 +778,13 @@ Lemma step_makeif:
Proof.
intros. functional induction (makeif a s1 s2).
- exploit eval_simpl_expr_sound; eauto. rewrite e0. intro EQ; subst v.
- assert (bool_val v1 (typeof a) m = Some true) by (apply static_bool_val_sound; auto).
+ assert (bool_val v1 (typeof a) m = Some true) by (apply static_bool_val_sound; auto).
replace b with true by congruence. constructor.
- exploit eval_simpl_expr_sound; eauto. rewrite e0. intro EQ; subst v.
- assert (bool_val v1 (typeof a) m = Some false) by (apply static_bool_val_sound; auto).
+ assert (bool_val v1 (typeof a) m = Some false) by (apply static_bool_val_sound; auto).
replace b with false by congruence. constructor.
-- apply star_one. eapply step_ifthenelse; eauto.
-- apply star_one. eapply step_ifthenelse; eauto.
+- apply star_one. eapply step_ifthenelse; eauto.
+- apply star_one. eapply step_ifthenelse; eauto.
Qed.
Lemma step_make_set:
@@ -801,7 +801,7 @@ Proof.
intros. change (PTree.set id v le) with (set_opttemp (Some id) v le). econstructor.
econstructor. constructor. eauto.
simpl. unfold sem_cast. simpl. eauto. constructor.
- simpl. econstructor; eauto.
+ simpl. econstructor; eauto.
(* nonvolatile case *)
intros [A B]. subst t. constructor. eapply eval_Elvalue; eauto.
Qed.
@@ -823,9 +823,9 @@ Proof.
econstructor. constructor. eauto.
simpl. unfold sem_cast. simpl. eauto.
econstructor; eauto. rewrite H3; eauto. constructor.
- simpl. econstructor; eauto.
+ simpl. econstructor; eauto.
(* nonvolatile case *)
- intros [A B]. subst t. econstructor; eauto. congruence.
+ intros [A B]. subst t. econstructor; eauto. congruence.
Qed.
Fixpoint Kseqlist (sl: list statement) (k: cont) :=
@@ -846,7 +846,7 @@ Lemma push_seq:
star step1 tge (State f (makeseq sl) k e le m)
E0 (State f Sskip (Kseqlist sl k) e le m).
Proof.
- intros. unfold makeseq. generalize Sskip. revert sl k.
+ intros. unfold makeseq. generalize Sskip. revert sl k.
induction sl; simpl; intros.
apply star_refl.
eapply star_right. apply IHsl. constructor. traceEq.
@@ -868,16 +868,16 @@ Proof.
intros. inv H1.
(* not volatile *)
exploit deref_loc_translated; eauto. unfold chunk_for_volatile_type; rewrite H3.
- intros [A B]. subst t.
+ intros [A B]. subst t.
exists le; split. apply star_refl.
split. eapply eval_Elvalue; eauto.
auto.
(* volatile *)
intros. exists (PTree.set t0 v le); split.
- simpl. eapply star_two. econstructor. eapply step_make_set; eauto. traceEq.
+ simpl. eapply star_two. econstructor. eapply step_make_set; eauto. traceEq.
split. constructor. apply PTree.gss.
split. auto.
- intros. apply PTree.gso. congruence.
+ intros. apply PTree.gso. congruence.
Qed.
(** Matching between continuations *)
@@ -980,7 +980,7 @@ Lemma match_cont_call:
match_cont k tk ->
match_cont (Csem.call_cont k) (call_cont tk).
Proof.
- induction 1; simpl; auto. constructor. econstructor; eauto.
+ induction 1; simpl; auto. constructor. econstructor; eauto.
Qed.
(** Matching between states *)
@@ -1029,14 +1029,14 @@ Proof.
| Some ls' =>
exists tls', select_switch_case n tls = Some tls' /\ tr_lblstmts ls' tls'
end).
- { induction 1; simpl; intros.
+ { induction 1; simpl; intros.
auto.
- destruct c; auto. destruct (zeq z n); auto.
+ destruct c; auto. destruct (zeq z n); auto.
econstructor; split; eauto. constructor; auto. }
intros. unfold Csem.select_switch, select_switch.
- specialize (CASE n ls tls H).
- destruct (Csem.select_switch_case n ls) as [ls'|].
- destruct CASE as [tls' [P Q]]. rewrite P. auto.
+ specialize (CASE n ls tls H).
+ destruct (Csem.select_switch_case n ls) as [ls'|].
+ destruct CASE as [tls' [P Q]]. rewrite P. auto.
rewrite CASE. apply DFL; auto.
Qed.
@@ -1066,22 +1066,22 @@ Fixpoint nolabel_list (sl: list statement) : Prop :=
Lemma nolabel_list_app:
forall sl2 sl1, nolabel_list sl1 -> nolabel_list sl2 -> nolabel_list (sl1 ++ sl2).
Proof.
- induction sl1; simpl; intros. auto. tauto.
+ induction sl1; simpl; intros. auto. tauto.
Qed.
Lemma makeseq_nolabel:
forall sl, nolabel_list sl -> nolabel (makeseq sl).
Proof.
assert (forall sl s, nolabel s -> nolabel_list sl -> nolabel (makeseq_rec s sl)).
- induction sl; simpl; intros. auto. destruct H0. apply IHsl; auto.
+ induction sl; simpl; intros. auto. destruct H0. apply IHsl; auto.
red. intros; simpl. rewrite H. apply H0.
- intros. unfold makeseq. apply H; auto. red. auto.
+ intros. unfold makeseq. apply H; auto. red. auto.
Qed.
Lemma makeif_nolabel:
forall a s1 s2, nolabel s1 -> nolabel s2 -> nolabel (makeif a s1 s2).
Proof.
- intros. functional induction (makeif a s1 s2); auto.
+ intros. functional induction (makeif a s1 s2); auto.
red; simpl; intros. rewrite H; auto.
red; simpl; intros. rewrite H; auto.
Qed.
@@ -1089,21 +1089,21 @@ Qed.
Lemma make_set_nolabel:
forall t a, nolabel (make_set t a).
Proof.
- unfold make_set; intros; red; intros.
+ unfold make_set; intros; red; intros.
destruct (chunk_for_volatile_type (typeof a)); auto.
Qed.
Lemma make_assign_nolabel:
forall l r, nolabel (make_assign l r).
Proof.
- unfold make_assign; intros; red; intros.
+ unfold make_assign; intros; red; intros.
destruct (chunk_for_volatile_type (typeof l)); auto.
Qed.
Lemma tr_rvalof_nolabel:
forall ty a sl a' tmp, tr_rvalof ty a sl a' tmp -> nolabel_list sl.
Proof.
- destruct 1; simpl; intuition. apply make_set_nolabel.
+ destruct 1; simpl; intuition. apply make_set_nolabel.
Qed.
Lemma nolabel_do_set:
@@ -1115,8 +1115,8 @@ Qed.
Lemma nolabel_final:
forall dst a, nolabel_list (final dst a).
Proof.
- destruct dst; simpl; intros. auto. auto. apply nolabel_do_set.
-Qed.
+ destruct dst; simpl; intros. auto. auto. apply nolabel_do_set.
+Qed.
Ltac NoLabelTac :=
match goal with
@@ -1179,7 +1179,7 @@ Lemma tr_find_label_if:
tr_if r Sskip Sbreak s ->
forall k, find_label lbl s k = None.
Proof.
- intros. inv H.
+ intros. inv H.
assert (nolabel (makeseq (sl ++ makeif a Sskip Sbreak :: nil))).
apply makeseq_nolabel.
apply nolabel_list_app.
@@ -1223,28 +1223,28 @@ Proof.
exploit (IHs1 (Csem.Kseq s2 k)); eauto. constructor; eauto.
destruct (Csem.find_label lbl s1 (Csem.Kseq s2 k)) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; auto.
- intro EQ. rewrite EQ. eapply IHs2; eauto.
+ intro EQ. rewrite EQ. eapply IHs2; eauto.
(* if *)
- rename s' into sr.
+ rename s' into sr.
rewrite (tr_find_label_expression _ _ _ H2).
exploit (IHs1 k); eauto.
destruct (Csem.find_label lbl s1 k) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
intro EQ. rewrite EQ. eapply IHs2; eauto.
(* while *)
- rename s' into sr.
+ rename s' into sr.
rewrite (tr_find_label_if _ _ H1); auto.
- exploit (IHs (Kwhile2 e s k)); eauto. econstructor; eauto.
+ exploit (IHs (Kwhile2 e s k)); eauto. econstructor; eauto.
destruct (Csem.find_label lbl s (Kwhile2 e s k)) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
- intro EQ. rewrite EQ. auto.
+ intro EQ. rewrite EQ. auto.
(* dowhile *)
rename s' into sr.
rewrite (tr_find_label_if _ _ H1); auto.
exploit (IHs (Kdowhile1 e s k)); eauto. econstructor; eauto.
destruct (Csem.find_label lbl s (Kdowhile1 e s k)) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
- intro EQ. rewrite EQ. auto.
+ intro EQ. rewrite EQ. auto.
(* for skip *)
rename s' into sr.
rewrite (tr_find_label_if _ _ H4); auto.
@@ -1256,8 +1256,8 @@ Proof.
(* for not skip *)
rename s' into sr.
rewrite (tr_find_label_if _ _ H3); auto.
- exploit (IHs1 (Csem.Kseq (Csyntax.Sfor Csyntax.Sskip e s2 s3) k)); eauto.
- econstructor; eauto. econstructor; eauto.
+ exploit (IHs1 (Csem.Kseq (Csyntax.Sfor Csyntax.Sskip e s2 s3) k)); eauto.
+ econstructor; eauto. econstructor; eauto.
destruct (Csem.find_label lbl s1
(Csem.Kseq (Csyntax.Sfor Csyntax.Sskip e s2 s3) k)) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
@@ -1265,7 +1265,7 @@ Proof.
exploit (IHs3 (Csem.Kfor3 e s2 s3 k)); eauto. econstructor; eauto.
destruct (Csem.find_label lbl s3 (Csem.Kfor3 e s2 s3 k)) as [[s'' k''] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
- intro EQ'. rewrite EQ'.
+ intro EQ'. rewrite EQ'.
exploit (IHs2 (Csem.Kfor4 e s2 s3 k)); eauto. econstructor; eauto.
(* break, continue, return 0 *)
auto. auto. auto.
@@ -1274,7 +1274,7 @@ Proof.
(* switch *)
rewrite (tr_find_label_expression _ _ _ H1). apply tr_find_label_ls. auto. constructor; auto.
(* labeled stmt *)
- destruct (ident_eq lbl l). exists ts0; exists tk; auto. apply IHs; auto.
+ destruct (ident_eq lbl l). exists ts0; exists tk; auto. apply IHs; auto.
(* goto *)
auto.
@@ -1390,7 +1390,7 @@ Proof.
induction 1; intros; inv MS.
(* expr *)
assert (tr_expr le dest r sl a tmps).
- inv H9. contradiction. auto.
+ inv H9. contradiction. auto.
exploit tr_simple_rvalue; eauto. destruct dest.
(* for val *)
intros [SL1 [TY1 EV1]]. subst sl.
@@ -1412,15 +1412,15 @@ Proof.
inv P. inv H2. inv H7; try congruence.
exploit tr_simple_lvalue; eauto. intros [SL [TY EV]]. subst sl0; simpl.
econstructor; split.
- left. eapply plus_two. constructor. eapply step_make_set; eauto. traceEq.
+ left. eapply plus_two. constructor. eapply step_make_set; eauto. traceEq.
econstructor; eauto.
change (final dst' (Etempvar t0 (Csyntax.typeof l)) ++ sl2) with (nil ++ (final dst' (Etempvar t0 (Csyntax.typeof l)) ++ sl2)).
- apply S. apply tr_val_gen. auto.
+ apply S. apply tr_val_gen. auto.
intros. constructor. rewrite H5; auto. apply PTree.gss.
- intros. apply PTree.gso. red; intros; subst; elim H5; auto.
+ intros. apply PTree.gso. red; intros; subst; elim H5; auto.
auto.
(* seqand true *)
- exploit tr_top_leftcontext; eauto. clear H9.
+ exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
(* for val *)
@@ -1432,7 +1432,7 @@ Proof.
apply push_seq. reflexivity. reflexivity.
rewrite <- Kseqlist_app.
eapply match_exprstates; eauto.
- apply S. apply tr_paren_val with (a1 := a2); auto.
+ apply S. apply tr_paren_val with (a1 := a2); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
(* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
@@ -1457,7 +1457,7 @@ Proof.
apply S. apply tr_paren_set with (a1 := a2) (t := sd_temp sd); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
(* seqand false *)
- exploit tr_top_leftcontext; eauto. clear H9.
+ exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
(* for val *)
@@ -1466,19 +1466,19 @@ Proof.
econstructor; split.
left. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
- apply star_one. constructor. constructor. reflexivity. reflexivity.
+ apply star_one. constructor. constructor. reflexivity. reflexivity.
eapply match_exprstates; eauto.
change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
- intros. constructor. rewrite H2. apply PTree.gss. auto.
+ intros. constructor. rewrite H2. apply PTree.gss. auto.
intros. apply PTree.gso. congruence.
- auto.
+ auto.
(* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
left. eapply plus_left. constructor.
apply step_makeif with (b := false) (v1 := v); auto. congruence.
- reflexivity.
+ reflexivity.
eapply match_exprstates; eauto.
change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
auto. auto.
@@ -1493,7 +1493,7 @@ Proof.
eapply match_exprstates; eauto.
apply S. econstructor; eauto. intros. constructor. auto. auto.
(* seqor true *)
- exploit tr_top_leftcontext; eauto. clear H9.
+ exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
(* for val *)
@@ -1502,19 +1502,19 @@ Proof.
econstructor; split.
left. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
- apply star_one. constructor. constructor. reflexivity. reflexivity.
+ apply star_one. constructor. constructor. reflexivity. reflexivity.
eapply match_exprstates; eauto.
change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
- intros. constructor. rewrite H2. apply PTree.gss. auto.
+ intros. constructor. rewrite H2. apply PTree.gss. auto.
intros. apply PTree.gso. congruence.
- auto.
+ auto.
(* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
left. eapply plus_left. constructor.
apply step_makeif with (b := true) (v1 := v); auto. congruence.
- reflexivity.
+ reflexivity.
eapply match_exprstates; eauto.
change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
auto. auto.
@@ -1529,7 +1529,7 @@ Proof.
eapply match_exprstates; eauto.
apply S. econstructor; eauto. intros. constructor. auto. auto.
(* seqand false *)
- exploit tr_top_leftcontext; eauto. clear H9.
+ exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
(* for val *)
@@ -1541,7 +1541,7 @@ Proof.
apply push_seq. reflexivity. reflexivity.
rewrite <- Kseqlist_app.
eapply match_exprstates; eauto.
- apply S. apply tr_paren_val with (a1 := a2); auto.
+ apply S. apply tr_paren_val with (a1 := a2); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
(* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
@@ -1566,9 +1566,9 @@ Proof.
apply S. apply tr_paren_set with (a1 := a2) (t := sd_temp sd); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
(* condition *)
- exploit tr_top_leftcontext; eauto. clear H9.
+ exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
- inv P. inv H2.
+ inv P. inv H2.
(* for value *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist. destruct b.
@@ -1592,22 +1592,22 @@ Proof.
econstructor; split.
left. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
- apply push_seq.
+ apply push_seq.
reflexivity. traceEq.
rewrite <- Kseqlist_app.
econstructor. eauto. apply S.
- econstructor; eauto. apply tr_expr_monotone with tmp2; eauto.
- econstructor; eauto.
+ econstructor; eauto. apply tr_expr_monotone with tmp2; eauto.
+ econstructor; eauto.
auto. auto.
econstructor; split.
left. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
- apply push_seq.
+ apply push_seq.
reflexivity. traceEq.
rewrite <- Kseqlist_app.
econstructor. eauto. apply S.
- econstructor; eauto. apply tr_expr_monotone with tmp3; eauto.
- econstructor; eauto.
+ econstructor; eauto. apply tr_expr_monotone with tmp3; eauto.
+ econstructor; eauto.
auto. auto.
(* for set *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
@@ -1615,25 +1615,25 @@ Proof.
econstructor; split.
left. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
- apply push_seq.
+ apply push_seq.
reflexivity. traceEq.
rewrite <- Kseqlist_app.
econstructor. eauto. apply S.
- econstructor; eauto. apply tr_expr_monotone with tmp2; eauto.
- econstructor; eauto.
+ econstructor; eauto. apply tr_expr_monotone with tmp2; eauto.
+ econstructor; eauto.
auto. auto.
econstructor; split.
left. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
- apply push_seq.
+ apply push_seq.
reflexivity. traceEq.
rewrite <- Kseqlist_app.
econstructor. eauto. apply S.
- econstructor; eauto. apply tr_expr_monotone with tmp3; eauto.
- econstructor; eauto.
+ econstructor; eauto. apply tr_expr_monotone with tmp3; eauto.
+ econstructor; eauto.
auto. auto.
(* assign *)
- exploit tr_top_leftcontext; eauto. clear H12.
+ exploit tr_top_leftcontext; eauto. clear H12.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H4.
(* for effects *)
@@ -1642,44 +1642,44 @@ Proof.
subst; simpl Kseqlist.
econstructor; split.
left. eapply plus_left. constructor.
- apply star_one. eapply step_make_assign; eauto.
+ apply star_one. eapply step_make_assign; eauto.
rewrite <- TY2; eauto. traceEq.
econstructor. auto. change sl2 with (nil ++ sl2). apply S.
constructor. auto. auto. auto.
(* for value *)
exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
exploit tr_simple_lvalue. eauto.
- eapply tr_expr_invariant with (le' := PTree.set t0 v le). eauto.
+ eapply tr_expr_invariant with (le' := PTree.set t0 v le). eauto.
intros. apply PTree.gso. intuition congruence.
intros [SL1 [TY1 EV1]].
subst; simpl Kseqlist.
econstructor; split.
- left. eapply plus_left. constructor.
- eapply star_left. constructor. eauto.
+ left. eapply plus_left. constructor.
+ eapply star_left. constructor. eauto.
eapply star_left. constructor.
- apply star_one. eapply step_make_assign; eauto.
- constructor. apply PTree.gss. reflexivity. reflexivity. traceEq.
+ apply star_one. eapply step_make_assign; eauto.
+ constructor. apply PTree.gss. reflexivity. reflexivity. traceEq.
econstructor. auto. apply S.
- apply tr_val_gen. auto. intros. econstructor; eauto. constructor.
- rewrite H4; auto. apply PTree.gss.
+ apply tr_val_gen. auto. intros. econstructor; eauto. constructor.
+ rewrite H4; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
auto. auto.
(* assignop *)
- exploit tr_top_leftcontext; eauto. clear H15.
+ exploit tr_top_leftcontext; eauto. clear H15.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H6.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
- exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
+ exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
intros. apply INV. NOTIN. intros [? [? EV1']].
- exploit tr_simple_rvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
+ exploit tr_simple_rvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
intros. apply INV. NOTIN. simpl. intros [SL2 [TY2 EV2]].
- subst; simpl Kseqlist.
+ subst; simpl Kseqlist.
econstructor; split.
left. eapply star_plus_trans. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
- eapply plus_two. simpl. econstructor. eapply step_make_assign; eauto.
- econstructor. eexact EV3. eexact EV2.
+ eapply plus_two. simpl. econstructor. eapply step_make_assign; eauto.
+ econstructor. eexact EV3. eexact EV2.
rewrite TY3; rewrite <- TY1; rewrite <- TY2; rewrite comp_env_preserved; auto.
reflexivity. traceEq.
econstructor. auto. change sl2 with (nil ++ sl2). apply S.
@@ -1687,132 +1687,132 @@ Proof.
(* for value *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
- exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
+ exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
intros. apply INV. NOTIN. intros [? [? EV1']].
- exploit tr_simple_rvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
+ exploit tr_simple_rvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
intros. apply INV. NOTIN. simpl. intros [SL2 [TY2 EV2]].
- exploit tr_simple_lvalue. eauto.
- eapply tr_expr_invariant with (le := le) (le' := PTree.set t v3 le'). eauto.
+ exploit tr_simple_lvalue. eauto.
+ eapply tr_expr_invariant with (le := le) (le' := PTree.set t v3 le'). eauto.
intros. rewrite PTree.gso. apply INV. NOTIN. intuition congruence.
intros [? [? EV1'']].
subst; simpl Kseqlist.
econstructor; split.
- left. rewrite app_ass. rewrite Kseqlist_app.
+ left. rewrite app_ass. rewrite Kseqlist_app.
eapply star_plus_trans. eexact EXEC.
simpl. eapply plus_four. econstructor. econstructor.
- econstructor. eexact EV3. eexact EV2.
+ econstructor. eexact EV3. eexact EV2.
rewrite TY3; rewrite <- TY1; rewrite <- TY2; rewrite comp_env_preserved; eauto.
- econstructor. eapply step_make_assign; eauto.
- constructor. apply PTree.gss.
+ econstructor. eapply step_make_assign; eauto.
+ constructor. apply PTree.gss.
reflexivity. traceEq.
econstructor. auto. apply S.
- apply tr_val_gen. auto. intros. econstructor; eauto. constructor.
- rewrite H10; auto. apply PTree.gss.
- intros. rewrite PTree.gso. apply INV.
- red; intros; elim H10; auto.
+ apply tr_val_gen. auto. intros. econstructor; eauto. constructor.
+ rewrite H10; auto. apply PTree.gss.
+ intros. rewrite PTree.gso. apply INV.
+ red; intros; elim H10; auto.
intuition congruence.
auto. auto.
(* assignop stuck *)
- exploit tr_top_leftcontext; eauto. clear H12.
+ exploit tr_top_leftcontext; eauto. clear H12.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H4.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
- subst; simpl Kseqlist.
+ subst; simpl Kseqlist.
econstructor; split.
- right; split. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
+ right; split. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
simpl. omega.
constructor.
(* for value *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
- subst; simpl Kseqlist.
+ subst; simpl Kseqlist.
econstructor; split.
- right; split. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
+ right; split. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
simpl. omega.
constructor.
(* postincr *)
- exploit tr_top_leftcontext; eauto. clear H14.
+ exploit tr_top_leftcontext; eauto. clear H14.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H5.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
- exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
+ exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
intros. apply INV. NOTIN. intros [? [? EV1']].
subst; simpl Kseqlist.
econstructor; split.
- left. rewrite app_ass; rewrite Kseqlist_app.
- eapply star_plus_trans. eexact EXEC.
+ left. rewrite app_ass; rewrite Kseqlist_app.
+ eapply star_plus_trans. eexact EXEC.
eapply plus_two. simpl. constructor.
- eapply step_make_assign; eauto.
- unfold transl_incrdecr. destruct id; simpl in H2.
+ eapply step_make_assign; eauto.
+ unfold transl_incrdecr. destruct id; simpl in H2.
econstructor. eauto. constructor. rewrite TY3; rewrite <- TY1; rewrite comp_env_preserved. simpl; eauto.
econstructor. eauto. constructor. rewrite TY3; rewrite <- TY1; rewrite comp_env_preserved. simpl; eauto.
- destruct id; auto.
+ destruct id; auto.
reflexivity. traceEq.
econstructor. auto. change sl2 with (nil ++ sl2). apply S.
constructor. auto. auto. auto.
(* for value *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_lvalue. eauto.
- eapply tr_expr_invariant with (le' := PTree.set t v1 le). eauto.
+ eapply tr_expr_invariant with (le' := PTree.set t v1 le). eauto.
intros. apply PTree.gso. intuition congruence.
intros [SL2 [TY2 EV2]].
subst; simpl Kseqlist.
econstructor; split.
left. eapply plus_four. constructor.
- eapply step_make_set; eauto.
+ eapply step_make_set; eauto.
constructor.
- eapply step_make_assign; eauto.
- unfold transl_incrdecr. destruct id; simpl in H2.
+ eapply step_make_assign; eauto.
+ unfold transl_incrdecr. destruct id; simpl in H2.
econstructor. constructor. apply PTree.gss. constructor.
rewrite comp_env_preserved; simpl; eauto.
econstructor. constructor. apply PTree.gss. constructor.
rewrite comp_env_preserved; simpl; eauto.
- destruct id; auto.
+ destruct id; auto.
traceEq.
econstructor. auto. apply S.
apply tr_val_gen. auto. intros. econstructor; eauto.
- rewrite H5; auto. apply PTree.gss.
+ rewrite H5; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
auto. auto.
(* postincr stuck *)
- exploit tr_top_leftcontext; eauto. clear H11.
+ exploit tr_top_leftcontext; eauto. clear H11.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H3.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
- subst. simpl Kseqlist.
+ subst. simpl Kseqlist.
econstructor; split.
right; split. rewrite app_ass; rewrite Kseqlist_app. eexact EXEC.
simpl; omega.
constructor.
(* for value *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
- subst. simpl Kseqlist.
+ subst. simpl Kseqlist.
econstructor; split.
- left. eapply plus_two. constructor. eapply step_make_set; eauto.
+ left. eapply plus_two. constructor. eapply step_make_set; eauto.
traceEq.
constructor.
(* comma *)
- exploit tr_top_leftcontext; eauto. clear H9.
+ exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H1.
exploit tr_simple_rvalue; eauto. simpl; intro SL1.
subst sl0; simpl Kseqlist.
econstructor; split.
- right; split. apply star_refl. simpl. apply plus_lt_compat_r.
- apply (leftcontext_size _ _ _ H). simpl. omega.
- econstructor; eauto. apply S.
- eapply tr_expr_monotone; eauto.
- auto. auto.
+ right; split. apply star_refl. simpl. apply plus_lt_compat_r.
+ apply (leftcontext_size _ _ _ H). simpl. omega.
+ econstructor; eauto. apply S.
+ eapply tr_expr_monotone; eauto.
+ auto. auto.
(* paren *)
- exploit tr_top_leftcontext; eauto. clear H9.
+ exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
(* for value *)
@@ -1821,11 +1821,11 @@ Proof.
econstructor; split.
left. eapply plus_left. constructor. apply star_one.
econstructor. econstructor; eauto. rewrite <- TY1; eauto. traceEq.
- econstructor; eauto.
+ econstructor; eauto.
change sl2 with (final For_val (Etempvar t (Csyntax.typeof r)) ++ sl2). apply S.
constructor. auto. intros. constructor. rewrite H2; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
- auto.
+ auto.
(* for effects *)
econstructor; split.
right; split. apply star_refl. simpl. apply plus_lt_compat_r.
@@ -1835,18 +1835,18 @@ Proof.
apply S. constructor; auto. auto. auto.
(* for set *)
exploit tr_simple_rvalue; eauto. simpl. intros [b [SL1 [TY1 EV1]]]. subst sl1.
- simpl Kseqlist.
+ simpl Kseqlist.
econstructor; split.
- left. eapply plus_left. constructor. apply star_one. econstructor. econstructor; eauto.
+ left. eapply plus_left. constructor. apply star_one. econstructor. econstructor; eauto.
rewrite <- TY1; eauto. traceEq.
econstructor; eauto.
- apply S. constructor; auto.
- intros. constructor. rewrite H2. apply PTree.gss. auto.
+ apply S. constructor; auto.
+ intros. constructor. rewrite H2. apply PTree.gss. auto.
intros. apply PTree.gso. congruence.
- auto.
+ auto.
(* call *)
- exploit tr_top_leftcontext; eauto. clear H12.
+ exploit tr_top_leftcontext; eauto. clear H12.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H5.
(* for effects *)
@@ -1854,20 +1854,20 @@ Proof.
exploit tr_simple_exprlist; eauto. intros [SL2 EV2].
subst. simpl Kseqlist.
exploit functions_translated; eauto. intros [tfd [J K]].
- econstructor; split.
+ econstructor; split.
left. eapply plus_left. constructor. apply star_one.
econstructor; eauto. rewrite <- TY1; eauto.
exploit type_of_fundef_preserved; eauto. congruence.
traceEq.
constructor; auto. econstructor; eauto.
intros. change sl2 with (nil ++ sl2). apply S.
- constructor. auto. auto.
+ constructor. auto. auto.
(* for value *)
exploit tr_simple_rvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_exprlist; eauto. intros [SL2 EV2].
subst. simpl Kseqlist.
exploit functions_translated; eauto. intros [tfd [J K]].
- econstructor; split.
+ econstructor; split.
left. eapply plus_left. constructor. apply star_one.
econstructor; eauto. rewrite <- TY1; eauto.
exploit type_of_fundef_preserved; eauto. congruence.
@@ -1875,10 +1875,10 @@ Proof.
constructor; auto. econstructor; eauto.
intros. apply S.
destruct dst'; constructor.
- auto. intros. constructor. rewrite H5; auto. apply PTree.gss.
- auto. intros. constructor. rewrite H5; auto. apply PTree.gss.
+ auto. intros. constructor. rewrite H5; auto. apply PTree.gss.
+ auto. intros. constructor. rewrite H5; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
- auto.
+ auto.
(* builtin *)
exploit tr_top_leftcontext; eauto. clear H9.
@@ -1887,22 +1887,22 @@ Proof.
(* for effects *)
exploit tr_simple_exprlist; eauto. intros [SL EV].
subst. simpl Kseqlist.
- econstructor; split.
+ econstructor; split.
left. eapply plus_left. constructor. apply star_one.
econstructor; eauto.
- eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
+ eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
traceEq.
econstructor; eauto.
- change sl2 with (nil ++ sl2). apply S. constructor. simpl; auto. auto.
+ change sl2 with (nil ++ sl2). apply S. constructor. simpl; auto. auto.
(* for value *)
exploit tr_simple_exprlist; eauto. intros [SL EV].
subst. simpl Kseqlist.
- econstructor; split.
+ econstructor; split.
left. eapply plus_left. constructor. apply star_one.
econstructor; eauto.
- eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
+ eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
traceEq.
econstructor; eauto.
change sl2 with (nil ++ sl2). apply S.
@@ -1918,7 +1918,7 @@ Lemma tr_top_val_for_val_inv:
tr_top tge e le m For_val (Csyntax.Eval v ty) sl a tmps ->
sl = nil /\ typeof a = ty /\ eval_expr tge e le m a v.
Proof.
- intros. inv H. auto. inv H0. auto.
+ intros. inv H. auto. inv H0. auto.
Qed.
Lemma alloc_variables_preserved:
@@ -1934,7 +1934,7 @@ Lemma bind_parameters_preserved:
Csem.bind_parameters ge e m params args m' ->
bind_parameters tge e m params args m'.
Proof.
- induction 1; econstructor; eauto. inv H0.
+ induction 1; econstructor; eauto. inv H0.
- eapply assign_loc_value; eauto.
- inv H4. eapply assign_loc_value; eauto.
- rewrite <- comp_env_preserved in *. eapply assign_loc_copy; eauto.
@@ -1943,10 +1943,10 @@ Qed.
Lemma blocks_of_env_preserved:
forall e, blocks_of_env tge e = Csem.blocks_of_env ge e.
Proof.
- intros; unfold blocks_of_env, Csem.blocks_of_env.
- unfold block_of_binding, Csem.block_of_binding.
+ intros; unfold blocks_of_env, Csem.blocks_of_env.
+ unfold block_of_binding, Csem.block_of_binding.
rewrite comp_env_preserved. auto.
-Qed.
+Qed.
Lemma sstep_simulation:
forall S1 t S2, Csem.sstep ge S1 t S2 ->
@@ -1958,14 +1958,14 @@ Lemma sstep_simulation:
Proof.
induction 1; intros; inv MS.
(* do 1 *)
- inv H6. inv H0.
+ inv H6. inv H0.
econstructor; split.
- right; split. apply push_seq.
+ right; split. apply push_seq.
simpl. omega.
econstructor; eauto. constructor. auto.
(* do 2 *)
- inv H7. inv H6. inv H.
- econstructor; split.
+ inv H7. inv H6. inv H.
+ econstructor; split.
right; split. apply star_refl. simpl. omega.
econstructor; eauto. constructor.
@@ -1992,8 +1992,8 @@ Proof.
econstructor; eauto. econstructor; eauto.
(* ifthenelse *)
- inv H8.
- exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+ inv H8.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. eapply plus_two. constructor.
apply step_ifthenelse with (v1 := v) (b := b); auto. traceEq.
@@ -2007,68 +2007,68 @@ Proof.
reflexivity. traceEq. rewrite Kseqlist_app.
econstructor; eauto. simpl. econstructor; eauto. econstructor; eauto.
(* while false *)
- inv H8.
- exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+ inv H8.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (v1 := v) (b := false); auto.
- eapply star_two. constructor. apply step_break_loop1.
+ eapply star_two. constructor. apply step_break_loop1.
reflexivity. reflexivity. traceEq.
constructor; auto. constructor.
(* while true *)
- inv H8.
- exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+ inv H8.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
constructor.
reflexivity. traceEq.
- constructor; auto. constructor; auto.
+ constructor; auto. constructor; auto.
(* skip-or-continue while *)
assert (ts = Sskip \/ ts = Scontinue). destruct H; subst s0; inv H7; auto.
inv H8.
econstructor; split.
left. eapply plus_two. apply step_skip_or_continue_loop1; auto.
apply step_skip_loop2. traceEq.
- constructor; auto. constructor; auto.
+ constructor; auto. constructor; auto.
(* break while *)
- inv H6. inv H7.
+ inv H6. inv H7.
econstructor; split.
left. apply plus_one. apply step_break_loop1.
constructor; auto. constructor.
(* dowhile *)
- inv H6.
+ inv H6.
econstructor; split.
- left. apply plus_one. apply step_loop.
+ left. apply plus_one. apply step_loop.
constructor; auto. constructor; auto.
(* skip_or_continue dowhile *)
assert (ts = Sskip \/ ts = Scontinue). destruct H; subst s0; inv H7; auto.
inv H8. inv H4.
econstructor; split.
left. eapply plus_left. apply step_skip_or_continue_loop1. auto.
- apply push_seq.
+ apply push_seq.
traceEq.
rewrite Kseqlist_app.
- econstructor; eauto. simpl. econstructor; auto. econstructor; eauto.
+ econstructor; eauto. simpl. econstructor; auto. econstructor; eauto.
(* dowhile false *)
- inv H8.
- exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+ inv H8.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
- left. simpl. eapply plus_left. constructor.
- eapply star_right. apply step_makeif with (v1 := v) (b := false); auto.
- constructor.
+ left. simpl. eapply plus_left. constructor.
+ eapply star_right. apply step_makeif with (v1 := v) (b := false); auto.
+ constructor.
reflexivity. traceEq.
constructor; auto. constructor.
(* dowhile true *)
- inv H8.
- exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+ inv H8.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
- left. simpl. eapply plus_left. constructor.
- eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
- constructor.
+ left. simpl. eapply plus_left. constructor.
+ eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
+ constructor.
reflexivity. traceEq.
- constructor; auto. constructor; auto.
+ constructor; auto. constructor; auto.
(* break dowhile *)
inv H6. inv H7.
econstructor; split.
@@ -2076,14 +2076,14 @@ Proof.
constructor; auto. constructor.
(* for start *)
- inv H7. congruence.
- econstructor; split.
+ inv H7. congruence.
+ econstructor; split.
left; apply plus_one. constructor.
- econstructor; eauto. constructor; auto. econstructor; eauto.
+ econstructor; eauto. constructor; auto. econstructor; eauto.
(* for *)
- inv H6; try congruence. inv H2.
+ inv H6; try congruence. inv H2.
econstructor; split.
- left. eapply plus_left. apply step_loop.
+ left. eapply plus_left. apply step_loop.
eapply star_left. constructor. apply push_seq.
reflexivity. traceEq.
rewrite Kseqlist_app. econstructor; eauto. simpl. constructor; auto. econstructor; eauto.
@@ -2098,11 +2098,11 @@ Proof.
(* for true *)
inv H8. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
- left. simpl. eapply plus_left. constructor.
+ left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
constructor.
reflexivity. traceEq.
- constructor; auto. constructor; auto.
+ constructor; auto. constructor; auto.
(* skip_or_continue for3 *)
assert (ts = Sskip \/ ts = Scontinue). destruct H; subst x; inv H7; auto.
inv H8.
@@ -2110,21 +2110,21 @@ Proof.
left. apply plus_one. apply step_skip_or_continue_loop1. auto.
econstructor; eauto. econstructor; auto.
(* break for3 *)
- inv H6. inv H7.
+ inv H6. inv H7.
econstructor; split.
left. apply plus_one. apply step_break_loop1.
econstructor; eauto. constructor.
(* skip for4 *)
- inv H6. inv H7.
+ inv H6. inv H7.
econstructor; split.
left. apply plus_one. constructor.
- econstructor; eauto. constructor; auto.
+ econstructor; eauto. constructor; auto.
(* return none *)
inv H7. econstructor; split.
- left. apply plus_one. econstructor; eauto. rewrite blocks_of_env_preserved; eauto.
- constructor. apply match_cont_call; auto.
+ left. apply plus_one. econstructor; eauto. rewrite blocks_of_env_preserved; eauto.
+ constructor. apply match_cont_call; auto.
(* return some 1 *)
inv H6. inv H0. econstructor; split.
left; eapply plus_left. constructor. apply push_seq. traceEq.
@@ -2133,34 +2133,34 @@ Proof.
inv H9. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. eapply plus_two. constructor. econstructor. eauto.
- erewrite function_return_preserved; eauto. rewrite blocks_of_env_preserved; eauto.
+ erewrite function_return_preserved; eauto. rewrite blocks_of_env_preserved; eauto.
eauto. traceEq.
constructor. apply match_cont_call; auto.
(* skip return *)
- inv H8.
+ inv H8.
assert (is_call_cont tk). inv H9; simpl in *; auto.
econstructor; split.
- left. apply plus_one. apply step_skip_call; eauto. rewrite blocks_of_env_preserved; eauto.
+ left. apply plus_one. apply step_skip_call; eauto. rewrite blocks_of_env_preserved; eauto.
constructor. auto.
(* switch *)
- inv H6. inv H1.
- econstructor; split.
+ inv H6. inv H1.
+ econstructor; split.
left; eapply plus_left. constructor. apply push_seq. traceEq.
- econstructor; eauto. constructor; auto.
+ econstructor; eauto. constructor; auto.
(* expr switch *)
inv H8. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left; eapply plus_two. constructor. econstructor; eauto. traceEq.
econstructor; eauto.
- apply tr_seq_of_labeled_statement. apply tr_select_switch. auto.
+ apply tr_seq_of_labeled_statement. apply tr_select_switch. auto.
constructor; auto.
(* skip-or-break switch *)
assert (ts = Sskip \/ ts = Sbreak). destruct H; subst x; inv H7; auto.
inv H8.
econstructor; split.
- left; apply plus_one. apply step_skip_break_switch. auto.
+ left; apply plus_one. apply step_skip_break_switch. auto.
constructor; auto. constructor.
(* continue switch *)
@@ -2176,13 +2176,13 @@ Proof.
(* goto *)
inv H7.
- inversion H6; subst.
- exploit tr_find_label. eauto. apply match_cont_call. eauto.
- instantiate (1 := lbl). rewrite H.
- intros [ts' [tk' [P [Q R]]]].
- econstructor; split.
+ inversion H6; subst.
+ exploit tr_find_label. eauto. apply match_cont_call. eauto.
+ instantiate (1 := lbl). rewrite H.
+ intros [ts' [tk' [P [Q R]]]].
+ econstructor; split.
left. apply plus_one. econstructor; eauto.
- econstructor; eauto.
+ econstructor; eauto.
(* internal function *)
inv H7. inversion H3; subst.
@@ -2191,15 +2191,15 @@ Proof.
rewrite H6; rewrite H7; auto.
rewrite H6; rewrite H7. eapply alloc_variables_preserved; eauto.
rewrite H6. eapply bind_parameters_preserved; eauto.
- eauto.
- constructor; auto.
+ eauto.
+ constructor; auto.
(* external function *)
inv H5.
econstructor; split.
left; apply plus_one. econstructor; eauto.
- eapply external_call_symbols_preserved; eauto.
- exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
+ eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
constructor; auto.
(* return *)
@@ -2219,7 +2219,7 @@ Theorem simulation:
(star step1 tge S1' t S2' /\ measure S2 < measure S1)%nat)
/\ match_states S2 S2'.
Proof.
- intros S1 t S2 STEP. destruct STEP.
+ intros S1 t S2 STEP. destruct STEP.
apply estep_simulation; auto.
apply sstep_simulation; auto.
Qed.
@@ -2229,14 +2229,14 @@ Lemma transl_initial_states:
Csem.initial_state prog S ->
exists S', Clight.initial_state tprog S' /\ match_states S S'.
Proof.
- intros. inv H. generalize TRANSL; intros TR; monadInv TR. rewrite H4.
+ intros. inv H. generalize TRANSL; intros TR; monadInv TR. rewrite H4.
exploit transl_program_spec; eauto. intros MP.
exploit function_ptr_translated; eauto. intros [tf [FIND TR]].
econstructor; split.
econstructor.
exploit Genv.init_mem_match; eauto.
change (Genv.globalenv tprog) with (genv_genv tge). rewrite symbols_preserved.
- rewrite <- H4; simpl; eauto.
+ rewrite <- H4; simpl; eauto.
eexact FIND.
rewrite <- H3. apply type_of_fundef_preserved. auto.
constructor. auto. constructor.
diff --git a/cfrontend/SimplExprspec.v b/cfrontend/SimplExprspec.v
index 9f9fb450..7003c78a 100644
--- a/cfrontend/SimplExprspec.v
+++ b/cfrontend/SimplExprspec.v
@@ -76,14 +76,14 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
(forall tge e le' m,
(forall id, In id tmp -> le'!id = le!id) ->
eval_expr tge e le' m a v) ->
- tr_expr le For_val (Csyntax.Eval v ty)
+ tr_expr le For_val (Csyntax.Eval v ty)
nil a tmp
| tr_val_set: forall le sd v ty a any tmp,
typeof a = ty ->
(forall tge e le' m,
(forall id, In id tmp -> le'!id = le!id) ->
eval_expr tge e le' m a v) ->
- tr_expr le (For_set sd) (Csyntax.Eval v ty)
+ tr_expr le (For_set sd) (Csyntax.Eval v ty)
(do_set sd a) any tmp
| tr_sizeof: forall le dst ty' ty tmp,
tr_expr le dst (Csyntax.Esizeof ty' ty)
@@ -102,7 +102,7 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
a2 tmp
| tr_addrof: forall le dst e1 ty tmp sl1 a1,
tr_expr le For_val e1 sl1 a1 tmp ->
- tr_expr le dst (Csyntax.Eaddrof e1 ty)
+ tr_expr le dst (Csyntax.Eaddrof e1 ty)
(sl1 ++ final dst (Eaddrof a1 ty))
(Eaddrof a1 ty) tmp
| tr_unop: forall le dst op e1 ty tmp sl1 a1,
@@ -207,7 +207,7 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
| tr_assign_effects: forall le e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_expr le For_val e2 sl2 a2 tmp2 ->
- list_disjoint tmp1 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
incl tmp1 tmp -> incl tmp2 tmp ->
tr_expr le For_effects (Csyntax.Eassign e1 e2 ty)
(sl1 ++ sl2 ++ make_assign a1 a2 :: nil)
@@ -216,12 +216,12 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_expr le For_val e2 sl2 a2 tmp2 ->
incl tmp1 tmp -> incl tmp2 tmp ->
- list_disjoint tmp1 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
In t tmp -> ~In t tmp1 -> ~In t tmp2 ->
ty1 = Csyntax.typeof e1 ->
ty2 = Csyntax.typeof e2 ->
tr_expr le dst (Csyntax.Eassign e1 e2 ty)
- (sl1 ++ sl2 ++
+ (sl1 ++ sl2 ++
Sset t a2 ::
make_assign a1 (Etempvar t ty2) ::
final dst (Ecast (Etempvar t ty2) ty1))
@@ -255,7 +255,7 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
tr_rvalof ty1 a1 sl2 a2 tmp2 ->
ty1 = Csyntax.typeof e1 ->
incl tmp1 tmp -> incl tmp2 tmp ->
- list_disjoint tmp1 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
tr_expr le For_effects (Csyntax.Epostincr id e1 ty)
(sl1 ++ sl2 ++ make_assign a1 (transl_incrdecr id a2 ty1) :: nil)
any tmp
@@ -271,7 +271,7 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
| tr_comma: forall le dst e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 tmp,
tr_expr le For_effects e1 sl1 a1 tmp1 ->
tr_expr le dst e2 sl2 a2 tmp2 ->
- list_disjoint tmp1 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
incl tmp1 tmp -> incl tmp2 tmp ->
tr_expr le dst (Csyntax.Ecomma e1 e2 ty) (sl1 ++ sl2) a2 tmp
| tr_call_effects: forall le e1 el2 ty sl1 a1 tmp1 sl2 al2 tmp2 any tmp,
@@ -306,7 +306,7 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
(Etempvar t ty) tmp
| tr_paren_val: forall le e1 tycast ty sl1 a1 t tmp,
tr_expr le (For_set (SDbase tycast ty t)) e1 sl1 a1 tmp ->
- In t tmp ->
+ In t tmp ->
tr_expr le For_val (Csyntax.Eparen e1 tycast ty)
sl1
(Etempvar t ty) tmp
@@ -499,7 +499,7 @@ Proof.
intros until I. unfold bind. destruct (f z1).
congruence.
caseEq (g a g'); intros; inv H0.
- econstructor; econstructor; econstructor; econstructor; eauto.
+ econstructor; econstructor; econstructor; econstructor; eauto.
Qed.
Remark bind2_inversion:
@@ -508,9 +508,9 @@ Remark bind2_inversion:
exists x1, exists x2, exists z2, exists I1, exists I2,
f z1 = Res (x1,x2) z2 I1 /\ g x1 x2 z2 = Res y z3 I2.
Proof.
- unfold bind2. intros.
- exploit bind_inversion; eauto.
- intros [[x1 x2] [z2 [I1 [I2 [P Q]]]]]. simpl in Q.
+ unfold bind2. intros.
+ exploit bind_inversion; eauto.
+ intros [[x1 x2] [z2 [I1 [I2 [P Q]]]]]. simpl in Q.
exists x1; exists x2; exists z2; exists I1; exists I2; auto.
Qed.
@@ -551,7 +551,7 @@ Ltac monadInv H :=
| (@error _ _ _ = Res _ _ _) => monadInv1 H
| (bind ?F ?G ?Z = Res ?X ?Z' ?I) => monadInv1 H
| (bind2 ?F ?G ?Z = Res ?X ?Z' ?I) => monadInv1 H
- | (?F _ _ _ _ _ _ _ _ = Res _ _ _) =>
+ | (?F _ _ _ _ _ _ _ _ = Res _ _ _) =>
((progress simpl in H) || unfold F in H); monadInv1 H
| (?F _ _ _ _ _ _ _ = Res _ _ _) =>
((progress simpl in H) || unfold F in H); monadInv1 H
@@ -588,7 +588,7 @@ Lemma within_widen:
Ple (gen_next g2) (gen_next g2') ->
within id g1' g2'.
Proof.
- intros. destruct H. split.
+ intros. destruct H. split.
eapply Ple_trans; eauto.
eapply Plt_Ple_trans; eauto.
Qed.
@@ -609,14 +609,14 @@ Lemma contained_widen:
Ple (gen_next g2) (gen_next g2') ->
contained l g1' g2'.
Proof.
- intros; red; intros. eapply within_widen; eauto.
+ intros; red; intros. eapply within_widen; eauto.
Qed.
Lemma contained_cons:
forall id l g1 g2,
within id g1 g2 -> contained l g1 g2 -> contained (id :: l) g1 g2.
Proof.
- intros; red; intros. simpl in H1; destruct H1. subst id0. auto. auto.
+ intros; red; intros. simpl in H1; destruct H1. subst id0. auto. auto.
Qed.
Lemma contained_app:
@@ -630,7 +630,7 @@ Lemma contained_disjoint:
forall g1 l1 g2 l2 g3,
contained l1 g1 g2 -> contained l2 g2 g3 -> list_disjoint l1 l2.
Proof.
- intros; red; intros. red; intro; subst y.
+ intros; red; intros. red; intro; subst y.
exploit H; eauto. intros [A B]. exploit H0; eauto. intros [Csyntax D].
elim (Plt_strict x). apply Plt_Ple_trans with (gen_next g2); auto.
Qed.
@@ -665,7 +665,7 @@ Qed.
Lemma dest_for_set_seqbool_set:
forall sd ty g, dest_below (For_set sd) g -> dest_below (For_set (sd_seqbool_set ty sd)) g.
Proof.
- intros. assumption.
+ intros. assumption.
Qed.
Lemma dest_for_set_condition_val:
@@ -683,7 +683,7 @@ Qed.
Lemma sd_temp_notin:
forall sd g1 g2 l, dest_below (For_set sd) g1 -> contained l g1 g2 -> ~In (sd_temp sd) l.
Proof.
- intros. simpl in H. red; intros. exploit H0; eauto. intros [A B].
+ intros. simpl in H. red; intros. exploit H0; eauto. intros [A B].
elim (Plt_strict (sd_temp sd)). apply Plt_Ple_trans with (gen_next g1); auto.
Qed.
@@ -701,7 +701,7 @@ Hint Resolve gensym_within within_widen contained_widen
dest_for_set_condition_val dest_for_set_condition_set
sd_temp_notin dest_below_le
incl_refl incl_tl incl_app incl_appl incl_appr incl_same_head
- in_eq in_cons
+ in_eq in_cons
Ple_trans Ple_refl: gensym.
Hint Resolve dest_for_val_below dest_for_effect_below.
@@ -787,27 +787,27 @@ Proof.
(* val *)
simpl in H. destruct v; monadInv H; exists (@nil ident); split; auto with gensym.
Opaque makeif.
- intros. destruct dst; simpl in *; inv H2.
+ intros. destruct dst; simpl in *; inv H2.
constructor. auto. intros; constructor.
constructor.
constructor. auto. intros; constructor.
- intros. destruct dst; simpl in *; inv H2.
+ intros. destruct dst; simpl in *; inv H2.
constructor. auto. intros; constructor.
constructor.
constructor. auto. intros; constructor.
- intros. destruct dst; simpl in *; inv H2.
+ intros. destruct dst; simpl in *; inv H2.
constructor. auto. intros; constructor.
constructor.
constructor. auto. intros; constructor.
- intros. destruct dst; simpl in *; inv H2.
+ intros. destruct dst; simpl in *; inv H2.
constructor. auto. intros; constructor.
constructor.
constructor. auto. intros; constructor.
(* var *)
monadInv H; econstructor; split; auto with gensym. UseFinish. constructor.
(* field *)
- monadInv H0. exploit H; eauto. auto. intros [tmp [A B]]. UseFinish.
- econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
+ monadInv H0. exploit H; eauto. auto. intros [tmp [A B]]. UseFinish.
+ econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
(* valof *)
monadInv H0. exploit H; eauto. intros [tmp1 [A B]].
exploit transl_valof_meets_spec; eauto. intros [tmp2 [Csyntax D]]. UseFinish.
@@ -815,22 +815,22 @@ Opaque makeif.
intros; apply tr_expr_add_dest. econstructor; eauto with gensym.
eauto with gensym.
(* deref *)
- monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
- econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
+ monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
+ econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
(* addrof *)
- monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
+ monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
econstructor; split; eauto. intros; apply tr_expr_add_dest. econstructor; eauto.
(* unop *)
- monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
- econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
+ monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
+ econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
(* binop *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [Csyntax D]]. UseFinish.
- exists (tmp1 ++ tmp2); split.
+ exists (tmp1 ++ tmp2); split.
intros; apply tr_expr_add_dest. econstructor; eauto with gensym.
eauto with gensym.
(* cast *)
- monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
+ monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
(* seqand *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
@@ -841,7 +841,7 @@ Opaque makeif.
exists (x0 :: tmp1 ++ tmp2); split.
intros; eapply tr_seqand_val; eauto with gensym.
apply list_disjoint_cons_r; eauto with gensym.
- apply contained_cons. eauto with gensym.
+ apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* for effects *)
exploit H0; eauto with gensym. intros [tmp2 [Csyntax D]].
@@ -880,18 +880,18 @@ Opaque makeif.
intros; eapply tr_seqor_set; eauto with gensym.
apply list_disjoint_cons_r; eauto with gensym.
apply contained_app; eauto with gensym.
-(* condition *)
+(* condition *)
monadInv H2. exploit H; eauto. intros [tmp1 [A B]].
destruct dst; monadInv EQ0.
(* for value *)
exploit H0; eauto with gensym. intros [tmp2 [C D]].
exploit H1; eauto with gensym. intros [tmp3 [E F]].
simpl add_dest in *.
- exists (x0 :: tmp1 ++ tmp2 ++ tmp3); split.
+ exists (x0 :: tmp1 ++ tmp2 ++ tmp3); split.
simpl; intros; eapply tr_condition_val; eauto with gensym.
apply list_disjoint_cons_r; eauto with gensym.
apply list_disjoint_cons_r; eauto with gensym.
- apply contained_cons. eauto with gensym.
+ apply contained_cons. eauto with gensym.
apply contained_app. eauto with gensym.
apply contained_app; eauto with gensym.
(* for effects *)
@@ -899,7 +899,7 @@ Opaque makeif.
exploit H1; eauto. intros [tmp3 [E F]].
simpl add_dest in *.
exists (tmp1 ++ tmp2 ++ tmp3); split.
- intros; eapply tr_condition_effects; eauto with gensym.
+ intros; eapply tr_condition_effects; eauto with gensym.
apply contained_app; eauto with gensym.
(* for test *)
exploit H0; eauto with gensym. intros [tmp2 [C D]].
@@ -909,7 +909,7 @@ Opaque makeif.
intros; eapply tr_condition_set; eauto with gensym.
apply list_disjoint_cons_r; eauto with gensym.
apply list_disjoint_cons_r; eauto with gensym.
- apply contained_cons; eauto with gensym.
+ apply contained_cons; eauto with gensym.
apply contained_app; eauto with gensym.
apply contained_app; eauto with gensym.
(* sizeof *)
@@ -923,19 +923,19 @@ Opaque makeif.
exploit H0; eauto. intros [tmp2 [Csyntax D]].
destruct dst; monadInv EQ2; simpl add_dest in *.
(* for value *)
- exists (x1 :: tmp1 ++ tmp2); split.
+ exists (x1 :: tmp1 ++ tmp2); split.
intros. eapply tr_assign_val with (dst := For_val); eauto with gensym.
- apply contained_cons. eauto with gensym.
+ apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* for effects *)
- exists (tmp1 ++ tmp2); split.
+ exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
(* for set *)
exists (x1 :: tmp1 ++ tmp2); split.
- repeat rewrite app_ass. simpl.
+ repeat rewrite app_ass. simpl.
intros. eapply tr_assign_val with (dst := For_set sd); eauto with gensym.
- apply contained_cons. eauto with gensym.
+ apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* assignop *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
@@ -945,43 +945,43 @@ Opaque makeif.
(* for value *)
exists (x2 :: tmp1 ++ tmp2 ++ tmp3); split.
intros. eapply tr_assignop_val with (dst := For_val); eauto with gensym.
- apply contained_cons. eauto with gensym.
+ apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* for effects *)
- exists (tmp1 ++ tmp2 ++ tmp3); split.
+ exists (tmp1 ++ tmp2 ++ tmp3); split.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
(* for set *)
exists (x2 :: tmp1 ++ tmp2 ++ tmp3); split.
repeat rewrite app_ass. simpl.
intros. eapply tr_assignop_val with (dst := For_set sd); eauto with gensym.
- apply contained_cons. eauto with gensym.
+ apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* postincr *)
monadInv H0. exploit H; eauto. intros [tmp1 [A B]].
destruct dst; monadInv EQ0; simpl add_dest in *.
(* for value *)
- exists (x0 :: tmp1); split.
+ exists (x0 :: tmp1); split.
econstructor; eauto with gensym.
- apply contained_cons; eauto with gensym.
+ apply contained_cons; eauto with gensym.
(* for effects *)
exploit transl_valof_meets_spec; eauto. intros [tmp2 [Csyntax D]].
- exists (tmp1 ++ tmp2); split.
+ exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
- eauto with gensym.
+ eauto with gensym.
(* for set *)
repeat rewrite app_ass; simpl.
- exists (x0 :: tmp1); split.
+ exists (x0 :: tmp1); split.
econstructor; eauto with gensym.
- apply contained_cons; eauto with gensym.
+ apply contained_cons; eauto with gensym.
(* comma *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto with gensym. intros [tmp2 [Csyntax D]].
- exists (tmp1 ++ tmp2); split.
+ exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
- destruct dst; simpl; eauto with gensym.
+ destruct dst; simpl; eauto with gensym.
apply list_disjoint_cons_r; eauto with gensym.
- simpl. eapply incl_tran. 2: apply add_dest_incl. auto with gensym.
+ simpl. eapply incl_tran. 2: apply add_dest_incl. auto with gensym.
destruct dst; simpl; auto with gensym.
apply contained_app; eauto with gensym.
(* call *)
@@ -989,44 +989,44 @@ Opaque makeif.
exploit H0; eauto. intros [tmp2 [Csyntax D]].
destruct dst; monadInv EQ2; simpl add_dest in *.
(* for value *)
- exists (x1 :: tmp1 ++ tmp2); split.
+ exists (x1 :: tmp1 ++ tmp2); split.
econstructor; eauto with gensym. congruence.
- apply contained_cons. eauto with gensym.
+ apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* for effects *)
- exists (tmp1 ++ tmp2); split.
+ exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
(* for set *)
- exists (x1 :: tmp1 ++ tmp2); split.
+ exists (x1 :: tmp1 ++ tmp2); split.
repeat rewrite app_ass. econstructor; eauto with gensym. congruence.
- apply contained_cons. eauto with gensym.
+ apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* builtin *)
monadInv H0. exploit H; eauto. intros [tmp1 [A B]].
destruct dst; monadInv EQ0; simpl add_dest in *.
(* for value *)
- exists (x0 :: tmp1); split.
+ exists (x0 :: tmp1); split.
econstructor; eauto with gensym. congruence.
- apply contained_cons; eauto with gensym.
+ apply contained_cons; eauto with gensym.
(* for effects *)
- exists tmp1; split.
+ exists tmp1; split.
econstructor; eauto with gensym.
auto.
(* for set *)
- exists (x0 :: tmp1); split.
+ exists (x0 :: tmp1); split.
repeat rewrite app_ass. econstructor; eauto with gensym. congruence.
- apply contained_cons; eauto with gensym.
+ apply contained_cons; eauto with gensym.
(* loc *)
monadInv H.
(* paren *)
- monadInv H0.
+ monadInv H0.
(* nil *)
monadInv H; exists (@nil ident); split; auto with gensym. constructor.
(* cons *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [Csyntax D]].
- exists (tmp1 ++ tmp2); split.
+ exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
eauto with gensym.
Qed.
@@ -1038,7 +1038,7 @@ Lemma transl_expr_meets_spec:
exists tmps, forall ge e le m, tr_top ge e le m dst r sl a tmps.
Proof.
intros. exploit (proj1 transl_meets_spec); eauto. intros [tmps [A B]].
- exists (add_dest dst tmps); intros. apply tr_top_base. auto.
+ exists (add_dest dst tmps); intros. apply tr_top_base. auto.
Qed.
Lemma transl_expression_meets_spec:
@@ -1046,8 +1046,8 @@ Lemma transl_expression_meets_spec:
transl_expression r g = Res (s, a) g' I ->
tr_expression r s a.
Proof.
- intros. monadInv H. exploit transl_expr_meets_spec; eauto.
- intros [tmps A]. econstructor; eauto.
+ intros. monadInv H. exploit transl_expr_meets_spec; eauto.
+ intros [tmps A]. econstructor; eauto.
Qed.
Lemma transl_expr_stmt_meets_spec:
@@ -1055,8 +1055,8 @@ Lemma transl_expr_stmt_meets_spec:
transl_expr_stmt r g = Res s g' I ->
tr_expr_stmt r s.
Proof.
- intros. monadInv H. exploit transl_expr_meets_spec; eauto.
- intros [tmps A]. econstructor; eauto.
+ intros. monadInv H. exploit transl_expr_meets_spec; eauto.
+ intros [tmps A]. econstructor; eauto.
Qed.
Lemma transl_if_meets_spec:
@@ -1064,8 +1064,8 @@ Lemma transl_if_meets_spec:
transl_if r s1 s2 g = Res s g' I ->
tr_if r s1 s2 s.
Proof.
- intros. monadInv H. exploit transl_expr_meets_spec; eauto.
- intros [tmps A]. econstructor; eauto.
+ intros. monadInv H. exploit transl_expr_meets_spec; eauto.
+ intros [tmps A]. econstructor; eauto.
Qed.
Lemma transl_stmt_meets_spec:
@@ -1076,11 +1076,11 @@ Proof.
generalize transl_expression_meets_spec transl_expr_stmt_meets_spec transl_if_meets_spec; intros T1 T2 T3.
Opaque transl_expression transl_expr_stmt.
clear transl_stmt_meets_spec.
- induction s; simpl; intros until I; intros TR;
+ induction s; simpl; intros until I; intros TR;
try (monadInv TR); try (constructor; eauto).
destruct (is_Sskip s1); monadInv EQ4.
apply tr_for_1; eauto.
- apply tr_for_2; eauto.
+ apply tr_for_2; eauto.
destruct o; monadInv TR; constructor; eauto.
clear transl_lblstmt_meets_spec.
@@ -1112,7 +1112,7 @@ Lemma transl_function_spec:
tr_function f tf.
Proof.
unfold transl_function; intros. monadInv H.
- constructor; auto. simpl. eapply transl_stmt_meets_spec; eauto.
+ constructor; auto. simpl. eapply transl_stmt_meets_spec; eauto.
Qed.
Lemma transl_fundef_spec:
@@ -1122,7 +1122,7 @@ Lemma transl_fundef_spec:
Proof.
unfold transl_fundef; intros.
destruct fd; monadInv H.
-+ constructor. eapply transl_function_spec; eauto.
++ constructor. eapply transl_function_spec; eauto.
+ constructor.
Qed.
@@ -1146,9 +1146,9 @@ Theorem transl_program_spec:
transl_program p = OK tp ->
match_program tr_fundef (fun v1 v2 => v1 = v2) nil (Csyntax.prog_main p) p tp.
Proof.
- unfold transl_program; intros.
+ unfold transl_program; intros.
destruct (transl_globdefs (Csyntax.prog_defs p) (initial_generator tt)) eqn:E; simpl in H; inv H.
- split; auto. exists l; split. eapply transl_globdefs_spec; eauto.
+ split; auto. exists l; split. eapply transl_globdefs_spec; eauto.
rewrite <- app_nil_end; auto.
Qed.
diff --git a/cfrontend/SimplLocals.v b/cfrontend/SimplLocals.v
index 7fc69324..c4b1054d 100644
--- a/cfrontend/SimplLocals.v
+++ b/cfrontend/SimplLocals.v
@@ -276,7 +276,7 @@ Definition transf_function (f: function) : res function :=
fn_params := f.(fn_params);
fn_vars := vars';
fn_temps := temps';
- fn_body := add_debug_params f.(fn_params)
+ fn_body := add_debug_params f.(fn_params)
(store_params cenv f.(fn_params)
(add_debug_vars vars' body')) |}.
diff --git a/cfrontend/SimplLocalsproof.v b/cfrontend/SimplLocalsproof.v
index 73092ab9..a47036bf 100644
--- a/cfrontend/SimplLocalsproof.v
+++ b/cfrontend/SimplLocalsproof.v
@@ -45,13 +45,13 @@ Let tge := globalenv tprog.
Lemma comp_env_preserved:
genv_cenv tge = genv_cenv ge.
Proof.
- monadInv TRANSF. unfold tge; rewrite <- H0; auto.
+ monadInv TRANSF. unfold tge; rewrite <- H0; auto.
Qed.
Lemma transf_programs:
AST.transform_partial_program transf_fundef (program_of_program prog) = OK (program_of_program tprog).
Proof.
- monadInv TRANSF. rewrite EQ. destruct x; reflexivity.
+ monadInv TRANSF. rewrite EQ. destruct x; reflexivity.
Qed.
Lemma symbols_preserved:
@@ -84,7 +84,7 @@ Lemma function_ptr_translated:
forall (b: block) (f: fundef),
Genv.find_funct_ptr ge b = Some f ->
exists tf, Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = OK tf.
-Proof.
+Proof.
exact (Genv.find_funct_ptr_transf_partial _ _ transf_programs).
Qed.
@@ -93,7 +93,7 @@ Lemma type_of_fundef_preserved:
transf_fundef fd = OK tfd -> type_of_fundef tfd = type_of_fundef fd.
Proof.
intros. destruct fd; monadInv H; auto.
- monadInv EQ. simpl; unfold type_of_function; simpl. auto.
+ monadInv EQ. simpl; unfold type_of_function; simpl. auto.
Qed.
(** Matching between environments before and after *)
@@ -162,18 +162,18 @@ Lemma match_envs_invariant:
(forall b b' delta, f' b = Some(b', delta) -> Ple tlo b' /\ Plt b' thi -> f' b = f b) ->
match_envs f' cenv e le m' lo hi te tle tlo thi.
Proof.
- intros until m'; intros ME LD INCR INV1 INV2.
- destruct ME; constructor; eauto.
+ intros until m'; intros ME LD INCR INV1 INV2.
+ destruct ME; constructor; eauto.
(* vars *)
intros. generalize (me_vars0 id); intros MV; inv MV.
eapply match_var_lifted; eauto.
- rewrite <- MAPPED; eauto.
+ rewrite <- MAPPED; eauto.
eapply match_var_not_lifted; eauto.
eapply match_var_not_local; eauto.
(* temps *)
- intros. exploit me_temps0; eauto. intros [[v' [A B]] C]. split; auto. exists v'; eauto.
+ intros. exploit me_temps0; eauto. intros [[v' [A B]] C]. split; auto. exists v'; eauto.
(* mapped *)
- intros. exploit me_mapped0; eauto. intros [b [A B]]. exists b; split; auto.
+ intros. exploit me_mapped0; eauto. intros [b [A B]]. exists b; split; auto.
(* flat *)
intros. eapply me_flat0; eauto. rewrite <- H0. symmetry. eapply INV2; eauto.
Qed.
@@ -189,14 +189,14 @@ Lemma match_envs_extcall:
Ple hi (Mem.nextblock m) -> Ple thi (Mem.nextblock tm) ->
match_envs f' cenv e le m' lo hi te tle tlo thi.
Proof.
- intros. eapply match_envs_invariant; eauto.
- intros. eapply Mem.load_unchanged_on; eauto.
- red in H2. intros. destruct (f b) as [[b' delta]|] eqn:?.
+ intros. eapply match_envs_invariant; eauto.
+ intros. eapply Mem.load_unchanged_on; eauto.
+ red in H2. intros. destruct (f b) as [[b' delta]|] eqn:?.
eapply H1; eauto.
- destruct (f' b) as [[b' delta]|] eqn:?; auto.
+ destruct (f' b) as [[b' delta]|] eqn:?; auto.
exploit H2; eauto. unfold Mem.valid_block. intros [A B].
xomegaContradiction.
- intros. destruct (f b) as [[b'' delta']|] eqn:?. eauto.
+ intros. destruct (f b) as [[b'' delta']|] eqn:?. eauto.
exploit H2; eauto. unfold Mem.valid_block. intros [A B].
xomegaContradiction.
Qed.
@@ -229,7 +229,7 @@ Inductive val_casted: val -> type -> Prop :=
Remark cast_int_int_idem:
forall sz sg i, cast_int_int sz sg (cast_int_int sz sg i) = cast_int_int sz sg i.
Proof.
- intros. destruct sz; simpl; auto.
+ intros. destruct sz; simpl; auto.
destruct sg; [apply Int.sign_ext_idem|apply Int.zero_ext_idem]; compute; intuition congruence.
destruct sg; [apply Int.sign_ext_idem|apply Int.zero_ext_idem]; compute; intuition congruence.
destruct (Int.eq i Int.zero); auto.
@@ -253,7 +253,7 @@ Proof.
destruct (cast_float_int s f); inv H1. constructor. apply (cast_int_int_idem I16 s).
constructor. auto.
constructor.
- constructor. auto.
+ constructor. auto.
destruct (cast_single_int s f); inv H1. constructor. auto.
destruct (cast_float_int s f); inv H1. constructor; auto.
constructor; auto.
@@ -303,7 +303,7 @@ Proof.
destruct si; inversion H0; clear H0; subst chunk; simpl in *; congruence.
destruct si; inversion H0; clear H0; subst chunk; simpl in *; congruence.
clear H1. inv H0. auto.
- inversion H0; clear H0; subst chunk. simpl in *.
+ inversion H0; clear H0; subst chunk. simpl in *.
destruct (Int.eq n Int.zero); subst n; reflexivity.
inv H0; auto.
inv H0; auto.
@@ -338,7 +338,7 @@ Lemma forall2_val_casted_inject:
forall f vl vl', Val.inject_list f vl vl' ->
forall tyl, list_forall2 val_casted vl tyl -> list_forall2 val_casted vl' tyl.
Proof.
- induction 1; intros tyl F; inv F; constructor; eauto. eapply val_casted_inject; eauto.
+ induction 1; intros tyl F; inv F; constructor; eauto. eapply val_casted_inject; eauto.
Qed.
Inductive val_casted_list: list val -> typelist -> Prop :=
@@ -353,9 +353,9 @@ Lemma val_casted_list_params:
val_casted_list vl (type_of_params params) ->
list_forall2 val_casted vl (map snd params).
Proof.
- induction params; simpl; intros.
+ induction params; simpl; intros.
inv H. constructor.
- destruct a as [id ty]. inv H. constructor; auto.
+ destruct a as [id ty]. inv H. constructor; auto.
Qed.
(** Correctness of [make_cast] *)
@@ -369,7 +369,7 @@ Proof.
intros.
assert (DFL: eval_expr tge e le m (Ecast a tto) v2).
econstructor; eauto.
- unfold sem_cast, make_cast in *.
+ unfold sem_cast, make_cast in *.
destruct (classify_cast (typeof a) tto); auto.
destruct v1; inv H0; auto.
destruct sz2; auto. destruct v1; inv H0; auto.
@@ -390,8 +390,8 @@ Lemma cast_typeconv:
val_casted v ty ->
sem_cast v ty (typeconv ty) = Some v.
Proof.
- induction 1; simpl; auto.
-- destruct sz; auto.
+ induction 1; simpl; auto.
+- destruct sz; auto.
- unfold sem_cast. simpl. rewrite dec_eq_true; auto.
- unfold sem_cast. simpl. rewrite dec_eq_true; auto.
Qed.
@@ -405,7 +405,7 @@ Lemma step_Sdebug_temp:
Proof.
intros. unfold Sdebug_temp. eapply step_builtin with (optid := None).
econstructor. constructor. eauto. simpl. eapply cast_typeconv; eauto. constructor.
- simpl. constructor.
+ simpl. constructor.
Qed.
Lemma step_Sdebug_var:
@@ -415,9 +415,9 @@ Lemma step_Sdebug_var:
E0 (State f Sskip k e le m).
Proof.
intros. unfold Sdebug_var. eapply step_builtin with (optid := None).
- econstructor. constructor. constructor. eauto.
- simpl. reflexivity. constructor.
- simpl. constructor.
+ econstructor. constructor. constructor. eauto.
+ simpl. reflexivity. constructor.
+ simpl. constructor.
Qed.
Lemma step_Sset_debug:
@@ -427,16 +427,16 @@ Lemma step_Sset_debug:
plus step2 tge (State f (Sset_debug id ty a) k e le m)
E0 (State f Sskip k e (PTree.set id v' le) m).
Proof.
- intros; unfold Sset_debug.
+ intros; unfold Sset_debug.
assert (forall k, step2 tge (State f (Sset id (make_cast a ty)) k e le m)
E0 (State f Sskip k e (PTree.set id v' le) m)).
{ intros. apply step_set. eapply make_cast_correct; eauto. }
destruct (Compopts.debug tt).
-- eapply plus_left. constructor.
+- eapply plus_left. constructor.
eapply star_left. apply H1.
eapply star_left. constructor.
apply star_one. apply step_Sdebug_temp with (v := v').
- apply PTree.gss. eapply cast_val_is_casted; eauto.
+ apply PTree.gss. eapply cast_val_is_casted; eauto.
reflexivity. reflexivity. reflexivity.
- apply plus_one. apply H1.
Qed.
@@ -448,12 +448,12 @@ Lemma step_add_debug_vars:
E0 (State f s k e le m).
Proof.
unfold add_debug_vars. destruct (Compopts.debug tt).
-- induction vars; simpl; intros.
+- induction vars; simpl; intros.
+ apply star_refl.
+ destruct a as [id ty].
- exploit H; eauto. intros (b & TE).
- simpl. eapply star_left. constructor.
- eapply star_left. eapply step_Sdebug_var; eauto.
+ exploit H; eauto. intros (b & TE).
+ simpl. eapply star_left. constructor.
+ eapply star_left. eapply step_Sdebug_var; eauto.
eapply star_left. constructor.
apply IHvars; eauto.
reflexivity. reflexivity. reflexivity.
@@ -466,11 +466,11 @@ Remark bind_parameter_temps_inv:
~In id (var_names params) ->
le'!id = le!id.
Proof.
- induction params; simpl; intros.
+ induction params; simpl; intros.
destruct args; inv H. auto.
- destruct a as [id1 ty1]. destruct args; try discriminate.
- transitivity ((PTree.set id1 v le)!id).
- eapply IHparams; eauto. apply PTree.gso. intuition.
+ destruct a as [id1 ty1]. destruct args; try discriminate.
+ transitivity ((PTree.set id1 v le)!id).
+ eapply IHparams; eauto. apply PTree.gso. intuition.
Qed.
Lemma step_add_debug_params:
@@ -485,8 +485,8 @@ Proof.
- induction params as [ | [id ty] params ]; simpl; intros until le1; intros NR CAST BIND; inv CAST; inv NR.
+ apply star_refl.
+ assert (le!id = Some a1). { erewrite bind_parameter_temps_inv by eauto. apply PTree.gss. }
- eapply star_left. constructor.
- eapply star_left. eapply step_Sdebug_temp; eauto.
+ eapply star_left. constructor.
+ eapply star_left. eapply step_Sdebug_temp; eauto.
eapply star_left. constructor.
eapply IHparams; eauto.
reflexivity. reflexivity. reflexivity.
@@ -511,18 +511,18 @@ Proof.
constructor; eauto; intros.
(* vars *)
destruct (peq id0 id). subst id0.
- eapply match_var_lifted with (v := v); eauto.
+ eapply match_var_lifted with (v := v); eauto.
exploit Mem.load_store_same; eauto. erewrite val_casted_load_result; eauto.
- apply PTree.gss.
+ apply PTree.gss.
generalize (me_vars0 id0); intros MV; inv MV.
eapply match_var_lifted; eauto.
- rewrite <- LOAD0. eapply Mem.load_store_other; eauto.
+ rewrite <- LOAD0. eapply Mem.load_store_other; eauto.
rewrite PTree.gso; auto.
- eapply match_var_not_lifted; eauto.
+ eapply match_var_not_lifted; eauto.
eapply match_var_not_local; eauto.
(* temps *)
exploit me_temps0; eauto. intros [[tv1 [A B]] C]. split; auto.
- rewrite PTree.gsspec. destruct (peq id0 id).
+ rewrite PTree.gsspec. destruct (peq id0 id).
subst id0. exists tv; split; auto. rewrite C; auto.
exists tv1; auto.
Qed.
@@ -536,18 +536,18 @@ Lemma match_envs_set_temp:
check_temp cenv id = OK x ->
match_envs f cenv e (PTree.set id v le) m lo hi te (PTree.set id tv tle) tlo thi.
Proof.
- intros. unfold check_temp in H1.
+ intros. unfold check_temp in H1.
destruct (VSet.mem id cenv) eqn:?; monadInv H1.
destruct H. constructor; eauto; intros.
(* vars *)
generalize (me_vars0 id0); intros MV; inv MV.
eapply match_var_lifted; eauto. rewrite PTree.gso. eauto. congruence.
- eapply match_var_not_lifted; eauto.
- eapply match_var_not_local; eauto.
+ eapply match_var_not_lifted; eauto.
+ eapply match_var_not_local; eauto.
(* temps *)
- rewrite PTree.gsspec in *. destruct (peq id0 id).
+ rewrite PTree.gsspec in *. destruct (peq id0 id).
inv H. split. exists tv; auto. intros; congruence.
- eapply me_temps0; eauto.
+ eapply me_temps0; eauto.
Qed.
Lemma match_envs_set_opttemp:
@@ -591,7 +591,7 @@ Proof.
intros. destruct H. constructor; auto; intros.
(* vars *)
generalize (me_vars0 id0); intros MV; inv MV.
- eapply match_var_lifted; eauto. rewrite PTree.gso; auto. congruence.
+ eapply match_var_lifted; eauto. rewrite PTree.gso; auto. congruence.
eapply match_var_not_lifted; eauto.
eapply match_var_not_local; eauto.
(* temps *)
@@ -608,7 +608,7 @@ Remark add_local_variable_charact:
VSet.In id1 (add_local_variable atk (id, ty) cenv) <->
VSet.In id1 cenv \/ exists chunk, access_mode ty = By_value chunk /\ id = id1 /\ VSet.mem id atk = false.
Proof.
- intros. unfold add_local_variable. split; intros.
+ intros. unfold add_local_variable. split; intros.
destruct (access_mode ty) eqn:?; auto.
destruct (VSet.mem id atk) eqn:?; auto.
rewrite VSF.add_iff in H. destruct H; auto. right; exists m; auto.
@@ -622,7 +622,7 @@ Lemma cenv_for_gen_domain:
Proof.
induction vars; simpl; intros.
rewrite VSF.empty_iff in H. auto.
- destruct a as [id1 ty1]. rewrite add_local_variable_charact in H.
+ destruct a as [id1 ty1]. rewrite add_local_variable_charact in H.
destruct H as [A | [chunk [A [B C]]]]; auto.
Qed.
@@ -635,13 +635,13 @@ Lemma cenv_for_gen_by_value:
Proof.
induction vars; simpl; intros.
contradiction.
- destruct a as [id1 ty1]. simpl in H0. inv H0.
+ destruct a as [id1 ty1]. simpl in H0. inv H0.
rewrite add_local_variable_charact in H1.
destruct H; destruct H1 as [A | [chunk [A [B C]]]].
- inv H. elim H4. eapply cenv_for_gen_domain; eauto.
+ inv H. elim H4. eapply cenv_for_gen_domain; eauto.
inv H. exists chunk; auto.
eauto.
- subst id1. elim H4. change id with (fst (id, ty)). apply in_map; auto.
+ subst id1. elim H4. change id with (fst (id, ty)). apply in_map; auto.
Qed.
Lemma cenv_for_gen_compat:
@@ -664,9 +664,9 @@ Definition compat_cenv (atk: VSet.t) (cenv: compilenv) : Prop :=
Lemma compat_cenv_for:
forall f, compat_cenv (addr_taken_stmt f.(fn_body)) (cenv_for f).
Proof.
- intros; red; intros.
+ intros; red; intros.
assert (VSet.mem id (addr_taken_stmt (fn_body f)) = false).
- eapply cenv_for_gen_compat. eexact H0.
+ eapply cenv_for_gen_compat. eexact H0.
rewrite VSF.mem_iff in H. congruence.
Qed.
@@ -674,20 +674,20 @@ Lemma compat_cenv_union_l:
forall atk1 atk2 cenv,
compat_cenv (VSet.union atk1 atk2) cenv -> compat_cenv atk1 cenv.
Proof.
- intros; red; intros. eapply H; eauto. apply VSet.union_2; auto.
+ intros; red; intros. eapply H; eauto. apply VSet.union_2; auto.
Qed.
Lemma compat_cenv_union_r:
forall atk1 atk2 cenv,
compat_cenv (VSet.union atk1 atk2) cenv -> compat_cenv atk2 cenv.
Proof.
- intros; red; intros. eapply H; eauto. apply VSet.union_3; auto.
+ intros; red; intros. eapply H; eauto. apply VSet.union_3; auto.
Qed.
Lemma compat_cenv_empty:
forall cenv, compat_cenv VSet.empty cenv.
Proof.
- intros; red; intros. eapply VSet.empty_1; eauto.
+ intros; red; intros. eapply VSet.empty_1; eauto.
Qed.
Hint Resolve compat_cenv_union_l compat_cenv_union_r compat_cenv_empty: compat.
@@ -700,7 +700,7 @@ Lemma alloc_variables_nextblock:
Proof.
induction 1.
apply Ple_refl.
- eapply Ple_trans; eauto. exploit Mem.nextblock_alloc; eauto. intros EQ; rewrite EQ. apply Ple_succ.
+ eapply Ple_trans; eauto. exploit Mem.nextblock_alloc; eauto. intros EQ; rewrite EQ. apply Ple_succ.
Qed.
Lemma alloc_variables_range:
@@ -711,12 +711,12 @@ Proof.
induction 1; intros.
auto.
exploit IHalloc_variables; eauto. rewrite PTree.gsspec. intros [A|A].
- destruct (peq id id0). inv A.
+ destruct (peq id id0). inv A.
right. exploit Mem.alloc_result; eauto. exploit Mem.nextblock_alloc; eauto.
- generalize (alloc_variables_nextblock _ _ _ _ _ _ H0). intros A B C.
- subst b. split. apply Ple_refl. eapply Plt_le_trans; eauto. rewrite B. apply Plt_succ.
+ generalize (alloc_variables_nextblock _ _ _ _ _ _ H0). intros A B C.
+ subst b. split. apply Ple_refl. eapply Plt_le_trans; eauto. rewrite B. apply Plt_succ.
auto.
- right. exploit Mem.nextblock_alloc; eauto. intros B. rewrite B in A. xomega.
+ right. exploit Mem.nextblock_alloc; eauto. intros B. rewrite B in A. xomega.
Qed.
Lemma alloc_variables_injective:
@@ -726,9 +726,9 @@ Lemma alloc_variables_injective:
(forall id b ty, e!id = Some(b, ty) -> Plt b (Mem.nextblock m)) ->
(e'!id1 = Some(b1, ty1) -> e'!id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2).
Proof.
- induction 1; intros.
+ induction 1; intros.
eauto.
- eapply IHalloc_variables; eauto.
+ eapply IHalloc_variables; eauto.
repeat rewrite PTree.gsspec; intros.
destruct (peq id1 id); destruct (peq id2 id).
congruence.
@@ -752,10 +752,10 @@ Lemma match_alloc_variables:
/\ inject_incr j j'
/\ (forall b, Mem.valid_block m b -> j' b = j b)
/\ (forall b b' delta, j' b = Some(b', delta) -> Mem.valid_block tm b' -> j' b = j b)
- /\ (forall b b' delta, j' b = Some(b', delta) -> ~Mem.valid_block tm b' ->
+ /\ (forall b b' delta, j' b = Some(b', delta) -> ~Mem.valid_block tm b' ->
exists id, exists ty, e'!id = Some(b, ty) /\ te'!id = Some(b', ty) /\ delta = 0)
/\ (forall id ty, In (id, ty) vars ->
- exists b,
+ exists b,
e'!id = Some(b, ty)
/\ if VSet.mem id cenv
then te'!id = te!id /\ j' b = None
@@ -766,83 +766,83 @@ Proof.
(* base case *)
exists j; exists te; exists tm. simpl.
split. constructor.
- split. auto. split. auto. split. auto. split. auto.
+ split. auto. split. auto. split. auto. split. auto.
split. intros. elim H2. eapply Mem.mi_mappedblocks; eauto.
- split. tauto. auto.
-
+ split. tauto. auto.
+
(* inductive case *)
simpl in H1. inv H1. simpl.
destruct (VSet.mem id cenv) eqn:?. simpl.
(* variable is lifted out of memory *)
- exploit Mem.alloc_left_unmapped_inject; eauto.
+ exploit Mem.alloc_left_unmapped_inject; eauto.
intros [j1 [A [B [C D]]]].
exploit IHalloc_variables; eauto. instantiate (1 := te).
intros [j' [te' [tm' [J [K [L [M [N [Q [O P]]]]]]]]]].
exists j'; exists te'; exists tm'.
split. auto.
split. auto.
- split. eapply inject_incr_trans; eauto.
- split. intros. transitivity (j1 b). apply M. eapply Mem.valid_block_alloc; eauto.
- apply D. apply Mem.valid_not_valid_diff with m; auto. eapply Mem.fresh_block_alloc; eauto.
+ split. eapply inject_incr_trans; eauto.
+ split. intros. transitivity (j1 b). apply M. eapply Mem.valid_block_alloc; eauto.
+ apply D. apply Mem.valid_not_valid_diff with m; auto. eapply Mem.fresh_block_alloc; eauto.
split. intros. transitivity (j1 b). eapply N; eauto.
- destruct (eq_block b b1); auto. subst.
- assert (j' b1 = j1 b1). apply M. eapply Mem.valid_new_block; eauto.
+ destruct (eq_block b b1); auto. subst.
+ assert (j' b1 = j1 b1). apply M. eapply Mem.valid_new_block; eauto.
congruence.
- split. exact Q.
+ split. exact Q.
split. intros. destruct (ident_eq id0 id).
(* same var *)
subst id0.
assert (ty0 = ty).
destruct H1. congruence. elim H5. unfold var_names. change id with (fst (id, ty0)). apply in_map; auto.
- subst ty0.
- exploit P; eauto. intros [X Y]. rewrite Heqb. rewrite X. rewrite Y.
+ subst ty0.
+ exploit P; eauto. intros [X Y]. rewrite Heqb. rewrite X. rewrite Y.
exists b1. split. apply PTree.gss.
split. auto.
- rewrite M. auto. eapply Mem.valid_new_block; eauto.
+ rewrite M. auto. eapply Mem.valid_new_block; eauto.
(* other vars *)
- eapply O; eauto. destruct H1. congruence. auto.
- intros. exploit (P id0). tauto. intros [X Y]. rewrite X; rewrite Y.
- split; auto. apply PTree.gso. intuition.
+ eapply O; eauto. destruct H1. congruence. auto.
+ intros. exploit (P id0). tauto. intros [X Y]. rewrite X; rewrite Y.
+ split; auto. apply PTree.gso. intuition.
(* variable is not lifted out of memory *)
exploit Mem.alloc_parallel_inject.
- eauto. eauto. apply Zle_refl. apply Zle_refl.
+ eauto. eauto. apply Zle_refl. apply Zle_refl.
intros [j1 [tm1 [tb1 [A [B [C [D E]]]]]]].
- exploit IHalloc_variables; eauto. instantiate (1 := PTree.set id (tb1, ty) te).
+ exploit IHalloc_variables; eauto. instantiate (1 := PTree.set id (tb1, ty) te).
intros [j' [te' [tm' [J [K [L [M [N [Q [O P]]]]]]]]]].
exists j'; exists te'; exists tm'.
- split. simpl. econstructor; eauto. rewrite comp_env_preserved; auto.
+ split. simpl. econstructor; eauto. rewrite comp_env_preserved; auto.
split. auto.
- split. eapply inject_incr_trans; eauto.
- split. intros. transitivity (j1 b). apply M. eapply Mem.valid_block_alloc; eauto.
- apply E. apply Mem.valid_not_valid_diff with m; auto. eapply Mem.fresh_block_alloc; eauto.
- split. intros. transitivity (j1 b). eapply N; eauto. eapply Mem.valid_block_alloc; eauto.
- destruct (eq_block b b1); auto. subst.
+ split. eapply inject_incr_trans; eauto.
+ split. intros. transitivity (j1 b). apply M. eapply Mem.valid_block_alloc; eauto.
+ apply E. apply Mem.valid_not_valid_diff with m; auto. eapply Mem.fresh_block_alloc; eauto.
+ split. intros. transitivity (j1 b). eapply N; eauto. eapply Mem.valid_block_alloc; eauto.
+ destruct (eq_block b b1); auto. subst.
assert (j' b1 = j1 b1). apply M. eapply Mem.valid_new_block; eauto.
rewrite H4 in H1. rewrite D in H1. inv H1. eelim Mem.fresh_block_alloc; eauto.
- split. intros. destruct (eq_block b' tb1).
+ split. intros. destruct (eq_block b' tb1).
subst b'. rewrite (N _ _ _ H1) in H1.
- destruct (eq_block b b1). subst b. rewrite D in H1; inv H1.
+ destruct (eq_block b b1). subst b. rewrite D in H1; inv H1.
exploit (P id); auto. intros [X Y]. exists id; exists ty.
rewrite X; rewrite Y. repeat rewrite PTree.gss. auto.
- rewrite E in H1; auto. elim H3. eapply Mem.mi_mappedblocks; eauto.
+ rewrite E in H1; auto. elim H3. eapply Mem.mi_mappedblocks; eauto.
eapply Mem.valid_new_block; eauto.
eapply Q; eauto. unfold Mem.valid_block in *.
- exploit Mem.nextblock_alloc. eexact A. exploit Mem.alloc_result. eexact A.
+ exploit Mem.nextblock_alloc. eexact A. exploit Mem.alloc_result. eexact A.
unfold block; xomega.
split. intros. destruct (ident_eq id0 id).
(* same var *)
subst id0.
assert (ty0 = ty).
destruct H1. congruence. elim H5. unfold var_names. change id with (fst (id, ty0)). apply in_map; auto.
- subst ty0.
- exploit P; eauto. intros [X Y]. rewrite Heqb. rewrite X. rewrite Y.
+ subst ty0.
+ exploit P; eauto. intros [X Y]. rewrite Heqb. rewrite X. rewrite Y.
exists b1. split. apply PTree.gss.
- exists tb1; split.
+ exists tb1; split.
apply PTree.gss.
- rewrite M. auto. eapply Mem.valid_new_block; eauto.
+ rewrite M. auto. eapply Mem.valid_new_block; eauto.
(* other vars *)
- exploit (O id0 ty0). destruct H1. congruence. auto.
+ exploit (O id0 ty0). destruct H1. congruence. auto.
rewrite PTree.gso; auto.
intros. exploit (P id0). tauto. intros [X Y]. rewrite X; rewrite Y.
split; apply PTree.gso; intuition.
@@ -875,7 +875,7 @@ Qed.
Definition env_initial_value (e: env) (m: mem) :=
forall id b ty chunk,
e!id = Some(b, ty) -> access_mode ty = By_value chunk -> Mem.load chunk m b 0 = Some Vundef.
-
+
Lemma alloc_variables_initial_value:
forall e m vars e' m',
alloc_variables ge e m vars e' m' ->
@@ -884,12 +884,12 @@ Lemma alloc_variables_initial_value:
Proof.
induction 1; intros.
auto.
- apply IHalloc_variables. red; intros. rewrite PTree.gsspec in H2.
- destruct (peq id0 id). inv H2.
- eapply Mem.load_alloc_same'; eauto.
- omega. rewrite Zplus_0_l. eapply sizeof_by_value; eauto.
- apply Zdivide_0.
- eapply Mem.load_alloc_other; eauto.
+ apply IHalloc_variables. red; intros. rewrite PTree.gsspec in H2.
+ destruct (peq id0 id). inv H2.
+ eapply Mem.load_alloc_same'; eauto.
+ omega. rewrite Zplus_0_l. eapply sizeof_by_value; eauto.
+ apply Zdivide_0.
+ eapply Mem.load_alloc_other; eauto.
Qed.
Lemma create_undef_temps_charact:
@@ -897,14 +897,14 @@ Lemma create_undef_temps_charact:
Proof.
induction vars; simpl; intros.
contradiction.
- destruct H. subst a. apply PTree.gss.
- destruct a as [id1 ty1]. rewrite PTree.gsspec. destruct (peq id id1); auto.
+ destruct H. subst a. apply PTree.gss.
+ destruct a as [id1 ty1]. rewrite PTree.gsspec. destruct (peq id id1); auto.
Qed.
Lemma create_undef_temps_inv:
forall vars id v, (create_undef_temps vars)!id = Some v -> v = Vundef /\ In id (var_names vars).
Proof.
- induction vars; simpl; intros.
+ induction vars; simpl; intros.
rewrite PTree.gempty in H; congruence.
destruct a as [id1 ty1]. rewrite PTree.gsspec in H. destruct (peq id id1).
inv H. auto.
@@ -924,16 +924,16 @@ Proof.
exploit list_in_map_inv. unfold var_names in H. apply H. eexact B.
intros [[id1 ty1] [P Q]]. simpl in P; subst id1.
right; symmetry; eapply create_undef_temps_charact; eauto.
- intros.
- exploit (H id l1 l2). tauto.
- exploit (H id l2 l1). tauto.
+ intros.
+ exploit (H id l1 l2). tauto.
+ exploit (H id l2 l1). tauto.
intuition congruence.
Qed.
Remark var_names_app:
forall vars1 vars2, var_names (vars1 ++ vars2) = var_names vars1 ++ var_names vars2.
Proof.
- intros. apply map_app.
+ intros. apply map_app.
Qed.
Remark filter_app:
@@ -949,8 +949,8 @@ Remark filter_charact:
forall (A: Type) (f: A -> bool) x l,
In x (List.filter f l) <-> In x l /\ f x = true.
Proof.
- induction l; simpl. tauto.
- destruct (f a) eqn:?.
+ induction l; simpl. tauto.
+ destruct (f a) eqn:?.
simpl. rewrite IHl. intuition congruence.
intuition congruence.
Qed.
@@ -959,8 +959,8 @@ Remark filter_norepet:
forall (A: Type) (f: A -> bool) l,
list_norepet l -> list_norepet (List.filter f l).
Proof.
- induction 1; simpl. constructor.
- destruct (f hd); auto. constructor; auto. rewrite filter_charact. tauto.
+ induction 1; simpl. constructor.
+ destruct (f hd); auto. constructor; auto. rewrite filter_charact. tauto.
Qed.
Remark filter_map:
@@ -979,12 +979,12 @@ Lemma create_undef_temps_lifted:
(create_undef_temps (add_lifted (cenv_for f) (fn_vars f) (fn_temps f))) ! id =
(create_undef_temps (add_lifted (cenv_for f) (fn_params f ++ fn_vars f) (fn_temps f))) ! id.
Proof.
- intros. apply create_undef_temps_exten.
- unfold add_lifted. rewrite filter_app.
- unfold var_names in *.
- repeat rewrite map_app. repeat rewrite in_app. intuition.
- exploit list_in_map_inv; eauto. intros [[id1 ty1] [P Q]]. simpl in P. subst id.
- rewrite filter_charact in Q. destruct Q.
+ intros. apply create_undef_temps_exten.
+ unfold add_lifted. rewrite filter_app.
+ unfold var_names in *.
+ repeat rewrite map_app. repeat rewrite in_app. intuition.
+ exploit list_in_map_inv; eauto. intros [[id1 ty1] [P Q]]. simpl in P. subst id.
+ rewrite filter_charact in Q. destruct Q.
elim H. change id1 with (fst (id1, ty1)). apply List.in_map. auto.
Qed.
@@ -998,11 +998,11 @@ Lemma vars_and_temps_properties:
Proof.
intros. rewrite list_norepet_app in H. destruct H as [A [B C]].
split. auto.
- split. unfold remove_lifted. unfold var_names. erewrite filter_map.
+ split. unfold remove_lifted. unfold var_names. erewrite filter_map.
instantiate (1 := fun a => negb (VSet.mem a cenv)). 2: auto.
apply filter_norepet. rewrite map_app. apply list_norepet_append; assumption.
- unfold add_lifted. rewrite var_names_app.
- unfold var_names at 2. erewrite filter_map.
+ unfold add_lifted. rewrite var_names_app.
+ unfold var_names at 2. erewrite filter_map.
instantiate (1 := fun a => VSet.mem a cenv). 2: auto.
change (map fst vars) with (var_names vars).
red; intros.
@@ -1029,9 +1029,9 @@ Theorem match_envs_alloc_variables:
/\ (forall b b' delta, j' b = Some(b', delta) -> Mem.valid_block tm b' -> j' b = j b)
/\ (forall id ty, In (id, ty) vars -> VSet.mem id cenv = false -> exists b, te!id = Some(b, ty)).
Proof.
- intros.
- exploit (match_alloc_variables cenv); eauto. instantiate (1 := empty_env).
- intros [j' [te [tm' [A [B [C [D [E [K [F G]]]]]]]]]].
+ intros.
+ exploit (match_alloc_variables cenv); eauto. instantiate (1 := empty_env).
+ intros [j' [te [tm' [A [B [C [D [E [K [F G]]]]]]]]]].
exists j'; exists te; exists tm'.
split. auto. split; auto.
constructor; intros.
@@ -1039,67 +1039,67 @@ Proof.
destruct (In_dec ident_eq id (var_names vars)).
unfold var_names in i. exploit list_in_map_inv; eauto.
intros [[id' ty] [EQ IN]]; simpl in EQ; subst id'.
- exploit F; eauto. intros [b [P R]].
+ exploit F; eauto. intros [b [P R]].
destruct (VSet.mem id cenv) eqn:?.
(* local var, lifted *)
- destruct R as [U V]. exploit H2; eauto. intros [chunk X].
+ destruct R as [U V]. exploit H2; eauto. intros [chunk X].
eapply match_var_lifted with (v := Vundef) (tv := Vundef); eauto.
rewrite U; apply PTree.gempty.
- eapply alloc_variables_initial_value; eauto.
+ eapply alloc_variables_initial_value; eauto.
red. unfold empty_env; intros. rewrite PTree.gempty in H4; congruence.
- apply create_undef_temps_charact with ty.
+ apply create_undef_temps_charact with ty.
unfold add_lifted. apply in_or_app. left.
rewrite filter_In. auto.
(* local var, not lifted *)
destruct R as [tb [U V]].
- eapply match_var_not_lifted; eauto.
+ eapply match_var_not_lifted; eauto.
(* non-local var *)
exploit G; eauto. unfold empty_env. rewrite PTree.gempty. intros [U V].
- eapply match_var_not_local; eauto.
+ eapply match_var_not_local; eauto.
destruct (VSet.mem id cenv) eqn:?; auto.
elim n; eauto.
(* temps *)
exploit create_undef_temps_inv; eauto. intros [P Q]. subst v.
- unfold var_names in Q. exploit list_in_map_inv; eauto.
- intros [[id1 ty] [EQ IN]]; simpl in EQ; subst id1.
- split; auto. exists Vundef; split; auto.
- apply create_undef_temps_charact with ty. unfold add_lifted.
+ unfold var_names in Q. exploit list_in_map_inv; eauto.
+ intros [[id1 ty] [EQ IN]]; simpl in EQ; subst id1.
+ split; auto. exists Vundef; split; auto.
+ apply create_undef_temps_charact with ty. unfold add_lifted.
apply in_or_app; auto.
(* injective *)
- eapply alloc_variables_injective. eexact H.
+ eapply alloc_variables_injective. eexact H.
rewrite PTree.gempty. congruence.
intros. rewrite PTree.gempty in H7. congruence.
- eauto. eauto. auto.
+ eauto. eauto. auto.
(* range *)
- exploit alloc_variables_range. eexact H. eauto.
+ exploit alloc_variables_range. eexact H. eauto.
rewrite PTree.gempty. intuition congruence.
(* trange *)
- exploit alloc_variables_range. eexact A. eauto.
+ exploit alloc_variables_range. eexact A. eauto.
rewrite PTree.gempty. intuition congruence.
(* mapped *)
destruct (In_dec ident_eq id (var_names vars)).
- unfold var_names in i. exploit list_in_map_inv; eauto.
+ unfold var_names in i. exploit list_in_map_inv; eauto.
intros [[id' ty'] [EQ IN]]; simpl in EQ; subst id'.
exploit F; eauto. intros [b [P Q]].
- destruct (VSet.mem id cenv).
+ destruct (VSet.mem id cenv).
rewrite PTree.gempty in Q. destruct Q; congruence.
destruct Q as [tb [U V]]. exists b; split; congruence.
exploit G; eauto. rewrite PTree.gempty. intuition congruence.
(* flat *)
- exploit alloc_variables_range. eexact A. eauto.
- rewrite PTree.gempty. intros [P|P]. congruence.
- exploit K; eauto. unfold Mem.valid_block. xomega.
- intros [id0 [ty0 [U [V W]]]]. split; auto.
+ exploit alloc_variables_range. eexact A. eauto.
+ rewrite PTree.gempty. intros [P|P]. congruence.
+ exploit K; eauto. unfold Mem.valid_block. xomega.
+ intros [id0 [ty0 [U [V W]]]]. split; auto.
destruct (ident_eq id id0). congruence.
assert (b' <> b').
eapply alloc_variables_injective with (e' := te) (id1 := id) (id2 := id0); eauto.
- rewrite PTree.gempty; congruence.
+ rewrite PTree.gempty; congruence.
intros until ty1; rewrite PTree.gempty; congruence.
congruence.
@@ -1127,13 +1127,13 @@ Proof.
intros. inv H.
- (* by value *)
exploit Mem.storev_mapped_inject; eauto. intros [tm' [A B]].
- exists tm'; split. eapply assign_loc_value; eauto.
+ exists tm'; split. eapply assign_loc_value; eauto.
split. auto.
intros. rewrite <- H5. eapply Mem.load_store_other; eauto.
left. inv H0. congruence.
- (* by copy *)
inv H0. inv H1.
- rename b' into bsrc. rename ofs'0 into osrc.
+ rename b' into bsrc. rename ofs'0 into osrc.
rename loc into bdst. rename ofs into odst.
rename loc' into bdst'. rename b2 into bsrc'.
rewrite <- comp_env_preserved in *.
@@ -1147,13 +1147,13 @@ Proof.
as [tm' SB].
simpl. red; intros; omegaContradiction.
exists tm'.
- split. eapply assign_loc_copy; eauto.
+ split. eapply assign_loc_copy; eauto.
intros; omegaContradiction.
intros; omegaContradiction.
rewrite e; right; omega.
- apply Mem.loadbytes_empty. omega.
- split. eapply Mem.storebytes_empty_inject; eauto.
- intros. rewrite <- H0. eapply Mem.load_storebytes_other; eauto.
+ apply Mem.loadbytes_empty. omega.
+ split. eapply Mem.storebytes_empty_inject; eauto.
+ intros. rewrite <- H0. eapply Mem.load_storebytes_other; eauto.
left. congruence.
+ (* general case size > 0 *)
exploit Mem.loadbytes_length; eauto. intros LEN.
@@ -1173,8 +1173,8 @@ Proof.
exploit Mem.address_inject. eauto. eexact PDST. eauto. intros EQ2.
exploit Mem.loadbytes_inject; eauto. intros [bytes2 [A B]].
exploit Mem.storebytes_mapped_inject; eauto. intros [tm' [C D]].
- exists tm'.
- split. eapply assign_loc_copy; try rewrite EQ1; try rewrite EQ2; eauto.
+ exists tm'.
+ split. eapply assign_loc_copy; try rewrite EQ1; try rewrite EQ2; eauto.
intros; eapply Mem.aligned_area_inject with (m := m); eauto.
apply alignof_blockcopy_1248.
apply sizeof_alignof_blockcopy_compat.
@@ -1185,7 +1185,7 @@ Proof.
apply Mem.range_perm_max with Cur; auto.
apply Mem.range_perm_max with Cur; auto.
split. auto.
- intros. rewrite <- H0. eapply Mem.load_storebytes_other; eauto.
+ intros. rewrite <- H0. eapply Mem.load_storebytes_other; eauto.
left. congruence.
Qed.
@@ -1193,7 +1193,7 @@ Lemma assign_loc_nextblock:
forall ge ty m b ofs v m',
assign_loc ge ty m b ofs v m' -> Mem.nextblock m' = Mem.nextblock m.
Proof.
- induction 1.
+ induction 1.
simpl in H0. eapply Mem.nextblock_store; eauto.
eapply Mem.nextblock_storebytes; eauto.
Qed.
@@ -1227,43 +1227,43 @@ Proof.
destruct (VSet.mem id cenv) eqn:?.
(* lifted to temp *)
eapply IHbind_parameters with (tle1 := PTree.set id v' tle1); eauto.
- eapply match_envs_assign_lifted; eauto.
+ eapply match_envs_assign_lifted; eauto.
inv MV; try congruence. rewrite ENV in H; inv H.
inv H0; try congruence.
unfold Mem.storev in H2. eapply Mem.store_unmapped_inject; eauto.
intros. repeat rewrite PTree.gsspec. destruct (peq id0 id). auto.
- apply TLE. intuition.
+ apply TLE. intuition.
(* still in memory *)
inv MV; try congruence. rewrite ENV in H; inv H.
- exploit assign_loc_inject; eauto.
+ exploit assign_loc_inject; eauto.
intros [tm1 [A [B C]]].
exploit IHbind_parameters. eauto. eauto. eauto.
instantiate (1 := PTree.set id v' tle1).
- apply match_envs_change_temp.
+ apply match_envs_change_temp.
eapply match_envs_invariant; eauto.
apply LE; auto. auto.
eauto.
- instantiate (1 := PTree.set id v' tle2).
+ instantiate (1 := PTree.set id v' tle2).
intros. repeat rewrite PTree.gsspec. destruct (peq id0 id). auto.
apply TLE. intuition.
intros. apply LE. auto.
- instantiate (1 := s).
+ instantiate (1 := s).
intros [tle [tm' [U [V [X [Y Z]]]]]].
exists tle; exists tm'; split.
eapply star_trans.
eapply star_left. econstructor.
- eapply star_left. econstructor.
- eapply eval_Evar_local. eauto.
+ eapply star_left. econstructor.
+ eapply eval_Evar_local. eauto.
eapply eval_Etempvar. erewrite bind_parameter_temps_inv; eauto.
- apply PTree.gss.
+ apply PTree.gss.
simpl. instantiate (1 := v'). apply cast_val_casted.
eapply val_casted_inject with (v := v1); eauto.
- simpl. eexact A.
+ simpl. eexact A.
apply star_one. constructor.
- reflexivity. reflexivity.
- eexact U.
+ reflexivity. reflexivity.
+ eexact U.
traceEq.
- rewrite (assign_loc_nextblock _ _ _ _ _ _ _ A) in Z. auto.
+ rewrite (assign_loc_nextblock _ _ _ _ _ _ _ A) in Z. auto.
Qed.
Lemma bind_parameters_nextblock:
@@ -1272,7 +1272,7 @@ Lemma bind_parameters_nextblock:
Proof.
induction 1.
auto.
- rewrite IHbind_parameters. eapply assign_loc_nextblock; eauto.
+ rewrite IHbind_parameters. eapply assign_loc_nextblock; eauto.
Qed.
Lemma bind_parameters_load:
@@ -1286,7 +1286,7 @@ Proof.
auto.
rewrite IHbind_parameters.
assert (b <> b0) by eauto.
- inv H1.
+ inv H1.
simpl in H5. eapply Mem.load_store_other; eauto.
eapply Mem.load_storebytes_other; eauto.
Qed.
@@ -1315,10 +1315,10 @@ Lemma free_list_perm':
Proof.
induction l; simpl; intros.
contradiction.
- destruct a as [[b1 lo1] hi1].
+ destruct a as [[b1 lo1] hi1].
destruct (Mem.free m b1 lo1 hi1) as [m1|] eqn:?; try discriminate.
- destruct H0. inv H0. eapply Mem.free_range_perm; eauto.
- red; intros. eapply Mem.perm_free_3; eauto. eapply IHl; eauto.
+ destruct H0. inv H0. eapply Mem.free_range_perm; eauto.
+ red; intros. eapply Mem.perm_free_3; eauto. eapply IHl; eauto.
Qed.
Lemma free_blocks_of_env_perm_2:
@@ -1327,7 +1327,7 @@ Lemma free_blocks_of_env_perm_2:
e!id = Some(b, ty) ->
Mem.range_perm m b 0 (sizeof ce ty) Cur Freeable.
Proof.
- intros. eapply free_list_perm'; eauto.
+ intros. eapply free_list_perm'; eauto.
unfold blocks_of_env. change (b, 0, sizeof ce ty) with (block_of_binding ce (id, (b, ty))).
apply in_map. apply PTree.elements_correct. auto.
Qed.
@@ -1350,18 +1350,18 @@ Proof.
induction l; simpl; intros.
- exists m; auto.
- destruct a as [[b lo] hi]. destruct H0.
- destruct (Mem.range_perm_free m b lo hi) as [m1 A]; auto.
+ destruct (Mem.range_perm_free m b lo hi) as [m1 A]; auto.
rewrite A. apply IHl; auto.
- intros. red; intros. eapply Mem.perm_free_1; eauto.
- exploit H1; eauto. intros [B|B]. auto. right; omega.
- eapply H; eauto.
+ intros. red; intros. eapply Mem.perm_free_1; eauto.
+ exploit H1; eauto. intros [B|B]. auto. right; omega.
+ eapply H; eauto.
Qed.
Lemma blocks_of_env_no_overlap:
forall (ge: genv) j cenv e le m lo hi te tle tlo thi tm,
match_envs j cenv e le m lo hi te tle tlo thi ->
Mem.inject j m tm ->
- (forall id b ty,
+ (forall id b ty,
e!id = Some(b, ty) -> Mem.range_perm m b 0 (sizeof ge ty) Cur Freeable) ->
forall l,
list_norepet (List.map fst l) ->
@@ -1373,20 +1373,20 @@ Proof.
- destruct a as [id [b ty]]. simpl in *. inv H. split.
+ apply IHl; auto.
+ intros. exploit list_in_map_inv; eauto. intros [[id' [b'' ty']] [A B]].
- simpl in A. inv A. rename b'' into b'.
- assert (TE: te!id = Some(b, ty)) by eauto.
- assert (TE': te!id' = Some(b', ty')) by eauto.
- exploit me_mapped. eauto. eexact TE. intros [b0 [INJ E]].
- exploit me_mapped. eauto. eexact TE'. intros [b0' [INJ' E']].
- destruct (zle (sizeof ge0 ty) 0); auto.
- destruct (zle (sizeof ge0 ty') 0); auto.
+ simpl in A. inv A. rename b'' into b'.
+ assert (TE: te!id = Some(b, ty)) by eauto.
+ assert (TE': te!id' = Some(b', ty')) by eauto.
+ exploit me_mapped. eauto. eexact TE. intros [b0 [INJ E]].
+ exploit me_mapped. eauto. eexact TE'. intros [b0' [INJ' E']].
+ destruct (zle (sizeof ge0 ty) 0); auto.
+ destruct (zle (sizeof ge0 ty') 0); auto.
assert (b0 <> b0').
{ eapply me_inj; eauto. red; intros; subst; elim H3.
change id' with (fst (id', (b', ty'))). apply List.in_map; auto. }
- assert (Mem.perm m b0 0 Max Nonempty).
+ assert (Mem.perm m b0 0 Max Nonempty).
{ apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable.
eapply PERMS; eauto. omega. auto with mem. }
- assert (Mem.perm m b0' 0 Max Nonempty).
+ assert (Mem.perm m b0' 0 Max Nonempty).
{ apply Mem.perm_cur_max. apply Mem.perm_implies with Freeable.
eapply PERMS; eauto. omega. auto with mem. }
exploit Mem.mi_no_overlap; eauto. intros [A|A]. auto. omegaContradiction.
@@ -1405,14 +1405,14 @@ Proof.
congruence.
destruct a as [[b lo] hi]. destruct (Mem.free m2 b lo hi) as [m21|] eqn:?; try discriminate.
eapply IHl with (m2 := m21); eauto.
- eapply Mem.free_right_inject; eauto.
+ eapply Mem.free_right_inject; eauto.
Qed.
Lemma blocks_of_env_translated:
forall e, blocks_of_env tge e = blocks_of_env ge e.
Proof.
- intros. unfold blocks_of_env, block_of_binding.
- rewrite comp_env_preserved; auto.
+ intros. unfold blocks_of_env, block_of_binding.
+ rewrite comp_env_preserved; auto.
Qed.
Theorem match_envs_free_blocks:
@@ -1424,7 +1424,7 @@ Theorem match_envs_free_blocks:
Mem.free_list tm (blocks_of_env tge te) = Some tm'
/\ Mem.inject j m' tm'.
Proof.
- intros.
+ intros.
Local Opaque ge tge.
assert (X: exists tm', Mem.free_list tm (blocks_of_env tge te) = Some tm').
{
@@ -1433,26 +1433,26 @@ Local Opaque ge tge.
intros. unfold blocks_of_env in H2.
exploit list_in_map_inv; eauto. intros [[id [b' ty]] [EQ IN]].
unfold block_of_binding in EQ; inv EQ.
- exploit me_mapped; eauto. eapply PTree.elements_complete; eauto.
- intros [b [A B]].
+ exploit me_mapped; eauto. eapply PTree.elements_complete; eauto.
+ intros [b [A B]].
change 0 with (0 + 0). replace (sizeof ge ty) with (sizeof ge ty + 0) by omega.
- eapply Mem.range_perm_inject; eauto.
+ eapply Mem.range_perm_inject; eauto.
eapply free_blocks_of_env_perm_2; eauto.
- (* no overlap *)
- unfold blocks_of_env; eapply blocks_of_env_no_overlap; eauto.
+ unfold blocks_of_env; eapply blocks_of_env_no_overlap; eauto.
intros. eapply free_blocks_of_env_perm_2; eauto.
apply PTree.elements_keys_norepet.
- intros. apply PTree.elements_complete; auto.
+ intros. apply PTree.elements_complete; auto.
}
destruct X as [tm' FREE].
exists tm'; split; auto.
- eapply free_list_right_inject; eauto.
- eapply Mem.free_list_left_inject; eauto.
- intros. unfold blocks_of_env in H3. exploit list_in_map_inv; eauto.
+ eapply free_list_right_inject; eauto.
+ eapply Mem.free_list_left_inject; eauto.
+ intros. unfold blocks_of_env in H3. exploit list_in_map_inv; eauto.
intros [[id [b' ty]] [EQ IN]]. unfold block_of_binding in EQ. inv EQ.
- exploit me_flat; eauto. apply PTree.elements_complete; eauto.
+ exploit me_flat; eauto. apply PTree.elements_complete; eauto.
intros [P Q]. subst delta. eapply free_blocks_of_env_perm_1 with (m := m); eauto.
- rewrite <- comp_env_preserved. omega.
+ rewrite <- comp_env_preserved. omega.
Qed.
(** Matching global environments *)
@@ -1472,7 +1472,7 @@ Lemma match_globalenvs_preserves_globals:
Proof.
intros. destruct H as [bound MG]. inv MG.
split; intros. eauto. split; intros. eauto. symmetry. eapply IMAGE; eauto.
-Qed.
+Qed.
(** Evaluation of expressions *)
@@ -1500,9 +1500,9 @@ Lemma deref_loc_inject:
Val.inject f (Vptr loc ofs) (Vptr loc' ofs') ->
exists tv, deref_loc ty tm loc' ofs' tv /\ Val.inject f v tv.
Proof.
- intros. inv H.
+ intros. inv H.
(* by value *)
- exploit Mem.loadv_inject; eauto. intros [tv [A B]].
+ exploit Mem.loadv_inject; eauto. intros [tv [A B]].
exists tv; split; auto. eapply deref_loc_value; eauto.
(* by reference *)
exists (Vptr loc' ofs'); split; auto. eapply deref_loc_reference; eauto.
@@ -1531,15 +1531,15 @@ Proof.
exists (Vsingle f0); split; auto. constructor.
exists (Vlong i); split; auto. constructor.
(* tempvar *)
- exploit me_temps; eauto. intros [[tv [A B]] C].
+ exploit me_temps; eauto. intros [[tv [A B]] C].
exists tv; split; auto. constructor; auto.
(* addrof *)
- exploit eval_simpl_lvalue; eauto.
+ exploit eval_simpl_lvalue; eauto.
destruct a; auto with compat.
- destruct a; auto. destruct (VSet.mem i cenv) eqn:?; auto.
+ destruct a; auto. destruct (VSet.mem i cenv) eqn:?; auto.
elim (H0 i). apply VSet.singleton_2. auto. apply VSet.mem_2. auto.
- intros [b' [ofs' [A B]]].
- exists (Vptr b' ofs'); split; auto. constructor; auto.
+ intros [b' [ofs' [A B]]].
+ exists (Vptr b' ofs'); split; auto. constructor; auto.
(* unop *)
exploit eval_simpl_expr; eauto. intros [tv1 [A B]].
exploit sem_unary_operation_inject; eauto. intros [tv [C D]].
@@ -1549,7 +1549,7 @@ Proof.
exploit eval_simpl_expr. eexact H0. eauto with compat. intros [tv2 [C D]].
exploit sem_binary_operation_inject; eauto. intros [tv [E F]].
exists tv; split; auto. econstructor; eauto.
- repeat rewrite typeof_simpl_expr; rewrite comp_env_preserved; auto.
+ repeat rewrite typeof_simpl_expr; rewrite comp_env_preserved; auto.
(* cast *)
exploit eval_simpl_expr; eauto. intros [tv1 [A B]].
exploit sem_cast_inject; eauto. intros [tv2 [C D]].
@@ -1561,15 +1561,15 @@ Proof.
(* rval *)
assert (EITHER: (exists id, exists ty, a = Evar id ty /\ VSet.mem id cenv = true)
\/ (match a with Evar id _ => VSet.mem id cenv = false | _ => True end)).
- destruct a; auto. destruct (VSet.mem i cenv) eqn:?; auto. left; exists i; exists t; auto.
+ destruct a; auto. destruct (VSet.mem i cenv) eqn:?; auto. left; exists i; exists t; auto.
destruct EITHER as [ [id [ty [EQ OPT]]] | NONOPT ].
(* a variable pulled out of memory *)
subst a. simpl. rewrite OPT.
- exploit me_vars; eauto. instantiate (1 := id). intros MV.
+ exploit me_vars; eauto. instantiate (1 := id). intros MV.
inv H; inv MV; try congruence.
rewrite ENV in H6; inv H6.
inv H0; try congruence.
- assert (chunk0 = chunk). simpl in H. congruence. subst chunk0.
+ assert (chunk0 = chunk). simpl in H. congruence. subst chunk0.
assert (v0 = v). unfold Mem.loadv in H2. rewrite Int.unsigned_zero in H2. congruence. subst v0.
exists tv; split; auto. constructor; auto.
simpl in H; congruence.
@@ -1577,12 +1577,12 @@ Proof.
(* any other l-value *)
exploit eval_simpl_lvalue; eauto. intros [loc' [ofs' [A B]]].
exploit deref_loc_inject; eauto. intros [tv [C D]].
- exists tv; split; auto. econstructor. eexact A. rewrite typeof_simpl_expr; auto.
+ exists tv; split; auto. econstructor. eexact A. rewrite typeof_simpl_expr; auto.
(* lvalues *)
destruct 1; simpl; intros.
(* local var *)
- rewrite H1.
+ rewrite H1.
exploit me_vars; eauto. instantiate (1 := id). intros MV. inv MV; try congruence.
rewrite ENV in H; inv H.
exists b'; exists Int.zero; split.
@@ -1592,25 +1592,25 @@ Proof.
rewrite H2.
exploit me_vars; eauto. instantiate (1 := id). intros MV. inv MV; try congruence.
exists l; exists Int.zero; split.
- apply eval_Evar_global. auto. rewrite <- H0. apply symbols_preserved.
+ apply eval_Evar_global. auto. rewrite <- H0. apply symbols_preserved.
destruct GLOB as [bound GLOB1]. inv GLOB1.
- econstructor; eauto.
+ econstructor; eauto.
(* deref *)
- exploit eval_simpl_expr; eauto. intros [tv [A B]].
- inversion B. subst.
- econstructor; econstructor; split; eauto. econstructor; eauto.
+ exploit eval_simpl_expr; eauto. intros [tv [A B]].
+ inversion B. subst.
+ econstructor; econstructor; split; eauto. econstructor; eauto.
(* field struct *)
rewrite <- comp_env_preserved in *.
- exploit eval_simpl_expr; eauto. intros [tv [A B]].
- inversion B. subst.
- econstructor; econstructor; split.
- eapply eval_Efield_struct; eauto. rewrite typeof_simpl_expr; eauto.
+ exploit eval_simpl_expr; eauto. intros [tv [A B]].
+ inversion B. subst.
+ econstructor; econstructor; split.
+ eapply eval_Efield_struct; eauto. rewrite typeof_simpl_expr; eauto.
econstructor; eauto. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut.
(* field union *)
rewrite <- comp_env_preserved in *.
- exploit eval_simpl_expr; eauto. intros [tv [A B]].
- inversion B. subst.
- econstructor; econstructor; split.
+ exploit eval_simpl_expr; eauto. intros [tv [A B]].
+ inversion B. subst.
+ econstructor; econstructor; split.
eapply eval_Efield_union; eauto. rewrite typeof_simpl_expr; eauto. auto.
Qed.
@@ -1628,7 +1628,7 @@ Proof.
exploit eval_simpl_expr; eauto with compat. intros [tv1 [A B]].
exploit sem_cast_inject; eauto. intros [tv2 [C D]].
exploit IHeval_exprlist; eauto with compat. intros [E [tvl [F G]]].
- split. constructor; auto. eapply cast_val_is_casted; eauto.
+ split. constructor; auto. eapply cast_val_is_casted; eauto.
exists (tv2 :: tvl); split. econstructor; eauto.
rewrite typeof_simpl_expr; auto.
econstructor; eauto.
@@ -1689,11 +1689,11 @@ Proof.
assert (f b1 = Some (b2, delta)). rewrite <- H; symmetry; eapply INJ2; eauto. xomega.
eapply IMAGE; eauto.
(* call *)
- eapply match_envs_invariant; eauto.
+ eapply match_envs_invariant; eauto.
intros. apply LOAD; auto. xomega.
intros. apply INJ1; auto; xomega.
intros. eapply INJ2; eauto; xomega.
- eapply IHmatch_cont; eauto.
+ eapply IHmatch_cont; eauto.
intros; apply LOAD; auto. inv H0; xomega.
intros; apply INJ1. inv H0; xomega.
intros; eapply INJ2; eauto. inv H0; xomega.
@@ -1711,7 +1711,7 @@ Proof.
intros. eapply match_cont_invariant; eauto.
intros. rewrite <- H4. inv H0.
(* scalar *)
- simpl in H6. eapply Mem.load_store_other; eauto. left. unfold block; xomega.
+ simpl in H6. eapply Mem.load_store_other; eauto. left. unfold block; xomega.
(* block copy *)
eapply Mem.load_storebytes_other; eauto. left. unfold block; xomega.
Qed.
@@ -1727,12 +1727,12 @@ Lemma match_cont_extcall:
Ple bound (Mem.nextblock m) -> Ple tbound (Mem.nextblock tm) ->
match_cont f' cenv k tk m' bound tbound.
Proof.
- intros. eapply match_cont_invariant; eauto.
- intros. eapply Mem.load_unchanged_on; eauto.
- red in H2. intros. destruct (f b) as [[b' delta] | ] eqn:?. auto.
- destruct (f' b) as [[b' delta] | ] eqn:?; auto.
+ intros. eapply match_cont_invariant; eauto.
+ intros. eapply Mem.load_unchanged_on; eauto.
+ red in H2. intros. destruct (f b) as [[b' delta] | ] eqn:?. auto.
+ destruct (f' b) as [[b' delta] | ] eqn:?; auto.
exploit H2; eauto. unfold Mem.valid_block. intros [A B]. xomegaContradiction.
- red in H2. intros. destruct (f b) as [[b'' delta''] | ] eqn:?. auto.
+ red in H2. intros. destruct (f b) as [[b'' delta''] | ] eqn:?. auto.
exploit H2; eauto. unfold Mem.valid_block. intros [A B]. xomegaContradiction.
Qed.
@@ -1765,7 +1765,7 @@ Lemma match_cont_is_call_cont:
is_call_cont k ->
is_call_cont tk.
Proof.
- intros. inv H; auto.
+ intros. inv H; auto.
Qed.
Lemma match_cont_call_cont:
@@ -1786,7 +1786,7 @@ Proof.
induction l; simpl; intros.
congruence.
destruct a. destruct p. destruct (Mem.free m b z0 z) as [m1|] eqn:?; try discriminate.
- transitivity (Mem.nextblock m1). eauto. eapply Mem.nextblock_free; eauto.
+ transitivity (Mem.nextblock m1). eauto. eapply Mem.nextblock_free; eauto.
Qed.
Remark free_list_load:
@@ -1798,7 +1798,7 @@ Proof.
induction l; simpl; intros.
inv H; auto.
destruct a. destruct p. destruct (Mem.free m b z0 z) as [m1|] eqn:?; try discriminate.
- transitivity (Mem.load chunk m1 b' 0). eauto.
+ transitivity (Mem.load chunk m1 b' 0). eauto.
eapply Mem.load_free. eauto. left. assert (Plt b' b) by eauto. unfold block; xomega.
Qed.
@@ -1813,10 +1813,10 @@ Lemma match_cont_free_env:
match_cont f cenv k tk m' (Mem.nextblock m') (Mem.nextblock tm').
Proof.
intros. apply match_cont_incr_bounds with lo tlo.
- eapply match_cont_invariant; eauto.
- intros. rewrite <- H7. eapply free_list_load; eauto.
- unfold blocks_of_env; intros. exploit list_in_map_inv; eauto.
- intros [[id [b1 ty]] [P Q]]. simpl in P. inv P.
+ eapply match_cont_invariant; eauto.
+ intros. rewrite <- H7. eapply free_list_load; eauto.
+ unfold blocks_of_env; intros. exploit list_in_map_inv; eauto.
+ intros [[id [b1 ty]] [P Q]]. simpl in P. inv P.
exploit me_range; eauto. eapply PTree.elements_complete; eauto. xomega.
rewrite (free_list_nextblock _ _ _ H3). inv H; xomega.
rewrite (free_list_nextblock _ _ _ H4). inv H; xomega.
@@ -1829,7 +1829,7 @@ Lemma match_cont_globalenv:
match_cont f cenv k tk m bound tbound ->
exists bound, match_globalenvs f bound.
Proof.
- induction 1; auto. exists hi; auto.
+ induction 1; auto. exists hi; auto.
Qed.
Hint Resolve match_cont_globalenv: compat.
@@ -1844,8 +1844,8 @@ Proof.
intros. exploit match_cont_globalenv; eauto. intros [bound1 MG]. destruct MG.
inv H1; simpl in H0; try discriminate. destruct (Int.eq_dec ofs1 Int.zero); try discriminate.
subst ofs1.
- assert (f b1 = Some(b1, 0)).
- apply DOMAIN. eapply FUNCTIONS; eauto.
+ assert (f b1 = Some(b1, 0)).
+ apply DOMAIN. eapply FUNCTIONS; eauto.
rewrite H1 in H2; inv H2.
rewrite Int.add_zero. simpl. rewrite dec_eq_true. apply function_ptr_translated; auto.
Qed.
@@ -1893,8 +1893,8 @@ Remark is_liftable_var_charact:
end.
Proof.
intros. destruct a; simpl; auto.
- destruct (VSet.mem i cenv) eqn:?.
- exists t; auto.
+ destruct (VSet.mem i cenv) eqn:?.
+ exists t; auto.
auto.
Qed.
@@ -1905,13 +1905,13 @@ Remark simpl_select_switch:
Proof.
intros cenv n.
assert (DFL:
- forall ls tls,
+ forall ls tls,
simpl_lblstmt cenv ls = OK tls ->
simpl_lblstmt cenv (select_switch_default ls) = OK (select_switch_default tls)).
{
- induction ls; simpl; intros; monadInv H.
+ induction ls; simpl; intros; monadInv H.
auto.
- simpl. destruct o. eauto. simpl; rewrite EQ, EQ1. auto.
+ simpl. destruct o. eauto. simpl; rewrite EQ, EQ1. auto.
}
assert (CASE:
forall ls tls,
@@ -1924,14 +1924,14 @@ Proof.
{
induction ls; simpl; intros; monadInv H; simpl.
auto.
- destruct o.
+ destruct o.
destruct (zeq z n).
econstructor; split; eauto. simpl; rewrite EQ, EQ1; auto.
- apply IHls. auto.
+ apply IHls. auto.
apply IHls. auto.
}
- intros; unfold select_switch.
- specialize (CASE _ _ H). destruct (select_switch_case n ls) as [ls'|].
+ intros; unfold select_switch.
+ specialize (CASE _ _ H). destruct (select_switch_case n ls) as [ls'|].
destruct CASE as [tls' [P Q]]. rewrite P, Q. auto.
rewrite CASE. apply DFL; auto.
Qed.
@@ -1956,7 +1956,7 @@ Proof.
compat_cenv (addr_taken_lblstmt ls) cenv ->
compat_cenv (addr_taken_lblstmt (select_switch_default ls)) cenv).
{
- induction ls; simpl; intros.
+ induction ls; simpl; intros.
eauto with compat.
destruct o; simpl; eauto with compat.
}
@@ -1967,11 +1967,11 @@ Proof.
{
induction ls; simpl; intros.
discriminate.
- destruct o. destruct (zeq z n). inv H0. auto. eauto with compat.
+ destruct o. destruct (zeq z n). inv H0. auto. eauto with compat.
eauto with compat.
}
- intros. specialize (CASE ls). unfold select_switch.
- destruct (select_switch_case n ls) as [ls'|]; eauto.
+ intros. specialize (CASE ls). unfold select_switch.
+ destruct (select_switch_case n ls) as [ls'|]; eauto.
Qed.
Remark addr_taken_seq_of_labeled_statement:
@@ -1997,7 +1997,7 @@ Lemma simpl_find_label:
| None =>
find_label lbl ts tk = None
| Some(s', k') =>
- exists ts', exists tk',
+ exists ts', exists tk',
find_label lbl ts tk = Some(ts', tk')
/\ compat_cenv (addr_taken_stmt s') cenv
/\ simpl_stmt cenv s' = OK ts'
@@ -2013,7 +2013,7 @@ with simpl_find_label_ls:
| None =>
find_label_ls lbl tls tk = None
| Some(s', k') =>
- exists ts', exists tk',
+ exists ts', exists tk',
find_label_ls lbl tls tk = Some(ts', tk')
/\ compat_cenv (addr_taken_stmt s') cenv
/\ simpl_stmt cenv s' = OK ts'
@@ -2025,8 +2025,8 @@ Proof.
(* skip *)
monadInv TS; auto.
(* var *)
- destruct (is_liftable_var cenv e); monadInv TS; auto.
- unfold Sset_debug. destruct (Compopts.debug tt); auto.
+ destruct (is_liftable_var cenv e); monadInv TS; auto.
+ unfold Sset_debug. destruct (Compopts.debug tt); auto.
(* set *)
monadInv TS; auto.
(* call *)
@@ -2035,23 +2035,23 @@ Proof.
monadInv TS; auto.
(* seq *)
monadInv TS.
- exploit (IHs1 (Kseq s2 k) x (Kseq x0 tk)); eauto with compat.
+ exploit (IHs1 (Kseq s2 k) x (Kseq x0 tk)); eauto with compat.
constructor; eauto with compat.
- destruct (find_label lbl s1 (Kseq s2 k)) as [[s' k']|].
+ destruct (find_label lbl s1 (Kseq s2 k)) as [[s' k']|].
intros [ts' [tk' [P [Q [R S]]]]]. exists ts'; exists tk'. simpl. rewrite P. auto.
intros E. simpl. rewrite E. eapply IHs2; eauto with compat.
(* ifthenelse *)
monadInv TS.
- exploit (IHs1 k x tk); eauto with compat.
- destruct (find_label lbl s1 k) as [[s' k']|].
- intros [ts' [tk' [P [Q [R S]]]]]. exists ts'; exists tk'. simpl. rewrite P. auto.
+ exploit (IHs1 k x tk); eauto with compat.
+ destruct (find_label lbl s1 k) as [[s' k']|].
+ intros [ts' [tk' [P [Q [R S]]]]]. exists ts'; exists tk'. simpl. rewrite P. auto.
intros E. simpl. rewrite E. eapply IHs2; eauto with compat.
(* loop *)
monadInv TS.
exploit (IHs1 (Kloop1 s1 s2 k) x (Kloop1 x x0 tk)); eauto with compat.
constructor; eauto with compat.
- destruct (find_label lbl s1 (Kloop1 s1 s2 k)) as [[s' k']|].
- intros [ts' [tk' [P [Q [R S]]]]]. exists ts'; exists tk'. simpl; rewrite P. auto.
+ destruct (find_label lbl s1 (Kloop1 s1 s2 k)) as [[s' k']|].
+ intros [ts' [tk' [P [Q [R S]]]]]. exists ts'; exists tk'. simpl; rewrite P. auto.
intros E. simpl; rewrite E. eapply IHs2; eauto with compat. econstructor; eauto with compat.
(* break *)
monadInv TS; auto.
@@ -2078,7 +2078,7 @@ Proof.
exploit (simpl_find_label s (Kseq (seq_of_labeled_statement ls) k)).
eauto. constructor. eapply simpl_seq_of_labeled_statement; eauto. eauto.
rewrite addr_taken_seq_of_labeled_statement. eauto with compat.
- eauto with compat.
+ eauto with compat.
destruct (find_label lbl s (Kseq (seq_of_labeled_statement ls) k)) as [[s' k']|].
intros [ts' [tk' [P [Q [R S]]]]]. exists ts'; exists tk'; split. simpl; rewrite P. auto. auto.
intros E. simpl; rewrite E. eapply IHls; eauto with compat.
@@ -2087,25 +2087,25 @@ Qed.
Lemma find_label_store_params:
forall s k params, find_label lbl (store_params cenv params s) k = find_label lbl s k.
Proof.
- induction params; simpl. auto.
- destruct a as [id ty]. destruct (VSet.mem id cenv); auto.
+ induction params; simpl. auto.
+ destruct a as [id ty]. destruct (VSet.mem id cenv); auto.
Qed.
Lemma find_label_add_debug_vars:
forall s k vars, find_label lbl (add_debug_vars vars s) k = find_label lbl s k.
Proof.
- unfold add_debug_vars. destruct (Compopts.debug tt); auto.
- induction vars; simpl; auto. destruct a as [id ty]; simpl. auto.
+ unfold add_debug_vars. destruct (Compopts.debug tt); auto.
+ induction vars; simpl; auto. destruct a as [id ty]; simpl. auto.
Qed.
Lemma find_label_add_debug_params:
forall s k vars, find_label lbl (add_debug_params vars s) k = find_label lbl s k.
Proof.
- unfold add_debug_params. destruct (Compopts.debug tt); auto.
- induction vars; simpl; auto. destruct a as [id ty]; simpl. auto.
+ unfold add_debug_params. destruct (Compopts.debug tt); auto.
+ induction vars; simpl; auto. destruct a as [id ty]; simpl. auto.
Qed.
-End FIND_LABEL.
+End FIND_LABEL.
Lemma step_simulation:
@@ -2120,40 +2120,40 @@ Proof.
intros [ty [P Q]]; subst a1; simpl in *.
exploit eval_simpl_expr; eauto with compat. intros [tv2 [A B]].
exploit sem_cast_inject; eauto. intros [tv [C D]].
- exploit me_vars; eauto. instantiate (1 := id). intros MV.
+ exploit me_vars; eauto. instantiate (1 := id). intros MV.
inv H.
(* local variable *)
econstructor; split.
- eapply step_Sset_debug. eauto. rewrite typeof_simpl_expr. eauto.
- econstructor; eauto with compat.
+ eapply step_Sset_debug. eauto. rewrite typeof_simpl_expr. eauto.
+ econstructor; eauto with compat.
eapply match_envs_assign_lifted; eauto. eapply cast_val_is_casted; eauto.
eapply match_cont_assign_loc; eauto. exploit me_range; eauto. xomega.
inv MV; try congruence. inv H2; try congruence. unfold Mem.storev in H3.
- eapply Mem.store_unmapped_inject; eauto. congruence.
+ eapply Mem.store_unmapped_inject; eauto. congruence.
erewrite assign_loc_nextblock; eauto.
(* global variable *)
inv MV; congruence.
(* not liftable *)
- intros P.
- exploit eval_simpl_lvalue; eauto with compat. intros [tb [tofs [E F]]].
+ intros P.
+ exploit eval_simpl_lvalue; eauto with compat. intros [tb [tofs [E F]]].
exploit eval_simpl_expr; eauto with compat. intros [tv2 [A B]].
exploit sem_cast_inject; eauto. intros [tv [C D]].
- exploit assign_loc_inject; eauto. intros [tm' [X [Y Z]]].
+ exploit assign_loc_inject; eauto. intros [tm' [X [Y Z]]].
econstructor; split.
- apply plus_one. econstructor. eexact E. eexact A. repeat rewrite typeof_simpl_expr. eexact C.
+ apply plus_one. econstructor. eexact E. eexact A. repeat rewrite typeof_simpl_expr. eexact C.
rewrite typeof_simpl_expr; auto. eexact X.
- econstructor; eauto with compat.
- eapply match_envs_invariant; eauto.
- eapply match_cont_invariant; eauto.
+ econstructor; eauto with compat.
+ eapply match_envs_invariant; eauto.
+ eapply match_cont_invariant; eauto.
erewrite assign_loc_nextblock; eauto.
erewrite assign_loc_nextblock; eauto.
(* set temporary *)
exploit eval_simpl_expr; eauto with compat. intros [tv [A B]].
econstructor; split.
- apply plus_one. econstructor. eauto.
- econstructor; eauto with compat.
- eapply match_envs_set_temp; eauto.
+ apply plus_one. econstructor. eauto.
+ econstructor; eauto with compat.
+ eapply match_envs_set_temp; eauto.
(* call *)
exploit eval_simpl_expr; eauto with compat. intros [tvf [A B]].
@@ -2162,9 +2162,9 @@ Proof.
econstructor; split.
apply plus_one. eapply step_call with (fd := tfd).
rewrite typeof_simpl_expr. eauto.
- eauto. eauto. eauto.
+ eauto. eauto. eauto.
erewrite type_of_fundef_preserved; eauto.
- econstructor; eauto.
+ econstructor; eauto.
intros. econstructor; eauto.
(* builtin *)
@@ -2172,13 +2172,13 @@ Proof.
exploit external_call_mem_inject; eauto. apply match_globalenvs_preserves_globals; eauto with compat.
intros [j' [tvres [tm' [P [Q [R [S [T [U V]]]]]]]]].
econstructor; split.
- apply plus_one. econstructor; eauto. eapply external_call_symbols_preserved; eauto.
+ apply plus_one. econstructor; eauto. eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
econstructor; eauto with compat.
- eapply match_envs_set_opttemp; eauto.
- eapply match_envs_extcall; eauto.
+ eapply match_envs_set_opttemp; eauto.
+ eapply match_envs_extcall; eauto.
eapply match_cont_extcall; eauto.
- inv MENV; xomega. inv MENV; xomega.
+ inv MENV; xomega. inv MENV; xomega.
eapply Ple_trans; eauto. eapply external_call_nextblock; eauto.
eapply Ple_trans; eauto. eapply external_call_nextblock; eauto.
@@ -2187,11 +2187,11 @@ Proof.
econstructor; eauto with compat. econstructor; eauto with compat.
(* skip sequence *)
- inv MCONT. econstructor; split. apply plus_one. econstructor. econstructor; eauto.
+ inv MCONT. econstructor; split. apply plus_one. econstructor. econstructor; eauto.
(* continue sequence *)
inv MCONT. econstructor; split. apply plus_one. econstructor. econstructor; eauto.
-
+
(* break sequence *)
inv MCONT. econstructor; split. apply plus_one. econstructor. econstructor; eauto.
@@ -2207,25 +2207,25 @@ Proof.
(* skip-or-continue loop *)
inv MCONT. econstructor; split.
- apply plus_one. econstructor. destruct H; subst x; simpl in *; intuition congruence.
+ apply plus_one. econstructor. destruct H; subst x; simpl in *; intuition congruence.
econstructor; eauto with compat. econstructor; eauto with compat.
(* break loop1 *)
- inv MCONT. econstructor; split. apply plus_one. eapply step_break_loop1.
+ inv MCONT. econstructor; split. apply plus_one. eapply step_break_loop1.
econstructor; eauto.
(* skip loop2 *)
- inv MCONT. econstructor; split. apply plus_one. eapply step_skip_loop2.
- econstructor; eauto with compat. simpl; rewrite H2; rewrite H4; auto.
+ inv MCONT. econstructor; split. apply plus_one. eapply step_skip_loop2.
+ econstructor; eauto with compat. simpl; rewrite H2; rewrite H4; auto.
(* break loop2 *)
- inv MCONT. econstructor; split. apply plus_one. eapply step_break_loop2.
+ inv MCONT. econstructor; split. apply plus_one. eapply step_break_loop2.
econstructor; eauto.
(* return none *)
exploit match_envs_free_blocks; eauto. intros [tm' [P Q]].
- econstructor; split. apply plus_one. econstructor; eauto.
- econstructor; eauto.
+ econstructor; split. apply plus_one. econstructor; eauto.
+ econstructor; eauto.
intros. eapply match_cont_call_cont. eapply match_cont_free_env; eauto.
(* return some *)
@@ -2234,7 +2234,7 @@ Proof.
exploit match_envs_free_blocks; eauto. intros [tm' [P Q]].
econstructor; split. apply plus_one. econstructor; eauto.
rewrite typeof_simpl_expr. monadInv TRF; simpl. eauto.
- econstructor; eauto.
+ econstructor; eauto.
intros. eapply match_cont_call_cont. eapply match_cont_free_env; eauto.
(* skip call *)
@@ -2242,29 +2242,29 @@ Proof.
econstructor; split. apply plus_one. econstructor; eauto.
eapply match_cont_is_call_cont; eauto.
monadInv TRF; auto.
- econstructor; eauto.
+ econstructor; eauto.
intros. apply match_cont_change_cenv with (cenv_for f); auto. eapply match_cont_free_env; eauto.
(* switch *)
exploit eval_simpl_expr; eauto with compat. intros [tv [A B]].
econstructor; split. apply plus_one. econstructor; eauto.
- rewrite typeof_simpl_expr. instantiate (1 := n).
+ rewrite typeof_simpl_expr. instantiate (1 := n).
unfold sem_switch_arg in *;
destruct (classify_switch (typeof a)); try discriminate;
inv B; inv H0; auto.
- econstructor; eauto.
+ econstructor; eauto.
erewrite simpl_seq_of_labeled_statement. reflexivity.
- eapply simpl_select_switch; eauto.
- econstructor; eauto. rewrite addr_taken_seq_of_labeled_statement.
+ eapply simpl_select_switch; eauto.
+ econstructor; eauto. rewrite addr_taken_seq_of_labeled_statement.
apply compat_cenv_select_switch. eauto with compat.
(* skip-break switch *)
- inv MCONT. econstructor; split.
- apply plus_one. eapply step_skip_break_switch. destruct H; subst x; simpl in *; intuition congruence.
+ inv MCONT. econstructor; split.
+ apply plus_one. eapply step_skip_break_switch. destruct H; subst x; simpl in *; intuition congruence.
econstructor; eauto with compat.
(* continue switch *)
- inv MCONT. econstructor; split.
+ inv MCONT. econstructor; split.
apply plus_one. eapply step_continue_switch.
econstructor; eauto with compat.
@@ -2272,18 +2272,18 @@ Proof.
econstructor; split. apply plus_one. econstructor. econstructor; eauto.
(* goto *)
- generalize TRF; intros TRF'. monadInv TRF'.
+ generalize TRF; intros TRF'. monadInv TRF'.
exploit (simpl_find_label j (cenv_for f) m lo tlo lbl (fn_body f) (call_cont k) x (call_cont tk)).
- eauto. eapply match_cont_call_cont. eauto.
+ eauto. eapply match_cont_call_cont. eauto.
apply compat_cenv_for.
- rewrite H. intros [ts' [tk' [A [B [C D]]]]].
+ rewrite H. intros [ts' [tk' [A [B [C D]]]]].
econstructor; split.
apply plus_one. econstructor; eauto. simpl.
rewrite find_label_add_debug_params. rewrite find_label_store_params. rewrite find_label_add_debug_vars. eexact A.
econstructor; eauto.
(* internal function *)
- monadInv TRFD. inv H.
+ monadInv TRFD. inv H.
generalize EQ; intro EQ'; monadInv EQ'.
assert (list_norepet (var_names (fn_params f ++ fn_vars f))).
unfold var_names. rewrite map_app. auto.
@@ -2295,10 +2295,10 @@ Proof.
assert (K: list_forall2 val_casted vargs (map snd (fn_params f))).
{ apply val_casted_list_params. unfold type_of_function in FUNTY. congruence. }
exploit store_params_correct.
- eauto.
+ eauto.
eapply list_norepet_append_left; eauto.
eexact K.
- apply val_inject_list_incr with j'; eauto.
+ apply val_inject_list_incr with j'; eauto.
eexact B. eexact C.
intros. apply (create_undef_temps_lifted id f). auto.
intros. destruct (create_undef_temps (fn_temps f))!id as [v|] eqn:?; auto.
@@ -2307,34 +2307,34 @@ Proof.
change (cenv_for_gen (addr_taken_stmt (fn_body f)) (fn_params f ++ fn_vars f))
with (cenv_for f) in *.
generalize (vars_and_temps_properties (cenv_for f) (fn_params f) (fn_vars f) (fn_temps f)).
- intros [X [Y Z]]. auto. auto.
- econstructor; split.
- eapply plus_left. econstructor.
+ intros [X [Y Z]]. auto. auto.
+ econstructor; split.
+ eapply plus_left. econstructor.
econstructor. exact Y. exact X. exact Z. simpl. eexact A. simpl. eexact Q.
simpl. eapply star_trans. eapply step_add_debug_params. auto. eapply forall2_val_casted_inject; eauto. eexact Q.
- eapply star_trans. eexact P. eapply step_add_debug_vars.
- unfold remove_lifted; intros. rewrite List.filter_In in H3. destruct H3.
- apply negb_true_iff in H4. eauto.
+ eapply star_trans. eexact P. eapply step_add_debug_vars.
+ unfold remove_lifted; intros. rewrite List.filter_In in H3. destruct H3.
+ apply negb_true_iff in H4. eauto.
reflexivity. reflexivity. traceEq.
- econstructor; eauto.
- eapply match_cont_invariant; eauto.
- intros. transitivity (Mem.load chunk m0 b 0).
- eapply bind_parameters_load; eauto. intros.
- exploit alloc_variables_range. eexact H1. eauto.
- unfold empty_env. rewrite PTree.gempty. intros [?|?]. congruence.
- red; intros; subst b'. xomega.
+ econstructor; eauto.
+ eapply match_cont_invariant; eauto.
+ intros. transitivity (Mem.load chunk m0 b 0).
+ eapply bind_parameters_load; eauto. intros.
+ exploit alloc_variables_range. eexact H1. eauto.
+ unfold empty_env. rewrite PTree.gempty. intros [?|?]. congruence.
+ red; intros; subst b'. xomega.
eapply alloc_variables_load; eauto.
- apply compat_cenv_for.
+ apply compat_cenv_for.
rewrite (bind_parameters_nextblock _ _ _ _ _ _ H2). xomega.
rewrite T; xomega.
(* external function *)
- monadInv TRFD. inv FUNTY.
- exploit external_call_mem_inject; eauto. apply match_globalenvs_preserves_globals.
- eapply match_cont_globalenv. eexact (MCONT VSet.empty).
+ monadInv TRFD. inv FUNTY.
+ exploit external_call_mem_inject; eauto. apply match_globalenvs_preserves_globals.
+ eapply match_cont_globalenv. eexact (MCONT VSet.empty).
intros [j' [tvres [tm' [P [Q [R [S [T [U V]]]]]]]]].
econstructor; split.
- apply plus_one. econstructor; eauto. eapply external_call_symbols_preserved; eauto.
+ apply plus_one. econstructor; eauto. eapply external_call_symbols_preserved; eauto.
exact symbols_preserved. exact public_preserved. exact varinfo_preserved.
econstructor; eauto.
intros. apply match_cont_incr_bounds with (Mem.nextblock m) (Mem.nextblock tm).
@@ -2343,9 +2343,9 @@ Proof.
eapply external_call_nextblock; eauto.
(* return *)
- specialize (MCONT (cenv_for f)). inv MCONT.
+ specialize (MCONT (cenv_for f)). inv MCONT.
econstructor; split.
- apply plus_one. econstructor.
+ apply plus_one. econstructor.
econstructor; eauto with compat.
eapply match_envs_set_opttemp; eauto.
Qed.
@@ -2355,20 +2355,20 @@ Lemma initial_states_simulation:
exists R, initial_state tprog R /\ match_states S R.
Proof.
intros. inv H.
- exploit function_ptr_translated; eauto. intros [tf [A B]].
+ exploit function_ptr_translated; eauto. intros [tf [A B]].
econstructor; split.
econstructor.
- eapply Genv.init_mem_transf_partial. eexact transf_programs. eauto.
- change (prog_main tprog) with (AST.prog_main tprog).
+ eapply Genv.init_mem_transf_partial. eexact transf_programs. eauto.
+ change (prog_main tprog) with (AST.prog_main tprog).
rewrite (transform_partial_program_main _ _ transf_programs).
instantiate (1 := b). rewrite <- H1. apply symbols_preserved.
eauto.
rewrite <- H3; apply type_of_fundef_preserved; auto.
- econstructor; eauto.
- intros. instantiate (1 := Mem.flat_inj (Mem.nextblock m0)).
- econstructor. instantiate (1 := Mem.nextblock m0).
- constructor; intros.
- unfold Mem.flat_inj. apply pred_dec_true; auto.
+ econstructor; eauto.
+ intros. instantiate (1 := Mem.flat_inj (Mem.nextblock m0)).
+ econstructor. instantiate (1 := Mem.nextblock m0).
+ constructor; intros.
+ unfold Mem.flat_inj. apply pred_dec_true; auto.
unfold Mem.flat_inj in H. destruct (plt b1 (Mem.nextblock m0)); inv H. auto.
eapply Genv.find_symbol_not_fresh; eauto.
eapply Genv.find_funct_ptr_not_fresh; eauto.
@@ -2383,7 +2383,7 @@ Lemma final_states_simulation:
match_states S R -> final_state S r -> final_state R r.
Proof.
intros. inv H0. inv H.
- specialize (MCONT VSet.empty). inv MCONT.
+ specialize (MCONT VSet.empty). inv MCONT.
inv RINJ. constructor.
Qed.