summaryrefslogtreecommitdiff
diff options
context:
space:
mode:
-rw-r--r--README17
-rw-r--r--arm/Asm.v2
-rw-r--r--arm/SelectOp.vp17
-rw-r--r--arm/SelectOpproof.v25
-rw-r--r--backend/CMtypecheck.ml18
-rw-r--r--backend/Cminor.v24
-rw-r--r--backend/PrintCminor.ml11
-rw-r--r--backend/RTL.v2
-rw-r--r--backend/Selection.v3
-rw-r--r--backend/Selectionproof.v10
-rw-r--r--cfrontend/C2C.ml82
-rw-r--r--cfrontend/Cexec.v2157
-rw-r--r--cfrontend/Clight.v77
-rw-r--r--cfrontend/Cminorgen.v89
-rw-r--r--cfrontend/Cminorgenproof.v311
-rw-r--r--cfrontend/Cparser.mlpack6
-rw-r--r--cfrontend/Csem.v271
-rw-r--r--cfrontend/Csharpminor.v61
-rw-r--r--cfrontend/Cshmgen.v100
-rw-r--r--cfrontend/Cshmgenproof.v519
-rw-r--r--cfrontend/Cstrategy.v625
-rw-r--r--cfrontend/Csyntax.v398
-rw-r--r--cfrontend/Initializers.v30
-rw-r--r--cfrontend/Initializersproof.v89
-rw-r--r--cfrontend/PrintClight.ml22
-rw-r--r--cfrontend/PrintCsyntax.ml55
-rw-r--r--cfrontend/SimplExpr.v58
-rw-r--r--cfrontend/SimplExprproof.v289
-rw-r--r--cfrontend/SimplExprspec.v103
-rw-r--r--common/Behaviors.v150
-rw-r--r--common/Events.v497
-rw-r--r--common/Smallstep.v303
-rw-r--r--common/Values.v7
-rwxr-xr-xconfigure18
-rw-r--r--cparser/.depend34
-rw-r--r--cparser/AddCasts.ml243
-rw-r--r--cparser/AddCasts.mli16
-rw-r--r--cparser/Elab.ml4
-rw-r--r--cparser/Makefile4
-rw-r--r--cparser/Parse.ml14
-rw-r--r--cparser/SimplExpr.ml568
-rw-r--r--cparser/SimplExpr.mli20
-rw-r--r--cparser/SimplVolatile.ml81
-rw-r--r--cparser/StructAssign.ml165
-rw-r--r--cparser/StructAssign.mli18
-rw-r--r--cparser/StructReturn.ml (renamed from cparser/StructByValue.ml)107
-rw-r--r--cparser/StructReturn.mli (renamed from cparser/StructByValue.mli)0
-rw-r--r--driver/Clflags.ml4
-rw-r--r--driver/Compiler.v13
-rw-r--r--driver/Complements.v10
-rw-r--r--driver/Driver.ml17
-rw-r--r--driver/Interp.ml36
-rw-r--r--ia32/Asm.v2
-rw-r--r--ia32/SelectOp.vp15
-rw-r--r--ia32/SelectOpproof.v25
-rw-r--r--powerpc/Asm.v8
-rw-r--r--powerpc/PrintAsm.ml49
-rw-r--r--powerpc/SelectOp.vp15
-rw-r--r--powerpc/SelectOpproof.v25
-rw-r--r--test/raytracer/Makefile2
-rw-r--r--test/regression/Makefile2
-rw-r--r--test/regression/Results/bool9
-rw-r--r--test/regression/bool.c29
-rw-r--r--test/spass/Makefile2
64 files changed, 4356 insertions, 3627 deletions
diff --git a/README b/README
index dfe6653..8673e73 100644
--- a/README
+++ b/README
@@ -33,12 +33,12 @@ SUPPORTED PLATFORMS:
natively at full speed on PowerPC-based Macs, and runs under
software emulation at reduced speed on Intel-based Macs.
-- PowerPC / Linux [somewhat experimental]
+- PowerPC / Linux [stable]
For PowerPC machines running the Linux operating system.
-- IA32 / Linux or MacOS or Windows+Cygwin [experimental]
+- IA32 / Linux or MacOS or Windows+Cygwin [stable]
For Intel/AMD x86 processors with SSE2 extensions
- (i.e. Pentium 4 and later), running either Linux, MacOS 10.6,
+ (i.e. Pentium 4 and later), running either Linux, MacOS 10.6 or 10.7,
or Windows with the Cygwin environment (http://www.cygwin.com/).
- ARM / Linux [experimental]
@@ -60,7 +60,7 @@ PREREQUISITES:
* The Caml functional language, version 3.10 or later.
Caml is free software, available from http://caml.inria.fr/
-* Under MacOS 10.5 and 10.6, some standard C include files in /usr/include/
+* Under MacOS, some standard C include files in /usr/include/
contain gcc-isms that cause errors when compiling with CompCert.
Symptoms include:
- references to undefined types uint16_t and uint32_t
@@ -84,6 +84,8 @@ where <target> is one of:
ppc-macosx (PowerPC, MacOS X)
ppc-linux (PowerPC, Linux)
+ ppc-eabi-unix (PowerPC, EABI with Unix tools)
+ ppc-eabi-diab (PowerPC, EABI with Diab tools)
arm-linux (ARM, Linux)
ia32-linux (x86 SSE2 32 bits, Linux)
ia32-macosx (x86 SSE2 32 bits, MacOS X)
@@ -172,9 +174,12 @@ Preprocessing options:
Language support options (use -fno-<opt> to turn off -f<opt>) :
-fbitfields Emulate bit fields in structs [off]
-flonglong Partial emulation of 'long long' types [on]
- -fstruct-passing Emulate passing structs and unions by value [off]
- -fstruct-assign Emulate assignment between structs or unions [off]
+ -flongdouble Treat 'long double' as 'double' [off]
+ -fstruct-return Emulate returning structs and unions by value [off]
-fvararg-calls Emulate calls to variable-argument functions [on]
+ -fpacked-structs Emulate packed structs [off]
+ -fall Activate all language support options above
+ -fnone Turn off all language support options above
Code generation options:
-fmadd Use fused multiply-add and multiply-sub instructions
-fsmall-data <n> Set maximal size <n> for allocation in small data area
diff --git a/arm/Asm.v b/arm/Asm.v
index 21b8c4c..5e16f05 100644
--- a/arm/Asm.v
+++ b/arm/Asm.v
@@ -750,7 +750,7 @@ Ltac Equalities :=
exploit external_call_determ. eexact H3. eexact H8. intros [A B].
split. auto. intros. destruct B; auto. subst. auto.
(* trace length *)
- inv H; simpl.
+ red; intros; inv H; simpl.
omega.
eapply external_call_trace_length; eauto.
eapply external_call_trace_length; eauto.
diff --git a/arm/SelectOp.vp b/arm/SelectOp.vp
index 432db94..7b8851c 100644
--- a/arm/SelectOp.vp
+++ b/arm/SelectOp.vp
@@ -68,11 +68,24 @@ Nondetfunction notint (e: expr) :=
| _ => Eop Onot (e:::Enil)
end.
-(** ** Boolean negation *)
+(** ** Boolean value and boolean negation *)
+
+Fixpoint boolval (e: expr) {struct e} : expr :=
+ let default := Eop (Ocmp (Ccompuimm Cne Int.zero)) (e ::: Enil) in
+ match e with
+ | Eop (Ointconst n) Enil =>
+ Eop (Ointconst (if Int.eq n Int.zero then Int.zero else Int.one)) Enil
+ | Eop (Ocmp cond) args =>
+ Eop (Ocmp cond) args
+ | Econdition e1 e2 e3 =>
+ Econdition e1 (boolval e2) (boolval e3)
+ | _ =>
+ default
+ end.
Fixpoint notbool (e: expr) {struct e} : expr :=
let default := Eop (Ocmp (Ccompuimm Ceq Int.zero)) (e ::: Enil) in
- match e with
+ match e with
| Eop (Ointconst n) Enil =>
Eop (Ointconst (if Int.eq n Int.zero then Int.one else Int.zero)) Enil
| Eop (Ocmp cond) args =>
diff --git a/arm/SelectOpproof.v b/arm/SelectOpproof.v
index fa41682..0a5ee64 100644
--- a/arm/SelectOpproof.v
+++ b/arm/SelectOpproof.v
@@ -141,6 +141,31 @@ Proof.
TrivialExists.
Qed.
+Theorem eval_boolval: unary_constructor_sound boolval Val.boolval.
+Proof.
+ assert (DFL:
+ forall le a x,
+ eval_expr ge sp e m le a x ->
+ exists v, eval_expr ge sp e m le (Eop (Ocmp (Ccompuimm Cne Int.zero)) (a ::: Enil)) v
+ /\ Val.lessdef (Val.boolval x) v).
+ intros. TrivialExists. simpl. destruct x; simpl; auto.
+
+ red. induction a; simpl; intros; eauto. destruct o; eauto.
+(* intconst *)
+ destruct e0; eauto. InvEval. TrivialExists. simpl. destruct (Int.eq i Int.zero); auto.
+(* cmp *)
+ inv H. simpl in H5.
+ destruct (eval_condition c vl m) as []_eqn.
+ TrivialExists. simpl. inv H5. rewrite Heqo. destruct b; auto.
+ simpl in H5. inv H5.
+ exists Vundef; split; auto. EvalOp; simpl. rewrite Heqo; auto.
+
+(* condition *)
+ inv H. destruct v1.
+ exploit IHa1; eauto. intros [v [A B]]. exists v; split; auto. eapply eval_Econdition; eauto.
+ exploit IHa2; eauto. intros [v [A B]]. exists v; split; auto. eapply eval_Econdition; eauto.
+Qed.
+
Theorem eval_notbool: unary_constructor_sound notbool Val.notbool.
Proof.
assert (DFL:
diff --git a/backend/CMtypecheck.ml b/backend/CMtypecheck.ml
index e3a6f70..244a73f 100644
--- a/backend/CMtypecheck.ml
+++ b/backend/CMtypecheck.ml
@@ -90,6 +90,7 @@ let type_unary_operation = function
| Ocast8unsigned -> tint, tint
| Ocast16unsigned -> tint, tint
| Onegint -> tint, tint
+ | Oboolval -> tint, tint
| Onotbool -> tint, tint
| Onotint -> tint, tint
| Onegf -> tfloat, tfloat
@@ -134,6 +135,7 @@ let name_of_unary_operation = function
| Ocast8unsigned -> "cast8unsigned"
| Ocast16unsigned -> "cast16unsigned"
| Onegint -> "negint"
+ | Oboolval -> "notbool"
| Onotbool -> "notbool"
| Onotint -> "notint"
| Onegf -> "negf"
@@ -293,6 +295,22 @@ let rec type_stmt env blk ret s =
with Error s ->
raise (Error (sprintf "In call:\n%s" s))
end
+ | Sbuiltin(optid, ef, el) ->
+ let sg = ef_sig ef in
+ let tel = type_exprlist env [] el in
+ begin try
+ unify_list (ty_of_sig_args sg.sig_args) tel;
+ let ty_res =
+ match sg.sig_res with
+ | None -> tint (*???*)
+ | Some t -> ty_of_typ t in
+ begin match optid with
+ | None -> ()
+ | Some id -> unify (type_var env id) ty_res
+ end
+ with Error s ->
+ raise (Error (sprintf "In builtin call:\n%s" s))
+ end
| Sseq(s1, s2) ->
type_stmt env blk ret s1;
type_stmt env blk ret s2
diff --git a/backend/Cminor.v b/backend/Cminor.v
index c9ee5b5..6d288a9 100644
--- a/backend/Cminor.v
+++ b/backend/Cminor.v
@@ -44,6 +44,7 @@ Inductive unary_operation : Type :=
| Ocast8signed: unary_operation (**r 8-bit sign extension *)
| Ocast16unsigned: unary_operation (**r 16-bit zero extension *)
| Ocast16signed: unary_operation (**r 16-bit sign extension *)
+ | Oboolval: unary_operation (**r 0 if null, 1 if non-null *)
| Onegint: unary_operation (**r integer opposite *)
| Onotbool: unary_operation (**r boolean negation *)
| Onotint: unary_operation (**r bitwise complement *)
@@ -103,6 +104,7 @@ Inductive stmt : Type :=
| Sstore : memory_chunk -> expr -> expr -> stmt
| Scall : option ident -> signature -> expr -> list expr -> stmt
| Stailcall: signature -> expr -> list expr -> stmt
+ | Sbuiltin : option ident -> external_function -> list expr -> stmt
| Sseq: stmt -> stmt -> stmt
| Sifthenelse: expr -> stmt -> stmt -> stmt
| Sloop: stmt -> stmt
@@ -228,6 +230,7 @@ Definition eval_unop (op: unary_operation) (arg: val) : option val :=
| Ocast8signed => Some (Val.sign_ext 8 arg)
| Ocast16unsigned => Some (Val.zero_ext 16 arg)
| Ocast16signed => Some (Val.sign_ext 16 arg)
+ | Oboolval => Some(Val.boolval arg)
| Onegint => Some (Val.negint arg)
| Onotbool => Some (Val.notbool arg)
| Onotint => Some (Val.notint arg)
@@ -401,6 +404,12 @@ Inductive step: state -> trace -> state -> Prop :=
step (State f (Stailcall sig a bl) k (Vptr sp Int.zero) e m)
E0 (Callstate fd vargs (call_cont k) m')
+ | step_builtin: forall f optid ef bl k sp e m vargs t vres m',
+ eval_exprlist sp e m bl vargs ->
+ external_call ef ge vargs m t vres m' ->
+ step (State f (Sbuiltin optid ef bl) k sp e m)
+ t (State f Sskip k sp (set_optvar optid vres e) m')
+
| step_seq: forall f s1 s2 k sp e m,
step (State f (Sseq s1 s2) k sp e m)
E0 (State f s1 (Kseq s2 k) sp e m)
@@ -505,9 +514,11 @@ Proof.
intros. subst. inv H0. exists s1; auto.
inversion H; subst; auto.
exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exists (State f Sskip k sp (set_optvar optid vres2 e) m2). econstructor; eauto.
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (Returnstate vres2 k m2). econstructor; eauto.
(* trace length *)
- inv H; simpl; try omega. eapply external_call_trace_length; eauto.
+ red; intros; inv H; simpl; try omega; eapply external_call_trace_length; eauto.
Qed.
(** * Alternate operational semantics (big-step) *)
@@ -608,6 +619,12 @@ with exec_stmt:
eval_funcall m fd vargs t m' vres ->
e' = set_optvar optid vres e ->
exec_stmt f sp e m (Scall optid sig a bl) t e' m' Out_normal
+ | exec_Sbuiltin:
+ forall f sp e m optid ef bl t m' vargs vres e',
+ eval_exprlist ge sp e m bl vargs ->
+ external_call ef ge vargs m t vres m' ->
+ e' = set_optvar optid vres e ->
+ exec_stmt f sp e m (Sbuiltin optid ef bl) t e' m' Out_normal
| exec_Sifthenelse:
forall f sp e m a s1 s2 v b t e' m' out,
eval_expr ge sp e m a v ->
@@ -884,6 +901,11 @@ Proof.
constructor. reflexivity. traceEq.
subst e'. constructor.
+(* builtin *)
+ econstructor; split.
+ apply star_one. econstructor; eauto.
+ subst e'. constructor.
+
(* ifthenelse *)
destruct (H2 k) as [S [A B]].
exists S; split.
diff --git a/backend/PrintCminor.ml b/backend/PrintCminor.ml
index 30884b1..110e735 100644
--- a/backend/PrintCminor.ml
+++ b/backend/PrintCminor.ml
@@ -21,6 +21,7 @@ open Datatypes
open BinPos
open Integers
open AST
+open PrintAST
open Cminor
(* Precedences and associativity -- like those of C *)
@@ -57,6 +58,7 @@ let name_of_unop = function
| Ocast16unsigned -> "int16u"
| Ocast16signed -> "int16s"
| Onegint -> "-"
+ | Oboolval -> "(_Bool)"
| Onotbool -> "!"
| Onotint -> "~"
| Onegf -> "-f"
@@ -193,6 +195,15 @@ let rec print_stmt p s =
print_expr e1
print_expr_list (true, el)
print_sig sg
+ | Sbuiltin(None, ef, el) ->
+ fprintf p "@[<hv 2>builtin %s@,(@[<hov 0>%a@])@;@]"
+ (name_of_external ef)
+ print_expr_list (true, el)
+ | Sbuiltin(Some id, ef, el) ->
+ fprintf p "@[<hv 2>%s =@ builtin %s@,(@[<hov 0>%a@]);@]@]"
+ (ident_name id)
+ (name_of_external ef)
+ print_expr_list (true, el)
| Sseq(Sskip, s2) ->
print_stmt p s2
| Sseq(s1, Sskip) ->
diff --git a/backend/RTL.v b/backend/RTL.v
index 6a20941..9b27a17 100644
--- a/backend/RTL.v
+++ b/backend/RTL.v
@@ -360,7 +360,7 @@ Proof.
exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (Returnstate s0 vres2 m2). econstructor; eauto.
(* trace length *)
- inv H; simpl; try omega.
+ red; intros; inv H; simpl; try omega.
eapply external_call_trace_length; eauto.
eapply external_call_trace_length; eauto.
Qed.
diff --git a/backend/Selection.v b/backend/Selection.v
index 9c037b8..2d6c901 100644
--- a/backend/Selection.v
+++ b/backend/Selection.v
@@ -112,6 +112,7 @@ Definition sel_unop (op: Cminor.unary_operation) (arg: expr) : expr :=
| Cminor.Ocast16unsigned => cast16unsigned arg
| Cminor.Ocast16signed => cast16signed arg
| Cminor.Onegint => negint arg
+ | Cminor.Oboolval => boolval arg
| Cminor.Onotbool => notbool arg
| Cminor.Onotint => notint arg
| Cminor.Onegf => negf arg
@@ -202,6 +203,8 @@ Fixpoint sel_stmt (ge: Cminor.genv) (s: Cminor.stmt) : stmt :=
| None => Scall optid sg (sel_expr fn) (sel_exprlist args)
| Some ef => Sbuiltin optid ef (sel_exprlist args)
end
+ | Cminor.Sbuiltin optid ef args =>
+ Sbuiltin optid ef (sel_exprlist args)
| Cminor.Stailcall sg fn args =>
Stailcall sg (sel_expr fn) (sel_exprlist args)
| Cminor.Sseq s1 s2 => Sseq (sel_stmt ge s1) (sel_stmt ge s2)
diff --git a/backend/Selectionproof.v b/backend/Selectionproof.v
index 54d59b1..9681c66 100644
--- a/backend/Selectionproof.v
+++ b/backend/Selectionproof.v
@@ -212,6 +212,7 @@ Proof.
apply eval_cast8signed; auto.
apply eval_cast16unsigned; auto.
apply eval_cast16signed; auto.
+ apply eval_boolval; auto.
apply eval_negint; auto.
apply eval_notbool; auto.
apply eval_notint; auto.
@@ -619,6 +620,15 @@ Proof.
eapply functions_translated; eauto.
apply sig_function_translated.
constructor; auto. apply call_cont_commut; auto.
+ (* Sbuiltin *)
+ exploit sel_exprlist_correct; eauto. intros [vargs' [P Q]].
+ exploit external_call_mem_extends; eauto.
+ intros [vres' [m2 [A [B [C D]]]]].
+ left; econstructor; split.
+ econstructor. eauto. eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+ constructor; auto.
+ destruct optid; simpl; auto. apply set_var_lessdef; auto.
(* Seq *)
left; econstructor; split. constructor. constructor; auto. constructor; auto.
(* Sifthenelse *)
diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml
index ee51914..d4faa2b 100644
--- a/cfrontend/C2C.ml
+++ b/cfrontend/C2C.ml
@@ -122,7 +122,9 @@ let name_for_string_literal env s =
id
let typeStringLiteral s =
- Tarray(Tint(I8, Unsigned), z_of_camlint(Int32.of_int(String.length s + 1)))
+ Tarray(Tint(I8, Unsigned, noattr),
+ z_of_camlint(Int32.of_int(String.length s + 1)),
+ noattr)
let global_for_string s id =
let init = ref [] in
@@ -162,8 +164,8 @@ let register_stub_function name tres targs =
| Tcons(_, tl) -> "i" :: letters_of_type tl in
let rec types_of_types = function
| Tnil -> Tnil
- | Tcons(Tfloat _, tl) -> Tcons(Tfloat F64, types_of_types tl)
- | Tcons(_, tl) -> Tcons(Tpointer Tvoid, types_of_types tl) in
+ | Tcons(Tfloat _, tl) -> Tcons(Tfloat(F64, noattr), types_of_types tl)
+ | Tcons(_, tl) -> Tcons(Tpointer(Tvoid, noattr), types_of_types tl) in
let stub_name =
name ^ "$" ^ String.concat "" (letters_of_type targs) in
let targs = types_of_types targs in
@@ -204,7 +206,8 @@ let register_inlined_memcpy sz al =
let al =
if al <= 4l then al else 4l in (* max alignment supported by CompCert *)
let name = Printf.sprintf "__builtin_memcpy_sz%ld_al%ld" sz al in
- let targs = Tcons(Tpointer Tvoid, Tcons(Tpointer Tvoid, Tnil))
+ let targs = Tcons(Tpointer(Tvoid, noattr),
+ Tcons(Tpointer(Tvoid, noattr), Tnil))
and tres = Tvoid
and ef = EF_memcpy(coqint_of_camlint sz, coqint_of_camlint al) in
register_special_external name ef targs tres;
@@ -233,10 +236,14 @@ let make_builtin_memcpy args =
let convertInt n = coqint_of_camlint(Int64.to_int32 n)
+(** Attributes *)
+
+let convertAttr a = List.mem AVolatile a
+
(** Types *)
let convertIkind = function
- | C.IBool -> unsupported "'_Bool' type"; (Unsigned, I8)
+ | C.IBool -> (Unsigned, IBool)
| C.IChar -> ((if (!Cparser.Machine.config).Cparser.Machine.char_signed
then Signed else Unsigned), I8)
| C.ISChar -> (Signed, I8)
@@ -258,12 +265,13 @@ let convertFkind = function
if not !Clflags.option_flongdouble then unsupported "'long double' type";
F64
-let int64_struct =
- let ty = Tint(I32,Unsigned) in
+let int64_struct a =
+ let ty = Tint(I32,Unsigned,noattr) in
Tstruct(intern_string "struct __int64",
- if Memdataaux.big_endian
- then Fcons(intern_string "hi", ty, Fcons(intern_string "lo", ty, Fnil))
- else Fcons(intern_string "lo", ty, Fcons(intern_string "hi", ty, Fnil)))
+ (if Memdataaux.big_endian
+ then Fcons(intern_string "hi", ty, Fcons(intern_string "lo", ty, Fnil))
+ else Fcons(intern_string "lo", ty, Fcons(intern_string "hi", ty, Fnil))),
+ a)
let convertTyp env t =
@@ -271,27 +279,27 @@ let convertTyp env t =
match Cutil.unroll env t with
| C.TVoid a -> Tvoid
| C.TInt((C.ILongLong|C.IULongLong), a) when !Clflags.option_flonglong ->
- int64_struct
+ int64_struct (convertAttr a)
| C.TInt(ik, a) ->
- let (sg, sz) = convertIkind ik in Tint(sz, sg)
+ let (sg, sz) = convertIkind ik in Tint(sz, sg, convertAttr a)
| C.TFloat(fk, a) ->
- Tfloat(convertFkind fk)
+ Tfloat(convertFkind fk, convertAttr a)
| C.TPtr(ty, a) ->
begin match Cutil.unroll env ty with
| C.TStruct(id, _) when List.mem id seen ->
- Tcomp_ptr(intern_string ("struct " ^ id.name))
+ Tcomp_ptr(intern_string ("struct " ^ id.name), convertAttr a)
| C.TUnion(id, _) when List.mem id seen ->
- Tcomp_ptr(intern_string ("union " ^ id.name))
+ Tcomp_ptr(intern_string ("union " ^ id.name), convertAttr a)
| _ ->
- Tpointer(convertTyp seen ty)
+ Tpointer(convertTyp seen ty, convertAttr a)
end
| C.TArray(ty, None, a) ->
(* Cparser verified that the type ty[] occurs only in
contexts that are safe for Clight, so just treat as ty[0]. *)
(* warning "array type of unspecified size"; *)
- Tarray(convertTyp seen ty, coqint_of_camlint 0l)
+ Tarray(convertTyp seen ty, coqint_of_camlint 0l, convertAttr a)
| C.TArray(ty, Some sz, a) ->
- Tarray(convertTyp seen ty, convertInt sz)
+ Tarray(convertTyp seen ty, convertInt sz, convertAttr a)
| C.TFun(tres, targs, va, a) ->
if va then unsupported "variadic function type";
if Cutil.is_composite_type env tres then
@@ -309,20 +317,18 @@ let convertTyp env t =
convertFields (id :: seen) (Env.find_struct env id)
with Env.Error e ->
error (Env.error_message e); Fnil in
- Tstruct(intern_string("struct " ^ id.name), flds)
+ Tstruct(intern_string("struct " ^ id.name), flds, convertAttr a)
| C.TUnion(id, a) ->
let flds =
try
convertFields (id :: seen) (Env.find_union env id)
with Env.Error e ->
error (Env.error_message e); Fnil in
- Tunion(intern_string("union " ^ id.name), flds)
+ Tunion(intern_string("union " ^ id.name), flds, convertAttr a)
and convertParams seen = function
| [] -> Tnil
| (id, ty) :: rem ->
- if Cutil.is_composite_type env ty then
- unsupported "function parameter of struct or union type";
Tcons(convertTyp seen ty, convertParams seen rem)
and convertFields seen ci =
@@ -358,10 +364,10 @@ let string_of_type ty =
let first_class_value env ty =
match Cutil.unroll env ty with
| C.TInt((C.ILongLong|C.IULongLong), _) -> false
- | C.TStruct _ -> false
- | C.TUnion _ -> false
| _ -> true
+(************ REMOVED
+
(* Handling of volatile *)
let is_volatile_access env e =
@@ -398,16 +404,15 @@ let volatile_write_fun ty =
let name = "__builtin_volatile_write_" ^ suffix in
register_special_external name (EF_vstore chunk) targs Tvoid;
Evalof(Evar(intern_string name, Tfunction(targs, Tvoid)), Tfunction(targs, Tvoid))
+****************************)
(** Expressions *)
-let ezero = Eval(Vint(coqint_of_camlint 0l), Tint(I32, Signed))
+let ezero = Eval(Vint(coqint_of_camlint 0l), type_int32s)
let check_assignop msg env e =
if not (first_class_value env e.etyp) then
- unsupported (msg ^ " on a l-value of type " ^ string_of_type e.etyp);
- if is_volatile_access env e then
- unsupported (msg ^ " on a volatile l-value")
+ unsupported (msg ^ " on a l-value of type " ^ string_of_type e.etyp)
let rec convertExpr env e =
let ty = convertTyp env e.etyp in
@@ -418,12 +423,7 @@ let rec convertExpr env e =
let l = convertLvalue env e in
if not (first_class_value env e.etyp) then
unsupported ("r-value of type " ^ string_of_type e.etyp);
- if is_volatile_access env e then
- Ecall(volatile_read_fun (typeof l),
- Econs(Eaddrof(l, Tpointer(typeof l)), Enil),
- ty)
- else
- Evalof(l, ty)
+ Evalof(l, ty)
| C.EConst(C.CInt(i, k, _)) ->
if k = C.ILongLong || k = C.IULongLong then
@@ -492,14 +492,8 @@ let rec convertExpr env e =
| C.EBinop(C.Oassign, e1, e2, _) ->
let e1' = convertLvalue env e1 in
let e2' = convertExpr env e2 in
- if not (first_class_value env e1.etyp) then
- unsupported ("assignment to a l-value of type " ^ string_of_type e1.etyp);
- if is_volatile_access env e1 then
- Ecall(volatile_write_fun (typeof e1'),
- Econs(Eaddrof(e1', Tpointer(typeof e1')), Econs(e2', Enil)),
- Tvoid) (* SimplVolatile guarantees that ret. value is unused *)
- else
- Eassign(e1', e2', ty)
+ check_assignop "assignment" env e1;
+ Eassign(e1', e2', ty)
| C.EBinop((C.Oadd_assign|C.Osub_assign|C.Omul_assign|C.Odiv_assign|
C.Omod_assign|C.Oand_assign|C.Oor_assign|C.Oxor_assign|
C.Oshl_assign|C.Oshr_assign) as op,
@@ -599,7 +593,7 @@ and convertLvalue env e =
let e1' = convertExpr env e1 in
let ty1 =
match typeof e1' with
- | Tpointer t -> t
+ | Tpointer(t, _) -> t
| _ -> error ("wrong type for ->" ^ id ^ " access"); Tvoid in
Efield(Ederef(e1', ty1), intern_string id, ty)
| C.EBinop(C.Oindex, e1, e2, _) ->
@@ -723,8 +717,6 @@ let convertFundef env fd =
let params =
List.map
(fun (id, ty) ->
- if Cutil.is_composite_type env ty then
- unsupported "function parameter of struct or union type";
(intern_string id.name, convertTyp env ty))
fd.fd_params in
let vars =
diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v
index 4bce535..b3c3f6b 100644
--- a/cfrontend/Cexec.v
+++ b/cfrontend/Cexec.v
@@ -13,6 +13,7 @@
(** Animating the CompCert C semantics *)
Require Import Axioms.
+Require Import Classical.
Require Import Coqlib.
Require Import Errors.
Require Import Maps.
@@ -28,28 +29,20 @@ Require Import Csyntax.
Require Import Csem.
Require Cstrategy.
-Lemma type_eq: forall (ty1 ty2: type), {ty1=ty2} + {ty1<>ty2}
-with typelist_eq: forall (tyl1 tyl2: typelist), {tyl1=tyl2} + {tyl1<>tyl2}
-with fieldlist_eq: forall (fld1 fld2: fieldlist), {fld1=fld2} + {fld1<>fld2}.
-Proof.
- assert (forall (x y: intsize), {x=y} + {x<>y}). decide equality.
- assert (forall (x y: signedness), {x=y} + {x<>y}). decide equality.
- assert (forall (x y: floatsize), {x=y} + {x<>y}). decide equality.
- generalize ident_eq zeq. intros E1 E2.
- decide equality.
- decide equality.
- generalize ident_eq. intros E1.
- decide equality.
-Defined.
-
-Opaque type_eq.
-
(** Error monad with options or lists *)
Notation "'do' X <- A ; B" := (match A with Some X => B | None => None end)
(at level 200, X ident, A at level 100, B at level 200)
: option_monad_scope.
+Notation "'do' X , Y <- A ; B" := (match A with Some (X, Y) => B | None => None end)
+ (at level 200, X ident, Y ident, A at level 100, B at level 200)
+ : option_monad_scope.
+
+Notation "'do' X , Y , Z <- A ; B" := (match A with Some (X, Y, Z) => B | None => None end)
+ (at level 200, X ident, Y ident, Z ident, A at level 100, B at level 200)
+ : option_monad_scope.
+
Notation " 'check' A ; B" := (if A then B else None)
(at level 200, A at level 100, B at level 200)
: option_monad_scope.
@@ -99,20 +92,558 @@ Proof.
destruct s; (left; congruence) || (right; congruence).
Qed.
-(** * Reduction of expressions *)
+(** * Events, volatile memory accesses, and external functions. *)
Section EXEC.
Variable ge: genv.
+Definition eventval_of_val (v: val) (t: typ) : option eventval :=
+ match v, t with
+ | Vint i, AST.Tint => Some (EVint i)
+ | Vfloat f, AST.Tfloat => Some (EVfloat f)
+ | Vptr b ofs, AST.Tint => do id <- Genv.invert_symbol ge b; Some (EVptr_global id ofs)
+ | _, _ => None
+ end.
+
+Fixpoint list_eventval_of_val (vl: list val) (tl: list typ) : option (list eventval) :=
+ match vl, tl with
+ | nil, nil => Some nil
+ | v1::vl, t1::tl =>
+ do ev1 <- eventval_of_val v1 t1;
+ do evl <- list_eventval_of_val vl tl;
+ Some (ev1 :: evl)
+ | _, _ => None
+ end.
+
+Definition val_of_eventval (ev: eventval) (t: typ) : option val :=
+ match ev, t with
+ | EVint i, AST.Tint => Some (Vint i)
+ | EVfloat f, AST.Tfloat => Some (Vfloat f)
+ | EVptr_global id ofs, AST.Tint => do b <- Genv.find_symbol ge id; Some (Vptr b ofs)
+ | _, _ => None
+ end.
+
+Lemma eventval_of_val_sound:
+ forall v t ev, eventval_of_val v t = Some ev -> eventval_match ge ev t v.
+Proof.
+ intros. destruct v; destruct t; simpl in H; inv H.
+ constructor.
+ constructor.
+ destruct (Genv.invert_symbol ge b) as [id|]_eqn; inv H1.
+ constructor. apply Genv.invert_find_symbol; auto.
+Qed.
+
+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 _ _ H). auto.
+Qed.
+
+Lemma list_eventval_of_val_sound:
+ forall vl tl evl, list_eventval_of_val vl tl = Some evl -> eventval_list_match ge evl tl vl.
+Proof with try discriminate.
+ induction vl; destruct tl; simpl; intros; inv H.
+ constructor.
+ destruct (eventval_of_val a t) as [ev1|]_eqn...
+ destruct (list_eventval_of_val vl tl) as [evl'|]_eqn...
+ inv H1. constructor. apply eventval_of_val_sound; auto. eauto.
+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.
+ rewrite (eventval_of_val_complete _ _ _ H). rewrite IHeventval_list_match. auto.
+Qed.
+
+Lemma val_of_eventval_sound:
+ forall ev t v, val_of_eventval ev t = Some v -> eventval_match ge ev t v.
+Proof.
+ intros. destruct ev; destruct t; simpl in H; inv H.
+ constructor.
+ constructor.
+ destruct (Genv.find_symbol ge i) as [b|]_eqn; inv H1.
+ constructor. auto.
+Qed.
+
+Lemma val_of_eventval_complete:
+ forall ev t v, eventval_match ge ev t v -> val_of_eventval ev t = Some v.
+Proof.
+ induction 1; simpl; auto. rewrite H; auto.
+Qed.
+
+(** Volatile memory accesses. *)
+
+Definition do_volatile_load (w: world) (chunk: memory_chunk) (m: mem) (b: block) (ofs: int)
+ : option (world * trace * val) :=
+ if block_is_volatile ge b then
+ do id <- Genv.invert_symbol ge b;
+ match nextworld_vload w chunk id ofs with
+ | None => None
+ | Some(res, w') =>
+ do vres <- val_of_eventval res (type_of_chunk chunk);
+ Some(w', Event_vload chunk id ofs res :: nil, Val.load_result chunk vres)
+ end
+ else
+ do v <- Mem.load chunk m b (Int.unsigned ofs);
+ Some(w, E0, v).
+
+Definition do_volatile_store (w: world) (chunk: memory_chunk) (m: mem) (b: block) (ofs: int) (v: val)
+ : option (world * trace * mem) :=
+ if block_is_volatile ge b then
+ do id <- Genv.invert_symbol ge b;
+ do ev <- eventval_of_val v (type_of_chunk chunk);
+ do w' <- nextworld_vstore w chunk id ofs ev;
+ Some(w', Event_vstore chunk id ofs ev :: nil, m)
+ else
+ do m' <- Mem.store chunk m b (Int.unsigned ofs) v;
+ Some(w, E0, m').
+
+Ltac mydestr :=
+ match goal with
+ | [ |- None = Some _ -> _ ] => intro X; discriminate
+ | [ |- Some _ = Some _ -> _ ] => intro X; inv X
+ | [ |- match ?x with Some _ => _ | None => _ end = Some _ -> _ ] => destruct x as []_eqn; mydestr
+ | [ |- match ?x with true => _ | false => _ end = Some _ -> _ ] => destruct x as []_eqn; mydestr
+ | [ |- match ?x with left _ => _ | right _ => _ end = Some _ -> _ ] => destruct x; mydestr
+ | _ => idtac
+ end.
+
+Lemma do_volatile_load_sound:
+ forall w chunk m b ofs w' t v,
+ 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.
+ split. constructor; auto. constructor.
+Qed.
+
+Lemma do_volatile_load_complete:
+ forall w chunk m b ofs w' t v,
+ volatile_load ge chunk m b ofs t v -> possible_trace w t w' ->
+ do_volatile_load w chunk m b ofs = Some(w', t, v).
+Proof.
+ unfold do_volatile_load; intros. inv H.
+ rewrite H1. rewrite (Genv.find_invert_symbol _ _ H2). inv H0. inv H8. inv H6. rewrite H9.
+ rewrite (val_of_eventval_complete _ _ _ H3). auto.
+ rewrite H1. rewrite H2. inv H0. auto.
+Qed.
+
+Lemma do_volatile_store_sound:
+ forall w chunk m b ofs v w' t m',
+ 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.
+ split. constructor; auto. constructor.
+Qed.
+
+Lemma do_volatile_store_complete:
+ forall w chunk m b ofs v w' t m',
+ volatile_store ge chunk m b ofs v t m' -> possible_trace w t w' ->
+ do_volatile_store w chunk m b ofs v = Some(w', t, m').
+Proof.
+ unfold do_volatile_store; intros. inv H.
+ rewrite H1. rewrite (Genv.find_invert_symbol _ _ H2).
+ rewrite (eventval_of_val_complete _ _ _ H3).
+ inv H0. inv H8. inv H6. rewrite H9. auto.
+ rewrite H1. rewrite H2. inv H0. auto.
+Qed.
+
+(** Accessing locations *)
+
+Definition do_deref_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: int) : option (world * trace * val) :=
+ match access_mode ty with
+ | By_value chunk =>
+ match type_is_volatile ty with
+ | false => do v <- Mem.loadv chunk m (Vptr b ofs); Some(w, E0, v)
+ | true => do_volatile_load w chunk m b ofs
+ end
+ | By_reference => Some(w, E0, Vptr b ofs)
+ | By_copy => Some(w, E0, Vptr b ofs)
+ | _ => None
+ end.
+
+Definition assign_copy_ok (ty: type) (b: block) (ofs: int) (b': block) (ofs': int) : Prop :=
+ (alignof ty | Int.unsigned ofs') /\ (alignof ty | Int.unsigned ofs) /\
+ (b' <> b \/ Int.unsigned ofs' = Int.unsigned ofs
+ \/ Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs
+ \/ Int.unsigned ofs + sizeof ty <= Int.unsigned ofs').
+
+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.
+ assert (alignof ty > 0). apply alignof_pos; auto.
+ destruct (Zdivide_dec (alignof ty) (Int.unsigned ofs')); auto...
+ destruct (Zdivide_dec (alignof ty) (Int.unsigned ofs)); auto...
+ assert (Y: {b' <> b \/
+ Int.unsigned ofs' = Int.unsigned ofs \/
+ Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs \/
+ Int.unsigned ofs + sizeof ty <= Int.unsigned ofs'} +
+ {~(b' <> b \/
+ Int.unsigned ofs' = Int.unsigned ofs \/
+ Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs \/
+ Int.unsigned ofs + sizeof ty <= Int.unsigned ofs')}).
+ destruct (eq_block b' b); auto.
+ destruct (zeq (Int.unsigned ofs') (Int.unsigned ofs)); auto.
+ destruct (zle (Int.unsigned ofs' + sizeof ty) (Int.unsigned ofs)); auto.
+ destruct (zle (Int.unsigned ofs + sizeof ty) (Int.unsigned ofs')); auto.
+ right; intuition omega.
+ destruct Y... left; intuition omega.
+Qed.
+
+Definition do_assign_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: int) (v: val): option (world * trace * mem) :=
+ match access_mode ty with
+ | By_value chunk =>
+ match type_is_volatile ty with
+ | false => do m' <- Mem.storev chunk m (Vptr b ofs) v; Some(w, E0, m')
+ | true => do_volatile_store w chunk m b ofs v
+ end
+ | By_copy =>
+ match v with
+ | Vptr b' ofs' =>
+ if check_assign_copy ty b ofs b' ofs' then
+ do bytes <- Mem.loadbytes m b' (Int.unsigned ofs') (sizeof ty);
+ do m' <- Mem.storebytes m b (Int.unsigned ofs) bytes;
+ Some(w, E0, m')
+ else None
+ | _ => None
+ end
+ | _ => None
+ end.
+
+Lemma do_deref_loc_sound:
+ forall w ty m b ofs w' t v,
+ do_deref_loc w ty m b ofs = Some(w', t, v) ->
+ 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) as []_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.
+Qed.
+
+Lemma do_deref_loc_complete:
+ forall w ty m b ofs w' t v,
+ deref_loc ge ty m b ofs t v -> possible_trace w t w' ->
+ do_deref_loc w ty m b ofs = Some(w', t, v).
+Proof.
+ unfold do_deref_loc; intros. inv H.
+ inv H0. rewrite H1; rewrite H2; rewrite H3; auto.
+ rewrite H1; rewrite H2. apply do_volatile_load_complete; auto.
+ inv H0. rewrite H1. auto.
+ inv H0. rewrite H1. auto.
+Qed.
+
+Lemma do_assign_loc_sound:
+ forall w ty m b ofs v w' t m',
+ do_assign_loc w ty m b ofs v = Some(w', t, m') ->
+ 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) as []_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]].
+ split. eapply assign_loc_copy; eauto. constructor.
+Qed.
+
+Lemma do_assign_loc_complete:
+ forall w ty m b ofs v w' t m',
+ assign_loc ge ty m b ofs v t m' -> possible_trace w t w' ->
+ do_assign_loc w ty m b ofs v = Some(w', t, m').
+Proof.
+ unfold do_assign_loc; intros. inv H.
+ inv H0. rewrite H1; rewrite H2; rewrite H3; auto.
+ 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.
+Qed.
+
+(** System calls and library functions *)
+
+Definition do_ef_external (name: ident) (sg: signature)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ do args <- list_eventval_of_val vargs (sig_args sg);
+ match nextworld_io w name args with
+ | None => None
+ | Some(res, w') =>
+ do vres <- val_of_eventval res (proj_sig_res sg);
+ Some(w', Event_syscall name args res :: E0, vres, m)
+ end.
+
+Definition do_ef_volatile_load (chunk: memory_chunk)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ match vargs with
+ | Vptr b ofs :: nil => do w',t,v <- do_volatile_load w chunk m b ofs; Some(w',t,v,m)
+ | _ => None
+ end.
+
+Definition do_ef_volatile_store (chunk: memory_chunk)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ match vargs with
+ | Vptr b ofs :: v :: nil => do w',t,m' <- do_volatile_store w chunk m b ofs v; Some(w',t,Vundef,m')
+ | _ => None
+ end.
+
+Definition do_ef_volatile_load_global (chunk: memory_chunk) (id: ident) (ofs: int)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ do b <- Genv.find_symbol ge id; do_ef_volatile_load chunk w (Vptr b ofs :: vargs) m.
+
+Definition do_ef_volatile_store_global (chunk: memory_chunk) (id: ident) (ofs: int)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ do b <- Genv.find_symbol ge id; do_ef_volatile_store chunk w (Vptr b ofs :: vargs) m.
+
+Definition do_ef_malloc
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ match vargs with
+ | Vint n :: nil =>
+ let (m', b) := Mem.alloc m (-4) (Int.unsigned n) in
+ do m'' <- Mem.store Mint32 m' b (-4) (Vint n);
+ Some(w, E0, Vptr b Int.zero, m'')
+ | _ => None
+ end.
+
+Definition do_ef_free
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ match vargs with
+ | Vptr b lo :: nil =>
+ do vsz <- Mem.load Mint32 m b (Int.unsigned lo - 4);
+ match vsz with
+ | Vint sz =>
+ check (zlt 0 (Int.unsigned sz));
+ do m' <- Mem.free m b (Int.unsigned lo - 4) (Int.unsigned lo + Int.unsigned sz);
+ Some(w, E0, Vundef, m')
+ | _ => None
+ end
+ | _ => None
+ end.
+
+Definition memcpy_args_ok
+ (sz al: Z) (bdst: block) (odst: Z) (bsrc: block) (osrc: Z) : Prop :=
+ (al = 1 \/ al = 2 \/ al = 4)
+ /\ sz > 0
+ /\ (al | sz) /\ (al | osrc) /\ (al | odst)
+ /\ (bsrc <> bdst \/ osrc = odst \/ osrc + sz <= odst \/ odst + sz <= osrc).
+
+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.
+ assert (X: {al = 1 \/ al = 2 \/ al = 4} + {~(al = 1 \/ al = 2 \/ al = 4)}).
+ destruct (zeq al 1); auto. destruct (zeq al 2); auto. destruct (zeq al 4); auto...
+ unfold memcpy_args_ok. destruct X...
+ assert (al > 0) by (intuition omega).
+ destruct (zlt 0 sz)...
+ destruct (Zdivide_dec al sz); auto...
+ destruct (Zdivide_dec al osrc); auto...
+ destruct (Zdivide_dec al odst); auto...
+ assert (Y: {bsrc <> bdst \/ osrc = odst \/ osrc + sz <= odst \/ odst + sz <= osrc}
+ +{~(bsrc <> bdst \/ osrc = odst \/ osrc + sz <= odst \/ odst + sz <= osrc)}).
+ destruct (eq_block bsrc bdst); auto.
+ destruct (zeq osrc odst); auto.
+ destruct (zle (osrc + sz) odst); auto.
+ destruct (zle (odst + sz) osrc); auto.
+ right; intuition omega.
+ destruct Y... left; intuition omega.
+Qed.
+
+Definition do_ef_memcpy (sz al: Z)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ match vargs with
+ | Vptr bdst odst :: Vptr bsrc osrc :: nil =>
+ if memcpy_check_args sz al bdst (Int.unsigned odst) bsrc (Int.unsigned osrc) then
+ do bytes <- Mem.loadbytes m bsrc (Int.unsigned osrc) sz;
+ do m' <- Mem.storebytes m bdst (Int.unsigned odst) bytes;
+ Some(w, E0, Vundef, m')
+ else None
+ | _ => None
+ end.
+
+Definition do_ef_annot (text: ident) (targs: list typ)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ do args <- list_eventval_of_val vargs targs;
+ Some(w, Event_annot text args :: E0, Vundef, m).
+
+Definition do_ef_annot_val (text: ident) (targ: typ)
+ (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
+ match vargs with
+ | varg :: nil =>
+ do arg <- eventval_of_val varg targ;
+ Some(w, Event_annot text (arg :: nil) :: E0, varg, m)
+ | _ => None
+ end.
+
+Definition do_external (ef: external_function):
+ world -> list val -> mem -> option (world * trace * val * mem) :=
+ match ef with
+ | EF_external name sg => do_ef_external name sg
+ | EF_builtin name sg => do_ef_external name sg
+ | EF_vload chunk => do_ef_volatile_load chunk
+ | EF_vstore chunk => do_ef_volatile_store chunk
+ | EF_vload_global chunk id ofs => do_ef_volatile_load_global chunk id ofs
+ | EF_vstore_global chunk id ofs => do_ef_volatile_store_global chunk id ofs
+ | EF_malloc => do_ef_malloc
+ | EF_free => do_ef_free
+ | EF_memcpy sz al => do_ef_memcpy sz al
+ | EF_annot text targs => do_ef_annot text targs
+ | EF_annot_val text targ => do_ef_annot_val text targ
+ end.
+
+Lemma do_ef_external_sound:
+ forall ef w vargs m w' t vres m',
+ do_external ef w vargs m = Some(w', t, vres, m') ->
+ external_call ef ge vargs m t vres m' /\ possible_trace w t w'.
+Proof with try congruence.
+ intros until m'.
+ assert (IO: forall name sg,
+ do_ef_external name sg w vargs m = Some(w', t, vres, m') ->
+ extcall_io_sem name sg ge vargs m t vres m' /\ possible_trace w t w').
+ intros until sg. unfold do_ef_external. mydestr. destruct p as [res w'']; mydestr.
+ split. econstructor. apply list_eventval_of_val_sound; auto.
+ apply val_of_eventval_sound; auto.
+ econstructor. constructor; eauto. constructor.
+
+ assert (VLOAD: forall chunk vargs,
+ do_ef_volatile_load chunk w vargs m = Some (w', t, vres, m') ->
+ volatile_load_sem chunk ge vargs m t vres m' /\ possible_trace w t w').
+ intros chunk 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.
+
+ assert (VSTORE: forall chunk vargs,
+ do_ef_volatile_store chunk w vargs m = Some (w', t, vres, m') ->
+ volatile_store_sem chunk ge vargs m t vres m' /\ possible_trace w t w').
+ intros chunk vargs'.
+ 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.
+
+ destruct ef; simpl.
+(* EF_external *)
+ auto.
+(* EF_builtin *)
+ auto.
+(* EF_vload *)
+ auto.
+(* EF_vstore *)
+ auto.
+(* EF_vload_global *)
+ rewrite volatile_load_global_charact.
+ unfold do_ef_volatile_load_global. destruct (Genv.find_symbol ge)...
+ intros. exploit VLOAD; eauto. intros [A B]. split; auto. exists b; auto.
+(* EF_vstore_global *)
+ rewrite volatile_store_global_charact.
+ unfold do_ef_volatile_store_global. destruct (Genv.find_symbol ge)...
+ intros. exploit VSTORE; eauto. intros [A B]. split; auto. exists b; auto.
+(* EF_malloc *)
+ unfold do_ef_malloc. destruct vargs... destruct v... destruct vargs...
+ 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.
+ 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.
+ split. econstructor; eauto; tauto. constructor.
+(* EF_annot *)
+ 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.
+ split. constructor. apply eventval_of_val_sound; auto.
+ econstructor. constructor; eauto. constructor.
+Qed.
+
+Lemma do_ef_external_complete:
+ forall ef w vargs m w' t vres m',
+ external_call ef ge vargs m t vres m' -> possible_trace w t w' ->
+ do_external ef w vargs m = Some(w', t, vres, m').
+Proof.
+ intros.
+ assert (IO: forall name sg,
+ extcall_io_sem name sg ge vargs m t vres m' ->
+ do_ef_external name sg w vargs m = Some (w', t, vres, m')).
+ intros. inv H1. inv H0. inv H8. inv H6.
+ unfold do_ef_external. rewrite (list_eventval_of_val_complete _ _ _ H2). rewrite H8.
+ rewrite (val_of_eventval_complete _ _ _ H3). auto.
+
+ assert (VLOAD: forall chunk vargs,
+ volatile_load_sem chunk ge vargs m t vres m' ->
+ do_ef_volatile_load chunk w vargs m = Some (w', t, vres, m')).
+ intros. inv H1; unfold do_ef_volatile_load.
+ exploit do_volatile_load_complete; eauto. intros EQ; rewrite EQ; auto.
+
+ assert (VSTORE: forall chunk vargs,
+ volatile_store_sem chunk ge vargs m t vres m' ->
+ do_ef_volatile_store chunk w vargs m = Some (w', t, vres, m')).
+ intros. inv H1; unfold do_ef_volatile_store.
+ exploit do_volatile_store_complete; eauto. intros EQ; rewrite EQ; auto.
+
+ destruct ef; simpl in *.
+(* EF_external *)
+ auto.
+(* EF_builtin *)
+ auto.
+(* EF_vload *)
+ auto.
+(* EF_vstore *)
+ auto.
+(* EF_vload_global *)
+ rewrite volatile_load_global_charact in H. destruct H as [b [P Q]].
+ unfold do_ef_volatile_load_global. rewrite P. auto.
+(* EF_vstore *)
+ rewrite volatile_store_global_charact in H. destruct H as [b [P Q]].
+ unfold do_ef_volatile_store_global. rewrite P. auto.
+(* EF_malloc *)
+ inv H; unfold do_ef_malloc.
+ inv H0. rewrite H1. rewrite H2. auto.
+(* EF_free *)
+ inv H; unfold do_ef_free.
+ inv H0. rewrite H1. rewrite zlt_true. rewrite H3. auto. omega.
+(* EF_memcpy *)
+ inv H; unfold do_ef_memcpy.
+ inv H0. rewrite pred_dec_true. rewrite H7; rewrite H8; auto.
+ red. tauto.
+(* EF_annot *)
+ 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.
+ rewrite (eventval_of_val_complete _ _ _ H1). auto.
+Qed.
+
+(** * Reduction of expressions *)
+
Inductive reduction: Type :=
| Lred (l': expr) (m': mem)
- | Rred (r': expr) (m': mem)
- | Callred (fd: fundef) (args: list val) (tyres: type) (m': mem).
+ | Rred (r': expr) (m': mem) (t: trace)
+ | Callred (fd: fundef) (args: list val) (tyres: type) (m': mem)
+ | Stuckred.
Section EXPRS.
Variable e: env.
+Variable w: world.
Fixpoint sem_cast_arguments (vtl: list (val * type)) (tl: typelist) : option (list val) :=
match vtl, tl with
@@ -122,41 +653,51 @@ Fixpoint sem_cast_arguments (vtl: list (val * type)) (tl: typelist) : option (li
| _, _ => None
end.
-(** The result of stepping an expression can be
-- [None] denoting that the expression is stuck;
-- [Some nil] meaning that the expression is fully reduced
- (it's [Eval] for a r-value and [Eloc] for a l-value);
-- [Some ll] meaning that the expression can reduce to any of
- the elements of [ll]. Each element is a pair of a context
- and a reduction inside this context (see type [reduction] above).
+(** The result of stepping an expression is a list [ll] of possible reducts.
+ Each element of [ll] is a pair of a context and the result of reducing
+ inside this context (see type [reduction] above).
+ The list [ll] is empty if the expression is fully reduced
+ (it's [Eval] for a r-value and [Eloc] for a l-value).
*)
-Definition reducts (A: Type): Type := option (list ((expr -> A) * reduction)).
+Definition reducts (A: Type): Type := list ((expr -> A) * reduction).
Definition topred (r: reduction) : reducts expr :=
- Some (((fun (x: expr) => x), r) :: nil).
+ ((fun (x: expr) => x), r) :: nil.
-Definition incontext {A B: Type} (ctx: A -> B) (r: reducts A) : reducts B :=
- match r with
- | None => None
- | Some l => Some (map (fun z => ((fun (x: expr) => ctx(fst z x)), snd z)) l)
- end.
+Definition stuck : reducts expr :=
+ ((fun (x: expr) => x), Stuckred) :: nil.
+
+Definition incontext {A B: Type} (ctx: A -> B) (ll: reducts A) : reducts B :=
+ map (fun z => ((fun (x: expr) => ctx(fst z x)), snd z)) ll.
Definition incontext2 {A1 A2 B: Type}
- (ctx1: A1 -> B) (r1: reducts A1)
- (ctx2: A2 -> B) (r2: reducts A2) : reducts B :=
- match r1, r2 with
- | None, _ => None
- | _, None => None
- | Some l1, Some l2 =>
- Some (map (fun z => ((fun (x: expr) => ctx1(fst z x)), snd z)) l1
- ++ map (fun z => ((fun (x: expr) => ctx2(fst z x)), snd z)) l2)
- end.
+ (ctx1: A1 -> B) (ll1: reducts A1)
+ (ctx2: A2 -> B) (ll2: reducts A2) : reducts B :=
+ incontext ctx1 ll1 ++ incontext ctx2 ll2.
+
+Notation "'do' X <- A ; B" := (match A with Some X => B | None => stuck end)
+ (at level 200, X ident, A at level 100, B at level 200)
+ : reducts_monad_scope.
+
+Notation "'do' X , Y <- A ; B" := (match A with Some (X, Y) => B | None => stuck end)
+ (at level 200, X ident, Y ident, A at level 100, B at level 200)
+ : reducts_monad_scope.
+
+Notation "'do' X , Y , Z <- A ; B" := (match A with Some (X, Y, Z) => B | None => stuck end)
+ (at level 200, X ident, Y ident, Z ident, A at level 100, B at level 200)
+ : reducts_monad_scope.
+
+Notation " 'check' A ; B" := (if A then B else stuck)
+ (at level 200, A at level 100, B at level 200)
+ : reducts_monad_scope.
+
+Local Open Scope reducts_monad_scope.
Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match k, a with
| LV, Eloc b ofs ty =>
- Some nil
+ nil
| LV, Evar x ty =>
match e!x with
| Some(b, ty') =>
@@ -173,47 +714,49 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
| Some(Vptr b ofs, ty') =>
topred (Lred (Eloc b ofs ty) m)
| Some _ =>
- None
+ stuck
| None =>
incontext (fun x => Ederef x ty) (step_expr RV r m)
end
- | LV, Efield l f ty =>
- match is_loc l with
- | Some(b, ofs, ty') =>
+ | LV, Efield r f ty =>
+ match is_val r with
+ | Some(Vptr b ofs, ty') =>
match ty' with
- | Tstruct id fList =>
+ | Tstruct id fList _ =>
match field_offset f fList with
- | Error _ => None
+ | Error _ => stuck
| OK delta => topred (Lred (Eloc b (Int.add ofs (Int.repr delta)) ty) m)
end
- | Tunion id fList =>
+ | Tunion id fList _ =>
topred (Lred (Eloc b ofs ty) m)
- | _ => None
+ | _ => stuck
end
+ | Some _ =>
+ stuck
| None =>
- incontext (fun x => Efield x f ty) (step_expr LV l m)
+ incontext (fun x => Efield x f ty) (step_expr RV r m)
end
| RV, Eval v ty =>
- Some nil
+ nil
| RV, Evalof l ty =>
match is_loc l with
| Some(b, ofs, ty') =>
check type_eq ty ty';
- do v <- load_value_of_type ty m b ofs;
- topred (Rred (Eval v ty) m)
+ do w',t,v <- do_deref_loc w ty m b ofs;
+ topred (Rred (Eval v ty) m t)
| None =>
incontext (fun x => Evalof x ty) (step_expr LV l m)
end
| RV, Eaddrof l ty =>
match is_loc l with
- | Some(b, ofs, ty') => topred (Rred (Eval (Vptr b ofs) ty) m)
+ | Some(b, ofs, ty') => topred (Rred (Eval (Vptr b ofs) ty) m E0)
| None => incontext (fun x => Eaddrof x ty) (step_expr LV l m)
end
| RV, Eunop op r1 ty =>
match is_val r1 with
| Some(v1, ty1) =>
do v <- sem_unary_operation op v1 ty1;
- topred (Rred (Eval v ty) m)
+ topred (Rred (Eval v ty) m E0)
| None =>
incontext (fun x => Eunop op x ty) (step_expr RV r1 m)
end
@@ -221,7 +764,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match is_val r1, is_val r2 with
| Some(v1, ty1), Some(v2, ty2) =>
do v <- sem_binary_operation op v1 ty1 v2 ty2 m;
- topred (Rred (Eval v ty) m)
+ topred (Rred (Eval v ty) m E0)
| _, _ =>
incontext2 (fun x => Ebinop op x r2 ty) (step_expr RV r1 m)
(fun x => Ebinop op r1 x ty) (step_expr RV r2 m)
@@ -230,7 +773,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match is_val r1 with
| Some(v1, ty1) =>
do v <- sem_cast v1 ty1 ty;
- topred (Rred (Eval v ty) m)
+ topred (Rred (Eval v ty) m E0)
| None =>
incontext (fun x => Ecast x ty) (step_expr RV r1 m)
end
@@ -238,19 +781,19 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match is_val r1 with
| Some(v1, ty1) =>
do b <- bool_val v1 ty1;
- topred (Rred (Eparen (if b then r2 else r3) ty) m)
+ topred (Rred (Eparen (if b then r2 else r3) ty) m E0)
| None =>
incontext (fun x => Econdition x r2 r3 ty) (step_expr RV r1 m)
end
| RV, Esizeof ty' ty =>
- topred (Rred (Eval (Vint (Int.repr (sizeof ty'))) ty) m)
+ topred (Rred (Eval (Vint (Int.repr (sizeof ty'))) ty) m E0)
| RV, Eassign l1 r2 ty =>
match is_loc l1, is_val r2 with
| Some(b, ofs, ty1), Some(v2, ty2) =>
check type_eq ty1 ty;
do v <- sem_cast v2 ty2 ty1;
- do m' <- store_value_of_type ty1 m b ofs v;
- topred (Rred (Eval v ty) m')
+ do w',t,m' <- do_assign_loc w ty1 m b ofs v;
+ topred (Rred (Eval v ty) m' t)
| _, _ =>
incontext2 (fun x => Eassign x r2 ty) (step_expr LV l1 m)
(fun x => Eassign l1 x ty) (step_expr RV r2 m)
@@ -259,11 +802,10 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match is_loc l1, is_val r2 with
| Some(b, ofs, ty1), Some(v2, ty2) =>
check type_eq ty1 ty;
- do v1 <- load_value_of_type ty1 m b ofs;
- do v <- sem_binary_operation op v1 ty1 v2 ty2 m;
- do v' <- sem_cast v tyres ty1;
- do m' <- store_value_of_type ty1 m b ofs v';
- topred (Rred (Eval v' ty) m')
+ do w',t,v1 <- do_deref_loc w ty1 m b ofs;
+ let r' := Eassign (Eloc b ofs ty1)
+ (Ebinop op (Eval v1 ty1) (Eval v2 ty2) tyres) ty1 in
+ topred (Rred r' m t)
| _, _ =>
incontext2 (fun x => Eassignop op x r2 tyres ty) (step_expr LV l1 m)
(fun x => Eassignop op l1 x tyres ty) (step_expr RV r2 m)
@@ -272,11 +814,14 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match is_loc l with
| Some(b, ofs, ty1) =>
check type_eq ty1 ty;
- do v1 <- load_value_of_type ty m b ofs;
- do v2 <- sem_incrdecr id v1 ty;
- do v3 <- sem_cast v2 (typeconv ty) ty;
- do m' <- store_value_of_type ty m b ofs v3;
- topred (Rred (Eval v1 ty) m')
+ 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)
+ (Ebinop op (Eval v1 ty) (Eval (Vint Int.one) type_int32s) (typeconv ty))
+ ty)
+ (Eval v1 ty) ty in
+ topred (Rred r' m t)
| None =>
incontext (fun x => Epostincr id x ty) (step_expr LV l m)
end
@@ -284,7 +829,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match is_val r1 with
| Some _ =>
check type_eq (typeof r2) ty;
- topred (Rred r2 m)
+ topred (Rred r2 m E0)
| None =>
incontext (fun x => Ecomma x r2 ty) (step_expr RV r1 m)
end
@@ -292,7 +837,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match is_val r1 with
| Some (v1, ty1) =>
do v <- sem_cast v1 ty1 ty;
- topred (Rred (Eval v ty) m)
+ topred (Rred (Eval v ty) m E0)
| None =>
incontext (fun x => Eparen x ty) (step_expr RV r1 m)
end
@@ -305,24 +850,224 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
do vargs <- sem_cast_arguments vtl tyargs;
check type_eq (type_of_fundef fd) (Tfunction tyargs tyres);
topred (Callred fd vargs ty m)
- | _ => None
+ | _ => stuck
end
| _, _ =>
incontext2 (fun x => Ecall x rargs ty) (step_expr RV r1 m)
(fun x => Ecall r1 x ty) (step_exprlist rargs m)
end
- | _, _ => None
+ | _, _ => stuck
end
with step_exprlist (rl: exprlist) (m: mem): reducts exprlist :=
match rl with
| Enil =>
- Some nil
+ nil
| Econs r1 rs =>
incontext2 (fun x => Econs x rs) (step_expr RV r1 m)
(fun x => Econs r1 x) (step_exprlist rs m)
end.
+(** Technical properties on safe expressions. *)
+
+Inductive imm_safe_t: kind -> expr -> mem -> Prop :=
+ | imm_safe_t_val: forall v ty m,
+ imm_safe_t RV (Eval v ty) m
+ | imm_safe_t_loc: forall b ofs ty m,
+ imm_safe_t LV (Eloc b ofs ty) m
+ | imm_safe_t_lred: forall to C l m l' m',
+ lred ge e l m l' m' ->
+ context LV to C ->
+ imm_safe_t to (C l) m
+ | imm_safe_t_rred: forall to C r m t r' m' w',
+ rred ge r m t r' m' -> possible_trace w t w' ->
+ context RV to C ->
+ imm_safe_t to (C r) m
+ | imm_safe_t_callred: forall to C r m fd args ty,
+ callred ge r fd args ty ->
+ context RV to C ->
+ imm_safe_t to (C r) m.
+
+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.
+ constructor.
+ constructor.
+ eapply imm_safe_lred; eauto.
+ eapply imm_safe_rred; eauto.
+ eapply imm_safe_callred; eauto.
+Qed.
+
+(*
+Definition not_stuck (a: expr) (m: mem) :=
+ forall a' k C, context k RV C -> a = C a' -> imm_safe_t k a' m.
+
+Lemma safe_expr_kind:
+ forall k C a m,
+ context k RV C ->
+ not_stuck (C a) m ->
+ k = Cstrategy.expr_kind a.
+Proof.
+ intros.
+ symmetry. eapply Cstrategy.imm_safe_kind. eapply imm_safe_t_imm_safe. eauto.
+Qed.
+*)
+
+Fixpoint exprlist_all_values (rl: exprlist) : Prop :=
+ match rl with
+ | Enil => True
+ | Econs (Eval v ty) rl' => exprlist_all_values rl'
+ | Econs _ _ => False
+ end.
+
+Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
+ match a with
+ | Eloc b ofs ty => False
+ | Evar x ty =>
+ exists b,
+ e!x = Some(b, ty)
+ \/ (e!x = None /\ Genv.find_symbol ge x = Some b /\ type_of_global ge b = Some ty)
+ | Ederef (Eval v ty1) ty =>
+ exists b, exists ofs, v = Vptr b ofs
+ | Efield (Eval v ty1) f ty =>
+ exists b, exists ofs, v = Vptr b ofs /\
+ match ty1 with
+ | Tstruct _ fList _ => exists delta, field_offset f fList = Errors.OK delta
+ | Tunion _ _ _ => True
+ | _ => False
+ end
+ | Eval v ty => False
+ | Evalof (Eloc b ofs ty') ty =>
+ ty' = ty /\ exists t, exists v, exists w', deref_loc ge ty m b ofs t v /\ possible_trace w t w'
+ | Eunop op (Eval v1 ty1) ty =>
+ exists v, sem_unary_operation op v1 ty1 = Some v
+ | Ebinop op (Eval v1 ty1) (Eval v2 ty2) ty =>
+ exists v, sem_binary_operation op v1 ty1 v2 ty2 m = Some v
+ | Ecast (Eval v1 ty1) ty =>
+ exists v, sem_cast v1 ty1 ty = Some v
+ | Econdition (Eval v1 ty1) r1 r2 ty =>
+ exists b, bool_val v1 ty1 = Some b
+ | Eassign (Eloc b ofs ty1) (Eval v2 ty2) ty =>
+ exists v, exists m', exists t, exists w',
+ ty = ty1 /\ sem_cast v2 ty2 ty1 = Some v /\ assign_loc ge ty1 m b ofs v t m' /\ possible_trace w t w'
+ | Eassignop op (Eloc b ofs ty1) (Eval v2 ty2) tyres ty =>
+ 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',
+ 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
+ | Eparen (Eval v1 ty1) ty =>
+ exists v, sem_cast v1 ty1 ty = Some v
+ | Ecall (Eval vf tyf) rargs ty =>
+ exprlist_all_values rargs ->
+ exists tyargs, exists tyres, exists fd, exists vl,
+ classify_fun tyf = fun_case_f tyargs tyres
+ /\ Genv.find_funct ge vf = Some fd
+ /\ cast_arguments rargs tyargs vl
+ /\ type_of_fundef fd = Tfunction tyargs tyres
+ | _ => True
+ end.
+
+Lemma lred_invert:
+ forall l m l' m', lred ge e l m l' m' -> invert_expr_prop l m.
+Proof.
+ induction 1; red; auto.
+ exists b; auto.
+ exists b; auto.
+ exists b; exists ofs; auto.
+ exists b; exists ofs; split; auto. exists delta; auto.
+ exists b; exists ofs; auto.
+Qed.
+
+Lemma rred_invert:
+ forall w' r m t r' m', rred ge r m t r' m' -> possible_trace w t w' -> invert_expr_prop r m.
+Proof.
+ induction 1; intros; red; auto.
+ split; auto; exists t; exists v; exists w'; auto.
+ exists v; auto.
+ exists v; auto.
+ exists v; auto.
+ exists b; auto.
+ exists v; exists m'; exists t; exists w'; auto.
+ exists t; exists v1; exists w'; auto.
+ exists t; exists v1; exists w'; auto.
+ exists v; auto.
+Qed.
+
+Lemma callred_invert:
+ forall r fd args ty m,
+ callred ge r fd args ty ->
+ invert_expr_prop r m.
+Proof.
+ intros. inv H. simpl.
+ intros. exists tyargs; exists tyres; exists fd; exists args; auto.
+Qed.
+
+Scheme context_ind2 := Minimality for context Sort Prop
+ with contextlist_ind2 := Minimality for contextlist Sort Prop.
+Combined Scheme context_contextlist_ind from context_ind2, contextlist_ind2.
+
+Lemma invert_expr_context:
+ (forall from to C, context from to C ->
+ forall a m,
+ invert_expr_prop a m ->
+ invert_expr_prop (C a) m)
+/\(forall from C, contextlist from C ->
+ forall a m,
+ invert_expr_prop a m ->
+ ~exprlist_all_values (C a)).
+Proof.
+ apply context_contextlist_ind; intros; try (exploit H0; [eauto|intros]); simpl.
+ auto.
+ destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ auto.
+ destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ destruct e1; auto; destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ destruct e1; auto; destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ destruct e1; auto; destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ destruct (C a); auto; contradiction.
+ destruct e1; 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.
+Qed.
+
+Lemma imm_safe_t_inv:
+ forall k a m,
+ imm_safe_t k a m ->
+ match a with
+ | Eloc _ _ _ => True
+ | Eval _ _ => True
+ | _ => invert_expr_prop a m
+ end.
+Proof.
+ destruct invert_expr_context as [A B].
+ intros. inv H.
+ auto.
+ auto.
+ assert (invert_expr_prop (C l) m).
+ eapply A; eauto. eapply lred_invert; eauto.
+ red in H. destruct (C l); auto; contradiction.
+ assert (invert_expr_prop (C r) m).
+ eapply A; eauto. eapply rred_invert; eauto.
+ red in H. destruct (C r); auto; contradiction.
+ assert (invert_expr_prop (C r) m).
+ eapply A; eauto. eapply callred_invert; eauto.
+ red in H. destruct (C r); auto; contradiction.
+Qed.
+
(** Soundness: if [step_expr] returns [Some ll], then every element
of [ll] is a reduct. *)
@@ -343,19 +1088,27 @@ Qed.
Hint Constructors context contextlist.
Hint Resolve context_compose contextlist_compose.
-Definition reduction_ok (a: expr) (m: mem) (rd: reduction) : Prop :=
- match rd with
- | Lred l' m' => lred ge e a m l' m'
- | Rred r' m' => rred a m r' m'
- | Callred fd args tyres m' => callred ge a fd args tyres /\ m' = m
+Definition reduction_ok (k: kind) (a: expr) (m: mem) (rd: reduction) : Prop :=
+ match k, rd with
+ | LV, Lred l' m' => lred ge e a m l' m'
+ | RV, Rred r' m' t => rred ge a m t r' m' /\ exists w', possible_trace w t w'
+ | RV, Callred fd args tyres m' => callred ge a fd args tyres /\ m' = m
+ | LV, Stuckred => ~imm_safe_t k a m
+ | RV, Stuckred => ~imm_safe_t k a m
+ | _, _ => False
end.
-Definition reduction_kind (rd: reduction): kind :=
- match rd with
- | Lred l' m' => LV
- | Rred r' m' => RV
- | Callred fd args tyres m' => RV
- end.
+Definition reducts_ok (k: kind) (a: expr) (m: mem) (ll: reducts expr) : Prop :=
+ (forall C rd,
+ In (C, rd) ll ->
+ exists a', exists k', context k' k C /\ a = C a' /\ reduction_ok k' a' m rd)
+ /\ (ll = nil -> match k with LV => is_loc a <> None | RV => is_val a <> None end).
+
+Definition list_reducts_ok (al: exprlist) (m: mem) (ll: reducts exprlist) : Prop :=
+ (forall C rd,
+ In (C, rd) ll ->
+ exists a', exists k', contextlist k' C /\ al = C a' /\ reduction_ok k' a' m rd)
+ /\ (ll = nil -> is_val_list al <> None).
Ltac monadInv :=
match goal with
@@ -378,36 +1131,59 @@ Proof.
inv H0. rewrite (is_val_inv _ _ _ Heqo). constructor. auto. eauto.
Qed.
-Definition reducts_ok (k: kind) (a: expr) (m: mem) (res: reducts expr) : Prop :=
- match res with
- | None => True
- | Some nil => match k with LV => is_loc a <> None | RV => is_val a <> None end
- | Some ll =>
- forall C rd,
- In (C, rd) ll ->
- context (reduction_kind rd) k C /\ exists a', a = C a' /\ reduction_ok a' m rd
- end.
-
-Definition list_reducts_ok (al: exprlist) (m: mem) (res: reducts exprlist) : Prop :=
- match res with
- | None => True
- | Some nil => is_val_list al <> None
- | Some ll =>
- forall C rd,
- In (C, rd) ll ->
- contextlist (reduction_kind rd) C /\ exists a', al = C a' /\ reduction_ok a' m rd
- end.
+Lemma sem_cast_arguments_complete:
+ forall al tyl vl,
+ cast_arguments al tyl vl ->
+ exists vtl, is_val_list al = Some vtl /\ sem_cast_arguments vtl tyl = Some vl.
+Proof.
+ induction 1.
+ exists (@nil (val * type)); auto.
+ destruct IHcast_arguments as [vtl [A B]].
+ exists ((v, ty) :: vtl); simpl. rewrite A; rewrite B; rewrite H. auto.
+Qed.
Lemma topred_ok:
forall k a m rd,
- reduction_ok a m rd ->
- k = reduction_kind rd ->
+ reduction_ok k a m rd ->
reducts_ok k a m (topred rd).
Proof.
- intros. unfold topred; red. simpl; intros. destruct H1; try contradiction.
- inv H1. split. auto. exists a; auto.
+ intros. unfold topred; split; simpl; intros.
+ destruct H0; try contradiction. inv H0. exists a; exists k; auto.
+ congruence.
+Qed.
+
+Lemma stuck_ok:
+ forall k a m,
+ ~imm_safe_t k a m ->
+ reducts_ok k a m stuck.
+Proof.
+ intros. unfold stuck; split; simpl; intros.
+ destruct H0; try contradiction. inv H0. exists a; exists k; intuition. red. destruct k; auto.
+ congruence.
Qed.
+Lemma wrong_kind_ok:
+ forall k a m,
+ k <> Cstrategy.expr_kind a ->
+ reducts_ok k a m stuck.
+Proof.
+ intros. apply stuck_ok. red; intros. exploit Cstrategy.imm_safe_kind; eauto.
+ eapply imm_safe_t_imm_safe; eauto.
+Qed.
+
+Lemma not_invert_ok:
+ forall k a m,
+ match a with
+ | Eloc _ _ _ => False
+ | Eval _ _ => False
+ | _ => invert_expr_prop a m -> False
+ end ->
+ reducts_ok k a m stuck.
+Proof.
+ 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',
reducts_ok k' a' m res ->
@@ -416,48 +1192,11 @@ Lemma incontext_ok:
match k' with LV => is_loc a' = None | RV => is_val a' = None end ->
reducts_ok k a m (incontext C res).
Proof.
- unfold reducts_ok; intros. destruct res; simpl. destruct l.
-(* res = Some nil *)
- destruct k'; congruence.
-(* res = Some nonempty-list *)
- simpl map at 1. hnf. intros.
+ unfold reducts_ok, incontext; intros. destruct H. split; intros.
exploit list_in_map_inv; eauto. intros [[C1 rd1] [P Q]]. inv P.
- exploit H; eauto. intros [U [a'' [V W]]].
- split. eapply context_compose; eauto. exists a''; split; auto. congruence.
-(* res = None *)
- auto.
-Qed.
-
-Remark incontext2_inv:
- forall {A1 A2 B: Type} (C1: A1 -> B) res1 (C2: A2 -> B) res2,
- match incontext2 C1 res1 C2 res2 with
- | None => res1 = None \/ res2 = None
- | Some nil => res1 = Some nil /\ res2 = Some nil
- | Some ll =>
- exists ll1, exists ll2,
- res1 = Some ll1 /\ res2 = Some ll2 /\
- forall C rd, In (C, rd) ll ->
- (exists C', C = (fun x => C1(C' x)) /\ In (C', rd) ll1)
- \/ (exists C', C = (fun x => C2(C' x)) /\ In (C', rd) ll2)
- end.
-Proof.
- intros. unfold incontext2. destruct res1 as [ll1|]; auto. destruct res2 as [ll2|]; auto.
- set (ll := map
- (fun z : (expr -> A1) * reduction =>
- (fun x : expr => C1 (fst z x), snd z)) ll1 ++
- map
- (fun z : (expr -> A2) * reduction =>
- (fun x : expr => C2 (fst z x), snd z)) ll2).
- destruct ll as []_eqn.
- destruct (app_eq_nil _ _ Heql).
- split. destruct ll1; auto || discriminate. destruct ll2; auto || discriminate.
- rewrite <- Heql. exists ll1; exists ll2. split. auto. split. auto.
- unfold ll; intros.
- rewrite in_app in H. destruct H.
- exploit list_in_map_inv; eauto. intros [[C' rd'] [P Q]]. inv P.
- left; exists C'; auto.
- exploit list_in_map_inv; eauto. intros [[C' rd'] [P Q]]. inv P.
- right; exists C'; auto.
+ exploit H; eauto. intros [a'' [k'' [U [V W]]]].
+ exists a''; exists k''. split. eapply context_compose; eauto. rewrite V; auto.
+ destruct res; simpl in H4; try congruence. destruct k'; intuition congruence.
Qed.
Lemma incontext2_ok:
@@ -470,17 +1209,16 @@ Lemma incontext2_ok:
\/ match k2 with LV => is_loc a2 = None | RV => is_val a2 = None end ->
reducts_ok k a m (incontext2 C1 res1 C2 res2).
Proof.
- unfold reducts_ok; intros.
- generalize (incontext2_inv C1 res1 C2 res2).
- destruct (incontext2 C1 res1 C2 res2) as [ll|]; auto.
- destruct ll.
- intros [EQ1 EQ2]. subst. destruct H5. destruct k1; congruence. destruct k2; congruence.
- intros [ll1 [ll2 [EQ1 [EQ2 IN]]]]. subst. intros.
- exploit IN; eauto. intros [[C' [P Q]] | [C' [P Q]]]; subst.
- destruct ll1; try contradiction. exploit H; eauto.
- intros [U [a' [V W]]]. split. eauto. exists a'; split. congruence. auto.
- destruct ll2; try contradiction. exploit H0; eauto.
- intros [U [a' [V W]]]. split. eauto. exists a'; split. congruence. auto.
+ unfold reducts_ok, incontext2, incontext; intros. destruct H; destruct H0; split; intros.
+ destruct (in_app_or _ _ _ H8).
+ 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. eapply context_compose; 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. eapply context_compose; eauto. rewrite H2; rewrite V; auto.
+ destruct res1; simpl in H8; try congruence. destruct res2; simpl in H8; try congruence.
+ destruct H5. destruct k1; intuition congruence. destruct k2; intuition congruence.
Qed.
Lemma incontext2_list_ok:
@@ -492,18 +1230,16 @@ Lemma incontext2_list_ok:
(incontext2 (fun x => Ecall x a2 ty) res1
(fun x => Ecall a1 x ty) res2).
Proof.
- unfold reducts_ok, list_reducts_ok; intros.
- set (C1 := fun x => Ecall x a2 ty). set (C2 := fun x => Ecall a1 x ty).
- generalize (incontext2_inv C1 res1 C2 res2).
- destruct (incontext2 C1 res1 C2 res2) as [ll|]; auto.
- destruct ll.
- intros [EQ1 EQ2]. subst. intuition congruence.
- intros [ll1 [ll2 [EQ1 [EQ2 IN]]]]. subst. intros.
- exploit IN; eauto. intros [[C' [P Q]] | [C' [P Q]]]; subst.
- destruct ll1; try contradiction. exploit H; eauto.
- intros [U [a' [V W]]]. split. unfold C1. auto. exists a'; split. unfold C1; congruence. auto.
- destruct ll2; try contradiction. exploit H0; eauto.
- intros [U [a' [V W]]]. split. unfold C2. auto. exists a'; split. unfold C2; congruence. auto.
+ unfold reducts_ok, incontext2, incontext; intros. destruct H; destruct H0; split; intros.
+ 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.
+ 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.
+ destruct res1; simpl in H4; try congruence. destruct res2; simpl in H4; try congruence.
+ tauto.
Qed.
Lemma incontext2_list_ok':
@@ -514,165 +1250,201 @@ Lemma incontext2_list_ok':
(incontext2 (fun x => Econs x a2) res1
(fun x => Econs a1 x) res2).
Proof.
- unfold reducts_ok, list_reducts_ok; intros.
- set (C1 := fun x => Econs x a2). set (C2 := fun x => Econs a1 x).
- generalize (incontext2_inv C1 res1 C2 res2).
- destruct (incontext2 C1 res1 C2 res2) as [ll|]; auto.
- destruct ll.
- intros [EQ1 EQ2]. subst.
- simpl. destruct (is_val a1); try congruence. destruct (is_val_list a2); congruence.
- intros [ll1 [ll2 [EQ1 [EQ2 IN]]]]. subst. intros.
- exploit IN; eauto. intros [[C' [P Q]] | [C' [P Q]]]; subst.
- destruct ll1; try contradiction. exploit H; eauto.
- intros [U [a' [V W]]]. split. unfold C1. auto. exists a'; split. unfold C1; congruence. auto.
- destruct ll2; try contradiction. exploit H0; eauto.
- intros [U [a' [V W]]]. split. unfold C2. auto. exists a'; split. unfold C2; congruence. auto.
-Qed.
-
-Ltac mysimpl :=
+ unfold reducts_ok, list_reducts_ok, incontext2, incontext; intros.
+ destruct H; destruct H0. split; intros.
+ 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.
+ 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.
+ destruct res1; simpl in H3; try congruence. destruct res2; simpl in H3; try congruence.
+ simpl. destruct (is_val a1). destruct (is_val_list a2). congruence. intuition congruence. intuition congruence.
+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.
+ 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.
+Qed.
+
+Ltac myinv :=
match goal with
- | [ |- reducts_ok _ _ _ (match ?x with Some _ => _ | None => None end) ] =>
- destruct x as []_eqn; [mysimpl|exact I]
- | [ |- reducts_ok _ _ _ (match ?x with left _ => _ | right _ => None end) ] =>
- destruct x as []_eqn; [subst;mysimpl|exact I]
- | _ =>
- idtac
+ | [ H: False |- _ ] => destruct H
+ | [ H: _ /\ _ |- _ ] => destruct H; myinv
+ | [ H: exists _, _ |- _ ] => destruct H; myinv
+ | _ => idtac
end.
Theorem step_expr_sound:
forall a k m, reducts_ok k a m (step_expr k a m)
with step_exprlist_sound:
forall al m, list_reducts_ok al m (step_exprlist al m).
-Proof with try (exact I).
- induction a; destruct k; intros; simpl...
+Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence; fail)).
+ induction a; intros; simpl; destruct k; try (apply wrong_kind_ok; simpl; congruence).
(* Eval *)
- congruence.
+ split; intros. tauto. simpl; congruence.
(* Evar *)
- destruct (e!x) as [[b ty'] | ]_eqn; mysimpl.
- apply topred_ok; auto. apply red_var_local; auto.
- apply topred_ok; auto. apply red_var_global; auto.
+ destruct (e!x) as [[b ty']|]_eqn.
+ destruct (type_eq ty ty')...
+ subst. apply topred_ok; auto. apply red_var_local; auto.
+ destruct (Genv.find_symbol ge x) as [b|]_eqn...
+ destruct (type_of_global ge b) as [ty'|]_eqn...
+ destruct (type_eq ty ty')...
+ subst. apply topred_ok; auto. apply red_var_global; auto.
(* Efield *)
- destruct (is_loc a) as [[[b ofs] ty'] | ]_eqn.
+ destruct (is_val a) as [[v ty'] | ]_eqn.
+ rewrite (is_val_inv _ _ _ Heqo).
+ destruct v...
destruct ty'...
(* top struct *)
destruct (field_offset f f0) as [delta|]_eqn...
- rewrite (is_loc_inv _ _ _ _ Heqo). apply topred_ok; auto. apply red_field_struct; auto.
+ apply topred_ok; auto. apply red_field_struct; auto.
(* top union *)
- rewrite (is_loc_inv _ _ _ _ Heqo). apply topred_ok; auto. apply red_field_union; auto.
+ apply topred_ok; auto. apply red_field_union; auto.
(* in depth *)
eapply incontext_ok; eauto.
(* Evalof *)
- destruct (is_loc a) as [[[b ofs] ty'] | ]_eqn.
+ destruct (is_loc a) as [[[b ofs] ty'] | ]_eqn. rewrite (is_loc_inv _ _ _ _ Heqo).
(* top *)
- mysimpl. apply topred_ok; auto. rewrite (is_loc_inv _ _ _ _ Heqo). apply red_rvalof; auto.
+ destruct (type_eq ty ty')... subst ty'.
+ destruct (do_deref_loc w ty m b ofs) as [[[w' t] v] | ]_eqn.
+ exploit do_deref_loc_sound; eauto. intros [A B].
+ apply topred_ok; auto. red. split. apply red_rvalof; auto. exists w'; auto.
+ apply not_invert_ok; simpl; intros; myinv. exploit do_deref_loc_complete; eauto. congruence.
(* depth *)
eapply incontext_ok; eauto.
(* Ederef *)
- destruct (is_val a) as [[v ty'] | ]_eqn.
+ destruct (is_val a) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
- destruct v... mysimpl. apply topred_ok; auto. rewrite (is_val_inv _ _ _ Heqo). apply red_deref; auto.
+ destruct v... apply topred_ok; auto. apply red_deref; auto.
(* depth *)
eapply incontext_ok; eauto.
(* Eaddrof *)
- destruct (is_loc a) as [[[b ofs] ty'] | ]_eqn.
+ destruct (is_loc a) as [[[b ofs] ty'] | ]_eqn. rewrite (is_loc_inv _ _ _ _ Heqo).
(* top *)
- apply topred_ok; auto. rewrite (is_loc_inv _ _ _ _ Heqo). apply red_addrof; auto.
+ apply topred_ok; auto. split. apply red_addrof; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
(* unop *)
- destruct (is_val a) as [[v ty'] | ]_eqn.
+ destruct (is_val a) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
- mysimpl. apply topred_ok; auto. rewrite (is_val_inv _ _ _ Heqo). apply red_unop; auto.
+ destruct (sem_unary_operation op v ty') as [v'|]_eqn...
+ 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).
(* top *)
- mysimpl. apply topred_ok; auto.
- rewrite (is_val_inv _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0). apply red_binop; auto.
+ destruct (sem_binary_operation 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.
(* cast *)
- destruct (is_val a) as [[v ty'] | ]_eqn.
+ destruct (is_val a) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
- mysimpl. apply topred_ok; auto.
- rewrite (is_val_inv _ _ _ Heqo). apply red_cast; auto.
+ destruct (sem_cast v ty' ty) as [v'|]_eqn...
+ apply topred_ok; auto. split. apply red_cast; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
(* condition *)
- destruct (is_val a1) as [[v ty'] | ]_eqn.
+ destruct (is_val a1) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
- mysimpl. apply topred_ok; auto.
- rewrite (is_val_inv _ _ _ Heqo). eapply red_condition; eauto.
+ destruct (bool_val v ty') as [v'|]_eqn...
+ apply topred_ok; auto. split. eapply red_condition; eauto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
(* sizeof *)
- apply topred_ok; auto. apply red_sizeof.
+ apply topred_ok; auto. split. apply red_sizeof. 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 *)
- mysimpl. apply topred_ok; auto.
- rewrite (is_loc_inv _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0). apply red_assign; auto.
+ destruct (type_eq ty1 ty)... subst ty1.
+ destruct (sem_cast v2 ty2 ty) as [v|]_eqn...
+ destruct (do_assign_loc w ty m b ofs v) as [[[w' t] m']|]_eqn.
+ exploit do_assign_loc_sound; eauto. intros [P Q].
+ apply topred_ok; auto. split. apply red_assign; auto. exists w'; auto.
+ apply not_invert_ok; simpl; intros; myinv. exploit do_assign_loc_complete; eauto. congruence.
(* depth *)
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.
+ 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 *)
- mysimpl. apply topred_ok; auto.
- rewrite (is_loc_inv _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0). eapply red_assignop; eauto.
+ destruct (type_eq ty1 ty)... subst ty1.
+ destruct (do_deref_loc w ty m b ofs) as [[[w' t] v] | ]_eqn.
+ exploit do_deref_loc_sound; eauto. intros [A B].
+ apply topred_ok; auto. red. split. apply red_assignop; auto. exists w'; auto.
+ apply not_invert_ok; simpl; intros; myinv. exploit do_deref_loc_complete; eauto. congruence.
(* depth *)
eapply incontext2_ok; eauto.
eapply incontext2_ok; eauto.
(* postincr *)
- destruct (is_loc a) as [[[b ofs] ty'] | ]_eqn.
+ destruct (is_loc a) as [[[b ofs] ty'] | ]_eqn. rewrite (is_loc_inv _ _ _ _ Heqo).
(* top *)
- mysimpl. apply topred_ok; auto.
- rewrite (is_loc_inv _ _ _ _ Heqo). eapply red_postincr; eauto.
+ destruct (type_eq ty' ty)... subst ty'.
+ destruct (do_deref_loc w ty m b ofs) as [[[w' t] v] | ]_eqn.
+ exploit do_deref_loc_sound; eauto. intros [A B].
+ apply topred_ok; auto. red. split. apply red_postincr; auto. exists w'; auto.
+ apply not_invert_ok; simpl; intros; myinv. exploit do_deref_loc_complete; eauto. congruence.
(* depth *)
eapply incontext_ok; eauto.
(* comma *)
- destruct (is_val a1) as [[v ty'] | ]_eqn.
+ destruct (is_val a1) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
- mysimpl. apply topred_ok; auto.
- rewrite (is_val_inv _ _ _ Heqo). apply red_comma; auto.
+ 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|]_eqn...
- mysimpl. apply topred_ok; auto.
- rewrite (is_val_inv _ _ _ Heqo). red. split; auto. eapply red_Ecall; eauto.
- eapply sem_cast_arguments_sound; eauto.
+ destruct (Genv.find_funct ge vf) as [fd|]_eqn...
+ destruct (sem_cast_arguments vtl tyargs) as [vargs|]_eqn...
+ destruct (type_eq (type_of_fundef fd) (Tfunction tyargs tyres))...
+ apply topred_ok; auto. red. split; auto. eapply red_Ecall; 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.
+ exploit sem_cast_arguments_complete; eauto. intros [vtl' [P Q]]. congruence.
+ apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv. congruence.
+ apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv. congruence.
(* depth *)
eapply incontext2_list_ok; eauto.
eapply incontext2_list_ok; eauto.
(* loc *)
- congruence.
+ split; intros. tauto. simpl; congruence.
(* paren *)
- destruct (is_val a) as [[v ty'] | ]_eqn.
+ destruct (is_val a) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
- mysimpl. apply topred_ok; auto.
- rewrite (is_val_inv _ _ _ Heqo). apply red_paren; auto.
+ destruct (sem_cast v ty' ty) as [v'|]_eqn...
+ apply topred_ok; auto. split. apply red_paren; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
induction al; simpl; intros.
(* nil *)
- congruence.
+ split; intros. tauto. simpl; congruence.
(* cons *)
eapply incontext2_list_ok'; eauto.
Qed.
-
Lemma step_exprlist_val_list:
- forall m al, is_val_list al <> None -> step_exprlist al m = Some nil.
+ forall m al, is_val_list al <> None -> step_exprlist al m = nil.
Proof.
induction al; simpl; intros.
auto.
@@ -682,8 +1454,7 @@ Proof.
rewrite IHal. auto. congruence.
Qed.
-(** Completeness, part 1: if [step_expr] returns [Some ll],
- then [ll] contains all possible reducts. *)
+(** Completeness part 1: [step_expr] contains all possible non-error reducts. *)
Lemma lred_topred:
forall l1 m1 l2 m2,
@@ -704,46 +1475,35 @@ Proof.
Qed.
Lemma rred_topred:
- forall r1 m1 r2 m2,
- rred r1 m1 r2 m2 ->
- step_expr RV r1 m1 = topred (Rred r2 m2).
+ forall w' r1 m1 t r2 m2,
+ rred ge r1 m1 t r2 m2 -> possible_trace w t w' ->
+ step_expr RV r1 m1 = topred (Rred r2 m2 t).
Proof.
- induction 1; simpl.
+ induction 1; simpl; intros.
(* valof *)
- rewrite dec_eq_true; auto. rewrite H; auto.
+ rewrite dec_eq_true; auto. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0). auto.
(* addrof *)
- auto.
+ inv H. auto.
(* unop *)
- rewrite H; auto.
+ inv H0. rewrite H; auto.
(* binop *)
- rewrite H; auto.
+ inv H0. rewrite H; auto.
(* cast *)
- rewrite H; auto.
+ inv H0. rewrite H; auto.
(* condition *)
- rewrite H; auto.
+ inv H0. rewrite H; auto.
(* sizeof *)
- auto.
+ inv H. auto.
(* assign *)
- rewrite dec_eq_true; auto. rewrite H; rewrite H0; auto.
+ rewrite dec_eq_true; auto. rewrite H. rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ H0 H1). auto.
(* assignop *)
- rewrite dec_eq_true; auto. rewrite H; rewrite H0; rewrite H1; rewrite H2; auto.
+ rewrite dec_eq_true; auto. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0). auto.
(* postincr *)
- rewrite dec_eq_true; auto. rewrite H; rewrite H0; rewrite H1; rewrite H2; auto.
+ rewrite dec_eq_true; auto. subst. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H1). auto.
(* comma *)
- rewrite H; rewrite dec_eq_true; auto.
+ inv H0. rewrite dec_eq_true; auto.
(* paren *)
- rewrite H; auto.
-Qed.
-
-Lemma sem_cast_arguments_complete:
- forall al tyl vl,
- cast_arguments al tyl vl ->
- exists vtl, is_val_list al = Some vtl /\ sem_cast_arguments vtl tyl = Some vl.
-Proof.
- induction 1.
- exists (@nil (val * type)); auto.
- destruct IHcast_arguments as [vtl [A B]].
- exists ((v, ty) :: vtl); simpl. rewrite A; rewrite B; rewrite H. auto.
+ inv H0. rewrite H; auto.
Qed.
Lemma callred_topred:
@@ -757,12 +1517,7 @@ Proof.
Qed.
Definition reducts_incl {A B: Type} (C: A -> B) (res1: reducts A) (res2: reducts B) : Prop :=
- match res1, res2 with
- | Some ll1, Some ll2 =>
- forall C1 rd, In (C1, rd) ll1 -> In ((fun x => C(C1 x)), rd) ll2
- | None, Some ll2 => False
- | _, None => True
- end.
+ forall C1 rd, In (C1, rd) res1 -> In ((fun x => C(C1 x)), rd) res2.
Lemma reducts_incl_trans:
forall (A1 A2: Type) (C: A1 -> A2) res1 res2, reducts_incl C res1 res2 ->
@@ -770,14 +1525,14 @@ Lemma reducts_incl_trans:
reducts_incl C' res2 res3 ->
reducts_incl (fun x => C'(C x)) res1 res3.
Proof.
- unfold reducts_incl; intros. destruct res1; destruct res2; destruct res3; auto. contradiction.
+ unfold reducts_incl; intros. auto.
Qed.
Lemma reducts_incl_nil:
forall (A B: Type) (C: A -> B) res,
- reducts_incl C (Some nil) res.
+ reducts_incl C nil res.
Proof.
- intros; red. destruct res; auto. intros; contradiction.
+ intros; red. intros; contradiction.
Qed.
Lemma reducts_incl_val:
@@ -805,29 +1560,29 @@ Lemma reducts_incl_incontext:
forall (A B: Type) (C: A -> B) res,
reducts_incl C res (incontext C res).
Proof.
- intros; unfold reducts_incl. destruct res; simpl; auto.
- intros. set (f := fun z : (expr -> A) * reduction => (fun x : expr => C (fst z x), snd z)).
- change (In (f (C1, rd)) (map f l)). apply in_map. auto.
+ unfold reducts_incl, incontext. intros.
+ set (f := fun z : (expr -> A) * reduction => (fun x : expr => C (fst z x), snd z)).
+ change (In (f (C1, rd)) (map f res)). apply in_map. auto.
Qed.
Lemma reducts_incl_incontext2_left:
forall (A1 A2 B: Type) (C1: A1 -> B) res1 (C2: A2 -> B) res2,
reducts_incl C1 res1 (incontext2 C1 res1 C2 res2).
Proof.
- intros. destruct res1; simpl; auto. destruct res2; auto.
- intros. rewrite in_app_iff. left.
+ unfold reducts_incl, incontext2, incontext. intros.
+ rewrite in_app_iff. left.
set (f := fun z : (expr -> A1) * reduction => (fun x : expr => C1 (fst z x), snd z)).
- change (In (f (C0, rd)) (map f l)). apply in_map; auto.
+ change (In (f (C0, rd)) (map f res1)). apply in_map; auto.
Qed.
Lemma reducts_incl_incontext2_right:
forall (A1 A2 B: Type) (C1: A1 -> B) res1 (C2: A2 -> B) res2,
reducts_incl C2 res2 (incontext2 C1 res1 C2 res2).
Proof.
- intros. destruct res1; destruct res2; simpl; auto.
- intros. rewrite in_app_iff. right.
+ unfold reducts_incl, incontext2, incontext. intros.
+ rewrite in_app_iff. right.
set (f := fun z : (expr -> A2) * reduction => (fun x : expr => C2 (fst z x), snd z)).
- change (In (f (C0, rd)) (map f l0)). apply in_map; auto.
+ change (In (f (C0, rd)) (map f res2)). apply in_map; auto.
Qed.
Hint Resolve reducts_incl_val reducts_incl_loc reducts_incl_listval
@@ -849,7 +1604,7 @@ Proof.
destruct (is_val (C a)) as [[v ty']|]_eqn; eauto.
(* field *)
eapply reducts_incl_trans with (C' := fun x => Efield x f ty); eauto.
- destruct (is_loc (C a)) as [[[b ofs] ty']|]_eqn; eauto.
+ destruct (is_val (C a)) as [[v ty']|]_eqn; eauto.
(* valof *)
eapply reducts_incl_trans with (C' := fun x => Evalof x ty); eauto.
destruct (is_loc (C a)) as [[[b ofs] ty']|]_eqn; eauto.
@@ -912,625 +1667,83 @@ Proof.
apply step_exprlist_context; eauto. eauto.
Qed.
-(** Completeness, part 2: given a safe expression, [step_expr] does not return [None]. *)
+(** Completeness part 2: if we can reduce to [Stuckstate], [step_expr]
+ contains at least one [Stuckred] reduction. *)
-Lemma topred_none:
- forall rd, topred rd <> None.
+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; unfold topred; congruence.
+ 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]]]].
+ 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.
Qed.
-Lemma incontext_none:
- forall (A B: Type) (C: A -> B) rds,
- rds <> None -> incontext C rds <> None.
-Proof.
- unfold incontext; intros. destruct rds; congruence.
-Qed.
-
-Lemma incontext2_none:
- forall (A1 A2 B: Type) (C1: A1 -> B) rds1 (C2: A2 -> B) rds2,
- rds1 <> None -> rds2 <> None -> incontext2 C1 rds1 C2 rds2 <> None.
-Proof.
- unfold incontext2; intros. destruct rds1; destruct rds2; congruence.
-Qed.
-
-Lemma safe_expr_kind:
- forall k C a m,
- context k RV C ->
- not_stuck ge e (C a) m ->
- k = Cstrategy.expr_kind a.
-Proof.
- intros.
- symmetry. eapply Cstrategy.not_imm_stuck_kind; eauto.
-Qed.
-
-Lemma safe_inversion:
- forall k C a m,
+Lemma not_imm_safe_stuck_red:
+ forall m a k C,
context k RV C ->
- not_stuck ge e (C a) m ->
- match a with
- | Eloc _ _ _ => True
- | Eval _ _ => True
- | _ => Cstrategy.invert_expr_prop ge e a m
- end.
+ ~imm_safe_t k a m ->
+ exists C', In (C', Stuckred) (step_expr RV (C a) m).
Proof.
- intros. eapply Cstrategy.not_imm_stuck_inv; eauto.
-Qed.
-
-Lemma is_val_list_all_values:
- forall al vtl, is_val_list al = Some vtl -> Cstrategy.exprlist_all_values al.
-Proof.
- 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.
-Qed.
-
-Theorem step_expr_defined:
- forall a k m C,
- context k RV C ->
- not_stuck ge e (C a) m ->
- step_expr k a m <> None
-with step_exprlist_defined:
- forall al m C,
- Cstrategy.contextlist' C ->
- not_stuck ge e (C al) m ->
- step_exprlist al m <> None.
-Proof.
- induction a; intros k m C CTX NS;
- rewrite (safe_expr_kind _ _ _ _ CTX NS);
- rewrite (safe_expr_kind _ _ _ _ CTX NS) in CTX;
- simpl in *;
- generalize (safe_inversion _ _ _ _ CTX NS); intro INV.
-(* val *)
- congruence.
-(* var *)
- red in INV. destruct INV as [b [P | [P [Q R]]]].
- rewrite P; rewrite dec_eq_true. apply topred_none.
- rewrite P; rewrite Q; rewrite R; rewrite dec_eq_true. apply topred_none.
-(* field *)
- destruct (is_loc a) as [[[b ofs] ty']|]_eqn.
- rewrite (is_loc_inv _ _ _ _ Heqo) in INV. red in INV.
- destruct ty'; try contradiction. destruct INV as [delta EQ]. rewrite EQ. apply topred_none.
- apply topred_none.
- apply incontext_none. apply IHa with (C := fun x => C(Efield x f ty)); eauto.
-(* valof *)
- destruct (is_loc a) as [[[b ofs] ty']|]_eqn.
- rewrite (is_loc_inv _ _ _ _ Heqo) in INV. red in INV. destruct INV as [EQ [v LD]]. subst.
- rewrite dec_eq_true; rewrite LD; apply topred_none.
- apply incontext_none. apply IHa with (C := fun x => C(Evalof x ty)); eauto.
-(* deref *)
- destruct (is_val a) as [[v ty']|]_eqn.
- rewrite (is_val_inv _ _ _ Heqo) in INV. red in INV. destruct INV as [b [ofs EQ]]. subst.
- apply topred_none.
- apply incontext_none. apply IHa with (C := fun x => C(Ederef x ty)); eauto.
-(* addrof *)
- destruct (is_loc a) as [[[b ofs] ty']|]_eqn.
- apply topred_none.
- apply incontext_none. apply IHa with (C := fun x => C(Eaddrof x ty)); eauto.
-(* unop *)
- destruct (is_val a) as [[v1 ty1]|]_eqn.
- rewrite (is_val_inv _ _ _ Heqo) in INV. red in INV. destruct INV as [v EQ].
- rewrite EQ; apply topred_none.
- apply incontext_none. apply IHa with (C := fun x => C(Eunop op x ty)); eauto.
-(* binop *)
- destruct (is_val a1) as [[v1 ty1]|]_eqn.
- destruct (is_val a2) as [[v2 ty2]|]_eqn.
- rewrite (is_val_inv _ _ _ Heqo) in INV.
- rewrite (is_val_inv _ _ _ Heqo0) in INV. red in INV. destruct INV as [v EQ].
- rewrite EQ; apply topred_none.
- apply incontext2_none. apply IHa1 with (C := fun x => C(Ebinop op x a2 ty)); eauto. apply IHa2 with (C := fun x => C(Ebinop op a1 x ty)); eauto.
- apply incontext2_none. apply IHa1 with (C := fun x => C(Ebinop op x a2 ty)); eauto. apply IHa2 with (C := fun x => C(Ebinop op a1 x ty)); eauto.
-(* cast *)
- destruct (is_val a) as [[v1 ty1]|]_eqn.
- rewrite (is_val_inv _ _ _ Heqo) in INV. red in INV. destruct INV as [v EQ].
- rewrite EQ; apply topred_none.
- apply incontext_none. apply IHa with (C := fun x => C(Ecast x ty)); eauto.
-(* condition *)
- destruct (is_val a1) as [[v1 ty1]|]_eqn.
- rewrite (is_val_inv _ _ _ Heqo) in INV. red in INV. destruct INV as [v EQ].
- rewrite EQ; apply topred_none.
- apply incontext_none. apply IHa1 with (C := fun x => C(Econdition x a2 a3 ty)); eauto.
-(* sizeof *)
- apply topred_none.
-(* assign *)
- destruct (is_loc a1) as [[[b ofs] ty1]|]_eqn.
- destruct (is_val a2) as [[v2 ty2]|]_eqn.
- rewrite (is_loc_inv _ _ _ _ Heqo) in INV.
- rewrite (is_val_inv _ _ _ Heqo0) in INV. red in INV.
- destruct INV as [v [m' [P [Q R]]]].
- subst. rewrite dec_eq_true; rewrite Q; rewrite R; apply topred_none.
- apply incontext2_none. apply IHa1 with (C := fun x => C(Eassign x a2 ty)); eauto. apply IHa2 with (C := fun x => C(Eassign a1 x ty)); eauto.
- apply incontext2_none. apply IHa1 with (C := fun x => C(Eassign x a2 ty)); eauto. apply IHa2 with (C := fun x => C(Eassign a1 x ty)); eauto.
-(* assignop *)
- destruct (is_loc a1) as [[[b ofs] ty1]|]_eqn.
- destruct (is_val a2) as [[v2 ty2]|]_eqn.
- rewrite (is_loc_inv _ _ _ _ Heqo) in INV.
- rewrite (is_val_inv _ _ _ Heqo0) in INV. red in INV.
- destruct INV as [v1 [v [v' [m' [P [Q [R [S T]]]]]]]].
- subst. rewrite dec_eq_true; rewrite Q; rewrite R; rewrite S; rewrite T; apply topred_none.
- apply incontext2_none. apply IHa1 with (C := fun x => C(Eassignop op x a2 tyres ty)); eauto. apply IHa2 with (C := fun x => C(Eassignop op a1 x tyres ty)); eauto.
- apply incontext2_none. apply IHa1 with (C := fun x => C(Eassignop op x a2 tyres ty)); eauto. apply IHa2 with (C := fun x => C(Eassignop op a1 x tyres ty)); eauto.
-(* postincr *)
- destruct (is_loc a) as [[[b ofs] ty1]|]_eqn.
- rewrite (is_loc_inv _ _ _ _ Heqo) in INV. red in INV.
- destruct INV as [v1 [v2 [v3 [m' [P [Q [R [S T]]]]]]]].
- subst. rewrite dec_eq_true; rewrite Q; rewrite R; rewrite S; rewrite T; apply topred_none.
- apply incontext_none. apply IHa with (C := fun x => C(Epostincr id x ty)); eauto.
-(* comma *)
- destruct (is_val a1) as [[v1 ty1]|]_eqn.
- rewrite (is_val_inv _ _ _ Heqo) in INV. red in INV. rewrite INV.
- rewrite dec_eq_true; apply topred_none.
- apply incontext_none. apply IHa1 with (C := fun x => C(Ecomma x a2 ty)); eauto.
-(* call *)
- destruct (is_val a) as [[vf tyf]|]_eqn.
- destruct (is_val_list rargs) as [vtl|]_eqn.
- rewrite (is_val_inv _ _ _ Heqo) in INV. red in INV.
- destruct INV as [tyargs [tyres [fd [vl [P [Q [R S]]]]]]].
- eapply is_val_list_all_values; eauto.
- rewrite P; rewrite Q.
- exploit sem_cast_arguments_complete; eauto. intros [vtl' [U V]].
- assert (vtl' = vtl) by congruence. subst. rewrite V. rewrite S. rewrite dec_eq_true.
- apply topred_none.
- apply incontext2_none. apply IHa with (C := fun x => C(Ecall x rargs ty)); eauto.
- apply step_exprlist_defined with (C := fun x => C(Ecall a x ty)); auto.
- apply Cstrategy.contextlist'_intro with (rl0 := Enil). auto.
- apply incontext2_none. apply IHa with (C := fun x => C(Ecall x rargs ty)); eauto.
- apply step_exprlist_defined with (C := fun x => C(Ecall a x ty)); auto.
- apply Cstrategy.contextlist'_intro with (rl0 := Enil). auto.
-(* loc *)
- congruence.
-(* paren *)
- destruct (is_val a) as [[v1 ty1]|]_eqn.
- rewrite (is_val_inv _ _ _ Heqo) in INV. red in INV. destruct INV as [v EQ].
- rewrite EQ; apply topred_none.
- apply incontext_none. apply IHa with (C := fun x => C(Eparen x ty)); eauto.
-
- induction al; intros; simpl.
-(* nil *)
- congruence.
-(* cons *)
- apply incontext2_none.
- apply step_expr_defined with (C := fun x => C(Econs x al)); eauto.
- apply Cstrategy.contextlist'_head; auto.
- apply IHal with (C := fun x => C(Econs r1 x)); auto.
- apply Cstrategy.contextlist'_tail; auto.
+ 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.
+ destruct H1 as [C' IN].
+ specialize (step_expr_context _ _ _ H a m). unfold reducts_incl.
+ intro.
+ exists (fun x => (C (C' x))). apply H1; auto.
Qed.
-(** Connections between [not_stuck] and [step_expr]. *)
+(** Connections between [imm_safe_t] and [imm_safe] *)
-Lemma step_expr_not_imm_stuck:
+Lemma imm_safe_imm_safe_t:
forall k a m,
- step_expr k a m <> None ->
- not_imm_stuck ge e k a m.
-Proof.
- intros. generalize (step_expr_sound a k m). unfold reducts_ok.
- destruct (step_expr k a m) as [ll|]. destruct ll.
-(* value or location *)
- destruct k; destruct a; simpl; intros; try congruence.
- apply not_stuck_loc.
- apply Csem.not_stuck_val.
-(* reducible *)
- intros R. destruct p as [C1 rd1]. destruct (R C1 rd1) as [P [a' [U V]]]; auto with coqlib.
- subst a. red in V. destruct rd1.
- eapply not_stuck_lred; eauto.
- eapply not_stuck_rred; eauto.
- destruct V. subst m'. eapply not_stuck_callred; eauto.
-(* stuck *)
- congruence.
-Qed.
-
-Lemma step_expr_not_stuck:
- forall a m,
- step_expr RV a m <> None ->
- not_stuck ge e a m.
-Proof.
- intros; red; intros. subst a.
- apply step_expr_not_imm_stuck.
- generalize (step_expr_context _ _ C H0 e' m). unfold reducts_incl.
- destruct (step_expr k e' m). congruence.
- destruct (step_expr RV (C e') m). tauto. congruence.
-Qed.
-
-Lemma not_stuck_step_expr:
- forall a m,
- not_stuck ge e a m ->
- step_expr RV a m <> None.
-Proof.
- intros. apply step_expr_defined with (C := fun x => x); auto.
-Qed.
-
-End EXPRS.
-
-(** * External functions and events. *)
-
-Section EVENTS.
-
-Variable F V: Type.
-Variable genv: Genv.t F V.
-
-Definition eventval_of_val (v: val) (t: typ) : option eventval :=
- match v, t with
- | Vint i, AST.Tint => Some (EVint i)
- | Vfloat f, AST.Tfloat => Some (EVfloat f)
- | Vptr b ofs, AST.Tint => do id <- Genv.invert_symbol genv b; Some (EVptr_global id ofs)
- | _, _ => None
- end.
-
-Fixpoint list_eventval_of_val (vl: list val) (tl: list typ) : option (list eventval) :=
- match vl, tl with
- | nil, nil => Some nil
- | v1::vl, t1::tl =>
- do ev1 <- eventval_of_val v1 t1;
- do evl <- list_eventval_of_val vl tl;
- Some (ev1 :: evl)
- | _, _ => None
- end.
-
-Definition val_of_eventval (ev: eventval) (t: typ) : option val :=
- match ev, t with
- | EVint i, AST.Tint => Some (Vint i)
- | EVfloat f, AST.Tfloat => Some (Vfloat f)
- | EVptr_global id ofs, AST.Tint => do b <- Genv.find_symbol genv id; Some (Vptr b ofs)
- | _, _ => None
- end.
-
-Lemma eventval_of_val_sound:
- forall v t ev, eventval_of_val v t = Some ev -> eventval_match genv ev t v.
-Proof.
- intros. destruct v; destruct t; simpl in H; inv H.
- constructor.
- constructor.
- destruct (Genv.invert_symbol genv b) as [id|]_eqn; inv H1.
- constructor. apply Genv.invert_find_symbol; auto.
-Qed.
-
-Lemma eventval_of_val_complete:
- forall ev t v, eventval_match genv ev t v -> eventval_of_val v t = Some ev.
+ imm_safe ge e k a m ->
+ imm_safe_t k a m \/
+ 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.
- induction 1; simpl; auto.
- rewrite (Genv.find_invert_symbol _ _ H). auto.
-Qed.
-
-Lemma list_eventval_of_val_sound:
- forall vl tl evl, list_eventval_of_val vl tl = Some evl -> eventval_list_match genv evl tl vl.
-Proof with try discriminate.
- induction vl; destruct tl; simpl; intros; inv H.
- constructor.
- destruct (eventval_of_val a t) as [ev1|]_eqn...
- destruct (list_eventval_of_val vl tl) as [evl'|]_eqn...
- inv H1. constructor. apply eventval_of_val_sound; auto. eauto.
-Qed.
-
-Lemma list_eventval_of_val_complete:
- forall evl tl vl, eventval_list_match genv evl tl vl -> list_eventval_of_val vl tl = Some evl.
-Proof.
- induction 1; simpl. auto.
- rewrite (eventval_of_val_complete _ _ _ H). rewrite IHeventval_list_match. auto.
-Qed.
-
-Lemma val_of_eventval_sound:
- forall ev t v, val_of_eventval ev t = Some v -> eventval_match genv ev t v.
-Proof.
- intros. destruct ev; destruct t; simpl in H; inv H.
- constructor.
- constructor.
- destruct (Genv.find_symbol genv i) as [b|]_eqn; inv H1.
- constructor. auto.
-Qed.
-
-Lemma val_of_eventval_complete:
- forall ev t v, eventval_match genv ev t v -> val_of_eventval ev t = Some v.
-Proof.
- induction 1; simpl; auto. rewrite H; auto.
-Qed.
-
-(** System calls and library functions *)
-
-Definition do_ef_external (name: ident) (sg: signature)
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- do args <- list_eventval_of_val vargs (sig_args sg);
- match nextworld_io w name args with
- | None => None
- | Some(res, w') =>
- do vres <- val_of_eventval res (proj_sig_res sg);
- Some(w', Event_syscall name args res :: E0, vres, m)
- end.
-
-Definition do_ef_volatile_load (chunk: memory_chunk)
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- match vargs with
- | Vptr b ofs :: nil =>
- if block_is_volatile genv b then
- do id <- Genv.invert_symbol genv b;
- match nextworld_vload w chunk id ofs with
- | None => None
- | Some(res, w') =>
- do vres <- val_of_eventval res (type_of_chunk chunk);
- Some(w', Event_vload chunk id ofs res :: nil, Val.load_result chunk vres, m)
- end
- else
- do v <- Mem.load chunk m b (Int.unsigned ofs);
- Some(w, E0, v, m)
- | _ => None
- end.
-
-Definition do_ef_volatile_store (chunk: memory_chunk)
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- match vargs with
- | Vptr b ofs :: v :: nil =>
- if block_is_volatile genv b then
- do id <- Genv.invert_symbol genv b;
- do ev <- eventval_of_val v (type_of_chunk chunk);
- do w' <- nextworld_vstore w chunk id ofs ev;
- Some(w', Event_vstore chunk id ofs ev :: nil, Vundef, m)
- else
- do m' <- Mem.store chunk m b (Int.unsigned ofs) v;
- Some(w, E0, Vundef, m')
- | _ => None
- end.
-
-Definition do_ef_volatile_load_global (chunk: memory_chunk) (id: ident) (ofs: int)
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- match Genv.find_symbol genv id with
- | Some b => do_ef_volatile_load chunk w (Vptr b ofs :: vargs) m
- | None => None
- end.
-
-Definition do_ef_volatile_store_global (chunk: memory_chunk) (id: ident) (ofs: int)
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- match Genv.find_symbol genv id with
- | Some b => do_ef_volatile_store chunk w (Vptr b ofs :: vargs) m
- | None => None
- end.
-
-Definition do_ef_malloc
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- match vargs with
- | Vint n :: nil =>
- let (m', b) := Mem.alloc m (-4) (Int.unsigned n) in
- do m'' <- Mem.store Mint32 m' b (-4) (Vint n);
- Some(w, E0, Vptr b Int.zero, m'')
- | _ => None
- end.
-
-Definition do_ef_free
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- match vargs with
- | Vptr b lo :: nil =>
- do vsz <- Mem.load Mint32 m b (Int.unsigned lo - 4);
- match vsz with
- | Vint sz =>
- check (zlt 0 (Int.unsigned sz));
- do m' <- Mem.free m b (Int.unsigned lo - 4) (Int.unsigned lo + Int.unsigned sz);
- Some(w, E0, Vundef, m')
- | _ => None
- end
- | _ => None
- end.
-
-Definition memcpy_args_ok
- (sz al: Z) (bdst: block) (odst: Z) (bsrc: block) (osrc: Z) : Prop :=
- (al = 1 \/ al = 2 \/ al = 4)
- /\ sz > 0
- /\ (al | sz) /\ (al | osrc) /\ (al | odst)
- /\ (bsrc <> bdst \/ osrc = odst \/ osrc + sz <= odst \/ odst + sz <= osrc).
-
-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.
- assert (X: {al = 1 \/ al = 2 \/ al = 4} + {~(al = 1 \/ al = 2 \/ al = 4)}).
- destruct (zeq al 1); auto. destruct (zeq al 2); auto. destruct (zeq al 4); auto...
- unfold memcpy_args_ok. destruct X...
- assert (al > 0) by (intuition omega).
- destruct (zlt 0 sz)...
- destruct (Zdivide_dec al sz); auto...
- destruct (Zdivide_dec al osrc); auto...
- destruct (Zdivide_dec al odst); auto...
- assert (Y: {bsrc <> bdst \/ osrc = odst \/ osrc + sz <= odst \/ odst + sz <= osrc}
- +{~(bsrc <> bdst \/ osrc = odst \/ osrc + sz <= odst \/ odst + sz <= osrc)}).
- destruct (eq_block bsrc bdst); auto.
- destruct (zeq osrc odst); auto.
- destruct (zle (osrc + sz) odst); auto.
- destruct (zle (odst + sz) osrc); auto.
- right; intuition omega.
- destruct Y... left; intuition omega.
+ 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.
+ 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.
-Definition do_ef_memcpy (sz al: Z)
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- match vargs with
- | Vptr bdst odst :: Vptr bsrc osrc :: nil =>
- if memcpy_check_args sz al bdst (Int.unsigned odst) bsrc (Int.unsigned osrc) then
- do bytes <- Mem.loadbytes m bsrc (Int.unsigned osrc) sz;
- do m' <- Mem.storebytes m bdst (Int.unsigned odst) bytes;
- Some(w, E0, Vundef, m')
- else None
- | _ => None
- end.
-
-Definition do_ef_annot (text: ident) (targs: list typ)
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- do args <- list_eventval_of_val vargs targs;
- Some(w, Event_annot text args :: E0, Vundef, m).
-
-Definition do_ef_annot_val (text: ident) (targ: typ)
- (w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
- match vargs with
- | varg :: nil =>
- do arg <- eventval_of_val varg targ;
- Some(w, Event_annot text (arg :: nil) :: E0, varg, m)
- | _ => None
- end.
-
-Definition do_external (ef: external_function):
- world -> list val -> mem -> option (world * trace * val * mem) :=
- match ef with
- | EF_external name sg => do_ef_external name sg
- | EF_builtin name sg => do_ef_external name sg
- | EF_vload chunk => do_ef_volatile_load chunk
- | EF_vstore chunk => do_ef_volatile_store chunk
- | EF_vload_global chunk id ofs => do_ef_volatile_load_global chunk id ofs
- | EF_vstore_global chunk id ofs => do_ef_volatile_store_global chunk id ofs
- | EF_malloc => do_ef_malloc
- | EF_free => do_ef_free
- | EF_memcpy sz al => do_ef_memcpy sz al
- | EF_annot text targs => do_ef_annot text targs
- | EF_annot_val text targ => do_ef_annot_val text targ
- end.
-
-Ltac mydestr :=
- match goal with
- | [ |- None = Some _ -> _ ] => intro X; discriminate
- | [ |- Some _ = Some _ -> _ ] => intro X; inv X
- | [ |- match ?x with Some _ => _ | None => _ end = Some _ -> _ ] => destruct x as []_eqn; mydestr
- | [ |- match ?x with true => _ | false => _ end = Some _ -> _ ] => destruct x as []_eqn; mydestr
- | [ |- match ?x with left _ => _ | right _ => _ end = Some _ -> _ ] => destruct x; mydestr
- | _ => idtac
- end.
-
-Lemma do_ef_external_sound:
- forall ef w vargs m w' t vres m',
- do_external ef w vargs m = Some(w', t, vres, m') ->
- external_call ef genv vargs m t vres m' /\ possible_trace w t w'.
-Proof with try congruence.
- intros until m'.
- assert (IO: forall name sg,
- do_ef_external name sg w vargs m = Some(w', t, vres, m') ->
- extcall_io_sem name sg genv vargs m t vres m' /\ possible_trace w t w').
- intros until sg. unfold do_ef_external. mydestr. destruct p as [res w'']; mydestr.
- split. econstructor. apply list_eventval_of_val_sound; auto.
- apply val_of_eventval_sound; auto.
- econstructor. constructor; eauto. constructor.
-
- assert (VLOAD: forall chunk vargs,
- do_ef_volatile_load chunk w vargs m = Some (w', t, vres, m') ->
- volatile_load_sem chunk genv vargs m t vres m' /\ possible_trace w t w').
- intros chunk vargs'.
- unfold do_ef_volatile_load. destruct vargs'... destruct v... destruct vargs'...
- mydestr. destruct p as [res w'']; mydestr.
- split. constructor. apply Genv.invert_find_symbol; auto. auto.
- apply val_of_eventval_sound; auto.
- econstructor. constructor; eauto. constructor.
- split. constructor; auto. constructor.
+(** A state can "crash the world" if it can make an observable transition
+ whose trace is not accepted by the external world. *)
- assert (VSTORE: forall chunk vargs,
- do_ef_volatile_store chunk w vargs m = Some (w', t, vres, m') ->
- volatile_store_sem chunk genv vargs m t vres m' /\ possible_trace w t w').
- intros chunk vargs'.
- unfold do_ef_volatile_store. destruct vargs'... destruct v... destruct vargs'... destruct vargs'...
- mydestr.
- split. constructor. apply Genv.invert_find_symbol; auto. auto.
- apply eventval_of_val_sound; auto.
- econstructor. constructor; eauto. constructor.
- split. constructor; auto. constructor.
+Definition can_crash_world (w: world) (S: state) : Prop :=
+ exists t, exists S', Csem.step ge S t S' /\ forall w', ~possible_trace w t w'.
- destruct ef; simpl.
-(* EF_external *)
- auto.
-(* EF_builtin *)
- auto.
-(* EF_vload *)
- auto.
-(* EF_vstore *)
- auto.
-(* EF_vload_global *)
- rewrite volatile_load_global_charact.
- unfold do_ef_volatile_load_global. destruct (Genv.find_symbol genv)...
- intros. exploit VLOAD; eauto. intros [A B]. split; auto. exists b; auto.
-(* EF_vstore_global *)
- rewrite volatile_store_global_charact.
- unfold do_ef_volatile_store_global. destruct (Genv.find_symbol genv)...
- intros. exploit VSTORE; eauto. intros [A B]. split; auto. exists b; auto.
-(* EF_malloc *)
- unfold do_ef_malloc. destruct vargs... destruct v... destruct vargs...
- 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.
- 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.
- split. econstructor; eauto; tauto. constructor.
-(* EF_annot *)
- 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.
- split. constructor. apply eventval_of_val_sound; auto.
- econstructor. constructor; eauto. constructor.
-Qed.
-
-Lemma do_ef_external_complete:
- forall ef w vargs m w' t vres m',
- external_call ef genv vargs m t vres m' -> possible_trace w t w' ->
- do_external ef w vargs m = Some(w', t, vres, m').
+Theorem not_imm_safe_t:
+ forall K C a m f k,
+ context K RV C ->
+ ~imm_safe_t K a m ->
+ 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.
- assert (IO: forall name sg,
- extcall_io_sem name sg genv vargs m t vres m' ->
- do_ef_external name sg w vargs m = Some (w', t, vres, m')).
- intros. inv H1. inv H0. inv H8. inv H6.
- unfold do_ef_external. rewrite (list_eventval_of_val_complete _ _ _ H2). rewrite H8.
- rewrite (val_of_eventval_complete _ _ _ H3). auto.
-
- assert (VLOAD: forall chunk vargs,
- volatile_load_sem chunk genv vargs m t vres m' ->
- do_ef_volatile_load chunk w vargs m = Some (w', t, vres, m')).
- intros. inv H1; unfold do_ef_volatile_load.
- inv H0. inv H9. inv H7.
- rewrite H3. rewrite (Genv.find_invert_symbol _ _ H2). rewrite H10.
- rewrite (val_of_eventval_complete _ _ _ H4). auto.
- inv H0. rewrite H2. rewrite H3. auto.
-
- assert (VSTORE: forall chunk vargs,
- volatile_store_sem chunk genv vargs m t vres m' ->
- do_ef_volatile_store chunk w vargs m = Some (w', t, vres, m')).
- intros. inv H1; unfold do_ef_volatile_store.
- inv H0. inv H9. inv H7.
- rewrite H3. rewrite (Genv.find_invert_symbol _ _ H2).
- rewrite (eventval_of_val_complete _ _ _ H4). rewrite H10. auto.
- inv H0. rewrite H2. rewrite H3. auto.
-
- destruct ef; simpl in *.
-(* EF_external *)
- auto.
-(* EF_builtin *)
- auto.
-(* EF_vload *)
- auto.
-(* EF_vstore *)
- auto.
-(* EF_vload_global *)
- rewrite volatile_load_global_charact in H. destruct H as [b [P Q]].
- unfold do_ef_volatile_load_global. rewrite P. auto.
-(* EF_vstore *)
- rewrite volatile_store_global_charact in H. destruct H as [b [P Q]].
- unfold do_ef_volatile_store_global. rewrite P. auto.
-(* EF_malloc *)
- inv H; unfold do_ef_malloc.
- inv H0. rewrite H1. rewrite H2. auto.
-(* EF_free *)
- inv H; unfold do_ef_free.
- inv H0. rewrite H1. rewrite zlt_true. rewrite H3. auto. omega.
-(* EF_memcpy *)
- inv H; unfold do_ef_memcpy.
- inv H0. rewrite pred_dec_true. rewrite H7; rewrite H8; auto.
- red. tauto.
-(* EF_annot *)
- 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.
- rewrite (eventval_of_val_complete _ _ _ H1). auto.
+ intros. destruct (classic (imm_safe ge e K a m)).
+ 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.
+ left. left. eapply step_stuck; eauto.
Qed.
-End EVENTS.
+End EXPRS.
(** * Transitions over states. *)
@@ -1560,7 +1773,7 @@ Proof.
rewrite H; rewrite IHalloc_variables; auto.
Qed.
-Function sem_bind_parameters (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
@@ -1568,37 +1781,39 @@ Function sem_bind_parameters (e: env) (m: mem) (l: list (ident * type)) (lv: lis
match PTree.get id e with
| Some (b, ty') =>
check (type_eq ty ty');
- do m1 <- store_value_of_type ty m b Int.zero v1;
- sem_bind_parameters e m1 params lv
+ do w', t, m1 <- do_assign_loc w ty m b Int.zero v1;
+ match t with nil => sem_bind_parameters w e m1 params lv | _ => None end
| None => None
end
| _, _ => None
end.
-Lemma sem_bind_parameters_sound : forall e m l lv m',
- sem_bind_parameters e m l lv = Some m' ->
- bind_parameters e m l lv m'.
+Lemma sem_bind_parameters_sound : forall w e m l lv 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 e m l lv); try discriminate.
- inversion H; constructor; auto.
- econstructor; eauto.
+ intros; functional induction (sem_bind_parameters w e m l lv); try discriminate.
+ inversion H; constructor; auto.
+ exploit do_assign_loc_sound; eauto. intros [A B]. econstructor; eauto.
Qed.
-Lemma sem_bind_parameters_complete : forall e m l lv m',
- bind_parameters e m l lv m' ->
- sem_bind_parameters e m l lv = Some m'.
+Lemma sem_bind_parameters_complete : forall w e m l lv m',
+ bind_parameters ge e m l lv m' ->
+ sem_bind_parameters w e m l lv = Some m'.
Proof.
induction 1; simpl; auto.
- rewrite H. rewrite dec_eq_true.
- destruct (store_value_of_type ty m b Int.zero v1); try discriminate.
- inv H0; auto.
+ rewrite H. rewrite dec_eq_true.
+ assert (possible_trace w E0 w) by constructor.
+ rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ H0 H2).
+ simpl. auto.
Qed.
Definition expr_final_state (f: function) (k: cont) (e: env) (C_rd: (expr -> expr) * reduction) :=
match snd C_rd with
| Lred a m => (E0, ExprState f (fst C_rd a) k e m)
- | Rred a m => (E0, ExprState f (fst C_rd a) k e m)
+ | Rred a m t => (t, ExprState f (fst C_rd a) k e m)
| Callred fd vargs ty m => (E0, Callstate fd vargs (Kcall f e (fst C_rd) ty k) m)
+ | Stuck => (E0, Stuckstate)
end.
Local Open Scope list_monad_scope.
@@ -1637,10 +1852,7 @@ Definition do_step (w: world) (s: state) : list (trace * state) :=
end
| None =>
- match step_expr e RV a m with
- | None => nil
- | Some ll => map (expr_final_state f k e) ll
- end
+ map (expr_final_state f k e) (step_expr e w RV a m)
end
| State f (Sdo x) k e m => ret(ExprState f x (Kdo k) e m)
@@ -1674,9 +1886,9 @@ Definition do_step (w: world) (s: state) : list (trace * state) :=
ret (Returnstate Vundef (call_cont k) m')
| State f (Sreturn (Some x)) k e m => ret (ExprState f x (Kreturn k) e m)
| State f Sskip ((Kstop | Kcall _ _ _ _ _) as k) e m =>
- check (type_eq f.(fn_return) Tvoid);
- do m' <- Mem.free_list m (blocks_of_env e);
- ret (Returnstate Vundef k m')
+ check type_eq (f.(fn_return)) Tvoid;
+ do m' <- Mem.free_list m (blocks_of_env e);
+ ret (Returnstate Vundef k m')
| State f (Sswitch x sl) k e m => ret (ExprState f x (Kswitch1 sl k) e m)
| State f (Sskip|Sbreak) (Kswitch2 k) e m => ret (State f Sskip k e m)
@@ -1692,10 +1904,10 @@ Definition do_step (w: world) (s: state) : list (trace * state) :=
| Callstate (Internal f) vargs k m =>
check (list_norepet_dec ident_eq (var_names (fn_params f) ++ var_names (fn_vars f)));
let (e,m1) := do_alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) in
- do m2 <- sem_bind_parameters e m1 f.(fn_params) vargs;
+ do m2 <- sem_bind_parameters w e m1 f.(fn_params) vargs;
ret (State f f.(fn_body) k e m2)
| Callstate (External ef _ _) vargs k m =>
- match do_external _ _ ge ef w vargs m with
+ match do_external ef w vargs m with
| None => nil
| Some(w',t,v,m') => (t, Returnstate v k m') :: nil
end
@@ -1721,9 +1933,11 @@ Ltac myinv :=
Hint Extern 3 => exact I.
-Lemma do_step_sound:
- forall w S t S', In (t, S') (do_step w S) -> Csem.step ge S t S'.
-Proof with try (right; econstructor; eauto; fail).
+Theorem do_step_sound:
+ forall w S t S',
+ In (t, S') (do_step w S) ->
+ Csem.step ge S t S' \/ (t = E0 /\ S' = Stuckstate /\ can_crash_world w S).
+Proof with try (left; right; econstructor; eauto; fail).
intros until S'. destruct S; simpl.
(* State *)
destruct s; myinv...
@@ -1742,82 +1956,87 @@ Proof with try (right; econstructor; eauto; fail).
destruct k; myinv...
destruct v; myinv...
(* expression reduces *)
- destruct (step_expr e RV r m) as [ll|]_eqn; try contradiction. intros.
- exploit list_in_map_inv; eauto. intros [[C rd] [A B]].
- generalize (step_expr_sound e r RV m). unfold reducts_ok. rewrite Heqr0.
- destruct ll; try contradiction. intros SOUND.
- exploit SOUND; eauto. intros [CTX [a [EQ RD]]]. subst r.
- unfold expr_final_state in A. simpl in A. left.
- destruct rd; inv A; simpl in RD.
- apply step_lred. auto. apply step_expr_not_stuck; congruence. auto.
- apply step_rred. auto. apply step_expr_not_stuck; congruence. auto.
- destruct RD; subst m'. apply Csem.step_call. auto. apply step_expr_not_stuck; congruence. auto.
+ intros. exploit list_in_map_inv; eauto. intros [[C rd] [A B]].
+ 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.
+ (* lred *)
+ left; left; apply step_lred; auto.
+ (* stuck lred *)
+ exploit not_imm_safe_t; eauto. intros [R | R]; eauto.
+ (* rred *)
+ destruct RD. left; left; apply step_rred; auto.
+ (* callred *)
+ destruct RD; subst m'. left; left; apply step_call; eauto.
+ (* stuck rred *)
+ exploit not_imm_safe_t; eauto. intros [R | R]; eauto.
(* callstate *)
destruct fd; myinv.
(* internal *)
destruct (do_alloc_variables empty_env m (fn_params f ++ fn_vars f)) as [e m1]_eqn.
- myinv. right; apply step_internal_function with m1. auto.
+ 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. apply sem_bind_parameters_sound; auto.
+ apply do_alloc_variables_sound. eapply sem_bind_parameters_sound; eauto.
(* external *)
- destruct p as [[[w' tr] v] m']. myinv. right; constructor.
+ destruct p as [[[w' tr] v] m']. myinv. left; right; constructor.
eapply do_ef_external_sound; eauto.
(* returnstate *)
destruct k; myinv...
+(* stuckstate *)
+ contradiction.
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.
- assert (forall b from to C, context from to C -> (C = fun x => x) \/ is_val (C b) = None).
+ 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.
- destruct (H0 a0 _ _ _ H10). subst. inv H8; auto. auto.
- destruct (H0 a0 _ _ _ H10). subst. inv H8; auto. auto.
- destruct (H0 a0 _ _ _ H10). 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.
+ destruct (H0 a0 _ _ _ H9) as [[A B] | A]. subst. inv H8; auto. auto.
+ destruct (H0 a0 _ _ _ H8) as [[A B] | A]. subst. destruct a0; auto. elim H9. constructor. auto.
Qed.
-Lemma do_step_complete:
+Theorem do_step_complete:
forall w S t S' w', possible_trace w t w' -> Csem.step ge S t S' -> In (t, S') (do_step w S).
Proof with (unfold ret; auto with coqlib).
- intros until w'; intro PT; intros.
+ intros until w'; intros PT H.
destruct H.
(* Expression step *)
inversion H; subst; exploit estep_not_val; eauto; intro NOTVAL.
(* lred *)
unfold do_step; rewrite NOTVAL.
- destruct (step_expr e RV (C a) m) as [ll|]_eqn.
change (E0, ExprState f (C a') k e m') with (expr_final_state f k e (C, Lred a' m')).
apply in_map.
- generalize (step_expr_context e _ _ _ H2 a m). unfold reducts_incl.
- rewrite Heqr. destruct (step_expr e LV a m) as [ll'|]_eqn; try tauto.
- intro. replace C with (fun x => C x). apply H3.
- rewrite (lred_topred _ _ _ _ _ H0) in Heqr0. inv Heqr0. auto with coqlib.
+ generalize (step_expr_context e w _ _ _ H1 a m). unfold reducts_incl.
+ intro. replace C with (fun x => C x). apply H2.
+ rewrite (lred_topred _ _ _ _ _ _ H0). unfold topred; auto with coqlib.
apply extensionality; auto.
- exploit not_stuck_step_expr; eauto.
(* rred *)
unfold do_step; rewrite NOTVAL.
- destruct (step_expr e RV (C a) m) as [ll|]_eqn.
- change (E0, ExprState f (C a') k e m') with (expr_final_state f k e (C, Rred a' m')).
+ change (t, ExprState f (C a') k e m') with (expr_final_state f k e (C, Rred a' m' t)).
apply in_map.
- generalize (step_expr_context e _ _ _ H2 a m). unfold reducts_incl.
- rewrite Heqr. destruct (step_expr e RV a m) as [ll'|]_eqn; try tauto.
- intro. replace C with (fun x => C x). apply H3.
- rewrite (rred_topred _ _ _ _ _ H0) in Heqr0. inv Heqr0. auto with coqlib.
+ generalize (step_expr_context e w _ _ _ H1 a m). unfold reducts_incl.
+ intro. replace C with (fun x => C x). apply H2.
+ rewrite (rred_topred _ _ _ _ _ _ _ _ H0 PT). unfold topred; auto with coqlib.
apply extensionality; auto.
- exploit not_stuck_step_expr; eauto.
(* callred *)
unfold do_step; rewrite NOTVAL.
- destruct (step_expr e RV (C a) m) as [ll|]_eqn.
change (E0, Callstate fd vargs (Kcall f e C ty k) m) with (expr_final_state f k e (C, Callred fd vargs ty m)).
apply in_map.
- generalize (step_expr_context e _ _ _ H2 a m). unfold reducts_incl.
- rewrite Heqr. destruct (step_expr e RV a m) as [ll'|]_eqn; try tauto.
- intro. replace C with (fun x => C x). apply H3.
- rewrite (callred_topred _ _ _ _ _ _ H0) in Heqr0. inv Heqr0. auto with coqlib.
+ generalize (step_expr_context e w _ _ _ H1 a m). unfold reducts_incl.
+ intro. replace C with (fun x => C x). apply H2.
+ rewrite (callred_topred _ _ _ _ _ _ _ H0). unfold topred; auto with coqlib.
apply extensionality; auto.
- exploit not_stuck_step_expr; eauto.
+(* stuck *)
+ exploit not_imm_safe_stuck_red. eauto. red; intros; elim H1. eapply imm_safe_t_imm_safe. eauto.
+ instantiate (1 := w). intros [C' IN].
+ simpl do_step. rewrite NOTVAL.
+ change (E0, Stuckstate) with (expr_final_state f k e (C', Stuckred)).
+ apply in_map. auto.
(* Statement step *)
inv H; simpl...
@@ -1835,13 +2054,13 @@ Proof with (unfold ret; auto with coqlib).
destruct H0; subst x...
rewrite H0...
rewrite H0; rewrite H1...
- rewrite pred_dec_true; auto. rewrite H2. red in H0. destruct k; try contradiction...
+ rewrite H1. rewrite dec_eq_true. rewrite H2. red in H0. destruct k; try contradiction...
destruct H0; subst x...
rewrite H0...
(* Call step *)
rewrite pred_dec_true; auto. rewrite (do_alloc_variables_complete _ _ _ _ _ H1).
- rewrite (sem_bind_parameters_complete _ _ _ _ _ H2)...
+ rewrite (sem_bind_parameters_complete _ _ _ _ _ _ H2)...
exploit do_ef_external_complete; eauto. intro EQ; rewrite EQ. auto with coqlib.
Qed.
@@ -1854,7 +2073,7 @@ Definition do_initial_state (p: program): option (genv * state) :=
do m0 <- Genv.init_mem p;
do b <- Genv.find_symbol ge p.(prog_main);
do f <- Genv.find_funct_ptr ge b;
- check (type_eq (type_of_fundef f) (Tfunction Tnil (Tint I32 Signed)));
+ check (type_eq (type_of_fundef f) (Tfunction Tnil type_int32s));
Some (ge, Callstate f nil Kstop m0).
Definition at_final_state (S: state): option int :=
diff --git a/cfrontend/Clight.v b/cfrontend/Clight.v
index 76f6ff6..f0073a1 100644
--- a/cfrontend/Clight.v
+++ b/cfrontend/Clight.v
@@ -91,6 +91,7 @@ Inductive statement : Type :=
| Sskip : statement (**r do nothing *)
| Sassign : expr -> expr -> statement (**r assignment [lvalue = rvalue] *)
| Sset : ident -> expr -> statement (**r assignment [tempvar = rvalue] *)
+ | Svolread : ident -> expr -> statement (**r volatile read [tempvar = volatile lvalue] *)
| Scall: option ident -> expr -> list expr -> statement (**r function call *)
| Ssequence : statement -> statement -> statement (**r sequence *)
| Sifthenelse : expr -> statement -> statement -> statement (**r conditional *)
@@ -179,7 +180,6 @@ Definition empty_env: env := (PTree.empty (block * type)).
Definition temp_env := PTree.t val.
-
(** Selection of the appropriate case of a [switch], given the value [n]
of the selector expression. *)
@@ -260,8 +260,8 @@ Inductive eval_expr: expr -> val -> Prop :=
sem_cast v1 (typeof a) ty = Some v ->
eval_expr (Ecast a ty) v
| eval_Elvalue: forall a loc ofs v,
- eval_lvalue a loc ofs ->
- load_value_of_type (typeof a) m loc ofs = Some v ->
+ eval_lvalue a loc ofs -> type_is_volatile (typeof a) = false ->
+ deref_loc ge (typeof a) m loc ofs E0 v ->
eval_expr a v
(** [eval_lvalue ge e m a b ofs] defines the evaluation of expression [a]
@@ -280,14 +280,14 @@ with eval_lvalue: expr -> block -> int -> Prop :=
| eval_Ederef: forall a ty l ofs,
eval_expr a (Vptr l ofs) ->
eval_lvalue (Ederef a ty) l ofs
- | eval_Efield_struct: forall a i ty l ofs id fList delta,
- eval_lvalue a l ofs ->
- typeof a = Tstruct id fList ->
+ | eval_Efield_struct: forall a i ty l ofs id fList att delta,
+ eval_expr a (Vptr l ofs) ->
+ typeof a = Tstruct id fList att ->
field_offset i fList = OK delta ->
eval_lvalue (Efield a i ty) l (Int.add ofs (Int.repr delta))
- | eval_Efield_union: forall a i ty l ofs id fList,
- eval_lvalue a l ofs ->
- typeof a = Tunion id fList ->
+ | eval_Efield_union: forall a i ty l ofs id fList att,
+ eval_expr a (Vptr l ofs) ->
+ typeof a = Tunion id fList att ->
eval_lvalue (Efield a i ty) l ofs.
Scheme eval_expr_ind2 := Minimality for eval_expr Sort Prop
@@ -415,19 +415,26 @@ with find_label_ls (lbl: label) (sl: labeled_statements) (k: cont)
Inductive step: state -> trace -> state -> Prop :=
- | step_assign: forall f a1 a2 k e le m loc ofs v2 v m',
+ | step_assign: forall f a1 a2 k e le m loc ofs v2 v t m',
eval_lvalue e le m a1 loc ofs ->
eval_expr e le m a2 v2 ->
sem_cast v2 (typeof a2) (typeof a1) = Some v ->
- store_value_of_type (typeof a1) m loc ofs v = Some m' ->
+ assign_loc ge (typeof a1) m loc ofs v t m' ->
step (State f (Sassign a1 a2) k e le m)
- E0 (State f Sskip k e le m')
+ t (State f Sskip k e le m')
| step_set: forall f id a k e le m v,
eval_expr e le m a v ->
step (State f (Sset id a) k e le m)
E0 (State f Sskip k e (PTree.set id v le) m)
+ | step_vol_read: forall f id a k e le m loc ofs t v,
+ eval_lvalue e le m a loc ofs ->
+ deref_loc ge (typeof a) m loc ofs t v ->
+ type_is_volatile (typeof a) = true ->
+ step (State f (Svolread id a) k e le m)
+ t (State f Sskip k e (PTree.set id v le) m)
+
| step_call: forall f optid a al k e le m tyargs tyres vf vargs fd,
classify_fun (typeof a) = fun_case_f tyargs tyres ->
eval_expr e le m a vf ->
@@ -558,7 +565,7 @@ Inductive step: state -> trace -> state -> Prop :=
| step_internal_function: forall f vargs k m e m1 m2,
list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) ->
alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 ->
- bind_parameters e m1 f.(fn_params) vargs m2 ->
+ bind_parameters ge e m1 f.(fn_params) vargs m2 ->
step (Callstate (Internal f) vargs k m)
E0 (State f f.(fn_body) k e (PTree.empty val) m2)
@@ -622,17 +629,23 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env ->
| exec_Sskip: forall e le m,
exec_stmt e le m Sskip
E0 le m Out_normal
- | exec_Sassign: forall e le m a1 a2 loc ofs v2 v m',
+ | exec_Sassign: forall e le m a1 a2 loc ofs v2 v t m',
eval_lvalue e le m a1 loc ofs ->
eval_expr e le m a2 v2 ->
sem_cast v2 (typeof a2) (typeof a1) = Some v ->
- store_value_of_type (typeof a1) m loc ofs v = Some m' ->
+ assign_loc ge (typeof a1) m loc ofs v t m' ->
exec_stmt e le m (Sassign a1 a2)
- E0 le m' Out_normal
+ t le m' Out_normal
| exec_Sset: forall e le m id a v,
eval_expr 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_Svol_read: forall e le m id a loc ofs t v,
+ eval_lvalue e le m a loc ofs ->
+ type_is_volatile (typeof a) = true ->
+ deref_loc ge (typeof a) m loc ofs t v ->
+ exec_stmt e le m (Svolread id a)
+ t (PTree.set id v le) m Out_normal
| exec_Scall_none: forall e le m a al tyargs tyres vf vargs f t m' vres,
classify_fun (typeof a) = fun_case_f tyargs tyres ->
eval_expr e le m a vf ->
@@ -756,7 +769,7 @@ with eval_funcall: mem -> fundef -> list val -> trace -> mem -> val -> Prop :=
| eval_funcall_internal: forall le m f vargs t e m1 m2 m3 out vres m4,
alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 ->
list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) ->
- bind_parameters e m1 f.(fn_params) vargs m2 ->
+ bind_parameters ge e m1 f.(fn_params) vargs m2 ->
exec_stmt e (PTree.empty val) m2 f.(fn_body) t le m3 out ->
outcome_result_value out f.(fn_return) vres ->
Mem.free_list m3 (blocks_of_env e) = Some m4 ->
@@ -858,7 +871,7 @@ with evalinf_funcall: mem -> fundef -> list val -> traceinf -> Prop :=
| evalinf_funcall_internal: forall m f vargs t e m1 m2,
alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 ->
list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) ->
- bind_parameters e m1 f.(fn_params) vargs m2 ->
+ bind_parameters ge e m1 f.(fn_params) vargs m2 ->
execinf_stmt e (PTree.empty val) m2 f.(fn_body) t ->
evalinf_funcall m (Internal f) vargs t.
@@ -877,7 +890,7 @@ Inductive initial_state (p: program): state -> Prop :=
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
- type_of_fundef f = Tfunction Tnil (Tint I32 Signed) ->
+ type_of_fundef f = Tfunction Tnil type_int32s ->
initial_state p (Callstate f nil Kstop m0).
(** A final state is a [Returnstate] with an empty continuation. *)
@@ -901,10 +914,19 @@ Proof.
assert (t1 = E0 -> exists s2, step (Genv.globalenv p) s t2 s2).
intros. subst. inv H0. exists s1; auto.
inversion H; subst; auto.
+ (* assign *)
+ inv H5; auto. exploit volatile_store_receptive; eauto. intros EQ. subst t2; econstructor; eauto.
+ (* volatile read *)
+ inv H3; auto. exploit volatile_load_receptive; eauto. intros [v2 LD].
+ econstructor. eapply step_vol_read; eauto. eapply deref_loc_volatile; eauto.
+ (* external *)
exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
exists (Returnstate vres2 k m2). econstructor; eauto.
(* trace length *)
- inv H; simpl; try omega. eapply external_call_trace_length; eauto.
+ red; intros. inv H; simpl; try omega.
+ inv H3; simpl; try omega. inv H5; simpl; omega.
+ inv H1; simpl; try omega. inv H4; simpl; omega.
+ eapply external_call_trace_length; eauto.
Qed.
(** Big-step execution of a whole program. *)
@@ -915,7 +937,7 @@ Inductive bigstep_program_terminates (p: program): trace -> int -> Prop :=
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
- type_of_fundef f = Tfunction Tnil (Tint I32 Signed) ->
+ type_of_fundef f = Tfunction Tnil type_int32s ->
eval_funcall ge m0 f nil t m1 (Vint r) ->
bigstep_program_terminates p t r.
@@ -925,7 +947,7 @@ Inductive bigstep_program_diverges (p: program): traceinf -> Prop :=
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
- type_of_fundef f = Tfunction Tnil (Tint I32 Signed) ->
+ type_of_fundef f = Tfunction Tnil type_int32s ->
evalinf_funcall ge m0 f nil t ->
bigstep_program_diverges p t.
@@ -942,9 +964,9 @@ Let ge : genv := Genv.globalenv prog.
Definition exec_stmt_eval_funcall_ind
(PS: env -> temp_env -> mem -> statement -> trace -> temp_env -> mem -> outcome -> Prop)
(PF: mem -> fundef -> list val -> trace -> mem -> val -> Prop) :=
- fun a b c d e f g h i j k l m n o p q r s t u v w x =>
- conj (exec_stmt_ind2 ge PS PF a b c d e f g h i j k l m n o p q r s t u v w x)
- (eval_funcall_ind2 ge PS PF a b c d e f g h i j k l m n o p q r s t u v w x).
+ fun a b c d e f g h i j k l m n o p q r s t u v w x y =>
+ conj (exec_stmt_ind2 ge PS PF a b c d e f g h i j k l m n o p q r s t u v w x y)
+ (eval_funcall_ind2 ge PS PF a b c d e f g h i j k l m n o p q r s t u v w x y).
Inductive outcome_state_match
(e: env) (le: temp_env) (m: mem) (f: function) (k: cont): outcome -> state -> Prop :=
@@ -994,6 +1016,9 @@ Proof.
(* set *)
econstructor; split. apply star_one. econstructor; eauto. constructor.
+(* set volatile *)
+ econstructor; split. apply star_one. econstructor; eauto. constructor.
+
(* call none *)
econstructor; split.
eapply star_left. econstructor; eauto.
diff --git a/cfrontend/Cminorgen.v b/cfrontend/Cminorgen.v
index 3a8b857..a849a9a 100644
--- a/cfrontend/Cminorgen.v
+++ b/cfrontend/Cminorgen.v
@@ -62,11 +62,11 @@ Definition for_temp (id: ident) : ident := xI id.
global variables, stored in the global symbols with the same names. *)
Inductive var_info: Type :=
- | Var_local: memory_chunk -> var_info
- | Var_stack_scalar: memory_chunk -> Z -> var_info
- | Var_stack_array: Z -> var_info
- | Var_global_scalar: memory_chunk -> var_info
- | Var_global_array: var_info.
+ | Var_local (chunk: memory_chunk)
+ | Var_stack_scalar (chunk: memory_chunk) (ofs: Z)
+ | Var_stack_array (ofs sz al: Z)
+ | Var_global_scalar (chunk: memory_chunk)
+ | Var_global_array.
Definition compilenv := PMap.t var_info.
@@ -140,6 +140,7 @@ Definition make_unop (op: unary_operation) (e: expr): expr :=
Inductive approx : Type :=
| Any (**r any value *)
+ | Int1 (**r [0] or [1] *)
| Int7 (**r [[0,127]] *)
| Int8s (**r [[-128,127]] *)
| Int8u (**r [[0,255]] *)
@@ -153,41 +154,51 @@ Module Approx.
Definition bge (x y: approx) : bool :=
match x, y with
| Any, _ => true
- | Int7, Int7 => true
- | Int8s, (Int7 | Int8s) => true
- | Int8u, (Int7 | Int8u) => true
- | Int15, (Int7 | Int8u | Int15) => true
- | Int16s, (Int7 | Int8s | Int8u | Int15 | Int16s) => true
- | Int16u, (Int7 | Int8u | Int15 | Int16u) => true
+ | Int1, Int1 => true
+ | Int7, (Int1 | Int7) => true
+ | Int8s, (Int1 | Int7 | Int8s) => true
+ | Int8u, (Int1 | Int7 | Int8u) => true
+ | Int15, (Int1 | Int7 | Int8u | Int15) => true
+ | Int16s, (Int1 | Int7 | Int8s | Int8u | Int15 | Int16s) => true
+ | Int16u, (Int1 | Int7 | Int8u | Int15 | Int16u) => true
| Float32, Float32 => true
| _, _ => false
end.
Definition lub (x y: approx) : approx :=
match x, y with
+ | Int1, Int1 => Int1
+ | Int1, Int7 => Int7
+ | Int1, Int8u => Int8u
+ | Int1, Int8s => Int8s
+ | Int1, Int15 => Int15
+ | Int1, Int16u => Int16u
+ | Int1, Int16s => Int16s
+ | Int7, Int1 => Int7
| Int7, Int7 => Int7
| Int7, Int8u => Int8u
| Int7, Int8s => Int8s
| Int7, Int15 => Int15
| Int7, Int16u => Int16u
| Int7, Int16s => Int16s
- | Int8u, (Int7|Int8u) => Int8u
+ | Int8u, (Int1|Int7|Int8u) => Int8u
| Int8u, Int15 => Int15
| Int8u, Int16u => Int16u
| Int8u, Int16s => Int16s
- | Int8s, (Int7|Int8s) => Int8s
+ | Int8s, (Int1|Int7|Int8s) => Int8s
| Int8s, (Int15|Int16s) => Int16s
- | Int15, (Int7|Int8u|Int15) => Int15
+ | Int15, (Int1|Int7|Int8u|Int15) => Int15
| Int15, Int16u => Int16u
| Int15, (Int8s|Int16s) => Int16s
- | Int16u, (Int7|Int8u|Int15|Int16u) => Int16u
- | Int16s, (Int7|Int8u|Int8s|Int15|Int16s) => Int16s
+ | Int16u, (Int1|Int7|Int8u|Int15|Int16u) => Int16u
+ | Int16s, (Int1|Int7|Int8u|Int8s|Int15|Int16s) => Int16s
| Float32, Float32 => Float32
| _, _ => Any
end.
Definition of_int (n: int) :=
- if Int.eq_dec n (Int.zero_ext 7 n) then Int7
+ if Int.eq_dec n Int.zero || Int.eq_dec n Int.one then Int1
+ else if Int.eq_dec n (Int.zero_ext 7 n) then Int7
else if Int.eq_dec n (Int.zero_ext 8 n) then Int8u
else if Int.eq_dec n (Int.sign_ext 8 n) then Int8s
else if Int.eq_dec n (Int.zero_ext 15 n) then Int15
@@ -216,7 +227,8 @@ Definition unop (op: unary_operation) (a: approx) :=
| Ocast16unsigned => Int16u
| Ocast16signed => Int16s
| Osingleoffloat => Float32
- | Onotbool => Int7
+ | Oboolval => Int1
+ | Onotbool => Int1
| _ => Any
end.
@@ -226,17 +238,20 @@ Definition unop_is_redundant (op: unary_operation) (a: approx) :=
| Ocast8signed => bge Int8s a
| Ocast16unsigned => bge Int16u a
| Ocast16signed => bge Int16s a
+ | Oboolval => bge Int1 a
| Osingleoffloat => bge Float32 a
| _ => false
end.
Definition bitwise_and (a1 a2: approx) :=
- if bge Int8u a1 || bge Int8u a2 then Int8u
+ if bge Int1 a1 || bge Int1 a2 then Int1
+ else if bge Int8u a1 || bge Int8u a2 then Int8u
else if bge Int16u a1 || bge Int16u a2 then Int16u
else Any.
Definition bitwise_or (a1 a2: approx) :=
- if bge Int8u a1 && bge Int8u a2 then Int8u
+ if bge Int1 a1 && bge Int1 a2 then Int1
+ else if bge Int8u a1 && bge Int8u a2 then Int8u
else if bge Int16u a1 && bge Int16u a2 then Int16u
else Any.
@@ -244,9 +259,9 @@ Definition binop (op: binary_operation) (a1 a2: approx) :=
match op with
| Oand => bitwise_and a1 a2
| Oor | Oxor => bitwise_or a1 a2
- | Ocmp _ => Int7
- | Ocmpu _ => Int7
- | Ocmpf _ => Int7
+ | Ocmp _ => Int1
+ | Ocmpu _ => Int1
+ | Ocmpf _ => Int1
| _ => Any
end.
@@ -275,7 +290,7 @@ Definition var_addr (cenv: compilenv) (id: ident): res (expr * approx) :=
match PMap.get id cenv with
| Var_local chunk => Error(msg "Cminorgen.var_addr")
| Var_stack_scalar chunk ofs => OK (make_stackaddr ofs, Any)
- | Var_stack_array ofs => OK (make_stackaddr ofs, Any)
+ | Var_stack_array ofs sz al => OK (make_stackaddr ofs, Any)
| _ => OK (make_globaladdr id, Any)
end.
@@ -299,16 +314,17 @@ Definition var_set (cenv: compilenv)
(** A variant of [var_set] used for initializing function parameters.
The value to be stored already resides in the Cminor variable called [id]. *)
-Definition var_set_self (cenv: compilenv) (id: ident) (ty: typ) (k: stmt): res stmt :=
+Definition var_set_self (cenv: compilenv) (id: ident) (k: stmt): res stmt :=
match PMap.get id cenv with
| Var_local chunk =>
OK k
| Var_stack_scalar chunk ofs =>
OK (Sseq (make_store chunk (make_stackaddr ofs) (Evar (for_var id))) k)
- | Var_global_scalar chunk =>
- OK (Sseq (make_store chunk (make_globaladdr id) (Evar (for_var id))) k)
+ | Var_stack_array ofs sz al =>
+ OK (Sseq (Sbuiltin None (EF_memcpy sz (Zmin al 4))
+ (make_stackaddr ofs :: Evar (for_var id) :: nil)) k)
| _ =>
- Error(msg "Cminorgen.var_set_self.2")
+ Error(msg "Cminorgen.var_set_self")
end.
(** Translation of constants. *)
@@ -432,6 +448,9 @@ Fixpoint transl_stmt (ret: option typ) (cenv: compilenv)
do (te, a) <- transl_expr cenv e;
do tel <- transl_exprlist cenv el;
OK (Scall (option_map for_temp optid) sig te tel)
+ | Csharpminor.Sbuiltin optid ef el =>
+ do tel <- transl_exprlist cenv el;
+ OK (Sbuiltin (option_map for_temp optid) ef tel)
| Csharpminor.Sseq s1 s2 =>
do ts1 <- transl_stmt ret cenv xenv s1;
do ts2 <- transl_stmt ret cenv xenv s2;
@@ -515,6 +534,8 @@ Fixpoint addr_taken_stmt (s: Csharpminor.stmt): Identset.t :=
Identset.union (addr_taken_expr e1) (addr_taken_expr e2)
| Csharpminor.Scall optid sig e el =>
Identset.union (addr_taken_expr e) (addr_taken_exprlist el)
+ | Csharpminor.Sbuiltin optid ef el =>
+ addr_taken_exprlist el
| Csharpminor.Sseq s1 s2 =>
Identset.union (addr_taken_stmt s1) (addr_taken_stmt s2)
| Csharpminor.Sifthenelse e s1 s2 =>
@@ -554,9 +575,9 @@ Definition assign_variable
(cenv_stacksize: compilenv * Z) : compilenv * Z :=
let (cenv, stacksize) := cenv_stacksize in
match id_lv with
- | (id, Varray sz) =>
+ | (id, Varray sz al) =>
let ofs := align stacksize (array_alignment sz) in
- (PMap.set id (Var_stack_array ofs) cenv, ofs + Zmax 0 sz)
+ (PMap.set id (Var_stack_array ofs sz al) cenv, ofs + Zmax 0 sz)
| (id, Vscalar chunk) =>
if Identset.mem id atk then
let sz := size_chunk chunk in
@@ -589,7 +610,7 @@ Definition assign_global_variable
match info with
| (id, mkglobvar vk _ _ _) =>
PMap.set id (match vk with Vscalar chunk => Var_global_scalar chunk
- | Varray _ => Var_global_array
+ | Varray _ _ => Var_global_array
end)
ce
end.
@@ -605,13 +626,13 @@ Definition build_global_compilenv (p: Csharpminor.program) : compilenv :=
local variables.) *)
Fixpoint store_parameters
- (cenv: compilenv) (params: list (ident * memory_chunk))
+ (cenv: compilenv) (params: list (ident * var_kind))
{struct params} : res stmt :=
match params with
| nil => OK Sskip
- | (id, chunk) :: rem =>
+ | (id, vk) :: rem =>
do s <- store_parameters cenv rem;
- var_set_self cenv id (type_of_chunk chunk) s
+ var_set_self cenv id s
end.
(** Translation of a Csharpminor function. We must check that the
diff --git a/cfrontend/Cminorgenproof.v b/cfrontend/Cminorgenproof.v
index a6656e0..1a66ec9 100644
--- a/cfrontend/Cminorgenproof.v
+++ b/cfrontend/Cminorgenproof.v
@@ -61,6 +61,18 @@ Lemma functions_translated:
Genv.find_funct tge v = Some tf /\ transl_fundef gce f = OK tf.
Proof (Genv.find_funct_transf_partial2 (transl_fundef gce) transl_globvar _ TRANSL).
+Lemma var_info_translated:
+ forall b v,
+ Genv.find_var_info ge b = Some v ->
+ exists tv, Genv.find_var_info tge b = Some tv /\ transf_globvar transl_globvar v = OK tv.
+Proof (Genv.find_var_info_transf_partial2 (transl_fundef gce) transl_globvar _ TRANSL).
+
+Lemma var_info_rev_translated:
+ forall b tv,
+ Genv.find_var_info tge b = Some tv ->
+ exists v, Genv.find_var_info ge b = Some v /\ transf_globvar transl_globvar v = OK tv.
+Proof (Genv.find_var_info_rev_transf_partial2 (transl_fundef gce) transl_globvar _ TRANSL).
+
Lemma sig_preserved_body:
forall f tf cenv size,
transl_funbody cenv size f = OK tf ->
@@ -245,10 +257,10 @@ Inductive match_var (f: meminj) (id: ident)
val_inject f (Vptr b Int.zero) (Vptr sp (Int.repr ofs)) ->
match_var f id e m te sp (Var_stack_scalar chunk ofs)
| match_var_stack_array:
- forall ofs sz b,
- PTree.get id e = Some (b, Varray sz) ->
+ forall ofs sz al b,
+ PTree.get id e = Some (b, Varray sz al) ->
val_inject f (Vptr b Int.zero) (Vptr sp (Int.repr ofs)) ->
- match_var f id e m te sp (Var_stack_array ofs)
+ match_var f id e m te sp (Var_stack_array ofs sz al)
| match_var_global_scalar:
forall chunk,
PTree.get id e = None ->
@@ -463,8 +475,8 @@ Inductive alloc_condition: var_info -> var_kind -> block -> option (block * Z) -
alloc_condition (Var_local chunk) (Vscalar chunk) sp None
| alloc_cond_stack_scalar: forall chunk pos sp,
alloc_condition (Var_stack_scalar chunk pos) (Vscalar chunk) sp (Some(sp, pos))
- | alloc_cond_stack_array: forall pos sz sp,
- alloc_condition (Var_stack_array pos) (Varray sz) sp (Some(sp, pos)).
+ | alloc_cond_stack_array: forall pos sz al sp,
+ alloc_condition (Var_stack_array pos sz al) (Varray sz al) sp (Some(sp, pos)).
Lemma match_env_alloc_same:
forall f1 cenv e le m1 te sp lo lv m2 b f2 id info tv,
@@ -1156,6 +1168,7 @@ Qed.
Definition val_match_approx (a: approx) (v: val) : Prop :=
match a with
+ | Int1 => v = Val.zero_ext 1 v
| Int7 => v = Val.zero_ext 8 v /\ v = Val.sign_ext 8 v
| Int8u => v = Val.zero_ext 8 v
| Int8s => v = Val.sign_ext 8 v
@@ -1187,8 +1200,16 @@ Proof.
intros. rewrite H.
destruct v; simpl; auto. decEq. symmetry.
apply Int.sign_zero_ext_widen. compute; auto. split. omega. compute; auto.
+ assert (D: forall v, v = Val.zero_ext 1 v -> v = Val.zero_ext 8 v).
+ intros. rewrite H.
+ destruct v; simpl; auto. decEq. symmetry.
+ apply Int.zero_ext_widen. compute; auto. split. omega. compute; auto.
+ assert (E: forall v, v = Val.zero_ext 1 v -> v = Val.sign_ext 8 v).
+ intros. rewrite H.
+ destruct v; simpl; auto. decEq. symmetry.
+ apply Int.sign_zero_ext_widen. compute; auto. split. omega. compute; auto.
intros.
- unfold Approx.bge in H; destruct a1; try discriminate; destruct a2; simpl in *; try discriminate; intuition; auto.
+ unfold Approx.bge in H; destruct a1; try discriminate; destruct a2; simpl in *; try discriminate; intuition.
Qed.
Lemma approx_lub_ge_left:
@@ -1206,7 +1227,9 @@ Qed.
Lemma approx_of_int_sound:
forall n, val_match_approx (Approx.of_int n) (Vint n).
Proof.
- unfold Approx.of_int; intros.
+ unfold Approx.of_int; intros.
+ destruct (Int.eq_dec n Int.zero); simpl. subst; auto.
+ destruct (Int.eq_dec n Int.one); simpl. subst; auto.
destruct (Int.eq_dec n (Int.zero_ext 7 n)). simpl.
split.
decEq. rewrite e. symmetry. apply Int.zero_ext_widen. compute; auto. split. omega. compute; auto.
@@ -1249,7 +1272,8 @@ Proof.
destruct v1; simpl; auto. rewrite Int.sign_ext_idem; auto. compute; auto.
destruct v1; simpl; auto. rewrite Int.zero_ext_idem; auto. compute; auto.
destruct v1; simpl; auto. rewrite Int.sign_ext_idem; auto. compute; auto.
- destruct v1; simpl; auto. destruct (Int.eq i Int.zero); auto.
+ destruct v1; simpl; auto. destruct (Int.eq i Int.zero); simpl; auto.
+ destruct v1; simpl; auto. destruct (Int.eq i Int.zero); simpl; auto.
destruct v1; simpl; auto. rewrite Float.singleoffloat_idem; auto.
Qed.
@@ -1273,7 +1297,13 @@ Proof.
assert (Q: forall a v, val_match_approx a v -> Approx.bge Int16u a = true ->
v = Val.zero_ext 16 v).
intros. apply (val_match_approx_increasing Int16u a v); auto.
+ assert (R: forall a v, val_match_approx a v -> Approx.bge Int1 a = true ->
+ v = Val.zero_ext 1 v).
+ intros. apply (val_match_approx_increasing Int1 a v); auto.
+
intros; unfold Approx.bitwise_and.
+ destruct (Approx.bge Int1 a1) as []_eqn. simpl. apply Y; eauto. compute; auto.
+ destruct (Approx.bge Int1 a2) as []_eqn. simpl. apply X; eauto. compute; auto.
destruct (Approx.bge Int8u a1) as []_eqn. simpl. apply Y; eauto. compute; auto.
destruct (Approx.bge Int8u a2) as []_eqn. simpl. apply X; eauto. compute; auto.
destruct (Approx.bge Int16u a1) as []_eqn. simpl. apply Y; eauto. compute; auto.
@@ -1298,14 +1328,21 @@ Proof.
simpl. rewrite Int.and_idem. auto.
unfold Approx.bitwise_or.
- destruct (Approx.bge Int8u a1 && Approx.bge Int8u a2) as []_eqn.
+
+ destruct (Approx.bge Int1 a1 && Approx.bge Int1 a2) as []_eqn.
destruct (andb_prop _ _ Heqb).
simpl. apply X. compute; auto.
+ apply (val_match_approx_increasing Int1 a1 v1); auto.
+ apply (val_match_approx_increasing Int1 a2 v2); auto.
+
+ destruct (Approx.bge Int8u a1 && Approx.bge Int8u a2) as []_eqn.
+ destruct (andb_prop _ _ Heqb0).
+ simpl. apply X. compute; auto.
apply (val_match_approx_increasing Int8u a1 v1); auto.
apply (val_match_approx_increasing Int8u a2 v2); auto.
destruct (Approx.bge Int16u a1 && Approx.bge Int16u a2) as []_eqn.
- destruct (andb_prop _ _ Heqb0).
+ destruct (andb_prop _ _ Heqb1).
simpl. apply X. compute; auto.
apply (val_match_approx_increasing Int16u a1 v1); auto.
apply (val_match_approx_increasing Int16u a2 v2); auto.
@@ -1319,7 +1356,7 @@ Lemma approx_of_binop_sound:
val_match_approx a1 v1 -> val_match_approx a2 v2 ->
val_match_approx (Approx.binop op a1 a2) v.
Proof.
- assert (OB: forall ob, val_match_approx Int7 (Val.of_optbool ob)).
+ assert (OB: forall ob, val_match_approx Int1 (Val.of_optbool ob)).
destruct ob; simpl. destruct b; auto. auto.
destruct op; intros; simpl Approx.binop; simpl in H; try (exact I); inv H.
@@ -1356,6 +1393,20 @@ Proof.
(* cast16signed *)
assert (V: val_match_approx Int16s v) by (eapply val_match_approx_increasing; eauto).
simpl in *. congruence.
+(* boolval *)
+ assert (V: val_match_approx Int1 v) by (eapply val_match_approx_increasing; eauto).
+ simpl in *.
+ assert (v = Vundef \/ v = Vzero \/ v = Vone).
+ rewrite V. destruct v; simpl; auto.
+ assert (0 <= Int.unsigned (Int.zero_ext 1 i) < 2).
+ apply Int.zero_ext_range. compute; auto.
+ assert (Int.unsigned(Int.zero_ext 1 i) = 0 \/ Int.unsigned(Int.zero_ext 1 i) = 1) by omega.
+ destruct H2.
+ assert (Int.repr (Int.unsigned (Int.zero_ext 1 i)) = Int.repr 0) by congruence.
+ rewrite Int.repr_unsigned in H3. rewrite H3; auto.
+ assert (Int.repr (Int.unsigned (Int.zero_ext 1 i)) = Int.repr 1) by congruence.
+ rewrite Int.repr_unsigned in H3. rewrite H3; auto.
+ intuition; subst v; auto.
(* singleoffloat *)
assert (V: val_match_approx Float32 v) by (eapply val_match_approx_increasing; eauto).
simpl in *. congruence.
@@ -1401,6 +1452,7 @@ Proof.
inv H; inv H0; simpl; TrivialExists.
inv H; inv H0; simpl; TrivialExists.
inv H; inv H0; simpl; TrivialExists.
+ inv H; inv H0; simpl; TrivialExists. apply val_inject_val_of_bool.
inv H; inv H0; simpl; TrivialExists.
inv H; inv H0; simpl; TrivialExists. apply val_inject_val_of_bool.
inv H; inv H0; simpl; TrivialExists.
@@ -1946,9 +1998,9 @@ Proof.
apply match_env_extensional with te1; auto.
Qed.
-Lemma var_set_self_correct:
- forall cenv id ty s a f tf e le te sp lo hi m cs tm tv v m' fn k,
- var_set_self cenv id ty s = OK a ->
+Lemma var_set_self_correct_scalar:
+ forall cenv id s a f tf e le te sp lo hi m cs tm tv v m' fn k,
+ var_set_self cenv id s = OK a ->
match_callstack f m tm (Frame cenv tf e le te sp lo hi :: cs) (Mem.nextblock m) (Mem.nextblock tm) ->
val_inject f v tv ->
Mem.inject f m tm ->
@@ -1995,20 +2047,55 @@ Proof.
split. auto.
rewrite NEXTBLOCK. rewrite (nextblock_storev _ _ _ _ _ STORE').
eapply match_callstack_storev_mapped; eauto.
- (* var_global_scalar *)
- simpl in *.
- assert (chunk0 = chunk). exploit H4; eauto. congruence. subst chunk0.
- assert (Mem.storev chunk m (Vptr b Int.zero) v = Some m'). assumption.
- exploit match_callstack_match_globalenvs; eauto. intros [bnd MG]. inv MG.
- exploit make_store_correct.
- eapply make_globaladdr_correct; eauto.
- rewrite symbols_preserved; eauto. eauto. eauto. eauto. eauto. eauto.
- intros [tm' [tvrhs' [EVAL' [STORE' MEMINJ]]]].
- exists tm'.
- split. eapply star_three. constructor. eauto. constructor. traceEq.
- split. auto.
- rewrite NEXTBLOCK. rewrite (nextblock_storev _ _ _ _ _ STORE').
- eapply match_callstack_store_mapped; eauto.
+Qed.
+
+Lemma var_set_self_correct_array:
+ forall cenv id s a f tf e le te sp lo hi m cs tm tv b v sz al m' fn k,
+ var_set_self cenv id s = OK a ->
+ match_callstack f m tm (Frame cenv tf e le te sp lo hi :: cs) (Mem.nextblock m) (Mem.nextblock tm) ->
+ val_inject f v tv ->
+ Mem.inject f m tm ->
+ PTree.get id e = Some(b, Varray sz al) ->
+ extcall_memcpy_sem sz (Zmin al 4) ge
+ (Vptr b Int.zero :: v :: nil) m E0 Vundef m' ->
+ te!(for_var id) = Some tv ->
+ exists f', exists tm',
+ star step tge (State fn a k (Vptr sp Int.zero) te tm)
+ E0 (State fn s k (Vptr sp Int.zero) te tm') /\
+ Mem.inject f' m' tm' /\
+ match_callstack f' m' tm' (Frame cenv tf e le te sp lo hi :: cs) (Mem.nextblock m') (Mem.nextblock tm') /\
+ inject_incr f f'.
+Proof.
+ intros until k.
+ intros VS MCS VINJ MINJ KIND MEMCPY VAL.
+ assert (MV: match_var f id e m te sp cenv!!id).
+ inv MCS. inv MENV. auto.
+ inv MV; try congruence. rewrite KIND in H0; inv H0.
+ (* var_stack_array *)
+ unfold var_set_self in VS. rewrite <- H in VS. inv VS.
+ exploit match_callstack_match_globalenvs; eauto. intros [hi' MG].
+ assert (external_call (EF_memcpy sz0 (Zmin al0 4)) ge (Vptr b0 Int.zero :: v :: nil) m E0 Vundef m').
+ assumption.
+ exploit external_call_mem_inject; eauto.
+ eapply inj_preserves_globals; eauto.
+ intros [f' [vres' [tm' [EC' [VINJ' [MINJ' [UNMAPPED [OUTOFREACH [INCR SEPARATED]]]]]]]]].
+ exists f'; exists tm'.
+ split. eapply star_step. constructor.
+ eapply star_step. econstructor; eauto.
+ constructor. apply make_stackaddr_correct. constructor. constructor. eauto. constructor.
+ eapply external_call_symbols_preserved_2; eauto.
+ exact symbols_preserved.
+ eexact var_info_translated.
+ eexact var_info_rev_translated.
+ apply star_one. constructor. reflexivity. traceEq.
+ split. auto.
+ split. apply match_callstack_incr_bound with (Mem.nextblock m) (Mem.nextblock tm).
+ eapply match_callstack_external_call; eauto.
+ intros. eapply external_call_bounds; eauto.
+ omega. omega.
+ eapply external_call_nextblock_incr; eauto.
+ eapply external_call_nextblock_incr; eauto.
+ auto.
Qed.
(** * Correctness of stack allocation of local variables *)
@@ -2032,16 +2119,15 @@ Remark assign_variable_incr:
Proof.
intros until sz'; simpl.
destruct lv. case (Identset.mem id atk); intros.
- inv H. generalize (size_chunk_pos m). intro.
- generalize (align_le sz (size_chunk m) H). omega.
+ inv H. generalize (size_chunk_pos chunk). intro.
+ generalize (align_le sz (size_chunk chunk) H). omega.
inv H. omega.
intros. inv H.
- generalize (align_le sz (array_alignment z) (array_alignment_pos z)).
- assert (0 <= Zmax 0 z). apply Zmax_bound_l. omega.
+ generalize (align_le sz (array_alignment sz0) (array_alignment_pos sz0)).
+ assert (0 <= Zmax 0 sz0). apply Zmax_bound_l. omega.
omega.
Qed.
-
Remark assign_variables_incr:
forall atk vars cenv sz cenv' sz',
assign_variables atk vars (cenv, sz) = (cenv', sz') -> sz <= sz'.
@@ -2166,7 +2252,7 @@ Proof.
subst b0. congruence.
rewrite OTHER in H; eauto.
(* 2 info = Var_stack_array ofs *)
- intros dim LV EQ. injection EQ; clear EQ; intros. rewrite <- H0.
+ intros dim al LV EQ. injection EQ; clear EQ; intros. rewrite <- H.
assert (0 <= Zmax 0 dim). apply Zmax1.
generalize (align_le sz (array_alignment dim) (array_alignment_pos dim)). intro.
set (ofs := align sz (array_alignment dim)) in *.
@@ -2181,7 +2267,7 @@ Proof.
intros. generalize (RANGE _ _ H3). omega.
intros [f1 [MINJ1 [INCR1 [SAME OTHER]]]].
exists f1; split. auto. split. auto. split.
- eapply match_callstack_alloc_left; eauto.
+ subst cenv'. eapply match_callstack_alloc_left; eauto.
rewrite <- LV; auto.
rewrite SAME; constructor.
intros. rewrite (Mem.bounds_alloc _ _ _ _ _ ALLOC).
@@ -2317,14 +2403,12 @@ Proof.
eapply Mem.valid_block_inject_2; eauto.
intros. unfold te. apply set_locals_params_defined.
elim (in_app_or _ _ _ H6); intros.
- elim (list_in_map_inv _ _ _ H7). intros x [A B].
- apply in_or_app; left. unfold tparams. apply List.in_map. inversion A. apply List.in_map. auto.
+ apply in_or_app; left. unfold tparams. apply List.in_map.
+ change id with (fst (id, lv)). apply List.in_map. auto.
apply in_or_app; right. apply in_or_app; left. unfold tvars. apply List.in_map.
change id with (fst (id, lv)). apply List.in_map; auto.
(* norepet *)
- unfold fn_variables.
- rewrite List.map_app. rewrite list_map_compose. simpl.
- assumption.
+ unfold fn_variables. rewrite List.map_app. assumption.
(* undef *)
intros. unfold empty_env. apply PTree.gempty.
Qed.
@@ -2333,17 +2417,23 @@ Qed.
to store in memory the values of parameters that are stack-allocated. *)
Inductive vars_vals_match (f:meminj):
- list (ident * memory_chunk) -> list val -> env -> Prop :=
+ list (ident * var_kind) -> list val -> env -> Prop :=
| vars_vals_nil:
forall te,
vars_vals_match f nil nil te
- | vars_vals_cons:
+ | vars_vals_scalar:
forall te id chunk vars v vals tv,
te!(for_var id) = Some tv ->
val_inject f v tv ->
val_normalized v chunk ->
vars_vals_match f vars vals te ->
- vars_vals_match f ((id, chunk) :: vars) (v :: vals) te.
+ vars_vals_match f ((id, Vscalar chunk) :: vars) (v :: vals) te
+ | vars_vals_array:
+ forall te id sz al vars v vals tv,
+ te!(for_var id) = Some tv ->
+ val_inject f v tv ->
+ vars_vals_match f vars vals te ->
+ vars_vals_match f ((id, Varray sz al) :: vars) (v :: vals) te.
Lemma vars_vals_match_extensional:
forall f vars vals te,
@@ -2357,88 +2447,120 @@ Proof.
econstructor; eauto.
rewrite <- H. eauto with coqlib.
apply IHvars_vals_match. intros. eapply H3; eauto with coqlib.
+ econstructor; eauto.
+ rewrite <- H. eauto with coqlib.
+ apply IHvars_vals_match. intros. eapply H2; eauto with coqlib.
+Qed.
+
+Lemma vars_vals_match_incr:
+ forall f f', inject_incr f f' ->
+ forall vars vals te,
+ vars_vals_match f vars vals te ->
+ vars_vals_match f' vars vals te.
+Proof.
+ induction 2; intros; econstructor; eauto.
Qed.
Lemma store_parameters_correct:
forall e le te m1 params vl m2,
- bind_parameters e m1 params vl m2 ->
+ bind_parameters ge e m1 params vl m2 ->
forall s f cenv tf sp lo hi cs tm1 fn k,
vars_vals_match f params vl te ->
- list_norepet (List.map param_name params) ->
+ list_norepet (List.map variable_name params) ->
Mem.inject f m1 tm1 ->
match_callstack f m1 tm1 (Frame cenv tf e le te sp lo hi :: cs) (Mem.nextblock m1) (Mem.nextblock tm1) ->
store_parameters cenv params = OK s ->
- exists tm2,
+ exists f', exists tm2,
star step tge (State fn s k (Vptr sp Int.zero) te tm1)
E0 (State fn Sskip k (Vptr sp Int.zero) te tm2)
- /\ Mem.inject f m2 tm2
- /\ match_callstack f m2 tm2 (Frame cenv tf e le te sp lo hi :: cs) (Mem.nextblock m2) (Mem.nextblock tm2).
+ /\ Mem.inject f' m2 tm2
+ /\ match_callstack f' m2 tm2 (Frame cenv tf e le te sp lo hi :: cs) (Mem.nextblock m2) (Mem.nextblock tm2)
+ /\ inject_incr f f'.
Proof.
induction 1.
(* base case *)
intros; simpl. monadInv H3.
- exists tm1. split. constructor. tauto.
- (* inductive case *)
+ exists f; exists tm1. split. constructor. auto.
+ (* scalar case *)
intros until k. intros VVM NOREPET MINJ MATCH STOREP.
- monadInv STOREP.
- inv VVM.
- inv NOREPET.
- exploit var_set_self_correct; eauto.
+ monadInv STOREP. inv VVM. inv NOREPET.
+ exploit var_set_self_correct_scalar; eauto.
econstructor; eauto. econstructor; eauto.
intros [tm2 [EXEC1 [MINJ1 MATCH1]]].
exploit IHbind_parameters; eauto.
- intros [tm3 [EXEC2 [MINJ2 MATCH2]]].
- exists tm3.
+ intros [f' [tm3 [EXEC2 [MINJ2 [MATCH2 INCR2]]]]].
+ exists f'; exists tm3.
split. eapply star_trans; eauto.
auto.
-Qed.
+ (* array case *)
+ intros until k. intros VVM NOREPET MINJ MATCH STOREP.
+ monadInv STOREP. inv VVM. inv NOREPET.
+ exploit var_set_self_correct_array; eauto.
+ intros [f2 [tm2 [EXEC1 [MINJ1 [MATCH1 INCR1]]]]].
+ exploit IHbind_parameters. eapply vars_vals_match_incr; eauto. auto. eauto. eauto. eauto.
+ intros [f3 [tm3 [EXEC2 [MINJ2 [MATCH2 INCR2]]]]].
+ exists f3; exists tm3.
+ split. eapply star_trans; eauto.
+ split. auto. split. auto. eapply inject_incr_trans; eauto.
+Qed.
+
+Definition val_normalized' (v: val) (vk: var_kind) : Prop :=
+ match vk with
+ | Vscalar chunk => val_normalized v chunk
+ | Varray _ _ => True
+ end.
Lemma vars_vals_match_holds_1:
forall f params args targs,
- list_norepet (List.map param_name params) ->
+ list_norepet (List.map variable_name params) ->
val_list_inject f args targs ->
- list_forall2 val_normalized args (List.map param_chunk params) ->
+ list_forall2 val_normalized' args (List.map variable_kind params) ->
vars_vals_match f params args
- (set_params targs (List.map for_var (List.map param_name params))).
+ (set_params targs (List.map for_var (List.map variable_name params))).
Proof.
Opaque for_var.
induction params; simpl; intros.
inv H1. constructor.
inv H. inv H1. inv H0.
- destruct a as [id chunk]; simpl in *. econstructor.
- rewrite PTree.gss. reflexivity.
- auto. auto.
+ destruct a as [id vk]; simpl in *.
+ assert (R: vars_vals_match f params al
+ (PTree.set (for_var id) v'
+ (set_params vl' (map for_var (map variable_name params))))).
apply vars_vals_match_extensional
- with (set_params vl' (map for_var (map param_name params))).
+ with (set_params vl' (map for_var (map variable_name params))).
eapply IHparams; eauto.
Transparent for_var.
intros. apply PTree.gso. unfold for_var; red; intros. inv H0.
- elim H4. change id with (param_name (id, lv)). apply List.in_map; auto.
+ elim H4. change id with (variable_name (id, lv)). apply List.in_map; auto.
+
+ destruct vk; red in H6.
+ econstructor. rewrite PTree.gss. reflexivity. auto. auto. auto.
+ econstructor. rewrite PTree.gss. reflexivity. auto. auto.
Qed.
Lemma vars_vals_match_holds_2:
forall f params args e,
vars_vals_match f params args e ->
forall vl,
- (forall id1 id2, In id1 (List.map param_name params) -> In id2 vl -> for_var id1 <> id2) ->
+ (forall id1 id2, In id1 (List.map variable_name params) -> In id2 vl -> for_var id1 <> id2) ->
vars_vals_match f params args (set_locals vl e).
Proof.
induction vl; simpl; intros.
auto.
apply vars_vals_match_extensional with (set_locals vl e); auto.
intros. apply PTree.gso. apply H0.
- change id with (param_name (id, lv)). apply List.in_map. auto.
+ change id with (variable_name (id, lv)). apply List.in_map. auto.
auto.
Qed.
Lemma vars_vals_match_holds:
forall f params args targs vars temps,
- list_norepet (List.map param_name params ++ vars) ->
+ list_norepet (List.map variable_name params ++ vars) ->
val_list_inject f args targs ->
- list_forall2 val_normalized args (List.map param_chunk params) ->
+ list_forall2 val_normalized' args (List.map variable_kind params) ->
vars_vals_match f params args
(set_locals (List.map for_var vars ++ List.map for_temp temps)
- (set_params targs (List.map for_var (List.map param_name params)))).
+ (set_params targs (List.map for_var (List.map variable_name params)))).
Proof.
intros. rewrite list_norepet_app in H. destruct H as [A [B C]].
apply vars_vals_match_holds_2; auto. apply vars_vals_match_holds_1; auto.
@@ -2452,12 +2574,13 @@ Qed.
Remark bind_parameters_normalized:
forall e m params args m',
- bind_parameters e m params args m' ->
- list_forall2 val_normalized args (List.map param_chunk params).
+ bind_parameters ge e m params args m' ->
+ list_forall2 val_normalized' args (List.map variable_kind params).
Proof.
induction 1; simpl.
constructor.
constructor; auto.
+ constructor; auto. red; auto.
Qed.
(** The main result in this section: the behaviour of function entry
@@ -2470,7 +2593,7 @@ Lemma function_entry_ok:
forall fn m e m1 vargs m2 f cs tm cenv tf tm1 sp tvargs s fn' k,
list_norepet (fn_params_names fn ++ fn_vars_names fn) ->
alloc_variables empty_env m (fn_variables fn) e m1 ->
- bind_parameters e m1 fn.(Csharpminor.fn_params) vargs m2 ->
+ bind_parameters ge e m1 fn.(Csharpminor.fn_params) vargs m2 ->
match_callstack f m tm cs (Mem.nextblock m) (Mem.nextblock tm) ->
build_compilenv gce fn = (cenv, tf.(fn_stackspace)) ->
tf.(fn_stackspace) <= Int.max_unsigned ->
@@ -2504,8 +2627,9 @@ Proof.
exploit store_parameters_correct.
eauto. eauto. eapply list_norepet_append_left; eauto.
eexact MINJ1. eexact MATCH1. eauto.
- intros [tm2 [EXEC [MINJ2 MATCH2]]].
- exists f1; exists tm2. eauto.
+ intros [f2 [tm2 [EXEC [MINJ2 [MATCH2 INCR2]]]]].
+ exists f2; exists tm2.
+ split; eauto. split; auto. split; auto. eapply inject_incr_trans; eauto.
Qed.
(** * Semantic preservation for the translation *)
@@ -2888,8 +3012,8 @@ Proof.
Qed.
Remark find_label_var_set_self:
- forall id ty s0 s k,
- var_set_self cenv id ty s0 = OK s ->
+ forall id s0 s k,
+ var_set_self cenv id s0 = OK s ->
find_label lbl s k = find_label lbl s0 k.
Proof.
intros. unfold var_set_self in H.
@@ -3114,6 +3238,35 @@ Proof.
eapply match_Kcall with (cenv' := cenv); eauto.
red; auto.
+(* builtin *)
+ monadInv TR.
+ exploit transl_exprlist_correct; eauto.
+ intros [tvargs [EVAL2 VINJ2]].
+ exploit match_callstack_match_globalenvs; eauto. intros [hi' MG].
+ 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.
+ eapply external_call_symbols_preserved_2; eauto.
+ exact symbols_preserved.
+ eexact var_info_translated.
+ eexact var_info_rev_translated.
+ assert (MCS': match_callstack f' m' tm'
+ (Frame cenv tfn e le te sp lo hi :: cs)
+ (Mem.nextblock m') (Mem.nextblock tm')).
+ apply match_callstack_incr_bound with (Mem.nextblock m) (Mem.nextblock tm).
+ eapply match_callstack_external_call; eauto.
+ intros. eapply external_call_bounds; eauto.
+ omega. omega.
+ eapply external_call_nextblock_incr; eauto.
+ eapply external_call_nextblock_incr; eauto.
+ econstructor; eauto.
+Opaque PTree.set.
+ unfold set_optvar. destruct optid; simpl.
+ eapply match_callstack_set_temp; eauto.
+ auto.
+
(* seq *)
monadInv TR.
left; econstructor; split.
@@ -3256,8 +3409,8 @@ Proof.
apply plus_one. econstructor.
eapply external_call_symbols_preserved_2; eauto.
exact symbols_preserved.
- eexact (Genv.find_var_info_transf_partial2 (transl_fundef gce) transl_globvar _ TRANSL).
- eexact (Genv.find_var_info_rev_transf_partial2 (transl_fundef gce) transl_globvar _ TRANSL).
+ eexact var_info_translated.
+ eexact var_info_rev_translated.
econstructor; eauto.
apply match_callstack_incr_bound with (Mem.nextblock m) (Mem.nextblock tm).
eapply match_callstack_external_call; eauto.
diff --git a/cfrontend/Cparser.mlpack b/cfrontend/Cparser.mlpack
index e6bbdc6..b59e30f 100644
--- a/cfrontend/Cparser.mlpack
+++ b/cfrontend/Cparser.mlpack
@@ -16,12 +16,8 @@ cparser/Elab
cparser/Rename
cparser/Transform
cparser/Unblock
-cparser/SimplExpr
-cparser/AddCasts
-cparser/StructByValue
-cparser/StructAssign
+cparser/StructReturn
cparser/Bitfields
cparser/PackedStructs
-cparser/SimplVolatile
cparser/Parse
diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v
index 426a753..49b2062 100644
--- a/cfrontend/Csem.v
+++ b/cfrontend/Csem.v
@@ -41,6 +41,7 @@ Definition cast_int_int (sz: intsize) (sg: signedness) (i: int) : int :=
| I16, Signed => Int.sign_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.
Definition cast_int_float (si : signedness) (i: int) : float :=
@@ -92,6 +93,22 @@ Function sem_cast (v: val) (t1 t2: type) : option val :=
end
| _ => None
end
+ | cast_case_ip2bool =>
+ match v with
+ | Vint i => Some (Vint (cast_int_int IBool Signed i))
+ | Vptr _ _ => Some (Vint Int.one)
+ | _ => None
+ end
+ | cast_case_f2bool =>
+ match v with
+ | Vfloat f =>
+ Some(Vint(if Float.cmp Ceq f Float.zero then Int.zero else Int.one))
+ | _ => None
+ end
+ | cast_case_struct id1 fld1 id2 fld2 =>
+ if ident_eq id1 id2 && fieldlist_eq fld1 fld2 then Some v else None
+ | cast_case_union id1 fld1 id2 fld2 =>
+ if ident_eq id1 id2 && fieldlist_eq fld1 fld2 then Some v else None
| cast_case_void =>
Some v
| cast_case_default =>
@@ -105,11 +122,9 @@ Function sem_cast (v: val) (t1 t2: type) : option val :=
Function bool_val (v: val) (t: type) : option bool :=
match v, t with
- | Vint n, Tint sz sg => Some (negb (Int.eq n Int.zero))
- | Vint n, Tpointer t' => Some (negb (Int.eq n Int.zero))
- | Vptr b ofs, Tint sz sg => Some true
- | Vptr b ofs, Tpointer t' => Some true
- | Vfloat f, Tfloat sz => Some (negb(Float.cmp Ceq f Float.zero))
+ | Vint n, (Tint _ _ _ | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _) => Some (negb (Int.eq n Int.zero))
+ | Vptr b ofs, (Tint _ _ _ | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _) => Some true
+ | Vfloat f, Tfloat sz _ => Some (negb(Float.cmp Ceq f Float.zero))
| _, _ => None
end.
@@ -193,13 +208,13 @@ Function sem_add (v1:val) (t1:type) (v2: val) (t2:type) : option val :=
| Vfloat n1, Vint n2 => Some (Vfloat (Float.add n1 (cast_int_float sg n2)))
| _, _ => None
end
- | add_case_pi ty => (**r pointer plus integer *)
+ | add_case_pi ty _ => (**r pointer plus integer *)
match v1,v2 with
| Vptr b1 ofs1, Vint n2 =>
Some (Vptr b1 (Int.add ofs1 (Int.mul (Int.repr (sizeof ty)) n2)))
| _, _ => None
end
- | add_case_ip ty => (**r integer plus pointer *)
+ | add_case_ip ty _ => (**r integer plus pointer *)
match v1,v2 with
| Vint n1, Vptr b2 ofs2 =>
Some (Vptr b2 (Int.add ofs2 (Int.mul (Int.repr (sizeof ty)) n1)))
@@ -472,10 +487,40 @@ Definition sem_binary_operation
Definition sem_incrdecr (id: incr_or_decr) (v: val) (ty: type) :=
match id with
- | Incr => sem_add v ty (Vint Int.one) (Tint I32 Signed)
- | Decr => sem_sub v ty (Vint Int.one) (Tint I32 Signed)
+ | Incr => sem_add v ty (Vint Int.one) type_int32s
+ | Decr => sem_sub v ty (Vint Int.one) type_int32s
end.
+(** Common-sense relations between boolean operators *)
+
+Lemma cast_bool_bool_val:
+ forall v t,
+ sem_cast v t (Tint IBool Signed noattr) =
+ match bool_val v t with None => None | Some b => Some(Val.of_bool b) end.
+Proof.
+ intros. unfold sem_cast, bool_val. destruct t; simpl; destruct v; auto.
+ destruct (Int.eq i0 Int.zero); auto.
+ destruct (Float.cmp Ceq f0 Float.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
+Qed.
+
+Lemma notbool_bool_val:
+ forall v t,
+ sem_notbool v t =
+ match bool_val v t with None => None | Some b => Some(Val.of_bool (negb b)) end.
+Proof.
+ assert (CB: forall i s a, classify_bool (Tint i s a) = bool_case_ip).
+ intros. destruct i; auto. destruct s; auto.
+ intros. unfold sem_notbool, bool_val. destruct t; try rewrite CB; simpl; destruct v; auto.
+ destruct (Int.eq i0 Int.zero); auto.
+ destruct (Float.cmp Ceq f0 Float.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
+ destruct (Int.eq i Int.zero); auto.
+Qed.
+
(** * Operational semantics *)
(** The semantics uses two environments. The global environment
@@ -492,30 +537,58 @@ Definition env := PTree.t (block * type). (* map variable -> location & type *)
Definition empty_env: env := (PTree.empty (block * type)).
-(** [load_value_of_type ty m b ofs] computes the value of a datum
+(** [deref_loc ty m b ofs t v] computes the value of a datum
of type [ty] residing in memory [m] at block [b], offset [ofs].
If the type [ty] indicates an access by value, the corresponding
memory load is performed. If the type [ty] indicates an access by
- reference, the pointer [Vptr b ofs] is returned. *)
-
-Definition load_value_of_type (ty: type) (m: mem) (b: block) (ofs: int) : option val :=
- match access_mode ty with
- | By_value chunk => Mem.loadv chunk m (Vptr b ofs)
- | By_reference => Some (Vptr b ofs)
- | By_nothing => None
- end.
-
-(** Symmetrically, [store_value_of_type ty m b ofs v] returns the
+ reference, the pointer [Vptr b ofs] is returned. [v] is the value
+ returned, and [t] the trace of observables (nonempty if this is
+ a volatile access). *)
+
+Inductive deref_loc {F V: Type} (ge: Genv.t F V) (ty: type) (m: mem) (b: block) (ofs: int) : trace -> val -> Prop :=
+ | deref_loc_value: forall chunk v,
+ access_mode ty = By_value chunk ->
+ type_is_volatile ty = false ->
+ Mem.loadv chunk m (Vptr b ofs) = Some v ->
+ deref_loc ge ty m b ofs E0 v
+ | deref_loc_volatile: forall chunk t v,
+ access_mode ty = By_value chunk -> type_is_volatile ty = true ->
+ volatile_load ge chunk m b ofs t v ->
+ deref_loc ge ty m b ofs t v
+ | deref_loc_reference:
+ access_mode ty = By_reference ->
+ deref_loc ge ty m b ofs E0 (Vptr b ofs)
+ | deref_loc_copy:
+ access_mode ty = By_copy ->
+ deref_loc ge ty m b ofs E0 (Vptr b ofs).
+
+(** Symmetrically, [assign_loc ty m b ofs v t m'] returns the
memory state after storing the value [v] in the datum
of type [ty] residing in memory [m] at block [b], offset [ofs].
- This is allowed only if [ty] indicates an access by value. *)
-
-Definition store_value_of_type (ty_dest: type) (m: mem) (loc: block) (ofs: int) (v: val) : option mem :=
- match access_mode ty_dest with
- | By_value chunk => Mem.storev chunk m (Vptr loc ofs) v
- | By_reference => None
- | By_nothing => None
- end.
+ This is allowed only if [ty] indicates an access by value or by copy.
+ [m'] is the updated memory state and [t] the trace of observables
+ (nonempty if this is a volatile store). *)
+
+Inductive assign_loc {F V: Type} (ge: Genv.t F V) (ty: type) (m: mem) (b: block) (ofs: int):
+ val -> trace -> mem -> Prop :=
+ | assign_loc_value: forall v chunk m',
+ access_mode ty = By_value chunk ->
+ type_is_volatile ty = false ->
+ Mem.storev chunk m (Vptr b ofs) v = Some m' ->
+ assign_loc ge ty m b ofs v E0 m'
+ | assign_loc_volatile: forall v chunk t m',
+ access_mode ty = By_value chunk -> type_is_volatile ty = true ->
+ volatile_store ge chunk m b ofs v t m' ->
+ assign_loc ge ty m b ofs v t m'
+ | assign_loc_copy: forall b' ofs' bytes m',
+ access_mode ty = By_copy ->
+ (alignof ty | Int.unsigned ofs') -> (alignof ty | Int.unsigned ofs) ->
+ b' <> b \/ Int.unsigned ofs' = Int.unsigned ofs
+ \/ Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs
+ \/ Int.unsigned ofs + sizeof ty <= Int.unsigned ofs' ->
+ Mem.loadbytes m b' (Int.unsigned ofs') (sizeof ty) = Some bytes ->
+ Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' ->
+ assign_loc ge ty m b ofs (Vptr b' ofs') E0 m'.
(** Allocation of function-local variables.
[alloc_variables e1 m1 vars e2 m2] allocates one memory block
@@ -541,18 +614,18 @@ Inductive alloc_variables: env -> mem ->
in the memory blocks corresponding to the variables [params].
[m1] is the initial memory state and [m2] the final memory state. *)
-Inductive bind_parameters: env ->
+Inductive bind_parameters {F V: Type} (ge: Genv.t F V) (e: env):
mem -> list (ident * type) -> list val ->
mem -> Prop :=
| bind_parameters_nil:
- forall e m,
- bind_parameters e m nil nil m
+ forall m,
+ bind_parameters ge e m nil nil m
| bind_parameters_cons:
- forall e m id ty params v1 vl b m1 m2,
+ forall m id ty params v1 vl b m1 m2,
PTree.get id e = Some(b, ty) ->
- store_value_of_type ty m b Int.zero v1 = Some m1 ->
- bind_parameters e m1 params vl m2 ->
- bind_parameters e m ((id, ty) :: params) (v1 :: vl) m2.
+ assign_loc ge ty m b Int.zero v1 E0 m1 ->
+ bind_parameters ge e m1 params vl m2 ->
+ bind_parameters ge e m ((id, ty) :: params) (v1 :: vl) m2.
(** Return the list of blocks in the codomain of [e], with low and high bounds. *)
@@ -624,70 +697,69 @@ Inductive lred: expr -> mem -> expr -> mem -> Prop :=
| red_deref: forall b ofs ty1 ty m,
lred (Ederef (Eval (Vptr b ofs) ty1) ty) m
(Eloc b ofs ty) m
- | red_field_struct: forall b ofs id fList f ty m delta,
+ | red_field_struct: forall b ofs id fList a f ty m delta,
field_offset f fList = OK delta ->
- lred (Efield (Eloc b ofs (Tstruct id fList)) f ty) m
+ lred (Efield (Eval (Vptr b ofs) (Tstruct id fList a)) f ty) m
(Eloc b (Int.add ofs (Int.repr delta)) ty) m
- | red_field_union: forall b ofs id fList f ty m,
- lred (Efield (Eloc b ofs (Tunion id fList)) f ty) m
+ | red_field_union: forall b ofs id fList a f ty m,
+ lred (Efield (Eval (Vptr b ofs) (Tunion id fList a)) f ty) m
(Eloc b ofs ty) m.
(** Head reductions for r-values *)
-Inductive rred: expr -> mem -> expr -> mem -> Prop :=
- | red_rvalof: forall b ofs ty m v,
- load_value_of_type ty m b ofs = Some v ->
+Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop :=
+ | red_rvalof: forall b ofs ty m t v,
+ deref_loc ge ty m b ofs t v ->
rred (Evalof (Eloc b ofs ty) ty) m
- (Eval v ty) m
+ t (Eval v ty) m
| red_addrof: forall b ofs ty1 ty m,
rred (Eaddrof (Eloc b ofs ty1) ty) m
- (Eval (Vptr b ofs) ty) m
+ E0 (Eval (Vptr b ofs) ty) m
| red_unop: forall op v1 ty1 ty m v,
sem_unary_operation op v1 ty1 = Some v ->
rred (Eunop op (Eval v1 ty1) ty) m
- (Eval v ty) m
+ E0 (Eval v ty) m
| red_binop: forall op v1 ty1 v2 ty2 ty m v,
sem_binary_operation op v1 ty1 v2 ty2 m = Some v ->
rred (Ebinop op (Eval v1 ty1) (Eval v2 ty2) ty) m
- (Eval v ty) m
+ E0 (Eval v ty) m
| red_cast: forall ty v1 ty1 m v,
sem_cast v1 ty1 ty = Some v ->
rred (Ecast (Eval v1 ty1) ty) m
- (Eval v ty) m
+ E0 (Eval v ty) m
| red_condition: forall v1 ty1 r1 r2 ty b m,
bool_val v1 ty1 = Some b ->
rred (Econdition (Eval v1 ty1) r1 r2 ty) m
- (Eparen (if b then r1 else r2) ty) m
+ E0 (Eparen (if b then r1 else r2) ty) m
| red_sizeof: forall ty1 ty m,
rred (Esizeof ty1 ty) m
- (Eval (Vint (Int.repr (sizeof ty1))) ty) m
- | red_assign: forall b ofs ty1 v2 ty2 m v m',
+ E0 (Eval (Vint (Int.repr (sizeof ty1))) ty) m
+ | red_assign: forall b ofs ty1 v2 ty2 m v t m',
sem_cast v2 ty2 ty1 = Some v ->
- store_value_of_type ty1 m b ofs v = Some m' ->
+ assign_loc ge ty1 m b ofs v t m' ->
rred (Eassign (Eloc b ofs ty1) (Eval v2 ty2) ty1) m
- (Eval v ty1) m'
- | red_assignop: forall op b ofs ty1 v2 ty2 tyres m v1 v v' m',
- load_value_of_type ty1 m b ofs = Some v1 ->
- sem_binary_operation op v1 ty1 v2 ty2 m = Some v ->
- sem_cast v tyres ty1 = Some v' ->
- store_value_of_type ty1 m b ofs v' = Some m' ->
+ t (Eval v ty1) m'
+ | red_assignop: forall op b ofs ty1 v2 ty2 tyres m t v1,
+ deref_loc ge ty1 m b ofs t v1 ->
rred (Eassignop op (Eloc b ofs ty1) (Eval v2 ty2) tyres ty1) m
- (Eval v' ty1) m'
- | red_postincr: forall id b ofs ty m v1 v2 v3 m',
- load_value_of_type ty m b ofs = Some v1 ->
- sem_incrdecr id v1 ty = Some v2 ->
- sem_cast v2 (typeconv ty) ty = Some v3 ->
- store_value_of_type ty m b ofs v3 = Some m' ->
+ t (Eassign (Eloc b ofs ty1)
+ (Ebinop op (Eval v1 ty1) (Eval v2 ty2) tyres) ty1) m
+ | red_postincr: forall id b ofs ty m t v1 op,
+ deref_loc ge ty m b ofs t v1 ->
+ op = match id with Incr => Oadd | Decr => Osub end ->
rred (Epostincr id (Eloc b ofs ty) ty) m
- (Eval v1 ty) m'
+ t (Ecomma (Eassign (Eloc b ofs ty)
+ (Ebinop op (Eval v1 ty) (Eval (Vint Int.one) type_int32s) (typeconv ty))
+ ty)
+ (Eval v1 ty) ty) m
| red_comma: forall v ty1 r2 ty m,
typeof r2 = ty ->
rred (Ecomma (Eval v ty1) r2 ty) m
- r2 m
+ E0 r2 m
| red_paren: forall v1 ty1 ty m v,
sem_cast v1 ty1 ty = Some v ->
rred (Eparen (Eval v1 ty1) ty) m
- (Eval v ty) m.
+ E0 (Eval v ty) m.
(** Head reduction for function calls.
(More exactly, identification of function calls that can reduce.) *)
@@ -729,7 +801,7 @@ Inductive context: kind -> kind -> (expr -> expr) -> Prop :=
| ctx_deref: forall k C ty,
context k RV C -> context k LV (fun x => Ederef (C x) ty)
| ctx_field: forall k C f ty,
- context k LV C -> context k LV (fun x => Efield (C x) f ty)
+ context k RV C -> context k LV (fun x => Efield (C x) f ty)
| ctx_rvalof: forall k C ty,
context k LV C -> context k RV (fun x => Evalof (C x) ty)
| ctx_addrof: forall k C ty,
@@ -789,31 +861,31 @@ with contextlist: kind -> (expr -> exprlist) -> Prop :=
is not immediately stuck if it is a value (of the appropriate kind)
or it can reduce (at head or within). *)
-Inductive not_imm_stuck: kind -> expr -> mem -> Prop :=
- | not_stuck_val: forall v ty m,
- not_imm_stuck RV (Eval v ty) m
- | not_stuck_loc: forall b ofs ty m,
- not_imm_stuck LV (Eloc b ofs ty) m
- | not_stuck_lred: forall to C e m e' m',
+Inductive imm_safe: kind -> expr -> mem -> Prop :=
+ | imm_safe_val: forall v ty m,
+ imm_safe RV (Eval v ty) m
+ | imm_safe_loc: forall b ofs ty m,
+ imm_safe LV (Eloc b ofs ty) m
+ | imm_safe_lred: forall to C e m e' m',
lred e m e' m' ->
context LV to C ->
- not_imm_stuck to (C e) m
- | not_stuck_rred: forall to C e m e' m',
- rred e m e' m' ->
+ imm_safe to (C e) m
+ | imm_safe_rred: forall to C e m t e' m',
+ rred e m t e' m' ->
context RV to C ->
- not_imm_stuck to (C e) m
- | not_stuck_callred: forall to C e m fd args ty,
+ imm_safe to (C e) m
+ | imm_safe_callred: forall to C e m fd args ty,
callred e fd args ty ->
context RV to C ->
- not_imm_stuck to (C e) m.
+ imm_safe to (C e) m.
(* An expression is not stuck if none of the potential redexes contained within
is immediately stuck. *)
-
+(*
Definition not_stuck (e: expr) (m: mem) : Prop :=
forall k C e' ,
context k RV C -> e = C e' -> not_imm_stuck k e' m.
-
+*)
End EXPR.
(** ** Transition semantics. *)
@@ -899,7 +971,8 @@ Inductive state: Type :=
| Returnstate (**r returning from a function *)
(res: val)
(k: cont)
- (m: mem) : state.
+ (m: mem) : state
+ | Stuckstate. (**r undefined behavior occurred *)
(** Find the statement and manufacture the continuation
corresponding to a label. *)
@@ -959,24 +1032,26 @@ Inductive estep: state -> trace -> state -> Prop :=
| step_lred: forall C f a k e m a' m',
lred e a m a' m' ->
- not_stuck e (C a) m ->
context LV RV C ->
estep (ExprState f (C a) k e m)
E0 (ExprState f (C a') k e m')
- | step_rred: forall C f a k e m a' m',
- rred a m a' m' ->
- not_stuck e (C a) m ->
+ | step_rred: forall C f a k e m t a' m',
+ rred a m t a' m' ->
context RV RV C ->
estep (ExprState f (C a) k e m)
- E0 (ExprState f (C a') k e m')
+ t (ExprState f (C a') k e m')
| step_call: forall C f a k e m fd vargs ty,
callred a fd vargs ty ->
- not_stuck e (C a) m ->
context RV RV C ->
estep (ExprState f (C a) k e m)
- E0 (Callstate fd vargs (Kcall f e C ty k) m).
+ E0 (Callstate fd vargs (Kcall f e C ty k) m)
+
+ | step_stuck: forall C f a k e m K,
+ context K RV C -> ~(imm_safe e K a m) ->
+ estep (ExprState f (C a) k e m)
+ E0 Stuckstate.
Inductive sstep: state -> trace -> state -> Prop :=
@@ -1117,7 +1192,7 @@ Inductive sstep: state -> trace -> state -> Prop :=
| step_internal_function: forall f vargs k m e m1 m2,
list_norepet (var_names (fn_params f) ++ var_names (fn_vars f)) ->
alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 ->
- bind_parameters e m1 f.(fn_params) vargs m2 ->
+ bind_parameters ge e m1 f.(fn_params) vargs m2 ->
sstep (Callstate (Internal f) vargs k m)
E0 (State f f.(fn_body) k e m2)
@@ -1148,7 +1223,7 @@ Inductive initial_state (p: program): state -> Prop :=
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
- type_of_fundef f = Tfunction Tnil (Tint I32 Signed) ->
+ type_of_fundef f = Tfunction Tnil type_int32s ->
initial_state p (Callstate f nil Kstop m0).
(** A final state is a [Returnstate] with an empty continuation. *)
@@ -1162,3 +1237,17 @@ Inductive final_state: state -> int -> Prop :=
Definition semantics (p: program) :=
Semantics step (initial_state p) final_state (Genv.globalenv p).
+(** This semantics has the single-event property. *)
+
+Lemma semantics_single_events:
+ forall p, single_events (semantics p).
+Proof.
+ intros; red; intros. destruct H.
+ set (ge := globalenv (semantics p)) in *.
+ assert (DEREF: forall chunk m b ofs t v, deref_loc ge chunk m b ofs t v -> (length t <= 1)%nat).
+ intros. inv H0; simpl; try omega. inv H3; simpl; try omega.
+ assert (ASSIGN: forall chunk m b ofs t v m', assign_loc ge chunk m b ofs v t m' -> (length t <= 1)%nat).
+ intros. inv H0; simpl; try omega. inv H3; simpl; try omega.
+ inv H; simpl; try omega. inv H0; eauto; simpl; try omega.
+ inv H; simpl; try omega. eapply external_call_trace_length; eauto.
+Qed.
diff --git a/cfrontend/Csharpminor.v b/cfrontend/Csharpminor.v
index a1ed8b3..88eb3c7 100644
--- a/cfrontend/Csharpminor.v
+++ b/cfrontend/Csharpminor.v
@@ -68,6 +68,7 @@ Inductive stmt : Type :=
| Sset : ident -> expr -> stmt
| Sstore : memory_chunk -> expr -> expr -> stmt
| Scall : option ident -> signature -> expr -> list expr -> stmt
+ | Sbuiltin : option ident -> external_function -> list expr -> stmt
| Sseq: stmt -> stmt -> stmt
| Sifthenelse: expr -> stmt -> stmt -> stmt
| Sloop: stmt -> stmt
@@ -84,32 +85,36 @@ with lbl_stmt : Type :=
(** The variables can be either scalar variables
(whose type, size and signedness are given by a [memory_chunk]
- or array variables (of the indicated sizes). The only operation
- permitted on an array variable is taking its address. *)
+ or array variables (of the indicated sizes and alignment).
+ The only operation permitted on an array variable is taking its address. *)
Inductive var_kind : Type :=
- | Vscalar: memory_chunk -> var_kind
- | Varray: Z -> var_kind.
+ | Vscalar(chunk: memory_chunk)
+ | Varray(sz al: Z).
Definition sizeof (lv: var_kind) : Z :=
match lv with
| Vscalar chunk => size_chunk chunk
- | Varray sz => Zmax 0 sz
+ | Varray sz al => Zmax 0 sz
+ end.
+
+Definition type_of_kind (lv: var_kind) : typ :=
+ match lv with
+ | Vscalar chunk => type_of_chunk chunk
+ | Varray _ _ => Tint
end.
(** Functions are composed of a return type, a list of parameter names
- with associated memory chunks (parameters must be scalar), a list of
- local variables with associated [var_kind] description, and a
+ with associated [var_kind] descriptions, a list of
+ local variables with associated [var_kind] descriptions, and a
statement representing the function body. *)
-Definition param_name (p: ident * memory_chunk) := fst p.
-Definition param_chunk (p: ident * memory_chunk) := snd p.
Definition variable_name (v: ident * var_kind) := fst v.
Definition variable_kind (v: ident * var_kind) := snd v.
Record function : Type := mkfunction {
fn_return: option typ;
- fn_params: list (ident * memory_chunk);
+ fn_params: list (ident * var_kind);
fn_vars: list (ident * var_kind);
fn_temps: list ident;
fn_body: stmt
@@ -120,7 +125,7 @@ Definition fundef := AST.fundef function.
Definition program : Type := AST.program fundef var_kind.
Definition fn_sig (f: function) :=
- mksignature (List.map type_of_chunk (List.map param_chunk f.(fn_params)))
+ mksignature (List.map type_of_kind (List.map variable_kind f.(fn_params)))
f.(fn_return).
Definition funsig (fd: fundef) :=
@@ -129,13 +134,9 @@ Definition funsig (fd: fundef) :=
| External ef => ef_sig ef
end.
-Definition var_of_param (p: ident * memory_chunk) : ident * var_kind :=
- (fst p, Vscalar (snd p)).
-
-Definition fn_variables (f: function) :=
- List.map var_of_param f.(fn_params) ++ f.(fn_vars).
+Definition fn_variables (f: function) := f.(fn_params) ++ f.(fn_vars).
-Definition fn_params_names (f: function) := List.map param_name f.(fn_params).
+Definition fn_params_names (f: function) := List.map variable_name f.(fn_params).
Definition fn_vars_names (f: function) := List.map variable_name f.(fn_vars).
(** * Operational semantics *)
@@ -292,6 +293,10 @@ Definition block_of_binding (id_b_lv: ident * (block * var_kind)) :=
Definition blocks_of_env (e: env) : list (block * Z * Z) :=
List.map block_of_binding (PTree.elements e).
+Section RELSEM.
+
+Variable ge: genv.
+
(** Initialization of local variables that are parameters. The value
of the corresponding argument is stored into the memory block
bound to the parameter. *)
@@ -300,22 +305,26 @@ Definition val_normalized (v: val) (chunk: memory_chunk) : Prop :=
Val.load_result chunk v = v.
Inductive bind_parameters: env ->
- mem -> list (ident * memory_chunk) -> list val ->
+ mem -> list (ident * var_kind) -> list val ->
mem -> Prop :=
| bind_parameters_nil:
forall e m,
bind_parameters e m nil nil m
- | bind_parameters_cons:
+ | bind_parameters_scalar:
forall e m id chunk params v1 vl b m1 m2,
PTree.get id e = Some (b, Vscalar chunk) ->
val_normalized v1 chunk ->
Mem.store chunk m b 0 v1 = Some m1 ->
bind_parameters e m1 params vl m2 ->
- bind_parameters e m ((id, chunk) :: params) (v1 :: vl) m2.
-
-Section RELSEM.
+ bind_parameters e m ((id, Vscalar chunk) :: params) (v1 :: vl) m2
+ | bind_parameters_array:
+ forall e m id sz al params v1 vl b m1 m2,
+ PTree.get id e = Some (b, Varray sz al) ->
+ extcall_memcpy_sem sz (Zmin al 4)
+ ge (Vptr b Int.zero :: v1 :: nil) m E0 Vundef m1 ->
+ bind_parameters e m1 params vl m2 ->
+ bind_parameters e m ((id, Varray sz al) :: params) (v1 :: vl) m2.
-Variable ge: genv.
(* Evaluation of the address of a variable:
[eval_var_addr prg ge e id b] states that variable [id]
@@ -459,6 +468,12 @@ Inductive step: state -> trace -> state -> Prop :=
step (State f (Scall optid sig a bl) k e le m)
E0 (Callstate fd vargs (Kcall optid f e le k) m)
+ | step_builtin: forall f optid ef bl k e le m vargs t vres m',
+ eval_exprlist e le m bl vargs ->
+ external_call ef ge vargs m t vres m' ->
+ step (State f (Sbuiltin optid ef bl) k e le m)
+ t (State f Sskip k e (Cminor.set_optvar optid vres le) m')
+
| step_seq: forall f s1 s2 k e le m,
step (State f (Sseq s1 s2) k e le m)
E0 (State f s1 (Kseq s2 k) e le m)
diff --git a/cfrontend/Cshmgen.v b/cfrontend/Cshmgen.v
index bbd4cfe..c17d79e 100644
--- a/cfrontend/Cshmgen.v
+++ b/cfrontend/Cshmgen.v
@@ -46,20 +46,21 @@ Definition chunk_of_type (ty: type): res memory_chunk :=
Definition var_kind_of_type (ty: type): res var_kind :=
match ty with
- | Tint I8 Signed => OK(Vscalar Mint8signed)
- | Tint I8 Unsigned => OK(Vscalar Mint8unsigned)
- | Tint I16 Signed => OK(Vscalar Mint16signed)
- | Tint I16 Unsigned => OK(Vscalar Mint16unsigned)
- | Tint I32 _ => OK(Vscalar Mint32)
- | Tfloat F32 => OK(Vscalar Mfloat32)
- | Tfloat F64 => OK(Vscalar Mfloat64)
+ | Tint I8 Signed _ => OK(Vscalar Mint8signed)
+ | Tint I8 Unsigned _ => OK(Vscalar Mint8unsigned)
+ | Tint I16 Signed _ => OK(Vscalar Mint16signed)
+ | Tint I16 Unsigned _ => OK(Vscalar Mint16unsigned)
+ | Tint I32 _ _ => OK(Vscalar Mint32)
+ | Tint IBool _ _ => OK(Vscalar Mint8unsigned)
+ | Tfloat F32 _ => OK(Vscalar Mfloat32)
+ | Tfloat F64 _ => OK(Vscalar Mfloat64)
| Tvoid => Error (msg "Cshmgen.var_kind_of_type(void)")
- | Tpointer _ => OK(Vscalar Mint32)
- | Tarray _ _ => OK(Varray (Csyntax.sizeof ty))
+ | Tpointer _ _ => OK(Vscalar Mint32)
+ | Tarray _ _ _ => OK(Varray (Csyntax.sizeof ty) (Csyntax.alignof ty))
| Tfunction _ _ => Error (msg "Cshmgen.var_kind_of_type(function)")
- | Tstruct _ fList => OK(Varray (Csyntax.sizeof ty))
- | Tunion _ fList => OK(Varray (Csyntax.sizeof ty))
- | Tcomp_ptr _ => OK(Vscalar Mint32)
+ | Tstruct _ fList _ => OK(Varray (Csyntax.sizeof ty) (Csyntax.alignof ty))
+ | Tunion _ fList _ => OK(Varray (Csyntax.sizeof ty) (Csyntax.alignof ty))
+ | Tcomp_ptr _ _ => OK(Vscalar Mint32)
end.
(** * Csharpminor constructors *)
@@ -101,7 +102,7 @@ Definition make_intoffloat (e: expr) (sg: signedness) :=
*)
Definition make_boolean (e: expr) (ty: type) :=
match ty with
- | Tfloat _ => Ebinop (Ocmpf Cne) e (make_floatconst Float.zero)
+ | Tfloat _ _ => Ebinop (Ocmpf Cne) e (make_floatconst Float.zero)
| _ => e
end.
@@ -128,10 +129,10 @@ Definition make_add (e1: expr) (ty1: type) (e2: expr) (ty2: type) :=
| add_case_ff => OK (Ebinop Oaddf e1 e2)
| add_case_if sg => OK (Ebinop Oaddf (make_floatofint e1 sg) e2)
| add_case_fi sg => OK (Ebinop Oaddf e1 (make_floatofint e2 sg))
- | add_case_pi ty =>
+ | add_case_pi ty _ =>
let n := make_intconst (Int.repr (Csyntax.sizeof ty)) in
OK (Ebinop Oadd e1 (Ebinop Omul n e2))
- | add_case_ip ty =>
+ | add_case_ip ty _ =>
let n := make_intconst (Int.repr (Csyntax.sizeof ty)) in
OK (Ebinop Oadd e2 (Ebinop Omul n e1))
| add_default => Error (msg "Cshmgen.make_add")
@@ -218,6 +219,7 @@ Definition make_cast_int (e: expr) (sz: intsize) (si: signedness) :=
| I16, Signed => Eunop Ocast16signed e
| I16, Unsigned => Eunop Ocast16unsigned e
| I32, _ => e
+ | IBool, _ => Eunop Oboolval e
end.
Definition make_cast_float (e: expr) (sz: floatsize) :=
@@ -233,6 +235,10 @@ Definition make_cast (from to: type) (e: expr) :=
| cast_case_f2f sz2 => make_cast_float e sz2
| cast_case_i2f si1 sz2 => make_cast_float (make_floatofint e si1) sz2
| cast_case_f2i sz2 si2 => make_cast_int (make_intoffloat e si2) sz2 si2
+ | cast_case_ip2bool => Eunop Oboolval e
+ | cast_case_f2bool => Ebinop (Ocmpf Cne) e (make_floatconst Float.zero)
+ | cast_case_struct id1 fld1 id2 fld2 => e
+ | cast_case_union id1 fld1 id2 fld2 => e
| cast_case_void => e
| cast_case_default => e
end.
@@ -247,10 +253,30 @@ Definition make_load (addr: expr) (ty_res: type) :=
match (access_mode ty_res) with
| By_value chunk => OK (Eload chunk addr)
| By_reference => OK addr
+ | By_copy => OK addr
| By_nothing => Error (msg "Cshmgen.make_load")
end.
-(** [make_store addr ty_res rhs ty_rhs] stores the value of the
+(** [make_vol_load dst addr ty] loads a volatile value of type [ty] from
+ the memory location denoted by the Csharpminor expression [addr],
+ and stores its result in variable [dst]. *)
+
+Definition make_vol_load (dst: ident) (addr: expr) (ty: type) :=
+ match (access_mode ty) with
+ | By_value chunk => OK (Sbuiltin (Some dst) (EF_vload chunk) (addr :: nil))
+ | By_reference => OK (Sset dst addr)
+ | By_copy => OK (Sset dst addr)
+ | By_nothing => Error (msg "Cshmgen.make_vol_load")
+ end.
+
+(** [make_memcpy dst src ty] returns a [memcpy] builtin appropriate for
+ by-copy assignment of a value of Clight type [ty]. *)
+
+Definition make_memcpy (dst src: expr) (ty: type) :=
+ Sbuiltin None (EF_memcpy (Csyntax.sizeof ty) (Zmin (Csyntax.alignof ty) 4))
+ (dst :: src :: nil).
+
+(** [make_store addr ty rhs] stores the value of the
Csharpminor expression [rhs] into the memory location denoted by the
Csharpminor expression [addr].
[ty] is the type of the memory location. *)
@@ -258,6 +284,17 @@ Definition make_load (addr: expr) (ty_res: type) :=
Definition make_store (addr: expr) (ty: type) (rhs: expr) :=
match access_mode ty with
| By_value chunk => OK (Sstore chunk addr rhs)
+ | By_copy => OK (make_memcpy addr rhs ty)
+ | _ => Error (msg "Cshmgen.make_store")
+ end.
+
+(** [make_vol_store] is similar, but for a store to a location that
+ can be volatile. *)
+
+Definition make_vol_store (addr: expr) (ty: type) (rhs: expr) :=
+ match access_mode ty with
+ | By_value chunk => OK (Sbuiltin None (EF_vstore chunk) (addr :: rhs :: nil))
+ | By_copy => OK (make_memcpy addr rhs ty)
| _ => Error (msg "Cshmgen.make_store")
end.
@@ -282,6 +319,7 @@ Definition var_get (id: ident) (ty: type) :=
match access_mode ty with
| By_value chunk => OK (Evar id)
| By_reference => OK (Eaddrof id)
+ | By_copy => OK (Eaddrof id)
| _ => Error (MSG "Cshmgen.var_get " :: CTX id :: nil)
end.
@@ -292,6 +330,7 @@ Definition var_get (id: ident) (ty: type) :=
Definition var_set (id: ident) (ty: type) (rhs: expr) :=
match access_mode ty with
| By_value chunk => OK (Sassign id rhs)
+ | By_copy => OK (make_memcpy (Eaddrof id) rhs ty)
| _ => Error (MSG "Cshmgen.var_set " :: CTX id :: nil)
end.
@@ -368,14 +407,14 @@ Fixpoint transl_expr (a: Clight.expr) {struct a} : res expr :=
OK(make_intconst (Int.repr (Csyntax.sizeof ty)))
| Clight.Efield b i ty =>
match typeof b with
- | Tstruct _ fld =>
- do tb <- transl_lvalue b;
+ | Tstruct _ fld _ =>
+ do tb <- transl_expr b;
do ofs <- field_offset i fld;
make_load
(Ebinop Oadd tb (make_intconst (Int.repr ofs)))
ty
- | Tunion _ fld =>
- do tb <- transl_lvalue b;
+ | Tunion _ fld _ =>
+ do tb <- transl_expr b;
make_load tb ty
| _ =>
Error(msg "Cshmgen.transl_expr(field)")
@@ -395,12 +434,12 @@ with transl_lvalue (a: Clight.expr) {struct a} : res expr :=
transl_expr b
| Clight.Efield b i ty =>
match typeof b with
- | Tstruct _ fld =>
- do tb <- transl_lvalue b;
+ | Tstruct _ fld _ =>
+ do tb <- transl_expr b;
do ofs <- field_offset i fld;
OK (Ebinop Oadd tb (make_intconst (Int.repr ofs)))
- | Tunion _ fld =>
- transl_lvalue b
+ | Tunion _ fld _ =>
+ transl_expr b
| _ =>
Error(msg "Cshmgen.transl_lvalue(field)")
end
@@ -503,7 +542,11 @@ Fixpoint transl_statement (tyret: type) (nbrk ncnt: nat)
| Clight.Sskip =>
OK Sskip
| Clight.Sassign b c =>
- match is_variable b with
+ if type_is_volatile (typeof b) then
+ (do tb <- transl_lvalue b;
+ do tc <- transl_expr c;
+ make_vol_store tb (typeof b) (make_cast (typeof c) (typeof b) tc))
+ else match is_variable b with
| Some id =>
do tc <- transl_expr c;
var_set id (typeof b) (make_cast (typeof c) (typeof b) tc)
@@ -515,6 +558,9 @@ Fixpoint transl_statement (tyret: type) (nbrk ncnt: nat)
| Clight.Sset x b =>
do tb <- transl_expr b;
OK(Sset x tb)
+ | Clight.Svolread x b =>
+ do tb <- transl_lvalue b;
+ make_vol_load x tb (typeof b)
| Clight.Scall x b cl =>
match classify_fun (typeof b) with
| fun_case_f args res =>
@@ -583,13 +629,11 @@ with transl_lbl_stmt (tyret: type) (nbrk ncnt: nat)
Definition prefix_var_name (id: ident) : errmsg :=
MSG "In local variable " :: CTX id :: MSG ": " :: nil.
-Definition transl_params (l: list (ident * type)) :=
- AST.map_partial prefix_var_name chunk_of_type l.
Definition transl_vars (l: list (ident * type)) :=
AST.map_partial prefix_var_name var_kind_of_type l.
Definition transl_function (f: Clight.function) : res function :=
- do tparams <- transl_params (Clight.fn_params f);
+ do tparams <- transl_vars (Clight.fn_params f);
do tvars <- transl_vars (Clight.fn_vars f);
do tbody <- transl_statement f.(Clight.fn_return) 1%nat 0%nat (Clight.fn_body f);
OK (mkfunction
diff --git a/cfrontend/Cshmgenproof.v b/cfrontend/Cshmgenproof.v
index 0f7810d..1089b6b 100644
--- a/cfrontend/Cshmgenproof.v
+++ b/cfrontend/Cshmgenproof.v
@@ -32,27 +32,24 @@ Require Import Cshmgen.
(** * Properties of operations over types *)
-Remark type_of_chunk_of_type:
- forall ty chunk,
- chunk_of_type ty = OK chunk ->
- type_of_chunk chunk = typ_of_type ty.
+Remark type_of_kind_of_type:
+ forall t k,
+ var_kind_of_type t = OK k -> type_of_kind k = typ_of_type t.
Proof.
- intros. unfold chunk_of_type in H. destruct ty; simpl in H; try monadInv H.
- destruct i; destruct s; monadInv H; reflexivity.
- destruct f; monadInv H; reflexivity.
- reflexivity.
+ intros. destruct t; try (monadInv H); auto.
+ destruct i; destruct s; monadInv H; auto.
+ destruct f; monadInv H; auto.
Qed.
Remark transl_params_types:
forall p tp,
- transl_params p = OK tp ->
- map type_of_chunk (map param_chunk tp) = typlist_of_typelist (type_of_params p).
+ transl_vars p = OK tp ->
+ map type_of_kind (map variable_kind tp) = typlist_of_typelist (type_of_params p).
Proof.
induction p; simpl; intros.
inv H. auto.
- destruct a as [id ty]. generalize H; clear H. case_eq (chunk_of_type ty); intros.
- monadInv H0. simpl. f_equal; auto. apply type_of_chunk_of_type; auto.
- inv H0.
+ destruct a as [id ty]. destruct (var_kind_of_type ty) as []_eqn; monadInv H.
+ simpl. f_equal; auto. apply type_of_kind_of_type; auto.
Qed.
Lemma transl_fundef_sig1:
@@ -93,13 +90,24 @@ Proof.
destruct f; congruence.
Qed.
+Lemma var_kind_by_reference:
+ forall ty vk,
+ access_mode ty = By_reference \/ access_mode ty = By_copy ->
+ var_kind_of_type ty = OK vk ->
+ vk = Varray (Csyntax.sizeof ty) (Csyntax.alignof ty).
+Proof.
+ intros ty vk; destruct ty; simpl; try intuition congruence.
+ destruct i; try congruence; destruct s; intuition congruence.
+ destruct f; intuition congruence.
+Qed.
+
Lemma sizeof_var_kind_of_type:
forall ty vk,
var_kind_of_type ty = OK vk ->
Csharpminor.sizeof vk = Csyntax.sizeof ty.
Proof.
intros ty vk.
- assert (sizeof (Varray (Csyntax.sizeof ty)) = Csyntax.sizeof ty).
+ assert (sizeof (Varray (Csyntax.sizeof ty) (Csyntax.alignof ty)) = Csyntax.sizeof ty).
simpl. rewrite Zmax_spec. apply zlt_false.
generalize (Csyntax.sizeof_pos ty). omega.
destruct ty; try (destruct i; try destruct s); try (destruct f);
@@ -107,8 +115,8 @@ Proof.
Qed.
Remark cast_int_int_normalized:
- forall sz si chunk n,
- access_mode (Tint sz si) = By_value chunk ->
+ forall sz si a chunk n,
+ access_mode (Tint sz si a) = By_value chunk ->
val_normalized (Vint (cast_int_int sz si n)) chunk.
Proof.
unfold access_mode, cast_int_int, val_normalized; intros. destruct sz.
@@ -119,11 +127,12 @@ Proof.
rewrite Int.sign_ext_idem; auto. compute; auto.
rewrite Int.zero_ext_idem; auto. compute; auto.
inv H. auto.
+ inv H. destruct (Int.eq n Int.zero); auto.
Qed.
Remark cast_float_float_normalized:
- forall sz chunk n,
- access_mode (Tfloat sz) = By_value chunk ->
+ forall sz a chunk n,
+ access_mode (Tfloat sz a) = By_value chunk ->
val_normalized (Vfloat (cast_float_float sz n)) chunk.
Proof.
unfold access_mode, cast_float_float, val_normalized; intros.
@@ -154,6 +163,16 @@ Proof.
functional inversion H2; subst. eapply cast_float_float_normalized; eauto.
functional inversion H2; subst. eapply cast_float_float_normalized; eauto.
functional inversion H2; subst. eapply cast_int_int_normalized; eauto.
+ assert (chunk = Mint8unsigned).
+ functional inversion H2; subst; simpl in H0; try congruence.
+ subst chunk. destruct (Int.eq i Int.zero); red; auto.
+ assert (chunk = Mint8unsigned).
+ functional inversion H2; subst; simpl in H0; try congruence.
+ subst chunk. red; auto.
+ functional inversion H2; subst. simpl in H0. inv H0. red; auto.
+ functional inversion H2; subst. simpl in H0. inv H0. red; auto.
+ functional inversion H2; subst. simpl in H0. congruence.
+ functional inversion H2; subst. simpl in H0. congruence.
functional inversion H5; subst. simpl in H0. congruence.
Qed.
@@ -213,14 +232,6 @@ Proof.
inv H0. rewrite (IHl1 _ _ _ H H1). auto.
Qed.
-Lemma transl_params_names:
- forall vars tvars,
- transl_params vars = OK tvars ->
- List.map param_name tvars = var_names vars.
-Proof.
- exact (map_partial_names _ _ chunk_of_type).
-Qed.
-
Lemma transl_vars_names:
forall vars tvars,
transl_vars vars = OK tvars ->
@@ -232,13 +243,13 @@ Qed.
Lemma transl_names_norepet:
forall params vars sg tparams tvars temps body,
list_norepet (var_names params ++ var_names vars) ->
- transl_params params = OK tparams ->
+ transl_vars params = OK tparams ->
transl_vars vars = OK tvars ->
let f := Csharpminor.mkfunction sg tparams tvars temps body in
list_norepet (fn_params_names f ++ fn_vars_names f).
Proof.
- intros. unfold fn_params_names, fn_vars_names, f. simpl.
- rewrite (transl_params_names _ _ H0).
+ intros. unfold fn_params_names, fn_vars_names; simpl.
+ rewrite (transl_vars_names _ _ H0).
rewrite (transl_vars_names _ _ H1).
auto.
Qed.
@@ -251,33 +262,15 @@ Proof.
exact (map_partial_append _ _ var_kind_of_type).
Qed.
-Lemma transl_params_vars:
- forall params tparams,
- transl_params params = OK tparams ->
- transl_vars params =
- OK (List.map (fun id_chunk => (fst id_chunk, Vscalar (snd id_chunk))) tparams).
-Proof.
- induction params; intro tparams; simpl.
- intros. inversion H. reflexivity.
- destruct a as [id x].
- unfold chunk_of_type. caseEq (access_mode x); try congruence.
- intros chunk AM.
- caseEq (transl_params params); simpl; intros; try congruence.
- inv H0.
- rewrite (var_kind_by_value _ _ AM).
- rewrite (IHparams _ H). reflexivity.
-Qed.
-
Lemma transl_fn_variables:
forall params vars sg tparams tvars temps body,
- transl_params params = OK tparams ->
+ transl_vars params = OK tparams ->
transl_vars vars = OK tvars ->
let f := Csharpminor.mkfunction sg tparams tvars temps body in
transl_vars (params ++ vars) = OK (fn_variables f).
Proof.
intros.
- generalize (transl_params_vars _ _ H); intro.
- rewrite (transl_vars_append _ _ _ _ H1 H0).
+ rewrite (transl_vars_append _ _ _ _ H H0).
reflexivity.
Qed.
@@ -300,16 +293,20 @@ Lemma transl_expr_lvalue:
(exists tb, transl_lvalue a = OK tb /\
make_load tb (typeof a) = OK ta).
Proof.
- intros. inversion H; subst; clear H; simpl in H0.
+ intros until ta; intros EVAL TR. inv EVAL.
+ (* var local *)
left; exists id; exists ty; auto.
+ (* var global *)
left; exists id; exists ty; auto.
- monadInv H0. right. exists x; split; auto.
- rewrite H2 in H0. monadInv H0. right.
- exists (Ebinop Oadd x (make_intconst (Int.repr x0))). split; auto.
- simpl. rewrite H2. rewrite EQ. rewrite EQ1. auto.
- rewrite H2 in H0. monadInv H0. right.
- exists x; split; auto.
- simpl. rewrite H2. auto.
+ (* deref *)
+ monadInv TR. right. exists x; split; auto.
+ (* field struct *)
+ simpl in TR. rewrite H0 in TR. monadInv TR.
+ right. econstructor; split. simpl. rewrite H0.
+ rewrite EQ; rewrite EQ1; simpl; eauto. auto.
+ (* field union *)
+ simpl in TR. rewrite H0 in TR. monadInv TR.
+ right. econstructor; split. simpl. rewrite H0. rewrite EQ; simpl; eauto. auto.
Qed.
(** Properties of labeled statements *)
@@ -389,6 +386,7 @@ Proof.
destruct si; eauto with cshm.
destruct si; eauto with cshm.
auto.
+ econstructor. eauto. simpl. destruct (Int.eq n Int.zero); auto.
Qed.
Lemma make_cast_float_correct:
@@ -420,7 +418,19 @@ Proof.
rewrite H2. auto with cshm.
(* float -> int *)
rewrite H2. eauto with cshm.
- (* void *)
+ (* int/pointer -> bool *)
+ rewrite H2. econstructor; eauto. simpl. destruct (Int.eq i Int.zero); auto.
+ rewrite H2. econstructor; eauto.
+ (* float -> bool *)
+ rewrite H2. econstructor; eauto with cshm.
+ simpl. unfold Val.cmpf, Val.cmpf_bool. rewrite Float.cmp_ne_eq. rewrite H7; auto.
+ rewrite H2. econstructor; eauto with cshm.
+ simpl. unfold Val.cmpf, Val.cmpf_bool. rewrite Float.cmp_ne_eq. rewrite H7; auto.
+ (* struct -> struct *)
+ rewrite H2. auto.
+ (* union -> union *)
+ rewrite H2. auto.
+ (* any -> void *)
rewrite H5. auto.
Qed.
@@ -439,6 +449,10 @@ Proof.
intros. functional inversion H0; subst; simpl.
exists (Vint n); split; auto.
exists (Vint n); split; auto.
+ exists (Vint n); split; auto.
+ exists (Vint n); split; auto.
+ exists (Vptr b0 ofs); split; auto. constructor.
+ exists (Vptr b0 ofs); split; auto. constructor.
exists (Vptr b0 ofs); split; auto. constructor.
exists (Vptr b0 ofs); split; auto. constructor.
rewrite <- Float.cmp_ne_eq.
@@ -670,31 +684,128 @@ Lemma make_load_correct:
forall addr ty code b ofs v e le m,
make_load addr ty = OK code ->
eval_expr ge e le m addr (Vptr b ofs) ->
- load_value_of_type ty m b ofs = Some v ->
+ deref_loc ge ty m b ofs E0 v ->
+ type_is_volatile ty = false ->
eval_expr ge e le m code v.
Proof.
- unfold make_load, load_value_of_type.
- intros until m; intros MKLOAD EVEXP LDVAL.
- destruct (access_mode ty); inversion MKLOAD.
- (* access_mode ty = By_value m *)
- apply eval_Eload with (Vptr b ofs); auto.
- (* access_mode ty = By_reference *)
- subst code. inversion LDVAL. auto.
+ unfold make_load; intros until m; intros MKLOAD EVEXP DEREF NONVOL.
+ inv DEREF.
+ (* nonvolatile scalar *)
+ rewrite H in MKLOAD. inv MKLOAD. apply eval_Eload with (Vptr b ofs); auto.
+ (* volatile scalar *)
+ congruence.
+ (* by reference *)
+ rewrite H in MKLOAD. inv MKLOAD. auto.
+ (* by copy *)
+ rewrite H in MKLOAD. inv MKLOAD. auto.
+Qed.
+
+Lemma make_vol_load_correct:
+ forall id addr ty code b ofs t v e le m f k,
+ make_vol_load id addr ty = OK code ->
+ eval_expr ge e le m addr (Vptr b ofs) ->
+ deref_loc ge ty m b ofs t v ->
+ type_is_volatile ty = true ->
+ step ge (State f code k e le m) t (State f Sskip k e (PTree.set id v le) m).
+Proof.
+ unfold make_vol_load; intros until k; intros MKLOAD EVEXP DEREF VOL.
+ inv DEREF.
+ (* nonvolatile scalar *)
+ congruence.
+(**
+ rewrite H in MKLOAD. inv MKLOAD.
+ change (PTree.set id v le) with (Cminor.set_optvar (Some id) v le).
+ econstructor. constructor. eauto. constructor. constructor; auto. constructor; auto.
+*)
+ (* volatile scalar *)
+ rewrite H in MKLOAD. inv MKLOAD.
+ change (PTree.set id v le) with (Cminor.set_optvar (Some id) v le).
+ econstructor. constructor. eauto. constructor. constructor; auto.
+ (* by reference *)
+ rewrite H in MKLOAD. inv MKLOAD. constructor; auto.
+ (* by copy *)
+ rewrite H in MKLOAD. inv MKLOAD. constructor; auto.
+Qed.
+
+Remark capped_alignof_divides:
+ forall ty n,
+ (alignof ty | n) -> (Zmin (alignof ty) 4 | n).
+Proof.
+ intros. generalize (alignof_1248 ty).
+ intros [A|[A|[A|A]]]; rewrite A in *; auto.
+ apply Zdivides_trans with 8; auto. exists 2; auto.
Qed.
+Remark capped_alignof_124:
+ forall ty,
+ Zmin (alignof ty) 4 = 1 \/ Zmin (alignof ty) 4 = 2 \/ Zmin (alignof ty) 4 = 4.
+Proof.
+ intros. generalize (alignof_1248 ty).
+ intros [A|[A|[A|A]]]; rewrite A; auto.
+Qed.
+
+Lemma make_memcpy_correct:
+ forall f dst src ty k e le m b ofs v t m',
+ eval_expr ge e le m dst (Vptr b ofs) ->
+ eval_expr ge e le m src v ->
+ assign_loc ge ty m b ofs v t m' ->
+ access_mode ty = By_copy ->
+ step ge (State f (make_memcpy dst src ty) k e le m) t (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.
+ econstructor.
+ econstructor. eauto. econstructor. eauto. constructor.
+ econstructor; eauto.
+ apply capped_alignof_124.
+ apply sizeof_pos.
+ apply capped_alignof_divides. apply sizeof_alignof_compat.
+ apply capped_alignof_divides; auto.
+ apply capped_alignof_divides; auto.
+Qed.
+
Lemma make_store_correct:
- forall addr ty rhs code e le m b ofs v m' f k,
+ forall addr ty rhs code e le m b ofs v t m' f k,
make_store addr ty rhs = OK code ->
eval_expr ge e le m addr (Vptr b ofs) ->
eval_expr ge e le m rhs v ->
- store_value_of_type ty m b ofs v = Some m' ->
- step ge (State f code k e le m) E0 (State f Sskip k e le m').
+ assign_loc ge ty m b ofs v t m' ->
+ type_is_volatile ty = false ->
+ step ge (State f code k e le m) t (State f Sskip k e le m').
+Proof.
+ unfold make_store. intros until k; intros MKSTORE EV1 EV2 ASSIGN NONVOL.
+ inversion ASSIGN; subst.
+ (* nonvolatile scalar *)
+ rewrite H in MKSTORE; inv MKSTORE.
+ econstructor; eauto.
+ (* volatile scalar *)
+ congruence.
+ (* by copy *)
+ rewrite H in MKSTORE; inv MKSTORE.
+ eapply make_memcpy_correct; eauto.
+Qed.
+
+Lemma make_vol_store_correct:
+ forall addr ty rhs code e le m b ofs v t m' f k,
+ make_vol_store addr ty rhs = OK code ->
+ eval_expr ge e le m addr (Vptr b ofs) ->
+ eval_expr ge e le m rhs v ->
+ assign_loc ge ty m b ofs v t m' ->
+ type_is_volatile ty = true ->
+ step ge (State f code k e le m) t (State f Sskip k e le m').
Proof.
- unfold make_store, store_value_of_type.
- intros until k; intros MKSTORE EV1 EV2 STVAL.
- destruct (access_mode ty); inversion MKSTORE.
- (* access_mode ty = By_value m *)
- eapply step_store; eauto.
+ unfold make_vol_store. intros until k; intros MKSTORE EV1 EV2 ASSIGN VOL.
+ inversion ASSIGN; subst.
+ (* nonvolatile scalar *)
+ congruence.
+ (* volatile scalar *)
+ rewrite H in MKSTORE; inv MKSTORE.
+ change le with (Cminor.set_optvar None Vundef le) at 2.
+ econstructor. constructor. eauto. constructor. eauto. constructor.
+ constructor. auto.
+ (* by copy *)
+ rewrite H in MKSTORE; inv MKSTORE.
+ eapply make_memcpy_correct; eauto.
Qed.
End CONSTRUCTORS.
@@ -732,6 +843,49 @@ Lemma var_info_translated:
exists tv, Genv.find_var_info tge b = Some tv /\ transf_globvar transl_globvar v = OK tv.
Proof (Genv.find_var_info_transf_partial2 transl_fundef transl_globvar _ TRANSL).
+Lemma var_info_rev_translated:
+ forall b tv,
+ Genv.find_var_info tge b = Some tv ->
+ exists v, Genv.find_var_info ge b = Some v /\ transf_globvar transl_globvar v = OK tv.
+Proof (Genv.find_var_info_rev_transf_partial2 transl_fundef transl_globvar _ TRANSL).
+
+Lemma block_is_volatile_preserved:
+ forall b, block_is_volatile tge b = block_is_volatile ge b.
+Proof.
+ intros. unfold block_is_volatile.
+ destruct (Genv.find_var_info ge b) as []_eqn.
+ exploit var_info_translated; eauto. intros [tv [A B]]. rewrite A.
+ unfold transf_globvar in B. monadInv B. auto.
+ destruct (Genv.find_var_info tge b) as []_eqn.
+ exploit var_info_rev_translated; eauto. intros [tv [A B]]. congruence.
+ auto.
+Qed.
+
+Lemma deref_loc_preserved:
+ forall ty m b ofs t v,
+ deref_loc ge ty m b ofs t v -> deref_loc tge ty m b ofs t v.
+Proof.
+ intros. inv H.
+ eapply deref_loc_value; eauto.
+ eapply deref_loc_volatile; eauto.
+ eapply volatile_load_preserved with (ge1 := ge); auto.
+ exact symbols_preserved. exact block_is_volatile_preserved.
+ eapply deref_loc_reference; eauto.
+ eapply deref_loc_copy; eauto.
+Qed.
+
+Lemma assign_loc_preserved:
+ forall ty m b ofs v t m',
+ assign_loc ge ty m b ofs v t m' -> assign_loc tge ty m b ofs v t m'.
+Proof.
+ intros. inv H.
+ eapply assign_loc_value; eauto.
+ eapply assign_loc_volatile; eauto.
+ eapply volatile_store_preserved with (ge1 := ge); auto.
+ exact symbols_preserved. exact block_is_volatile_preserved.
+ eapply assign_loc_copy; eauto.
+Qed.
+
(** * Matching between environments *)
(** In this section, we define a matching relation between
@@ -861,32 +1015,43 @@ Qed.
Lemma bind_parameters_match:
forall e m1 vars vals m2,
- Csem.bind_parameters e m1 vars vals m2 ->
+ Csem.bind_parameters ge e m1 vars vals m2 ->
forall te tvars,
val_casted_list vals (type_of_params vars) ->
match_env e te ->
- transl_params vars = OK tvars ->
- Csharpminor.bind_parameters te m1 tvars vals m2.
+ transl_vars vars = OK tvars ->
+ Csharpminor.bind_parameters tge te m1 tvars vals m2.
Proof.
induction 1; intros.
(* base case *)
monadInv H1. constructor.
(* inductive case *)
simpl in H2. destruct H2.
- revert H4; simpl.
- caseEq (chunk_of_type ty); simpl; [intros chunk CHK | congruence].
- caseEq (transl_params params); simpl; [intros tparams TPARAMS | congruence].
- intro EQ; inversion EQ; clear EQ; subst tvars.
- generalize CHK. unfold chunk_of_type.
- caseEq (access_mode ty); intros; try discriminate.
- inversion CHK0; clear CHK0; subst m0.
- unfold store_value_of_type in H0. rewrite H4 in H0.
- apply bind_parameters_cons with b m1.
- exploit me_local; eauto. intros [vk [A B]].
- exploit var_kind_by_value; eauto. congruence.
+ simpl in H4. destruct (var_kind_of_type ty) as [vk|]_eqn; monadInv H4.
+ assert (A: (exists chunk, access_mode ty = By_value chunk /\ Mem.store chunk m b 0 v1 = Some m1)
+ \/ access_mode ty = By_copy).
+ inv H0; auto; left; econstructor; split; eauto. inv H7. auto.
+ destruct A as [[chunk [MODE STORE]] | MODE].
+ (* scalar case *)
+ assert (vk = Vscalar chunk). exploit var_kind_by_value; eauto. congruence.
+ subst vk.
+ apply bind_parameters_scalar with b m1.
+ exploit me_local; eauto. intros [vk [A B]]. congruence.
eapply val_casted_normalized; eauto.
assumption.
apply IHbind_parameters; auto.
+ (* array case *)
+ inv H0; try congruence.
+ exploit var_kind_by_reference; eauto. intros; subst vk.
+ apply bind_parameters_array with b m1.
+ exploit me_local; eauto. intros [vk [A B]]. congruence.
+ econstructor; eauto.
+ apply capped_alignof_124.
+ apply sizeof_pos.
+ apply capped_alignof_divides. apply sizeof_alignof_compat.
+ apply capped_alignof_divides; auto.
+ apply capped_alignof_divides; auto.
+ apply IHbind_parameters; auto.
Qed.
(* ** Correctness of variable accessors *)
@@ -896,16 +1061,21 @@ Qed.
Lemma var_get_correct:
forall e le m id ty loc ofs v code te,
Clight.eval_lvalue ge e le m (Clight.Evar id ty) loc ofs ->
- load_value_of_type ty m loc ofs = Some v ->
+ deref_loc ge ty m loc ofs E0 v ->
var_get id ty = OK code ->
match_env e te ->
eval_expr tge te le m code v.
Proof.
- intros. revert H0 H1. unfold load_value_of_type, var_get.
- case_eq (access_mode ty).
+ unfold var_get; intros.
+ destruct (access_mode ty) as [chunk| | | ]_eqn.
(* access mode By_value *)
- intros chunk ACC LOAD EQ. inv EQ.
- inv H.
+ assert (Mem.loadv chunk m (Vptr loc ofs) = Some v).
+ inv H0.
+ congruence.
+ inv H5. simpl. congruence.
+ congruence.
+ congruence.
+ inv H1. inv H.
(* local variable *)
exploit me_local; eauto. intros [vk [A B]].
assert (vk = Vscalar chunk).
@@ -922,7 +1092,20 @@ Proof.
eauto. eauto.
assumption.
(* access mode By_reference *)
- intros ACC EQ1 EQ2. inv EQ1; inv EQ2; inv H.
+ assert (v = Vptr loc ofs). inv H0; congruence.
+ inv H1. inv H.
+ (* local variable *)
+ exploit me_local; eauto. intros [vk [A B]].
+ eapply eval_Eaddrof.
+ eapply eval_var_addr_local. eauto.
+ (* global variable *)
+ exploit match_env_globals; eauto. intros [A B].
+ eapply eval_Eaddrof.
+ eapply eval_var_addr_global. auto.
+ rewrite symbols_preserved. eauto.
+ (* access mode By_copy *)
+ assert (v = Vptr loc ofs). inv H0; congruence.
+ inv H1. inv H.
(* local variable *)
exploit me_local; eauto. intros [vk [A B]].
eapply eval_Eaddrof.
@@ -939,19 +1122,19 @@ Qed.
(** Correctness of the code generated by [var_set]. *)
Lemma var_set_correct:
- forall e le m id ty loc ofs v m' code te rhs f k,
+ forall e le m id ty loc ofs v t m' code te rhs f k,
Clight.eval_lvalue ge e le m (Clight.Evar id ty) loc ofs ->
val_casted v ty ->
- store_value_of_type ty m loc ofs v = Some m' ->
+ assign_loc ge ty m loc ofs v t m' ->
+ type_is_volatile ty = false ->
var_set id ty rhs = OK code ->
match_env e te ->
eval_expr tge te le m rhs v ->
- step tge (State f code k te le m) E0 (State f Sskip k te le m').
+ step tge (State f code k te le m) t (State f Sskip k te le m').
Proof.
- intros. revert H1 H2. unfold store_value_of_type, var_set.
- caseEq (access_mode ty).
- (* access mode By_value *)
- intros chunk ACC STORE EQ. inv EQ.
+ intros. unfold var_set in H3.
+ inversion H1; subst; rewrite H6 in H3; inv H3.
+ (* scalar, non volatile *)
inv H.
(* local variable *)
exploit me_local; eauto. intros [vk [A B]].
@@ -968,65 +1151,20 @@ Proof.
econstructor. eapply eval_var_ref_global. auto.
rewrite symbols_preserved. eauto.
eauto. eauto.
- eapply val_casted_normalized; eauto. assumption.
- (* access mode By_reference *)
- congruence.
- (* access mode By_nothing *)
+ eapply val_casted_normalized; eauto. assumption.
+ (* scalar, volatile *)
congruence.
+ (* array *)
+ assert (eval_expr tge te le m (Eaddrof id) (Vptr loc ofs)).
+ inv H.
+ exploit me_local; eauto. intros [vk [A B]].
+ constructor. eapply eval_var_addr_local; eauto.
+ exploit match_env_globals; eauto. intros [A B].
+ constructor. eapply eval_var_addr_global; eauto.
+ rewrite symbols_preserved. eauto.
+ eapply make_memcpy_correct; eauto. eapply assign_loc_preserved; eauto.
Qed.
-(****************************
-Lemma call_dest_correct:
- forall e m lhs loc ofs optid te,
- Csem.eval_lvalue ge e m lhs loc ofs ->
- transl_lhs_call (Some lhs) = OK optid ->
- match_env e te ->
- exists id,
- optid = Some id
- /\ ofs = Int.zero
- /\ match access_mode (typeof lhs) with
- | By_value chunk => eval_var_ref tge te id loc chunk
- | _ => True
- end.
-Proof.
- intros. revert H0. simpl. caseEq (is_variable lhs); try congruence.
- intros id ISV EQ. inv EQ.
- exploit is_variable_correct; eauto. intro EQ.
- rewrite EQ in H. clear EQ.
- exists id. split; auto.
- inv H.
-(* local variable *)
- split. auto.
- exploit me_local; eauto. intros [vk [A B]].
- case_eq (access_mode (typeof lhs)); intros; auto.
- assert (vk = Vscalar m0).
- exploit var_kind_by_value; eauto. congruence.
- subst vk. apply eval_var_ref_local; auto.
-(* global variable *)
- split. auto.
- exploit match_env_globals; eauto. intros [A B].
- case_eq (access_mode (typeof lhs)); intros; auto.
- exploit B; eauto. intros [gv [C D]].
- eapply eval_var_ref_global; eauto.
- rewrite symbols_preserved. auto.
-Qed.
-
-Lemma set_call_dest_correct:
- forall ty m loc v m' e te id,
- store_value_of_type ty m loc Int.zero v = Some m' ->
- match access_mode ty with
- | By_value chunk => eval_var_ref tge te id loc chunk
- | _ => True
- end ->
- match_env e te ->
- exec_opt_assign tge te m (Some id) v m'.
-Proof.
- intros. generalize H. unfold store_value_of_type. case_eq (access_mode ty); intros; try congruence.
- rewrite H2 in H0.
- constructor. econstructor. eauto. auto.
-Qed.
-**************************)
-
(** * Proof of semantic preservation *)
(** ** Semantic preservation for expressions *)
@@ -1097,7 +1235,7 @@ Proof.
(* Case a is a variable *)
subst a. eapply var_get_correct; eauto.
(* Case a is another lvalue *)
- eapply make_load_correct; eauto.
+ eapply make_load_correct; eauto. eapply deref_loc_preserved; eauto.
(* var local *)
exploit (me_local _ _ MENV); eauto.
intros [vk [A B]].
@@ -1352,11 +1490,17 @@ Proof.
(* skip *)
auto.
(* assign *)
- simpl in TR. destruct (is_variable e); monadInv TR.
- unfold var_set in EQ0. destruct (access_mode (typeof e)); inv EQ0. auto.
- unfold make_store in EQ2. destruct (access_mode (typeof e)); inv EQ2. auto.
+ simpl in TR. destruct (type_is_volatile (typeof e)) as []_eqn.
+ monadInv TR. unfold make_vol_store, make_memcpy in EQ2.
+ destruct (access_mode (typeof e)); inv EQ2; auto.
+ destruct (is_variable e); monadInv TR.
+ unfold var_set, make_memcpy in EQ0.
+ destruct (access_mode (typeof e)); inv EQ0; auto.
+ unfold make_store, make_memcpy in EQ2. destruct (access_mode (typeof e)); inv EQ2; auto.
(* set *)
auto.
+(* vol load *)
+ unfold make_vol_load in EQ0. destruct (access_mode (typeof e)); inv EQ0; auto.
(* call *)
simpl in TR. destruct (classify_fun (typeof e)); monadInv TR. auto.
(* seq *)
@@ -1451,25 +1595,39 @@ Proof.
induction 1; intros T1 MST; inv MST.
(* assign *)
- revert TR. simpl. case_eq (is_variable a1); intros; monadInv TR.
- exploit is_variable_correct; eauto. intro EQ1. rewrite EQ1 in H.
- assert (ts' = ts /\ tk' = tk).
+ simpl in TR. destruct (type_is_volatile (typeof a1)) as []_eqn.
+ (* Case 1: volatile *)
+ monadInv TR.
+ assert (SAME: ts' = ts /\ tk' = tk).
+ inversion MTR. auto.
+ subst ts. unfold make_vol_store, make_memcpy in EQ2.
+ destruct (access_mode (typeof a1)); congruence.
+ destruct SAME; subst ts' tk'.
+ econstructor; split.
+ apply plus_one. eapply make_vol_store_correct; eauto.
+ eapply transl_lvalue_correct; eauto. eapply make_cast_correct; eauto.
+ eapply transl_expr_correct; eauto. eapply assign_loc_preserved; eauto.
+ eapply match_states_skip; eauto.
+ (* Case 2: variable *)
+ destruct (is_variable a1) as []_eqn; monadInv TR.
+ assert (SAME: ts' = ts /\ tk' = tk).
inversion MTR. auto.
- subst ts. unfold var_set in EQ0. destruct (access_mode (typeof a1)); congruence.
- destruct H4; subst ts' tk'.
+ subst ts. unfold var_set, make_memcpy in EQ0. destruct (access_mode (typeof a1)); congruence.
+ destruct SAME; subst ts' tk'.
+ exploit is_variable_correct; eauto. intro EQ1. rewrite EQ1 in H.
econstructor; split.
apply plus_one. eapply var_set_correct; eauto. exists v2; exists (typeof a2); auto.
eapply make_cast_correct; eauto. eapply transl_expr_correct; eauto.
eapply match_states_skip; eauto.
-
- assert (ts' = ts /\ tk' = tk).
+ (* Case 3: everything else *)
+ assert (SAME: ts' = ts /\ tk' = tk).
inversion MTR. auto.
- subst ts. unfold make_store in EQ2. destruct (access_mode (typeof a1)); congruence.
- destruct H4; subst ts' tk'.
+ subst ts. unfold make_store, make_memcpy in EQ2. destruct (access_mode (typeof a1)); congruence.
+ destruct SAME; subst ts' tk'.
econstructor; split.
apply plus_one. eapply make_store_correct; eauto.
- exploit transl_lvalue_correct; eauto.
- eapply make_cast_correct; eauto. eapply transl_expr_correct; eauto.
+ eapply transl_lvalue_correct; eauto. eapply make_cast_correct; eauto.
+ eapply transl_expr_correct; eauto. eapply assign_loc_preserved; eauto.
eapply match_states_skip; eauto.
(* set *)
@@ -1477,6 +1635,17 @@ Proof.
apply plus_one. econstructor. eapply transl_expr_correct; eauto.
eapply match_states_skip; eauto.
+(* vol read *)
+ monadInv TR.
+ assert (SAME: ts' = ts /\ tk' = tk).
+ inversion MTR. auto.
+ subst ts. unfold make_vol_load in EQ0. destruct (access_mode (typeof a)); congruence.
+ destruct SAME; subst ts' tk'.
+ econstructor; split.
+ apply plus_one. eapply make_vol_load_correct; eauto. eapply transl_lvalue_correct; eauto.
+ eapply deref_loc_preserved; eauto.
+ eapply match_states_skip; eauto.
+
(* call *)
revert TR. simpl. case_eq (classify_fun (typeof a)); try congruence.
intros targs tres CF TR. monadInv TR. inv MTR.
@@ -1759,8 +1928,8 @@ Proof.
monadInv TR. monadInv EQ.
exploit match_cont_is_call_cont; eauto. intros [A B].
exploit match_env_alloc_variables; eauto.
- apply match_env_empty.
- apply transl_fn_variables. eauto. eauto.
+ apply match_env_empty.
+ eapply transl_fn_variables; eauto.
intros [te1 [C D]].
econstructor; split.
apply plus_one. econstructor.
@@ -1806,7 +1975,7 @@ Proof.
rewrite symbols_preserved. replace (prog_main tprog) with (prog_main prog).
auto. symmetry. unfold transl_program in TRANSL.
eapply transform_partial_program2_main; eauto.
- assert (funsig tf = signature_of_type Tnil (Tint I32 Signed)).
+ assert (funsig tf = signature_of_type Tnil type_int32s).
eapply transl_fundef_sig2; eauto.
econstructor; split.
econstructor; eauto. eapply Genv.init_mem_transf_partial2; eauto.
diff --git a/cfrontend/Cstrategy.v b/cfrontend/Cstrategy.v
index 8b66ef9..e088c26 100644
--- a/cfrontend/Cstrategy.v
+++ b/cfrontend/Cstrategy.v
@@ -49,9 +49,9 @@ Fixpoint simple (a: expr) : Prop :=
| Eloc _ _ _ => True
| Evar _ _ => True
| Ederef r _ => simple r
- | Efield l1 _ _ => simple l1
+ | Efield r _ _ => simple r
| Eval _ _ => True
- | Evalof l _ => simple l
+ | Evalof l _ => simple l /\ type_is_volatile (typeof l) = false
| Eaddrof l _ => simple l
| Eunop _ r1 _ => simple r1
| Ebinop _ r1 r2 _ => simple r1 /\ simple r2
@@ -93,22 +93,22 @@ Inductive eval_simple_lvalue: expr -> block -> int -> Prop :=
| esl_deref: forall r ty b ofs,
eval_simple_rvalue r (Vptr b ofs) ->
eval_simple_lvalue (Ederef r ty) b ofs
- | esl_field_struct: forall l f ty b ofs id fList delta,
- eval_simple_lvalue l b ofs ->
- typeof l = Tstruct id fList -> field_offset f fList = OK delta ->
- eval_simple_lvalue (Efield l f ty) b (Int.add ofs (Int.repr delta))
- | esl_field_union: forall l f ty b ofs id fList,
- eval_simple_lvalue l b ofs ->
- typeof l = Tunion id fList ->
- eval_simple_lvalue (Efield l f ty) b ofs
+ | esl_field_struct: forall r f ty b ofs id fList a delta,
+ eval_simple_rvalue r (Vptr b ofs) ->
+ typeof r = Tstruct id fList a -> field_offset f fList = OK delta ->
+ eval_simple_lvalue (Efield r f ty) b (Int.add ofs (Int.repr delta))
+ | esl_field_union: forall r f ty b ofs id fList a,
+ eval_simple_rvalue r (Vptr b ofs) ->
+ typeof r = Tunion id fList a ->
+ eval_simple_lvalue (Efield r f ty) b ofs
with eval_simple_rvalue: expr -> val -> Prop :=
| esr_val: forall v ty,
eval_simple_rvalue (Eval v ty) v
| esr_rvalof: forall b ofs l ty v,
eval_simple_lvalue l b ofs ->
- ty = typeof l ->
- load_value_of_type ty m b ofs = Some v ->
+ ty = typeof l -> type_is_volatile ty = false ->
+ deref_loc ge ty m b ofs E0 v ->
eval_simple_rvalue (Evalof l ty) v
| esr_addrof: forall b ofs l ty,
eval_simple_lvalue l b ofs ->
@@ -151,7 +151,7 @@ Inductive leftcontext: kind -> kind -> (expr -> expr) -> Prop :=
| lctx_deref: forall k C ty,
leftcontext k RV C -> leftcontext k LV (fun x => Ederef (C x) ty)
| lctx_field: forall k C f ty,
- leftcontext k LV C -> leftcontext k LV (fun x => Efield (C x) f ty)
+ leftcontext k RV C -> leftcontext k LV (fun x => Efield (C x) f ty)
| lctx_rvalof: forall k C ty,
leftcontext k LV C -> leftcontext k RV (fun x => Evalof (C x) ty)
| lctx_addrof: forall k C ty,
@@ -222,6 +222,14 @@ Inductive estep: state -> trace -> state -> Prop :=
estep (ExprState f r k e m)
E0 (ExprState f (Eval v ty) k e m)
+ | step_rvalof_volatile: forall f C l ty k e m b ofs t v,
+ leftcontext RV RV C ->
+ eval_simple_lvalue e m l b ofs ->
+ deref_loc ge ty m b ofs t v ->
+ ty = typeof l -> type_is_volatile ty = true ->
+ estep (ExprState f (C (Evalof l ty)) k e m)
+ t (ExprState f (C (Eval v ty)) k e m)
+
| step_condition: forall f C r1 r2 r3 ty k e m v b,
leftcontext RV RV C ->
eval_simple_rvalue e m r1 v ->
@@ -229,38 +237,73 @@ Inductive estep: state -> trace -> state -> Prop :=
estep (ExprState f (C (Econdition r1 r2 r3 ty)) k e m)
E0 (ExprState f (C (Eparen (if b then r2 else r3) ty)) k e m)
- | step_assign: forall f C l r ty k e m b ofs v v' m',
+ | step_assign: forall f C l r ty k e m b ofs v v' t m',
leftcontext RV RV C ->
eval_simple_lvalue e m l b ofs ->
eval_simple_rvalue e m r v ->
sem_cast v (typeof r) (typeof l) = Some v' ->
- store_value_of_type (typeof l) m b ofs v' = Some m' ->
+ assign_loc ge (typeof l) m b ofs v' t m' ->
ty = typeof l ->
estep (ExprState f (C (Eassign l r ty)) k e m)
- E0 (ExprState f (C (Eval v' ty)) k e m')
+ t (ExprState f (C (Eval v' ty)) k e m')
- | step_assignop: forall f C op l r tyres ty k e m b ofs v1 v2 v3 v4 m',
+ | step_assignop: forall f C op l r tyres ty k e m b ofs v1 v2 v3 v4 t1 t2 m' t,
leftcontext RV RV C ->
eval_simple_lvalue e m l b ofs ->
- load_value_of_type (typeof l) m b ofs = Some v1 ->
+ deref_loc ge (typeof l) m b ofs t1 v1 ->
eval_simple_rvalue e m r v2 ->
sem_binary_operation op v1 (typeof l) v2 (typeof r) m = Some v3 ->
sem_cast v3 tyres (typeof l) = Some v4 ->
- store_value_of_type (typeof l) m b ofs v4 = Some m' ->
+ assign_loc ge (typeof l) m b ofs v4 t2 m' ->
ty = typeof l ->
+ t = t1 ** t2 ->
estep (ExprState f (C (Eassignop op l r tyres ty)) k e m)
- E0 (ExprState f (C (Eval v4 ty)) k e m')
+ t (ExprState f (C (Eval v4 ty)) k e m')
- | step_postincr: forall f C id l ty k e m b ofs v1 v2 v3 m',
+ | step_assignop_stuck: forall f C op l r tyres ty k e m b ofs v1 v2 t,
leftcontext RV RV C ->
eval_simple_lvalue e m l b ofs ->
- load_value_of_type ty m b ofs = Some v1 ->
+ deref_loc ge (typeof l) m b ofs t v1 ->
+ eval_simple_rvalue e m r v2 ->
+ match sem_binary_operation op v1 (typeof l) v2 (typeof r) m with
+ | None => True
+ | Some v3 =>
+ match sem_cast v3 tyres (typeof l) with
+ | None => True
+ | Some v4 => forall t2 m', ~(assign_loc ge (typeof l) m b ofs v4 t2 m')
+ end
+ end ->
+ ty = typeof l ->
+ estep (ExprState f (C (Eassignop op l r tyres ty)) k e m)
+ t Stuckstate
+
+ | step_postincr: forall f C id l ty k e m b ofs v1 v2 v3 t1 t2 m' t,
+ leftcontext RV RV C ->
+ eval_simple_lvalue e m l b ofs ->
+ deref_loc ge ty m b ofs t1 v1 ->
sem_incrdecr id v1 ty = Some v2 ->
sem_cast v2 (typeconv ty) ty = Some v3 ->
- store_value_of_type ty m b ofs v3 = Some m' ->
+ assign_loc ge ty m b ofs v3 t2 m' ->
+ ty = typeof l ->
+ t = t1 ** t2 ->
+ estep (ExprState f (C (Epostincr id l ty)) k e m)
+ t (ExprState f (C (Eval v1 ty)) k e m')
+
+ | step_postincr_stuck: forall f C id l ty k e m b ofs v1 t,
+ leftcontext RV RV C ->
+ eval_simple_lvalue e m l b ofs ->
+ deref_loc ge ty m b ofs t v1 ->
+ match sem_incrdecr id v1 ty with
+ | None => True
+ | Some v2 =>
+ match sem_cast v2 (typeconv ty) ty with
+ | None => True
+ | Some v3 => forall t2 m', ~(assign_loc ge (typeof l) m b ofs v3 t2 m')
+ end
+ end ->
ty = typeof l ->
estep (ExprState f (C (Epostincr id l ty)) k e m)
- E0 (ExprState f (C (Eval v1 ty)) k e m')
+ t Stuckstate
| step_comma: forall f C r1 r2 ty k e m v,
leftcontext RV RV C ->
@@ -326,53 +369,37 @@ Proof.
Qed.
Lemma star_safe:
- forall s1 s2 s3,
- safe s1 -> star Csem.step ge s1 E0 s2 -> (safe s2 -> star Csem.step ge s2 E0 s3) ->
- star Csem.step ge s1 E0 s3.
+ forall s1 s2 t s3,
+ safe s1 -> star Csem.step ge s1 E0 s2 -> (safe s2 -> star Csem.step ge s2 t s3) ->
+ star Csem.step ge s1 t s3.
Proof.
intros. eapply star_trans; eauto. apply H1. eapply safe_steps; eauto. auto.
Qed.
Lemma plus_safe:
- forall s1 s2 s3,
- safe s1 -> star Csem.step ge s1 E0 s2 -> (safe s2 -> plus Csem.step ge s2 E0 s3) ->
- plus Csem.step ge s1 E0 s3.
+ forall s1 s2 t s3,
+ safe s1 -> star Csem.step ge s1 E0 s2 -> (safe s2 -> plus Csem.step ge s2 t s3) ->
+ plus Csem.step ge s1 t s3.
Proof.
intros. eapply star_plus_trans; eauto. apply H1. eapply safe_steps; eauto. auto.
Qed.
-Remark not_stuck_val:
- forall e v ty m,
- not_stuck ge e (Eval v ty) m.
-Proof.
- intros; red; intros. inv H; try congruence. subst e'. constructor.
-Qed.
-
-Lemma safe_not_stuck:
- forall f a k e m,
- safe (ExprState f a k e m) ->
- not_stuck ge e a m.
-Proof.
- intros. exploit H. apply star_refl. intros [[r FIN] | [t [s' STEP]]].
- inv FIN.
- inv STEP.
- inv H0; auto; apply not_stuck_val.
- inv H0; apply not_stuck_val.
-Qed.
-
-Hint Resolve safe_not_stuck.
+Require Import Classical.
-Lemma safe_not_imm_stuck:
- forall k C f a K e m,
- safe (ExprState f (C a) K e m) ->
- context k RV C ->
- not_imm_stuck ge e k a m.
+Lemma safe_imm_safe:
+ forall f C a k e m K,
+ safe (ExprState f (C a) k e m) ->
+ context K RV C ->
+ imm_safe ge e K a m.
Proof.
- intros. exploit safe_not_stuck; eauto.
+ intros. destruct (classic (imm_safe ge e K a m)); auto.
+ 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.
Qed.
-(** Simple, non-stuck expressions are well-formed with respect to
- l-values and r-values. *)
+(** Safe expressions are well-formed with respect to l-values and r-values. *)
Definition expr_kind (a: expr) : kind :=
match a with
@@ -390,7 +417,7 @@ Proof.
Qed.
Lemma rred_kind:
- forall a m a' m', rred a m a' m' -> expr_kind a = RV.
+ forall a m t a' m', rred ge a m t a' m' -> expr_kind a = RV.
Proof.
induction 1; auto.
Qed.
@@ -407,8 +434,8 @@ Proof.
induction 1; intros; simpl; auto.
Qed.
-Lemma not_imm_stuck_kind:
- forall e k a m, not_imm_stuck ge e k a m -> expr_kind a = k.
+Lemma imm_safe_kind:
+ forall e k a m, imm_safe ge e k a m -> expr_kind a = k.
Proof.
induction 1.
auto.
@@ -424,10 +451,10 @@ Lemma safe_expr_kind:
safe (ExprState f (C a) k e m) ->
expr_kind a = from.
Proof.
- intros. eapply not_imm_stuck_kind. eapply safe_not_imm_stuck; eauto.
+ intros. eapply imm_safe_kind. eapply safe_imm_safe; eauto.
Qed.
-(** Painful inversion lemmas on particular states that are not stuck. *)
+(** Painful inversion lemmas on particular states that are safe. *)
Section INVERSION_LEMMAS.
@@ -449,15 +476,16 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
\/ (e!x = None /\ Genv.find_symbol ge x = Some b /\ type_of_global ge b = Some ty)
| Ederef (Eval v ty1) ty =>
exists b, exists ofs, v = Vptr b ofs
- | Efield (Eloc b ofs ty1) f ty =>
+ | Efield (Eval v ty1) f ty =>
+ exists b, exists ofs, v = Vptr b ofs /\
match ty1 with
- | Tstruct _ fList => exists delta, field_offset f fList = Errors.OK delta
- | Tunion _ _ => True
+ | Tstruct _ fList _ => exists delta, field_offset f fList = Errors.OK delta
+ | Tunion _ _ _ => True
| _ => False
end
| Eval v ty => False
| Evalof (Eloc b ofs ty') ty =>
- ty' = ty /\ exists v, load_value_of_type ty m b ofs = Some v
+ ty' = ty /\ exists t, exists v, deref_loc ge ty m b ofs t v
| Eunop op (Eval v1 ty1) ty =>
exists v, sem_unary_operation op v1 ty1 = Some v
| Ebinop op (Eval v1 ty1) (Eval v2 ty2) ty =>
@@ -467,22 +495,16 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
| Econdition (Eval v1 ty1) r1 r2 ty =>
exists b, bool_val v1 ty1 = Some b
| Eassign (Eloc b ofs ty1) (Eval v2 ty2) ty =>
- exists v, exists m',
- ty = ty1 /\ sem_cast v2 ty2 ty1 = Some v /\ store_value_of_type ty1 m b ofs v = Some m'
+ exists v, exists m', exists t,
+ ty = ty1 /\ sem_cast v2 ty2 ty1 = Some v /\ assign_loc ge ty1 m b ofs v t m'
| Eassignop op (Eloc b ofs ty1) (Eval v2 ty2) tyres ty =>
- exists v1, exists v, exists v', exists m',
- ty = ty1
- /\ load_value_of_type ty1 m b ofs = Some v1
- /\ sem_binary_operation op v1 ty1 v2 ty2 m = Some v
- /\ sem_cast v tyres ty1 = Some v'
- /\ store_value_of_type ty1 m b ofs v' = Some m'
+ exists t, exists v1,
+ ty = ty1
+ /\ deref_loc ge ty1 m b ofs t v1
| Epostincr id (Eloc b ofs ty1) ty =>
- exists v1, exists v2, exists v3, exists m',
+ exists t, exists v1,
ty = ty1
- /\ load_value_of_type ty m b ofs = Some v1
- /\ sem_incrdecr id v1 ty = Some v2
- /\ sem_cast v2 (typeconv ty) ty = Some v3
- /\ store_value_of_type ty m b ofs v3 = Some m'
+ /\ deref_loc ge ty m b ofs t v1
| Ecomma (Eval v ty1) r2 ty =>
typeof r2 = ty
| Eparen (Eval v1 ty1) ty =>
@@ -504,21 +526,22 @@ Proof.
exists b; auto.
exists b; auto.
exists b; exists ofs; auto.
- exists delta; auto.
+ exists b; exists ofs; split; auto. exists delta; auto.
+ exists b; exists ofs; auto.
Qed.
Lemma rred_invert:
- forall r m r' m', rred r m r' m' -> invert_expr_prop r m.
+ forall r m t r' m', rred ge r m t r' m' -> invert_expr_prop r m.
Proof.
induction 1; red; auto.
- split; auto; exists v; auto.
+ split; auto; exists t; exists v; auto.
exists v; auto.
exists v; auto.
exists v; auto.
exists b; auto.
- exists v; exists m'; auto.
- exists v1; exists v; exists v'; exists m'; auto.
- exists v1; exists v2; exists v3; exists m'; auto.
+ exists v; exists m'; exists t; auto.
+ exists t; exists v1; auto.
+ exists t; exists v1; auto.
exists v; auto.
Qed.
@@ -569,9 +592,9 @@ Proof.
red; intros. destruct e1; auto. elim (H0 a m); auto.
Qed.
-Lemma not_imm_stuck_inv:
+Lemma imm_safe_inv:
forall k a m,
- not_imm_stuck ge e k a m ->
+ imm_safe ge e k a m ->
match a with
| Eloc _ _ _ => True
| Eval _ _ => True
@@ -603,7 +626,7 @@ Lemma safe_inv:
| _ => invert_expr_prop a m
end.
Proof.
- intros. eapply not_imm_stuck_inv; eauto. eapply safe_not_imm_stuck; eauto.
+ intros. eapply imm_safe_inv; eauto. eapply safe_imm_safe; eauto.
Qed.
End INVERSION_LEMMAS.
@@ -619,16 +642,16 @@ Variable m: mem.
Lemma eval_simple_steps:
(forall a v, eval_simple_rvalue e m a v ->
- forall C, context RV RV C -> safe (ExprState f (C a) k e m) ->
+ forall C, context RV RV C ->
star Csem.step ge (ExprState f (C a) k e m)
E0 (ExprState f (C (Eval v (typeof a))) k e m))
/\ (forall a b ofs, eval_simple_lvalue e m a b ofs ->
- forall C, context LV RV C -> safe (ExprState f (C a) k e m) ->
+ forall C, context LV RV C ->
star Csem.step ge (ExprState f (C a) k e m)
E0 (ExprState f (C (Eloc b ofs (typeof a))) k e m)).
Proof.
-Ltac Steps REC C' := eapply star_safe; eauto; [apply (REC C'); eauto | intros].
+Ltac Steps REC C' := eapply star_trans; [apply (REC C'); eauto | idtac | simpl; reflexivity].
Ltac FinishR := apply star_one; left; apply step_rred; eauto; simpl; try (econstructor; eauto; fail).
Ltac FinishL := apply star_one; left; apply step_lred; eauto; simpl; try (econstructor; eauto; fail).
@@ -665,14 +688,14 @@ Qed.
Lemma eval_simple_rvalue_steps:
forall a v, eval_simple_rvalue e m a v ->
- forall C, context RV RV C -> safe (ExprState f (C a) k e m) ->
+ forall C, context RV RV C ->
star Csem.step ge (ExprState f (C a) k e m)
E0 (ExprState f (C (Eval v (typeof a))) k e m).
Proof (proj1 eval_simple_steps).
Lemma eval_simple_lvalue_steps:
forall a b ofs, eval_simple_lvalue e m a b ofs ->
- forall C, context LV RV C -> safe (ExprState f (C a) k e m) ->
+ forall C, context LV RV C ->
star Csem.step ge (ExprState f (C a) k e m)
E0 (ExprState f (C (Eloc b ofs (typeof a))) k e m).
Proof (proj2 eval_simple_steps).
@@ -727,14 +750,16 @@ Ltac StepR REC C' a :=
exists b; exists Int.zero.
intuition. apply esl_var_local; auto. apply esl_var_global; auto.
(* field *)
- StepL IHa (fun x => C(Efield x f0 ty)) a.
- exploit safe_inv. eexact SAFE0. eauto. simpl. case_eq (typeof a); intros; try contradiction.
- destruct H0 as [delta OFS]. exists b; exists (Int.add ofs (Int.repr delta)); econstructor; eauto.
+ StepR IHa (fun x => C(Efield x f0 ty)) a.
+ exploit safe_inv. eexact SAFE0. eauto. simpl.
+ intros [b [ofs [EQ TY]]]. subst v. destruct (typeof a) as []_eqn; try contradiction.
+ destruct TY as [delta OFS]. exists b; exists (Int.add ofs (Int.repr delta)); econstructor; eauto.
exists b; exists ofs; econstructor; eauto.
(* valof *)
- StepL IHa (fun x => C(Evalof x ty)) a.
- exploit safe_inv. eexact SAFE0. eauto. simpl. intros [TY [v LOAD]].
- exists v; econstructor; eauto.
+ destruct S. StepL IHa (fun x => C(Evalof x ty)) a.
+ exploit safe_inv. eexact SAFE0. eauto. simpl. intros [TY [t [v LOAD]]].
+ assert (t = E0). inv LOAD; auto. congruence. subst t.
+ exists v; econstructor; eauto. congruence.
(* deref *)
StepR IHa (fun x => C(Ederef x ty)) a.
exploit safe_inv. eexact SAFE0. eauto. simpl. intros [b [ofs EQ]].
@@ -875,7 +900,7 @@ Hint Resolve contextlist'_head contextlist'_tail.
Lemma eval_simple_list_steps:
forall rl vl, eval_simple_list' rl vl ->
- forall C, contextlist' C -> safe (ExprState f (C rl) k e m) ->
+ forall C, contextlist' C ->
star Csem.step ge (ExprState f (C rl) k e m)
E0 (ExprState f (C (rval_list vl rl)) k e m).
Proof.
@@ -883,9 +908,10 @@ Proof.
(* nil *)
apply star_refl.
(* cons *)
- eapply star_safe. eauto.
+ eapply star_trans.
eapply eval_simple_rvalue_steps with (C := fun x => C(Econs x rl)); eauto.
- intros. apply IHeval_simple_list' with (C := fun x => C(Econs (Eval v (typeof r)) x)); auto.
+ apply IHeval_simple_list' with (C := fun x => C(Econs (Eval v (typeof r)) x)); auto.
+ auto.
Qed.
Lemma simple_list_can_eval:
@@ -926,7 +952,8 @@ Variable m: mem.
Definition simple_side_effect (r: expr) : Prop :=
match r with
- | Econdition r1 r2 r3 ty => simple r1
+ | Evalof l _ => simple l /\ type_is_volatile (typeof l) = true
+ | Econdition r1 r2 r3 _ => simple r1
| Eassign l1 r2 _ => simple l1 /\ simple r2
| Eassignop _ l1 r2 _ _ => simple l1 /\ simple r2
| Epostincr _ l1 _ => simple l1
@@ -964,9 +991,11 @@ Ltac Base :=
right; exists (fun x => x); econstructor; split; [eauto | simpl; auto].
(* field *)
- Kind. Rec H LV C (fun x => Efield x f0 ty).
+ Kind. Rec H RV C (fun x => Efield x f0 ty).
(* rvalof *)
Kind. Rec H LV C (fun x => Evalof x ty).
+ destruct (type_is_volatile (typeof l)) as []_eqn.
+ Base. auto.
(* deref *)
Kind. Rec H RV C (fun x => Ederef x ty).
(* addrof *)
@@ -1020,95 +1049,185 @@ End DECOMPOSITION.
Lemma estep_simulation:
forall S t S',
- estep S t S' -> safe S -> plus Csem.step ge S t S'.
+ estep S t S' -> plus Csem.step ge S t S'.
Proof.
intros. inv H.
(* simple *)
exploit eval_simple_rvalue_steps; eauto. simpl; intros STEPS.
- exploit star_inv; eauto. intros [[EQ1 EQ2] | A]; auto.
- inversion EQ1. rewrite <- H3 in H2; contradiction.
+ exploit star_inv; eauto. intros [[EQ1 EQ2] | A]; eauto.
+ inversion EQ1. rewrite <- H2 in H1; contradiction.
+(* valof volatile *)
+ 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.
(* condition *)
- eapply plus_safe; eauto.
+ eapply plus_right.
eapply eval_simple_rvalue_steps with (C := fun x => C(Econdition x r2 r3 ty)); eauto.
- intros. apply plus_one. left; apply step_rred; eauto. constructor; auto.
+ left; apply step_rred; eauto. constructor; auto. auto.
(* assign *)
- eapply plus_safe; eauto.
+ eapply star_plus_trans.
eapply eval_simple_lvalue_steps with (C := fun x => C(Eassign x r (typeof l))); eauto.
- intros. eapply plus_safe; eauto.
+ eapply plus_right.
eapply eval_simple_rvalue_steps with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto.
- intros. apply plus_one; left; apply step_rred; eauto. econstructor; eauto.
+ left; apply step_rred; eauto. econstructor; eauto.
+ reflexivity. auto.
(* assignop *)
- eapply plus_safe; eauto.
+ eapply star_plus_trans.
eapply eval_simple_lvalue_steps with (C := fun x => C(Eassignop op x r tyres (typeof l))); eauto.
- intros. eapply plus_safe; eauto.
+ eapply star_plus_trans.
eapply eval_simple_rvalue_steps with (C := fun x => C(Eassignop op (Eloc b ofs (typeof l)) x tyres (typeof l))); eauto.
- intros. apply plus_one; left; apply step_rred; eauto. econstructor; eauto.
+ 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.
+ reflexivity. reflexivity. reflexivity. traceEq.
+(* assignop stuck *)
+ 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.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eassignop op (Eloc b ofs (typeof l)) x tyres (typeof l))); eauto.
+ eapply plus_left.
+ left; apply step_rred; auto. econstructor; eauto.
+ destruct (sem_binary_operation 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; 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.
+ reflexivity.
+ apply star_one.
+ left; eapply step_stuck with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto.
+ red; intros. exploit imm_safe_inv; eauto. simpl. intros [v3 A]. congruence.
+ reflexivity.
+ reflexivity. traceEq.
(* postincr *)
- eapply plus_safe; eauto.
+ eapply star_plus_trans.
+ eapply eval_simple_lvalue_steps with (C := fun x => C(Epostincr id x (typeof l))); eauto.
+ eapply plus_left.
+ 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.
+ 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.
+ reflexivity. reflexivity. reflexivity. traceEq.
+(* postincr stuck *)
+ eapply star_plus_trans.
eapply eval_simple_lvalue_steps with (C := fun x => C(Epostincr id x (typeof l))); eauto.
- intros. apply plus_one; left; apply step_rred; eauto. econstructor; eauto.
+ eapply plus_left.
+ left; apply step_rred; auto. econstructor; eauto.
+ set (op := match id with Incr => Oadd | Decr => Osub end).
+ assert (SEM: sem_binary_operation op v1 (typeof l) (Vint Int.one) type_int32s m =
+ sem_incrdecr id v1 (typeof l)).
+ destruct id; auto.
+ destruct (sem_incrdecr id v1 (typeof l)) as [v2|].
+ 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; eauto.
+ apply star_one.
+ left; eapply step_stuck with (C := fun x => C (Ecomma x (Eval v1 (typeof l)) (typeof l))); eauto.
+ red; intros. exploit imm_safe_inv; eauto. simpl. intros [v3 [m' [t' [A [B D]]]]].
+ rewrite B in H3. eelim H3; eauto.
+ reflexivity.
+ 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.
+ traceEq.
(* comma *)
- eapply plus_safe; eauto.
+ eapply plus_right.
eapply eval_simple_rvalue_steps with (C := fun x => C(Ecomma x r2 (typeof r2))); eauto.
- intros. apply plus_one; left; apply step_rred; eauto. econstructor; eauto.
+ left; apply step_rred; eauto. econstructor; eauto. auto.
(* paren *)
- eapply plus_safe; eauto.
+ eapply plus_right; eauto.
eapply eval_simple_rvalue_steps with (C := fun x => C(Eparen x ty)); eauto.
- intros. apply plus_one; left; apply step_rred; eauto. econstructor; eauto.
+ left; apply step_rred; eauto. econstructor; eauto. auto.
(* call *)
exploit eval_simple_list_implies; eauto. intros [vl' [A B]].
- eapply plus_safe; eauto.
+ eapply star_plus_trans.
eapply eval_simple_rvalue_steps with (C := fun x => C(Ecall x rargs ty)); eauto.
- intros. eapply plus_safe; eauto.
+ eapply plus_right.
eapply eval_simple_list_steps with (C := fun x => C(Ecall (Eval vf (typeof rf)) x ty)); eauto.
eapply contextlist'_intro with (rl0 := Enil); auto.
- intros. apply plus_one; left; apply Csem.step_call; eauto.
- econstructor; eauto.
+ left; apply Csem.step_call; eauto. econstructor; eauto.
+ traceEq. auto.
Qed.
Lemma can_estep:
forall f a k e m,
safe (ExprState f a k e m) ->
match a with Eval _ _ => False | _ => True end ->
- exists S, estep (ExprState f a k e m) E0 S.
+ exists t, exists S, estep (ExprState f a k e m) t S.
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; eapply step_expr; eauto.
+ econstructor; econstructor; eapply step_expr; eauto.
(* side effect *)
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.
(* 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. 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.
intros [b [ofs [E1 S1]]].
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' [A [B D]]]].
- econstructor; eapply step_assign; eauto.
+ exploit safe_inv. eexact S2. eauto. simpl. intros [v' [m' [t [A [B D]]]]].
+ 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.
intros [b [ofs [E1 S1]]].
exploit (simple_can_eval_rval f k e m b2 (fun x => C(Eassignop op (Eloc b ofs (typeof b1)) x tyres ty))); eauto.
intros [v [E2 S2]].
- exploit safe_inv. eexact S2. eauto. simpl. intros [v1 [v2 [v3 [m' [A [B [D [E F]]]]]]]].
- econstructor; eapply step_assignop; eauto.
+ exploit safe_inv. eexact S2. eauto. simpl. intros [t1 [v1 [A B]]].
+ destruct (sem_binary_operation op v1 (typeof b1) v (typeof b2) m) as [v3|]_eqn.
+ destruct (sem_cast v3 tyres (typeof b1)) as [v4|]_eqn.
+ 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.
+ rewrite Heqo. rewrite Heqo0. intros; red; intros. elim H2. exists t2; exists m'; auto.
+ econstructor; econstructor; eapply step_assignop_stuck; eauto.
+ rewrite Heqo. rewrite Heqo0. auto.
+ 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.
intros [b1 [ofs [E1 S1]]].
- exploit safe_inv. eexact S1. eauto. simpl. intros [v1 [v2 [v3 [m' [A [B [D [E F]]]]]]]].
- econstructor; eapply step_postincr; eauto.
+ exploit safe_inv. eexact S1. eauto. simpl. intros [t [v1 [A B]]].
+ destruct (sem_incrdecr id v1 ty) as [v2|]_eqn.
+ destruct (sem_cast v2 (typeconv ty) ty) as [v3|]_eqn.
+ 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.
+ rewrite Heqo. rewrite Heqo0. intros; red; intros. elim H0. exists t2; exists m'; congruence.
+ econstructor; econstructor; eapply step_postincr_stuck; eauto.
+ rewrite Heqo. rewrite Heqo0. auto.
+ 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.
intros [v1 [E1 S1]].
exploit safe_inv. eexact S1. eauto. simpl. intros EQ.
- econstructor; eapply step_comma; eauto.
+ econstructor; econstructor; eapply step_comma; eauto.
(* call *)
destruct Q.
exploit (simple_can_eval_rval f k e m b (fun x => C(Ecall x rargs ty))); eauto.
@@ -1122,19 +1241,19 @@ Proof.
apply (eval_simple_list_steps f k e m rargs vl E2 C'); auto.
simpl. intros X. exploit X. eapply rval_list_all_values.
intros [tyargs [tyres [fd [vargs [P [Q [U V]]]]]]].
- econstructor; eapply step_call; eauto. eapply can_eval_simple_list; eauto.
+ econstructor; econstructor; eapply step_call; eauto. eapply can_eval_simple_list; eauto.
(* paren *)
exploit (simple_can_eval_rval f k e m b (fun x => C(Eparen x ty))); eauto.
intros [v1 [E1 S1]].
exploit safe_inv. eexact S1. eauto. simpl. intros [v CAST].
- econstructor; eapply step_paren; eauto.
+ econstructor; econstructor; eapply step_paren; eauto.
Qed.
(** Simulation for all states *)
Theorem step_simulation:
forall S1 t S2,
- step S1 t S2 -> safe S1 -> plus Csem.step ge S1 t S2.
+ step S1 t S2 -> plus Csem.step ge S1 t S2.
Proof.
intros. inv H.
apply estep_simulation; auto.
@@ -1150,16 +1269,19 @@ Proof.
auto.
right. destruct STEP.
(* 2. Expression step. *)
- assert (exists S', estep S E0 S').
+ assert (exists t, exists S', estep S t S').
inv H0.
(* lred *)
- eapply can_estep; eauto. inv H3; auto.
+ eapply can_estep; eauto. inv H2; auto.
(* rred *)
- eapply can_estep; eauto. inv H3; auto. inv H1; auto.
+ eapply can_estep; eauto. inv H2; auto. inv H1; auto.
(* callred *)
- eapply can_estep; eauto. inv H3; auto. inv H1; auto.
- destruct H1 as [S'' ESTEP].
- exists E0; exists S''; left; auto.
+ eapply can_estep; eauto. inv H2; auto. inv H1; auto.
+ (* stuck *)
+ exploit (H Stuckstate). apply star_one. left. econstructor; eauto.
+ intros [[r F] | [t [S' R]]]. inv F. inv R. inv H0. inv H0.
+ destruct H1 as [t' [S'' ESTEP]].
+ exists t'; exists S''; left; auto.
(* 3. Other step. *)
exists t; exists S'; right; auto.
Qed.
@@ -1173,22 +1295,145 @@ Definition semantics (p: program) :=
(** This semantics is receptive to changes in events. *)
-Lemma semantics_receptive:
- forall p, receptive (semantics p).
+Remark deref_loc_trace:
+ forall F V (ge: Genv.t F V) ty m b ofs t v,
+ 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.
+Qed.
+
+Remark deref_loc_receptive:
+ forall F V (ge: Genv.t F V) ty m b ofs ev1 t1 v ev2,
+ deref_loc ge ty m b ofs (ev1 :: t1) v ->
+ match_traces ge (ev1 :: nil) (ev2 :: nil) ->
+ t1 = nil /\ exists v', deref_loc ge ty m b ofs (ev2 :: nil) v'.
+Proof.
+ intros.
+ assert (t1 = nil). exploit deref_loc_trace; eauto. destruct t1; simpl; tauto.
+ inv H. exploit volatile_load_receptive; eauto. intros [v' A].
+ split; auto; exists v'; econstructor; eauto.
+Qed.
+
+Remark assign_loc_trace:
+ forall F V (ge: Genv.t F V) ty m b ofs t v m',
+ 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.
+Qed.
+
+Remark assign_loc_receptive:
+ forall F V (ge: Genv.t F V) ty m b ofs ev1 t1 v m' ev2,
+ assign_loc ge ty m b ofs v (ev1 :: t1) m' ->
+ match_traces ge (ev1 :: nil) (ev2 :: nil) ->
+ ev1 :: t1 = ev2 :: nil.
+Proof.
+ intros.
+ assert (t1 = nil). exploit assign_loc_trace; eauto. destruct t1; simpl; tauto.
+ inv H. eapply volatile_store_receptive; eauto.
+Qed.
+
+Lemma semantics_strongly_receptive:
+ forall p, strongly_receptive (semantics p).
Proof.
intros. constructor; simpl; intros.
(* receptiveness *)
- assert (t1 = E0 -> exists s2, step (Genv.globalenv p) s t2 s2).
- intros. subst. inv H0. exists s1; auto.
inversion H; subst.
- inv H2; auto.
- inv H2; auto.
+ inv H1.
+ (* valof volatile *)
+ 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.
+ 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 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 (Genv.globalenv p) (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.
+ 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.
+ rewrite Heqo; rewrite Heqo0; auto.
+ 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 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 (Genv.globalenv p) (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.
+ 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.
+ rewrite Heqo; rewrite Heqo0; auto.
+ econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
+ rewrite Heqo; auto.
+ (* postincr *)
+ 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 id v1' (typeof l)) as [v2'|]_eqn.
+ destruct (sem_cast v2' (typeconv (typeof l)) (typeof l)) as [v3'|]_eqn.
+ destruct (classic (exists t2', exists m'', assign_loc (Genv.globalenv p) (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.
+ 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.
+ rewrite Heqo; rewrite Heqo0; auto.
+ 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 id v1' (typeof l)) as [v2'|]_eqn.
+ destruct (sem_cast v2' (typeconv (typeof l)) (typeof l)) as [v3'|]_eqn.
+ destruct (classic (exists t2', exists m'', assign_loc (Genv.globalenv p) (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.
+ 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.
+ rewrite Heqo; rewrite Heqo0; auto.
+ econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
+ rewrite Heqo; auto.
+ (* external calls *)
+ inv H1.
+ exploit external_call_trace_length; eauto. destruct t1; simpl; intros.
exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
- exists (Returnstate vres2 k m2). right; econstructor; eauto.
-(* trace length *)
- inv H.
- inv H0; simpl; omega.
- inv H0; simpl; try omega. eapply external_call_trace_length; eauto.
+ exists (Returnstate vres2 k m2); exists E0; right; econstructor; eauto.
+ omegaContradiction.
+(* well-behaved traces *)
+ red; intros. inv H; inv H0; simpl; auto.
+ (* valof volatile *)
+ exploit deref_loc_trace; eauto. destruct t; auto. destruct t; tauto.
+ (* assign *)
+ 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.
+ 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.
+ tauto.
+ (* postincr stuck *)
+ exploit deref_loc_trace; eauto. destruct t; auto. destruct t; tauto.
+ (* external calls *)
+ exploit external_call_trace_length; eauto.
+ destruct t; simpl; auto. destruct t; simpl; auto. intros; omegaContradiction.
Qed.
(** The main simulation result. *)
@@ -1200,8 +1445,6 @@ Proof.
apply backward_simulation_plus with (match_states := fun (S1 S2: state) => S1 = S2); simpl.
(* symbols *)
auto.
-(* trace length *)
- intros. eapply sr_traces. apply semantics_receptive. simpl. eauto.
(* initial states exist *)
intros. exists s1; auto.
(* initial states match *)
@@ -1269,11 +1512,19 @@ with eval_expr: env -> mem -> kind -> expr -> trace -> mem -> expr -> Prop :=
| eval_var: forall e m x ty,
eval_expr e m LV (Evar x ty) E0 m (Evar x ty)
| eval_field: forall e m a t m' a' f ty,
- eval_expr e m LV a t m' a' ->
+ eval_expr e m RV a t m' a' ->
eval_expr e m LV (Efield a f ty) t m' (Efield a' f ty)
| eval_valof: forall e m a t m' a' ty,
+ type_is_volatile (typeof a) = false ->
eval_expr e m LV a t m' a' ->
eval_expr e m RV (Evalof a ty) t m' (Evalof a' ty)
+ | eval_valof_volatile: forall e m a t1 m' a' ty b ofs t2 v,
+ type_is_volatile (typeof a) = true ->
+ eval_expr e m LV a t1 m' a' ->
+ eval_simple_lvalue ge e m' a' b ofs ->
+ deref_loc ge (typeof a) m' b ofs t2 v ->
+ ty = typeof a ->
+ eval_expr e m RV (Evalof a ty) (t1 ** t2) m' (Eval v ty)
| eval_deref: forall e m a t m' a' ty,
eval_expr e m RV a t m' a' ->
eval_expr e m LV (Ederef a ty) t m' (Ederef a' ty)
@@ -1297,34 +1548,34 @@ with eval_expr: env -> mem -> kind -> expr -> trace -> mem -> expr -> Prop :=
eval_expr e m RV (Econdition a1 a2 a3 ty) (t1**t2) m'' (Eval v ty)
| eval_sizeof: forall e m ty' ty,
eval_expr e m RV (Esizeof ty' ty) E0 m (Esizeof ty' ty)
- | eval_assign: forall e m l r ty t1 m1 l' t2 m2 r' b ofs v v' m3,
+ | eval_assign: forall e m l r ty t1 m1 l' t2 m2 r' b ofs v v' t3 m3,
eval_expr e m LV l t1 m1 l' -> eval_expr e m1 RV r t2 m2 r' ->
eval_simple_lvalue ge e m2 l' b ofs ->
eval_simple_rvalue ge e m2 r' v ->
sem_cast v (typeof r) (typeof l) = Some v' ->
- store_value_of_type (typeof l) m2 b ofs v' = Some m3 ->
+ assign_loc ge (typeof l) m2 b ofs v' t3 m3 ->
ty = typeof l ->
- eval_expr e m RV (Eassign l r ty) (t1**t2) m3 (Eval v' ty)
+ eval_expr e m RV (Eassign l r ty) (t1**t2**t3) m3 (Eval v' ty)
| eval_assignop: forall e m op l r tyres ty t1 m1 l' t2 m2 r' b ofs
- v1 v2 v3 v4 m3,
+ v1 v2 v3 v4 t3 t4 m3,
eval_expr e m LV l t1 m1 l' -> eval_expr e m1 RV r t2 m2 r' ->
eval_simple_lvalue ge e m2 l' b ofs ->
- load_value_of_type (typeof l) m2 b ofs = Some v1 ->
+ deref_loc ge (typeof l) m2 b ofs t3 v1 ->
eval_simple_rvalue ge e m2 r' v2 ->
sem_binary_operation op v1 (typeof l) v2 (typeof r) m2 = Some v3 ->
sem_cast v3 tyres (typeof l) = Some v4 ->
- store_value_of_type (typeof l) m2 b ofs v4 = Some m3 ->
+ assign_loc ge (typeof l) m2 b ofs v4 t4 m3 ->
ty = typeof l ->
- eval_expr e m RV (Eassignop op l r tyres ty) (t1**t2) m3 (Eval v4 ty)
- | eval_postincr: forall e m id l ty t m1 l' b ofs v1 v2 v3 m2,
- eval_expr e m LV l t m1 l' ->
+ eval_expr e m RV (Eassignop op l r tyres ty) (t1**t2**t3**t4) m3 (Eval v4 ty)
+ | eval_postincr: forall e m id l ty t1 m1 l' b ofs v1 v2 v3 m2 t2 t3,
+ eval_expr e m LV l t1 m1 l' ->
eval_simple_lvalue ge e m1 l' b ofs ->
- load_value_of_type ty m1 b ofs = Some v1 ->
+ deref_loc ge ty m1 b ofs t2 v1 ->
sem_incrdecr id v1 ty = Some v2 ->
sem_cast v2 (typeconv ty) ty = Some v3 ->
- store_value_of_type ty m1 b ofs v3 = Some m2 ->
+ assign_loc ge ty m1 b ofs v3 t3 m2 ->
ty = typeof l ->
- eval_expr e m RV (Epostincr id l ty) t m2 (Eval v1 ty)
+ eval_expr e m RV (Epostincr id l ty) (t1**t2**t3) m2 (Eval v1 ty)
| eval_comma: forall e m r1 r2 ty t1 m1 r1' v1 t2 m2 r2',
eval_expr e m RV r1 t1 m1 r1' ->
eval_simple_rvalue ge e m1 r1' v1 ->
@@ -1472,7 +1723,7 @@ with eval_funcall: mem -> fundef -> list val -> trace -> mem -> val -> Prop :=
| eval_funcall_internal: forall m f vargs t e m1 m2 m3 out vres m4,
list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) ->
alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 ->
- bind_parameters e m1 f.(fn_params) vargs m2 ->
+ bind_parameters ge e m1 f.(fn_params) vargs m2 ->
exec_stmt e m2 f.(fn_body) t m3 out ->
outcome_result_value out f.(fn_return) vres ->
Mem.free_list m3 (blocks_of_env e) = Some m4 ->
@@ -1497,7 +1748,7 @@ Combined Scheme bigstep_induction from
CoInductive evalinf_expr: env -> mem -> kind -> expr -> traceinf -> Prop :=
| evalinf_field: forall e m a t f ty,
- evalinf_expr e m LV a t ->
+ evalinf_expr e m RV a t ->
evalinf_expr e m LV (Efield a f ty) t
| evalinf_valof: forall e m a t ty,
evalinf_expr e m LV a t ->
@@ -1677,7 +1928,7 @@ with evalinf_funcall: mem -> fundef -> list val -> traceinf -> Prop :=
| evalinf_funcall_internal: forall m f vargs t e m1 m2,
list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) ->
alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 ->
- bind_parameters e m1 f.(fn_params) vargs m2 ->
+ bind_parameters ge e m1 f.(fn_params) vargs m2 ->
execinf_stmt e m2 f.(fn_body) t ->
evalinf_funcall m (Internal f) vargs t.
@@ -1780,9 +2031,16 @@ Proof.
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
simpl; intuition; eauto.
(* valof *)
- exploit (H0 (fun x => C(Evalof x ty))).
+ exploit (H1 (fun x => C(Evalof x ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
- simpl; intuition; eauto.
+ simpl; intuition; eauto. congruence.
+(* valof volatile *)
+ 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.
+ traceEq.
(* deref *)
exploit (H0 (fun x => C(Ederef x ty))).
eapply leftcontext_compose; eauto. repeat constructor. intros [A [B D]].
@@ -1825,7 +2083,7 @@ Proof.
simpl; intuition.
eapply star_trans. eexact D.
eapply star_right. eexact G.
- left. eapply step_assign; eauto. congruence. congruence. congruence.
+ left. eapply step_assign; eauto. congruence. rewrite B; eauto. congruence.
reflexivity. traceEq.
(* assignop *)
exploit (H0 (fun x => C(Eassignop op x r tyres ty))).
@@ -1836,7 +2094,7 @@ Proof.
eapply star_trans. eexact D.
eapply star_right. eexact G.
left. eapply step_assignop; eauto.
- rewrite B; eauto. rewrite B; rewrite F; eauto. congruence. congruence. congruence.
+ rewrite B; eauto. rewrite B; rewrite F; eauto. congruence. rewrite B; eauto. congruence.
reflexivity. traceEq.
(* postincr *)
exploit (H0 (fun x => C(Epostincr id x ty))).
@@ -2113,16 +2371,15 @@ Proof.
(* Out_normal *)
assert (fn_return f = Tvoid /\ vres = Vundef).
destruct (fn_return f); auto || contradiction.
- destruct H7. subst vres. right; apply step_skip_call; auto.
+ destruct H7 as [P Q]. subst vres. right; eapply step_skip_call; eauto.
(* Out_return None *)
assert (fn_return f = Tvoid /\ vres = Vundef).
destruct (fn_return f); auto || contradiction.
- destruct H8. subst vres.
+ destruct H8 as [P Q]. subst vres.
rewrite <- (is_call_cont_call_cont k H6). rewrite <- H7.
right; apply step_return_0; auto.
(* Out_return Some *)
- destruct H4.
- rewrite <- (is_call_cont_call_cont k H6). rewrite <- H7.
+ destruct H4. rewrite <- (is_call_cont_call_cont k H6). rewrite <- H7.
right; eapply step_return_2; eauto.
reflexivity. traceEq.
@@ -2510,7 +2767,7 @@ Inductive bigstep_program_terminates (p: program): trace -> int -> Prop :=
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
- type_of_fundef f = Tfunction Tnil (Tint I32 Signed) ->
+ type_of_fundef f = Tfunction Tnil type_int32s ->
eval_funcall ge m0 f nil t m1 (Vint r) ->
bigstep_program_terminates p t r.
@@ -2520,7 +2777,7 @@ Inductive bigstep_program_diverges (p: program): traceinf -> Prop :=
Genv.init_mem p = Some m0 ->
Genv.find_symbol ge p.(prog_main) = Some b ->
Genv.find_funct_ptr ge b = Some f ->
- type_of_fundef f = Tfunction Tnil (Tint I32 Signed) ->
+ type_of_fundef f = Tfunction Tnil type_int32s ->
evalinf_funcall ge m0 f nil t ->
bigstep_program_diverges p t.
diff --git a/cfrontend/Csyntax.v b/cfrontend/Csyntax.v
index c76d9b9..ffe08bf 100644
--- a/cfrontend/Csyntax.v
+++ b/cfrontend/Csyntax.v
@@ -30,7 +30,8 @@ Require Import AST.
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. *)
+ flag and a bit size: 8, 16, or 32 bits, or the special [IBool] size
+ standing for the C99 [_Bool] type. *)
Inductive signedness : Type :=
| Signed: signedness
@@ -39,7 +40,8 @@ Inductive signedness : Type :=
Inductive intsize : Type :=
| I8: intsize
| I16: intsize
- | I32: intsize.
+ | I32: intsize
+ | IBool: intsize.
(** Float types come in two sizes: 32 bits (single precision)
and 64-bit (double precision). *)
@@ -48,6 +50,15 @@ Inductive floatsize : Type :=
| F32: floatsize
| F64: floatsize.
+(** Every type carries a set of attributes. Currently, only one
+ attribute is modeled: [volatile]. *)
+
+Record attr : Type := mk_attr {
+ attr_volatile: bool
+}.
+
+Definition noattr := {| attr_volatile := false |}.
+
(** The syntax of type expressions. Some points to note:
- Array types [Tarray n] carry the size [n] of the array.
Arrays with unknown sizes are represented by pointer types.
@@ -84,15 +95,15 @@ Inductive floatsize : Type :=
*)
Inductive type : Type :=
- | Tvoid: type (**r the [void] type *)
- | Tint: intsize -> signedness -> type (**r integer types *)
- | Tfloat: floatsize -> type (**r floating-point types *)
- | Tpointer: type -> type (**r pointer types ([*ty]) *)
- | Tarray: type -> Z -> type (**r array types ([ty[len]]) *)
- | Tfunction: typelist -> type -> type (**r function types *)
- | Tstruct: ident -> fieldlist -> type (**r struct types *)
- | Tunion: ident -> fieldlist -> type (**r union types *)
- | Tcomp_ptr: ident -> type (**r pointer to named struct or union *)
+ | Tvoid: type (**r the [void] type *)
+ | Tint: intsize -> signedness -> attr -> type (**r integer types *)
+ | Tfloat: floatsize -> attr -> type (**r floating-point types *)
+ | Tpointer: type -> attr -> type (**r pointer types ([*ty]) *)
+ | Tarray: type -> Z -> attr -> type (**r array types ([ty[len]]) *)
+ | Tfunction: typelist -> type -> type (**r function types *)
+ | Tstruct: ident -> fieldlist -> attr -> type (**r struct types *)
+ | Tunion: ident -> fieldlist -> attr -> type (**r union types *)
+ | Tcomp_ptr: ident -> attr -> type (**r pointer to named struct or union *)
with typelist : Type :=
| Tnil: typelist
@@ -102,16 +113,51 @@ with fieldlist : Type :=
| Fnil: fieldlist
| Fcons: ident -> type -> fieldlist -> fieldlist.
+Lemma type_eq: forall (ty1 ty2: type), {ty1=ty2} + {ty1<>ty2}
+with typelist_eq: forall (tyl1 tyl2: typelist), {tyl1=tyl2} + {tyl1<>tyl2}
+with fieldlist_eq: forall (fld1 fld2: fieldlist), {fld1=fld2} + {fld1<>fld2}.
+Proof.
+ assert (forall (x y: intsize), {x=y} + {x<>y}). decide equality.
+ assert (forall (x y: signedness), {x=y} + {x<>y}). decide equality.
+ assert (forall (x y: floatsize), {x=y} + {x<>y}). decide equality.
+ assert (forall (x y: attr), {x=y} + {x<>y}). decide equality. apply bool_dec.
+ generalize ident_eq zeq. intros E1 E2.
+ decide equality.
+ decide equality.
+ generalize ident_eq. intros E1.
+ decide equality.
+Defined.
+
+Opaque type_eq typelist_eq fieldlist_eq.
+
+(** Extract the attributes of a type. *)
+
+Definition attr_of_type (ty: type) :=
+ match ty with
+ | Tvoid => noattr
+ | Tint sz si a => a
+ | Tfloat sz a => a
+ | Tpointer elt a => a
+ | Tarray elt sz a => a
+ | Tfunction args res => noattr
+ | Tstruct id fld a => a
+ | Tunion id fld a => a
+ | Tcomp_ptr id a => a
+ end.
+
+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. *)
Definition typeconv (ty: type) : type :=
match ty with
- | Tint I32 Unsigned => ty
- | Tint _ _ => Tint I32 Signed
- | Tarray t sz => Tpointer t
- | Tfunction _ _ => Tpointer ty
- | _ => ty
+ | Tint I32 Unsigned _ => ty
+ | Tint _ _ a => Tint I32 Signed a
+ | Tarray t sz a => Tpointer t a
+ | Tfunction _ _ => Tpointer ty noattr
+ | _ => ty
end.
(** ** Expressions *)
@@ -207,33 +253,31 @@ as [*(r1 + r2)].
*)
Definition Eindex (r1 r2: expr) (ty: type) :=
- Ederef (Ebinop Oadd r1 r2 (Tpointer ty)) ty.
+ Ederef (Ebinop Oadd r1 r2 (Tpointer ty noattr)) ty.
(** Pre-increment [++l] and pre-decrement [--l] are expressed as
[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)
- l (Eval (Vint Int.one) (Tint I32 Signed)) (typeconv ty) ty.
+ l (Eval (Vint Int.one) type_int32s) (typeconv ty) ty.
-(** Sequential ``and'' [r1 && r2] is viewed as two conditionals
- [r1 ? (r2 ? 1 : 0) : 0]. *)
+(** Sequential ``and'' [r1 && r2] is viewed as a conditional and a cast:
+ [r1 ? (_Bool) r2 : 0]. *)
Definition Eseqand (r1 r2: expr) (ty: type) :=
Econdition r1
- (Econdition r2 (Eval (Vint Int.one) (Tint I32 Signed))
- (Eval (Vint Int.zero) (Tint I32 Signed)) ty)
- (Eval (Vint Int.zero) (Tint I32 Signed))
+ (Ecast r2 type_bool)
+ (Eval (Vint Int.zero) type_int32s)
ty.
-(** Sequential ``or'' [r1 || r2] is viewed as two conditionals
- [r1 ? 1 : (r2 ? 1 : 0)]. *)
+(** Sequential ``or'' [r1 || r2] is viewed as a conditional and a cast:
+ [r1 ? 1 : (_Bool) r2]. *)
Definition Eseqor (r1 r2: expr) (ty: type) :=
Econdition r1
- (Eval (Vint Int.one) (Tint I32 Signed))
- (Econdition r2 (Eval (Vint Int.one) (Tint I32 Signed))
- (Eval (Vint Int.zero) (Tint I32 Signed)) ty)
+ (Eval (Vint Int.one) type_int32s)
+ (Ecast r2 type_bool)
ty.
(** Extract the type part of a type-annotated expression. *)
@@ -353,17 +397,18 @@ Definition type_of_fundef (f: fundef) : type :=
Fixpoint alignof (t: type) : Z :=
match t with
| Tvoid => 1
- | Tint I8 _ => 1
- | Tint I16 _ => 2
- | Tint I32 _ => 4
- | Tfloat F32 => 4
- | Tfloat F64 => 8
- | Tpointer _ => 4
- | Tarray t' n => alignof t'
+ | Tint I8 _ _ => 1
+ | Tint I16 _ _ => 2
+ | Tint I32 _ _ => 4
+ | Tint IBool _ _ => 1
+ | Tfloat F32 _ => 4
+ | Tfloat F64 _ => 8
+ | Tpointer _ _ => 4
+ | Tarray t' _ _ => alignof t'
| Tfunction _ _ => 1
- | Tstruct _ fld => alignof_fields fld
- | Tunion _ fld => alignof_fields fld
- | Tcomp_ptr _ => 4
+ | Tstruct _ fld _ => alignof_fields fld
+ | Tunion _ fld _ => alignof_fields fld
+ | Tcomp_ptr _ _ => 4
end
with alignof_fields (f: fieldlist) : Z :=
@@ -375,90 +420,47 @@ with alignof_fields (f: fieldlist) : Z :=
Scheme type_ind2 := Induction for type Sort Prop
with fieldlist_ind2 := Induction for fieldlist Sort Prop.
-Lemma alignof_power_of_2:
- forall t, exists n, alignof t = two_power_nat n
-with alignof_fields_power_of_2:
- forall f, exists n, alignof_fields f = two_power_nat n.
+Lemma alignof_1248:
+ forall t, alignof t = 1 \/ alignof t = 2 \/ alignof t = 4 \/ alignof t = 8
+with alignof_fields_1248:
+ forall f, alignof_fields f = 1 \/ alignof_fields f = 2 \/ alignof_fields f = 4 \/ alignof_fields f = 8.
Proof.
- induction t; simpl.
- exists 0%nat; auto.
- destruct i. exists 0%nat; auto. exists 1%nat; auto. exists 2%nat; auto.
- destruct f. exists 2%nat; auto. exists 3%nat; auto.
- exists 2%nat; auto.
- auto.
- exists 0%nat; auto.
- apply alignof_fields_power_of_2.
- apply alignof_fields_power_of_2.
- exists 2%nat; auto.
- induction f; simpl.
- exists 0%nat; auto.
+ induction t; simpl; auto.
+ destruct i; auto.
+ destruct f; auto.
+ induction f; simpl; auto.
rewrite Zmax_spec. destruct (zlt (alignof_fields f) (alignof t)); auto.
Qed.
Lemma alignof_pos:
forall t, alignof t > 0.
Proof.
- intros. destruct (alignof_power_of_2 t) as [p EQ]. rewrite EQ. apply two_power_nat_pos.
+ intros. generalize (alignof_1248 t). omega.
Qed.
Lemma alignof_fields_pos:
forall f, alignof_fields f > 0.
Proof.
- intros. destruct (alignof_fields_power_of_2 f) as [p EQ]. rewrite EQ. apply two_power_nat_pos.
-Qed.
-
-(*
-Fixpoint In_fieldlist (id: ident) (ty: type) (f: fieldlist) : Prop :=
- match f with
- | Fnil => False
- | Fcons id1 ty1 f1 => (id1 = id /\ ty1 = ty) \/ In_fieldlist id ty f1
- end.
-
-Remark divides_max_pow_two:
- forall a b,
- (two_power_nat b | Zmax (two_power_nat a) (two_power_nat b)).
-Proof.
- intros.
- rewrite Zmax_spec. destruct (zlt (two_power_nat b) (two_power_nat a)).
- repeat rewrite two_power_nat_two_p in *.
- destruct (zle (Z_of_nat a) (Z_of_nat b)).
- assert (two_p (Z_of_nat a) <= two_p (Z_of_nat b)). apply two_p_monotone; omega.
- omegaContradiction.
- exists (two_p (Z_of_nat a - Z_of_nat b)).
- rewrite <- two_p_is_exp. decEq. omega. omega. omega.
- apply Zdivide_refl.
+ intros. generalize (alignof_fields_1248 f). omega.
Qed.
-Lemma alignof_each_field:
- forall f id t, In_fieldlist id t f -> (alignof t | alignof_fields f).
-Proof.
- induction f; simpl; intros.
- contradiction.
- destruct (alignof_power_of_2 t) as [k1 EQ1].
- destruct (alignof_fields_power_of_2 f) as [k2 EQ2].
- destruct H as [[A B] | A]; subst; rewrite EQ1; rewrite EQ2.
- rewrite Zmax_comm. apply divides_max_pow_two.
- eapply Zdivide_trans. eapply IHf; eauto.
- rewrite EQ2. apply divides_max_pow_two.
-Qed.
-*)
-
(** Size of a type, in bytes. *)
Fixpoint sizeof (t: type) : Z :=
match t with
| Tvoid => 1
- | Tint I8 _ => 1
- | Tint I16 _ => 2
- | Tint I32 _ => 4
- | Tfloat F32 => 4
- | Tfloat F64 => 8
- | Tpointer _ => 4
- | Tarray t' n => sizeof t' * Zmax 1 n
+ | Tint I8 _ _ => 1
+ | Tint I16 _ _ => 2
+ | Tint I32 _ _ => 4
+ | Tint IBool _ _ => 1
+ | Tfloat F32 _ => 4
+ | Tfloat F64 _ => 8
+ | Tpointer _ _ => 4
+ | Tarray t' n _ => sizeof t' * Zmax 1 n
| Tfunction _ _ => 1
- | Tstruct _ fld => align (Zmax 1 (sizeof_struct fld 0)) (alignof t)
- | Tunion _ fld => align (Zmax 1 (sizeof_union fld)) (alignof t)
- | Tcomp_ptr _ => 4
+ | Tstruct _ fld _ => align (Zmax 1 (sizeof_struct fld 0)) (alignof t)
+ | Tunion _ fld _ => align (Zmax 1 (sizeof_union fld)) (alignof t)
+ | Tcomp_ptr _ _ => 4
end
with sizeof_struct (fld: fieldlist) (pos: Z) {struct fld} : Z :=
@@ -557,9 +559,9 @@ Proof.
Qed.
Lemma field_offset_in_range:
- forall sid fld fid ofs ty,
+ forall sid fld a fid ofs ty,
field_offset fid fld = OK ofs -> field_type fid fld = OK ty ->
- 0 <= ofs /\ ofs + sizeof ty <= sizeof (Tstruct sid fld).
+ 0 <= ofs /\ ofs + sizeof ty <= sizeof (Tstruct sid fld a).
Proof.
intros. exploit field_offset_rec_in_range; eauto. intros [A B].
split. auto. simpl. eapply Zle_trans. eauto.
@@ -638,33 +640,47 @@ type must be accessed:
- [By_value ch]: access by value, i.e. by loading from the address
of the l-value using the memory chunk [ch];
- [By_reference]: access by reference, i.e. by just returning
- the address of the l-value;
+ the address of the l-value (used for arrays and functions);
+- [By_copy]: access is by reference, assignment is by copy
+ (used for [struct] and [union] types)
- [By_nothing]: no access is possible, e.g. for the [void] type.
*)
Inductive mode: Type :=
| By_value: memory_chunk -> mode
| By_reference: mode
+ | By_copy: mode
| By_nothing: mode.
Definition access_mode (ty: type) : mode :=
match ty with
- | Tint I8 Signed => By_value Mint8signed
- | Tint I8 Unsigned => By_value Mint8unsigned
- | Tint I16 Signed => By_value Mint16signed
- | Tint I16 Unsigned => By_value Mint16unsigned
- | Tint I32 _ => By_value Mint32
- | Tfloat F32 => By_value Mfloat32
- | Tfloat F64 => By_value Mfloat64
+ | Tint I8 Signed _ => By_value Mint8signed
+ | Tint I8 Unsigned _ => By_value Mint8unsigned
+ | Tint I16 Signed _ => By_value Mint16signed
+ | Tint I16 Unsigned _ => By_value Mint16unsigned
+ | Tint I32 _ _ => By_value Mint32
+ | Tint IBool _ _ => By_value Mint8unsigned
+ | Tfloat F32 _ => By_value Mfloat32
+ | Tfloat F64 _ => By_value Mfloat64
| Tvoid => By_nothing
- | Tpointer _ => By_value Mint32
- | Tarray _ _ => By_reference
+ | Tpointer _ _ => By_value Mint32
+ | Tarray _ _ _ => By_reference
| Tfunction _ _ => By_reference
- | Tstruct _ fList => By_nothing
- | Tunion _ fList => By_nothing
- | Tcomp_ptr _ => By_nothing
+ | Tstruct _ _ _ => By_copy
+ | Tunion _ _ _ => By_copy
+ | Tcomp_ptr _ _ => By_nothing
end.
+(** For the purposes of the semantics and the compiler, a type denotes
+ a volatile access if it carries the [volatile] attribute and it is
+ accessed by value. *)
+
+Definition type_is_volatile (ty: type) : bool :=
+ match access_mode ty with
+ | By_value _ => attr_volatile (attr_of_type ty)
+ | _ => false
+ end.
+
(** Unroll the type of a structure or union field, substituting
[Tcomp_ptr] by a pointer to the structure. *)
@@ -676,14 +692,14 @@ Variable comp: type.
Fixpoint unroll_composite (ty: type) : type :=
match ty with
| Tvoid => ty
- | Tint _ _ => ty
- | Tfloat _ => ty
- | Tpointer t1 => Tpointer (unroll_composite t1)
- | Tarray t1 sz => Tarray (unroll_composite t1) sz
+ | Tint _ _ _ => ty
+ | Tfloat _ _ => ty
+ | Tpointer t1 a => Tpointer (unroll_composite t1) a
+ | Tarray t1 sz a => Tarray (unroll_composite t1) sz a
| Tfunction t1 t2 => Tfunction (unroll_composite_list t1) (unroll_composite t2)
- | Tstruct id fld => if ident_eq id cid then ty else Tstruct id (unroll_composite_fields fld)
- | Tunion id fld => if ident_eq id cid then ty else Tunion id (unroll_composite_fields fld)
- | Tcomp_ptr id => if ident_eq id cid then Tpointer comp else ty
+ | Tstruct id fld a => if ident_eq id cid then ty else Tstruct id (unroll_composite_fields fld) a
+ | Tunion id fld a => if ident_eq id cid then ty else Tunion id (unroll_composite_fields fld) a
+ | Tcomp_ptr id a => if ident_eq id cid then Tpointer comp a else ty
end
with unroll_composite_list (tl: typelist) : typelist :=
@@ -721,9 +737,9 @@ Opaque alignof.
sizeof_struct (unroll_composite_fields fld) pos = sizeof_struct fld pos));
simpl; intros; auto.
congruence.
- destruct H. rewrite <- (alignof_unroll_composite (Tstruct i f)).
+ destruct H. rewrite <- (alignof_unroll_composite (Tstruct i f a)).
simpl. destruct (ident_eq i cid); simpl. auto. rewrite H0; auto.
- destruct H. rewrite <- (alignof_unroll_composite (Tunion i f)).
+ destruct H. rewrite <- (alignof_unroll_composite (Tunion i f a)).
simpl. destruct (ident_eq i cid); simpl. auto. rewrite H; auto.
destruct (ident_eq i cid); auto.
destruct H0. split. congruence.
@@ -750,9 +766,9 @@ Inductive classify_neg_cases : Type :=
Definition classify_neg (ty: type) : classify_neg_cases :=
match ty with
- | Tint I32 Unsigned => neg_case_i Unsigned
- | Tint _ _ => neg_case_i Signed
- | Tfloat _ => neg_case_f
+ | Tint I32 Unsigned _ => neg_case_i Unsigned
+ | Tint _ _ _ => neg_case_i Signed
+ | Tfloat _ _ => neg_case_f
| _ => neg_default
end.
@@ -762,8 +778,8 @@ Inductive classify_notint_cases : Type :=
Definition classify_notint (ty: type) : classify_notint_cases :=
match ty with
- | Tint I32 Unsigned => notint_case_i Unsigned
- | Tint _ _ => notint_case_i Signed
+ | Tint I32 Unsigned _ => notint_case_i Unsigned
+ | Tint _ _ _ => notint_case_i Signed
| _ => notint_default
end.
@@ -778,9 +794,9 @@ Inductive classify_bool_cases : Type :=
Definition classify_bool (ty: type) : classify_bool_cases :=
match typeconv ty with
- | Tint _ _ => bool_case_ip
- | Tpointer _ => bool_case_ip
- | Tfloat _ => bool_case_f
+ | Tint _ _ _ => bool_case_ip
+ | Tpointer _ _ => bool_case_ip
+ | Tfloat _ _ => bool_case_f
| _ => bool_default
end.
@@ -789,20 +805,20 @@ Inductive classify_add_cases : Type :=
| add_case_ff (**r float, float *)
| add_case_if(s: signedness) (**r int, float *)
| add_case_fi(s: signedness) (**r float, int *)
- | add_case_pi(ty: type) (**r pointer, int *)
- | add_case_ip(ty: type) (**r int, pointer *)
+ | add_case_pi(ty: type)(a: attr) (**r pointer, int *)
+ | add_case_ip(ty: type)(a: attr) (**r int, pointer *)
| add_default.
Definition classify_add (ty1: type) (ty2: type) :=
match typeconv ty1, typeconv ty2 with
- | Tint I32 Unsigned, Tint _ _ => add_case_ii Unsigned
- | Tint _ _, Tint I32 Unsigned => add_case_ii Unsigned
- | Tint _ _, Tint _ _ => add_case_ii Signed
- | Tfloat _, Tfloat _ => add_case_ff
- | Tint _ sg, Tfloat _ => add_case_if sg
- | Tfloat _, Tint _ sg => add_case_fi sg
- | Tpointer ty, Tint _ _ => add_case_pi ty
- | Tint _ _, Tpointer ty => add_case_ip ty
+ | Tint I32 Unsigned _, Tint _ _ _ => add_case_ii Unsigned
+ | Tint _ _ _, Tint I32 Unsigned _ => add_case_ii Unsigned
+ | Tint _ _ _, Tint _ _ _ => add_case_ii Signed
+ | Tfloat _ _, Tfloat _ _ => add_case_ff
+ | Tint _ sg _, Tfloat _ _ => add_case_if sg
+ | Tfloat _ _, Tint _ sg _ => add_case_fi sg
+ | Tpointer ty a, Tint _ _ _ => add_case_pi ty a
+ | Tint _ _ _, Tpointer ty a => add_case_ip ty a
| _, _ => add_default
end.
@@ -817,14 +833,14 @@ Inductive classify_sub_cases : Type :=
Definition classify_sub (ty1: type) (ty2: type) :=
match typeconv ty1, typeconv ty2 with
- | Tint I32 Unsigned, Tint _ _ => sub_case_ii Unsigned
- | Tint _ _, Tint I32 Unsigned => sub_case_ii Unsigned
- | Tint _ _, Tint _ _ => sub_case_ii Signed
- | Tfloat _ , Tfloat _ => sub_case_ff
- | Tint _ sg, Tfloat _ => sub_case_if sg
- | Tfloat _, Tint _ sg => sub_case_fi sg
- | Tpointer ty , Tint _ _ => sub_case_pi ty
- | Tpointer ty , Tpointer _ => sub_case_pp ty
+ | Tint I32 Unsigned _, Tint _ _ _ => sub_case_ii Unsigned
+ | Tint _ _ _, Tint I32 Unsigned _ => sub_case_ii Unsigned
+ | Tint _ _ _, Tint _ _ _ => sub_case_ii Signed
+ | Tfloat _ _ , Tfloat _ _ => sub_case_ff
+ | Tint _ sg _, Tfloat _ _ => sub_case_if sg
+ | Tfloat _ _, Tint _ sg _ => sub_case_fi sg
+ | Tpointer ty _, Tint _ _ _ => sub_case_pi ty
+ | Tpointer ty _ , Tpointer _ _ => sub_case_pp ty
| _ ,_ => sub_default
end.
@@ -837,12 +853,12 @@ Inductive classify_mul_cases : Type:=
Definition classify_mul (ty1: type) (ty2: type) :=
match typeconv ty1, typeconv ty2 with
- | Tint I32 Unsigned, Tint _ _ => mul_case_ii Unsigned
- | Tint _ _, Tint I32 Unsigned => mul_case_ii Unsigned
- | Tint _ _, Tint _ _ => mul_case_ii Signed
- | Tfloat _ , Tfloat _ => mul_case_ff
- | Tint _ sg, Tfloat _ => mul_case_if sg
- | Tfloat _, Tint _ sg => mul_case_fi sg
+ | Tint I32 Unsigned _, Tint _ _ _ => mul_case_ii Unsigned
+ | Tint _ _ _, Tint I32 Unsigned _ => mul_case_ii Unsigned
+ | Tint _ _ _, Tint _ _ _ => mul_case_ii Signed
+ | Tfloat _ _ , Tfloat _ _ => mul_case_ff
+ | Tint _ sg _, Tfloat _ _ => mul_case_if sg
+ | Tfloat _ _, Tint _ sg _ => mul_case_fi sg
| _,_ => mul_default
end.
@@ -855,12 +871,12 @@ Inductive classify_div_cases : Type:=
Definition classify_div (ty1: type) (ty2: type) :=
match typeconv ty1, typeconv ty2 with
- | Tint I32 Unsigned, Tint _ _ => div_case_ii Unsigned
- | Tint _ _, Tint I32 Unsigned => div_case_ii Unsigned
- | Tint _ _, Tint _ _ => div_case_ii Signed
- | Tfloat _ , Tfloat _ => div_case_ff
- | Tint _ sg, Tfloat _ => div_case_if sg
- | Tfloat _, Tint _ sg => div_case_fi sg
+ | Tint I32 Unsigned _, Tint _ _ _ => div_case_ii Unsigned
+ | Tint _ _ _, Tint I32 Unsigned _ => div_case_ii Unsigned
+ | Tint _ _ _, Tint _ _ _ => div_case_ii Signed
+ | Tfloat _ _ , Tfloat _ _ => div_case_ff
+ | Tint _ sg _, Tfloat _ _ => div_case_if sg
+ | Tfloat _ _, Tint _ sg _ => div_case_fi sg
| _,_ => div_default
end.
@@ -873,9 +889,9 @@ Inductive classify_binint_cases : Type:=
Definition classify_binint (ty1: type) (ty2: type) :=
match typeconv ty1, typeconv ty2 with
- | Tint I32 Unsigned, Tint _ _ => binint_case_ii Unsigned
- | Tint _ _, Tint I32 Unsigned => binint_case_ii Unsigned
- | Tint _ _, Tint _ _ => binint_case_ii Signed
+ | Tint I32 Unsigned _, Tint _ _ _ => binint_case_ii Unsigned
+ | Tint _ _ _, Tint I32 Unsigned _ => binint_case_ii Unsigned
+ | Tint _ _ _, Tint _ _ _ => binint_case_ii Signed
| _,_ => binint_default
end.
@@ -887,8 +903,8 @@ Inductive classify_shift_cases : Type:=
Definition classify_shift (ty1: type) (ty2: type) :=
match typeconv ty1, typeconv ty2 with
- | Tint I32 Unsigned, Tint _ _ => shift_case_ii Unsigned
- | Tint _ _, Tint _ _ => shift_case_ii Signed
+ | Tint I32 Unsigned _, Tint _ _ _ => shift_case_ii Unsigned
+ | Tint _ _ _, Tint _ _ _ => shift_case_ii Signed
| _,_ => shift_default
end.
@@ -902,15 +918,15 @@ Inductive classify_cmp_cases : Type:=
Definition classify_cmp (ty1: type) (ty2: type) :=
match typeconv ty1, typeconv ty2 with
- | Tint I32 Unsigned , Tint _ _ => cmp_case_ii Unsigned
- | Tint _ _ , Tint I32 Unsigned => cmp_case_ii Unsigned
- | Tint _ _ , Tint _ _ => cmp_case_ii Signed
- | Tfloat _ , Tfloat _ => cmp_case_ff
- | Tint _ sg, Tfloat _ => cmp_case_if sg
- | Tfloat _, Tint _ sg => cmp_case_fi sg
- | Tpointer _ , Tpointer _ => cmp_case_pp
- | Tpointer _ , Tint _ _ => cmp_case_pp
- | Tint _ _, Tpointer _ => cmp_case_pp
+ | Tint I32 Unsigned _ , Tint _ _ _ => cmp_case_ii Unsigned
+ | Tint _ _ _ , Tint I32 Unsigned _ => cmp_case_ii Unsigned
+ | Tint _ _ _ , Tint _ _ _ => cmp_case_ii Signed
+ | Tfloat _ _ , Tfloat _ _ => cmp_case_ff
+ | Tint _ sg _, Tfloat _ _ => cmp_case_if sg
+ | Tfloat _ _, Tint _ sg _ => cmp_case_fi sg
+ | Tpointer _ _ , Tpointer _ _ => cmp_case_pp
+ | Tpointer _ _ , Tint _ _ _ => cmp_case_pp
+ | Tint _ _ _, Tpointer _ _ => cmp_case_pp
| _ , _ => cmp_default
end.
@@ -921,7 +937,7 @@ Inductive classify_fun_cases : Type:=
Definition classify_fun (ty: type) :=
match ty with
| Tfunction args res => fun_case_f args res
- | Tpointer (Tfunction args res) => fun_case_f args res
+ | Tpointer (Tfunction args res) _ => fun_case_f args res
| _ => fun_default
end.
@@ -931,17 +947,25 @@ Inductive classify_cast_cases : Type :=
| cast_case_f2f (sz2:floatsize) (**r float -> float *)
| cast_case_i2f (si1:signedness) (sz2:floatsize) (**r int -> float *)
| cast_case_f2i (sz2:intsize) (si2:signedness) (**r float -> int *)
+ | cast_case_ip2bool (**r int|pointer -> bool *)
+ | cast_case_f2bool (**r float -> bool *)
+ | cast_case_struct (id1: ident) (fld1: fieldlist) (id2: ident) (fld2: fieldlist) (**r struct -> struct *)
+ | cast_case_union (id1: ident) (fld1: fieldlist) (id2: ident) (fld2: fieldlist) (**r union -> union *)
| cast_case_void (**r any -> void *)
| cast_case_default.
Function classify_cast (tfrom tto: type) : classify_cast_cases :=
match tto, tfrom with
- | Tint I32 si2, (Tint _ _ | Tpointer _ | Tarray _ _ | Tfunction _ _) => cast_case_neutral
- | Tint sz2 si2, Tint sz1 si1 => cast_case_i2i sz2 si2
- | Tint sz2 si2, Tfloat sz1 => cast_case_f2i sz2 si2
- | Tfloat sz2, Tfloat sz1 => cast_case_f2f sz2
- | Tfloat sz2, Tint sz1 si1 => cast_case_i2f si1 sz2
- | Tpointer _, (Tint _ _ | Tpointer _ | Tarray _ _ | Tfunction _ _) => cast_case_neutral
+ | Tint I32 si2 _, (Tint _ _ _ | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _) => cast_case_neutral
+ | Tint IBool _ _, (Tint _ _ _ | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _) => cast_case_ip2bool
+ | Tint IBool _ _, Tfloat _ _ => cast_case_f2bool
+ | Tint sz2 si2 _, Tint sz1 si1 _ => cast_case_i2i sz2 si2
+ | Tint sz2 si2 _, Tfloat sz1 _ => cast_case_f2i sz2 si2
+ | Tfloat sz2 _, Tfloat sz1 _ => cast_case_f2f sz2
+ | Tfloat sz2 _, Tint sz1 si1 _ => cast_case_i2f si1 sz2
+ | Tpointer _ _, (Tint _ _ _ | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _) => cast_case_neutral
+ | Tstruct id2 fld2 _, Tstruct id1 fld1 _ => cast_case_struct id1 fld1 id2 fld2
+ | Tunion id2 fld2 _, Tunion id1 fld1 _ => cast_case_union id1 fld1 id2 fld2
| Tvoid, _ => cast_case_void
| _, _ => cast_case_default
end.
@@ -951,14 +975,14 @@ Function classify_cast (tfrom tto: type) : classify_cast_cases :=
Definition typ_of_type (t: type) : AST.typ :=
match t with
- | Tfloat _ => AST.Tfloat
+ | Tfloat _ _ => AST.Tfloat
| _ => AST.Tint
end.
Definition opttyp_of_type (t: type) : option AST.typ :=
match t with
| Tvoid => None
- | Tfloat _ => Some AST.Tfloat
+ | Tfloat _ _ => Some AST.Tfloat
| _ => Some AST.Tint
end.
diff --git a/cfrontend/Initializers.v b/cfrontend/Initializers.v
index 223d75c..e9c40a2 100644
--- a/cfrontend/Initializers.v
+++ b/cfrontend/Initializers.v
@@ -97,11 +97,11 @@ Fixpoint constval (a: expr) : res val :=
constval r
| Efield l f ty =>
match typeof l with
- | Tstruct id fList =>
+ | Tstruct id fList _ =>
do delta <- field_offset f fList;
do v <- constval l;
OK (Val.add v (Vint (Int.repr delta)))
- | Tunion id fList =>
+ | Tunion id fList _ =>
constval l
| _ =>
Error(msg "ill-typed field access")
@@ -128,14 +128,14 @@ Definition transl_init_single (ty: type) (a: expr) : res init_data :=
do v1 <- constval a;
do v2 <- do_cast v1 (typeof a) ty;
match v2, ty with
- | Vint n, Tint I8 sg => OK(Init_int8 n)
- | Vint n, Tint I16 sg => OK(Init_int16 n)
- | Vint n, Tint I32 sg => OK(Init_int32 n)
- | Vint n, Tpointer _ => OK(Init_int32 n)
- | Vfloat f, Tfloat F32 => OK(Init_float32 f)
- | Vfloat f, Tfloat F64 => OK(Init_float64 f)
- | Vptr (Zpos id) ofs, Tint I32 sg => OK(Init_addrof id ofs)
- | Vptr (Zpos id) ofs, Tpointer _ => OK(Init_addrof id ofs)
+ | Vint n, Tint I8 sg _ => OK(Init_int8 n)
+ | Vint n, Tint I16 sg _ => OK(Init_int16 n)
+ | Vint n, Tint I32 sg _ => OK(Init_int32 n)
+ | Vint n, Tpointer _ _ => OK(Init_int32 n)
+ | Vfloat f, Tfloat F32 _ => OK(Init_float32 f)
+ | Vfloat f, Tfloat F64 _ => OK(Init_float64 f)
+ | Vptr (Zpos id) ofs, Tint I32 sg _ => OK(Init_addrof id ofs)
+ | Vptr (Zpos id) ofs, Tpointer _ _ => OK(Init_addrof id ofs)
| Vundef, _ => Error(msg "undefined operation in initializer")
| _, _ => Error (msg "type mismatch in initializer")
end.
@@ -152,17 +152,17 @@ Fixpoint transl_init (ty: type) (i: initializer)
match i, ty with
| Init_single a, _ =>
do d <- transl_init_single ty a; OK (d :: nil)
- | Init_compound il, Tarray tyelt sz =>
+ | Init_compound il, Tarray tyelt sz _ =>
if zle sz 0
then OK (Init_space(sizeof tyelt) :: nil)
else transl_init_array tyelt il sz
- | Init_compound il, Tstruct _ Fnil =>
+ | Init_compound il, Tstruct _ Fnil _ =>
OK (Init_space (sizeof ty) :: nil)
- | Init_compound il, Tstruct id fl =>
+ | Init_compound il, Tstruct id fl _ =>
transl_init_struct id ty fl il 0
- | Init_compound il, Tunion _ Fnil =>
+ | Init_compound il, Tunion _ Fnil _ =>
OK (Init_space (sizeof ty) :: nil)
- | Init_compound il, Tunion id (Fcons _ ty1 _) =>
+ | Init_compound il, Tunion id (Fcons _ ty1 _) _ =>
transl_init_union id ty ty1 il
| _, _ =>
Error (msg "wrong type for compound initializer")
diff --git a/cfrontend/Initializersproof.v b/cfrontend/Initializersproof.v
index 6563a35..4d287bc 100644
--- a/cfrontend/Initializersproof.v
+++ b/cfrontend/Initializersproof.v
@@ -86,14 +86,14 @@ Inductive eval_simple_lvalue: expr -> block -> int -> Prop :=
| esl_deref: forall r ty b ofs,
eval_simple_rvalue r (Vptr b ofs) ->
eval_simple_lvalue (Ederef r ty) b ofs
- | esl_field_struct: forall l f ty b ofs id fList delta,
- eval_simple_lvalue l b ofs ->
- typeof l = Tstruct id fList -> field_offset f fList = OK delta ->
- eval_simple_lvalue (Efield l f ty) b (Int.add ofs (Int.repr delta))
- | esl_field_union: forall l f ty b ofs id fList,
- eval_simple_lvalue l b ofs ->
- typeof l = Tunion id fList ->
- eval_simple_lvalue (Efield l f ty) b ofs
+ | esl_field_struct: forall r f ty b ofs id fList a delta,
+ eval_simple_rvalue r (Vptr b ofs) ->
+ typeof r = Tstruct id fList a -> field_offset f fList = OK delta ->
+ eval_simple_lvalue (Efield r f ty) b (Int.add ofs (Int.repr delta))
+ | esl_field_union: forall r f ty b ofs id fList a,
+ eval_simple_rvalue r (Vptr b ofs) ->
+ typeof r = Tunion id fList a ->
+ eval_simple_lvalue (Efield r f ty) b ofs
with eval_simple_rvalue: expr -> val -> Prop :=
| esr_val: forall v ty,
@@ -101,7 +101,7 @@ with eval_simple_rvalue: expr -> val -> Prop :=
| esr_rvalof: forall b ofs l ty v,
eval_simple_lvalue l b ofs ->
ty = typeof l ->
- load_value_of_type ty m b ofs = Some v ->
+ deref_loc ge ty m b ofs E0 v ->
eval_simple_rvalue (Evalof l ty) v
| esr_addrof: forall b ofs l ty,
eval_simple_lvalue l b ofs ->
@@ -166,17 +166,17 @@ Proof.
Qed.
Lemma rred_simple:
- forall r m r' m', rred r m r' m' -> simple r -> simple r'.
+ 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.
Qed.
Lemma rred_compat:
- forall e r m r' m', rred r m r' m' ->
+ forall e r m r' m', rred ge r m E0 r' m' ->
simple r ->
m = m' /\ compat_eval RV e r r' m.
Proof.
- induction 1; simpl; intro SIMP; try contradiction; split; auto; split; auto; intros vx EV.
+ 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.
inv EV. econstructor; eauto. constructor.
@@ -227,12 +227,12 @@ 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; tauto.
+ induction 2; simpl; try tauto.
Qed.
Lemma compat_eval_steps:
- forall f r e m w r' m',
- star step ge (ExprState f r Kstop e m) w (ExprState f r' Kstop e m') ->
+ 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 ->
m' = m /\ compat_eval RV e r r' m.
Proof.
@@ -240,10 +240,10 @@ Proof.
(* base case *)
split. auto. red; auto.
(* inductive case *)
- destruct H.
+ destruct (app_eq_nil t1 t2); auto. subst. inv H.
(* expression step *)
assert (X: exists r1, s2 = ExprState f r1 Kstop e m /\ compat_eval RV e r r1 m /\ simple r1).
- inv H.
+ inv H3.
(* lred *)
assert (S: simple a) by (eapply simple_context_1; eauto).
exploit lred_compat; eauto. intros [A B]. subst m'0.
@@ -258,17 +258,19 @@ Proof.
eapply simple_context_2; eauto. eapply rred_simple; eauto.
(* callred *)
assert (S: simple a) by (eapply simple_context_1; eauto).
- inv H10; simpl in S; contradiction.
+ inv H9; simpl in S; contradiction.
+ (* stuckred *)
+ inv H2. destruct H; inv H.
destruct X as [r1 [A [B C]]]. subst s2.
exploit IHstar; eauto. intros [D E].
split. auto. destruct B; destruct E. split. congruence. auto.
(* statement steps *)
- inv H.
+ inv H3.
Qed.
Theorem eval_simple_steps:
- forall f r e m w v ty m',
- star step ge (ExprState f r Kstop e m) w (ExprState f (Eval v ty) Kstop e m') ->
+ forall f r e m v ty m',
+ star step ge (ExprState f r Kstop e m) E0 (ExprState f (Eval v ty) Kstop e m') ->
simple r ->
m' = m /\ ty = typeof r /\ eval_simple_rvalue e m r v.
Proof.
@@ -406,6 +408,12 @@ Proof.
rewrite H2 in H. inv H0. inv H. constructor.
rewrite H2 in H. inv H0. inv H. constructor.
rewrite H2 in H. inv H0. destruct (cast_float_int si2 f); inv H. inv H7. constructor.
+ rewrite H2 in H. inv H0. inv H. constructor.
+ rewrite H2 in H. inv H0. inv H. constructor.
+ rewrite H2 in H. inv H0. inv H. rewrite H7. constructor.
+ rewrite H2 in H. inv H0. inv H. rewrite H7. constructor.
+ rewrite H2 in H. destruct (ident_eq id1 id2 && fieldlist_eq fld1 fld2); inv H. auto.
+ rewrite H2 in H. destruct (ident_eq id1 id2 && fieldlist_eq fld1 fld2); inv H. auto.
rewrite H5 in H. inv H. auto.
Qed.
@@ -437,8 +445,7 @@ Proof.
(* val *)
destruct v; monadInv CV; constructor.
(* rval *)
- unfold load_value_of_type in H1. destruct (access_mode ty); try congruence. inv H1.
- eauto.
+ inv H1; rewrite H2 in CV; try congruence. eauto.
(* addrof *)
eauto.
(* unop *)
@@ -468,10 +475,10 @@ Proof.
(* deref *)
eauto.
(* field struct *)
- rewrite H0 in CV. monadInv CV. exploit IHeval_simple_lvalue; eauto. intro MV. inv MV.
+ rewrite H0 in CV. monadInv CV. exploit constval_rvalue; eauto. intro MV. inv MV.
simpl. replace x with delta by congruence. constructor. auto.
(* field union *)
- rewrite H0 in CV. auto.
+ rewrite H0 in CV. eauto.
Qed.
Lemma constval_simple:
@@ -487,8 +494,8 @@ Qed.
(** Soundness of [constval] with respect to the reduction semantics. *)
Theorem constval_steps:
- forall f r m w v v' ty m',
- star step ge (ExprState f r Kstop empty_env m) w (ExprState f (Eval v' ty) Kstop empty_env m') ->
+ forall f r m v v' ty m',
+ star step ge (ExprState f r Kstop empty_env m) E0 (ExprState f (Eval v' ty) Kstop empty_env m') ->
constval r = OK v ->
m' = m /\ ty = typeof r /\ match_val v v'.
Proof.
@@ -501,9 +508,9 @@ Qed.
(** Soundness for single initializers. *)
Theorem transl_init_single_steps:
- forall ty a data f m w v1 ty1 m' v chunk b ofs m'',
+ forall ty a data f m v1 ty1 m' v chunk b ofs m'',
transl_init_single ty a = OK data ->
- star step ge (ExprState f a Kstop empty_env m) w (ExprState f (Eval v1 ty1) Kstop empty_env m') ->
+ star step ge (ExprState f a Kstop empty_env m) E0 (ExprState f (Eval v1 ty1) Kstop empty_env m') ->
sem_cast v1 ty1 ty = Some v ->
access_mode ty = By_value chunk ->
Mem.store chunk m' b ofs v = Some m'' ->
@@ -583,9 +590,9 @@ Proof.
Qed.
Remark sizeof_struct_eq:
- forall id fl,
+ forall id fl a,
fl <> Fnil ->
- sizeof (Tstruct id fl) = align (sizeof_struct fl 0) (alignof (Tstruct id fl)).
+ sizeof (Tstruct id fl a) = align (sizeof_struct fl 0) (alignof (Tstruct id fl a)).
Proof.
intros. simpl. f_equal. rewrite Zmax_spec. apply zlt_false.
destruct fl. congruence. simpl.
@@ -595,9 +602,9 @@ Proof.
Qed.
Remark sizeof_union_eq:
- forall id fl,
+ forall id fl a,
fl <> Fnil ->
- sizeof (Tunion id fl) = align (sizeof_union fl) (alignof (Tunion id fl)).
+ sizeof (Tunion id fl a) = align (sizeof_union fl) (alignof (Tunion id fl a)).
Proof.
intros. simpl. f_equal. rewrite Zmax_spec. apply zlt_false.
destruct fl. congruence. simpl.
@@ -705,15 +712,15 @@ Inductive exec_init: mem -> block -> Z -> type -> initializer -> mem -> Prop :=
access_mode ty = By_value chunk ->
Mem.store chunk m' b ofs v = Some m'' ->
exec_init m b ofs ty (Init_single a) m''
- | exec_init_compound_array: forall m b ofs ty sz il m',
+ | exec_init_compound_array: forall m b ofs ty sz a il m',
exec_init_array m b ofs ty sz il m' ->
- exec_init m b ofs (Tarray ty sz) (Init_compound il) m'
- | exec_init_compound_struct: forall m b ofs id fl il m',
- exec_init_list m b ofs (fields_of_struct id (Tstruct id fl) fl 0) il m' ->
- exec_init m b ofs (Tstruct id fl) (Init_compound il) m'
- | exec_init_compound_union: forall m b ofs id id1 ty1 fl i m',
- exec_init m b ofs (unroll_composite id (Tunion id (Fcons id1 ty1 fl)) ty1) i m' ->
- exec_init m b ofs (Tunion id (Fcons id1 ty1 fl)) (Init_compound (Init_cons i Init_nil)) m'
+ exec_init m b ofs (Tarray ty sz a) (Init_compound il) m'
+ | exec_init_compound_struct: forall m b ofs id fl a il m',
+ exec_init_list m b ofs (fields_of_struct id (Tstruct id fl a) fl 0) il m' ->
+ exec_init m b ofs (Tstruct id fl a) (Init_compound il) m'
+ | exec_init_compound_union: forall m b ofs id id1 ty1 fl a i m',
+ exec_init m b ofs (unroll_composite id (Tunion id (Fcons id1 ty1 fl) a) ty1) i m' ->
+ exec_init m b ofs (Tunion id (Fcons id1 ty1 fl) a) (Init_compound (Init_cons i Init_nil)) m'
with exec_init_array: mem -> block -> Z -> type -> Z -> initializer_list -> mem -> Prop :=
| exec_init_array_nil: forall m b ofs ty,
diff --git a/cfrontend/PrintClight.ml b/cfrontend/PrintClight.ml
index 306224b..83a0747 100644
--- a/cfrontend/PrintClight.ml
+++ b/cfrontend/PrintClight.ml
@@ -128,6 +128,8 @@ let rec print_stmt p s =
fprintf p "@[<hv 2>%a =@ %a;@]" print_expr e1 print_expr e2
| Sset(id, e2) ->
fprintf p "@[<hv 2>%s =@ %a;@]" (temp_name id) print_expr e2
+ | Svolread(id, e2) ->
+ fprintf p "@[<hv 2>%s =@ %a; /*volatile*/@]" (temp_name id) print_expr e2
| Scall(None, e1, el) ->
fprintf p "@[<hv 2>%a@,(@[<hov 0>%a@]);@]"
print_expr e1
@@ -251,25 +253,6 @@ let print_fundef p (id, fd) =
(* Collect struct and union types *)
-let rec collect_type = function
- | Tvoid -> ()
- | Tint(sz, sg) -> ()
- | Tfloat sz -> ()
- | Tpointer t -> collect_type t
- | Tarray(t, n) -> collect_type t
- | Tfunction(args, res) -> collect_type_list args; collect_type res
- | Tstruct(id, fld) -> register_struct_union id fld; collect_fields fld
- | Tunion(id, fld) -> register_struct_union id fld; collect_fields fld
- | Tcomp_ptr _ -> ()
-
-and collect_type_list = function
- | Tnil -> ()
- | Tcons(hd, tl) -> collect_type hd; collect_type_list tl
-
-and collect_fields = function
- | Fnil -> ()
- | Fcons(id, hd, tl) -> collect_type hd; collect_fields tl
-
let rec collect_expr = function
| Econst_int _ -> ()
| Econst_float _ -> ()
@@ -293,6 +276,7 @@ let rec collect_stmt = function
| Sskip -> ()
| Sassign(e1, e2) -> collect_expr e1; collect_expr e2
| Sset(id, e2) -> collect_expr e2
+ | Svolread(id, e2) -> collect_expr e2
| Scall(optid, e1, el) -> collect_expr e1; collect_exprlist el
| Ssequence(s1, s2) -> collect_stmt s1; collect_stmt s2
| Sifthenelse(e, s1, s2) -> collect_expr e; collect_stmt s1; collect_stmt s2
diff --git a/cfrontend/PrintCsyntax.ml b/cfrontend/PrintCsyntax.ml
index f0e9ee5..d4e728e 100644
--- a/cfrontend/PrintCsyntax.ml
+++ b/cfrontend/PrintCsyntax.ml
@@ -15,6 +15,7 @@
(** Pretty-printer for Csyntax *)
+open Printf
open Format
open Camlcoq
open Datatypes
@@ -53,6 +54,7 @@ let name_inttype sz sg =
| I16, Unsigned -> "unsigned short"
| I32, Signed -> "int"
| I32, Unsigned -> "unsigned int"
+ | IBool, _ -> "_Bool"
let name_floattype sz =
match sz with
@@ -73,31 +75,38 @@ let register_struct_union id fld =
(* Declarator (identifier + type) *)
+let attributes a =
+ if attr_volatile a then " volatile" else ""
+
let name_optid id =
if id = "" then "" else " " ^ id
+(*
let parenthesize_if_pointer id =
if String.length id > 0 && id.[0] = '*' then "(" ^ id ^ ")" else id
+*)
let rec name_cdecl id ty =
match ty with
| Tvoid ->
"void" ^ name_optid id
- | Tint(sz, sg) ->
- name_inttype sz sg ^ name_optid id
- | Tfloat sz ->
- name_floattype sz ^ name_optid id
- | Tpointer t ->
- name_cdecl ("*" ^ id) t
- | Tarray(t, n) ->
- name_cdecl
- (sprintf "%s[%ld]" (parenthesize_if_pointer id) (camlint_of_coqint n))
- t
+ | Tint(sz, sg, a) ->
+ name_inttype sz sg ^ attributes a ^ name_optid id
+ | Tfloat(sz, a) ->
+ name_floattype sz ^ attributes a ^ name_optid id
+ | Tpointer(t, a) ->
+ let id' =
+ match t with
+ | Tfunction _ | Tarray _ -> sprintf "(*%s%s)" (attributes a) id
+ | _ -> sprintf "*%s%s" (attributes a) id in
+ name_cdecl id' t
+ | Tarray(t, n, a) ->
+ name_cdecl (sprintf "%s[%ld]" id (camlint_of_coqint n)) t
| Tfunction(args, res) ->
let b = Buffer.create 20 in
if id = ""
then Buffer.add_string b "(*)"
- else Buffer.add_string b (parenthesize_if_pointer id);
+ else Buffer.add_string b id;
Buffer.add_char b '(';
begin match args with
| Tnil ->
@@ -113,12 +122,12 @@ let rec name_cdecl id ty =
end;
Buffer.add_char b ')';
name_cdecl (Buffer.contents b) res
- | Tstruct(name, fld) ->
- extern_atom name ^ name_optid id
- | Tunion(name, fld) ->
- extern_atom name ^ name_optid id
- | Tcomp_ptr name ->
- extern_atom name ^ " *" ^ id
+ | Tstruct(name, fld, a) ->
+ extern_atom name ^ attributes a ^ name_optid id
+ | Tunion(name, fld, a) ->
+ extern_atom name ^ attributes a ^ name_optid id
+ | Tcomp_ptr(name, a) ->
+ extern_atom name ^ " *" ^ attributes a ^ id
(* Type *)
@@ -403,13 +412,13 @@ let print_globvar p (id, v) =
let rec collect_type = function
| Tvoid -> ()
- | Tint(sz, sg) -> ()
- | Tfloat sz -> ()
- | Tpointer t -> collect_type t
- | Tarray(t, n) -> collect_type t
+ | Tint _ -> ()
+ | Tfloat _ -> ()
+ | Tpointer(t, _) -> collect_type t
+ | Tarray(t, _, _) -> collect_type t
| Tfunction(args, res) -> collect_type_list args; collect_type res
- | Tstruct(id, fld) -> register_struct_union id fld; collect_fields fld
- | Tunion(id, fld) -> register_struct_union id fld; collect_fields fld
+ | Tstruct(id, fld, _) -> register_struct_union id fld; collect_fields fld
+ | Tunion(id, fld, _) -> register_struct_union id fld; collect_fields fld
| Tcomp_ptr _ -> ()
and collect_type_list = function
diff --git a/cfrontend/SimplExpr.v b/cfrontend/SimplExpr.v
index a2e810b..1dae78c 100644
--- a/cfrontend/SimplExpr.v
+++ b/cfrontend/SimplExpr.v
@@ -129,10 +129,35 @@ Function makeif (a: expr) (s1 s2: statement) : statement :=
Definition transl_incrdecr (id: incr_or_decr) (a: expr) (ty: type) : expr :=
match id with
- | Incr => Ebinop Oadd a (Econst_int Int.one (Tint I32 Signed)) (typeconv ty)
- | Decr => Ebinop Osub a (Econst_int Int.one (Tint I32 Signed)) (typeconv ty)
+ | Incr => Ebinop Oadd a (Econst_int Int.one type_int32s) (typeconv ty)
+ | Decr => Ebinop Osub a (Econst_int Int.one type_int32s) (typeconv ty)
end.
+(** Generate a [Sset] or [Svolread] operation as appropriate
+ to dereference a l-value [l] and store its result in temporary variable [id]. *)
+
+Definition make_set (id: ident) (l: expr) : statement :=
+ if type_is_volatile (typeof l)
+ then Svolread id l
+ else Sset id l.
+
+(** Translation of a "valof" operation.
+ If the l-value accessed is of volatile type, we go through a temporary. *)
+
+Definition transl_valof (ty: type) (l: expr) : mon (list statement * expr) :=
+ if type_is_volatile ty
+ then (do t <- gensym ty; ret (Svolread t l :: nil, Etempvar t ty))
+ else ret (nil, l).
+(*
+ match access_mode ty with
+ | By_value _ =>
+ if type_is_volatile ty
+ then (do t <- gensym ty; ret (Sset t l :: nil, Etempvar t ty))
+ else ret (nil, l)
+ | _ => ret (nil, l)
+ end.
+*)
+
(** Translation of expressions. Return a pair [(sl, a)] of
a list of statements [sl] and a pure expression [a].
- If the [dst] argument is [For_val], the statements [sl]
@@ -152,7 +177,7 @@ Inductive destination : Type :=
| For_effects
| For_test (tyl: list type) (s1 s2: statement).
-Definition dummy_expr := Econst_int Int.zero (Tint I32 Signed).
+Definition dummy_expr := Econst_int Int.zero type_int32s.
Definition finish (dst: destination) (sl: list statement) (a: expr) :=
match dst with
@@ -177,8 +202,8 @@ Fixpoint transl_expr (dst: destination) (a: C.expr) : mon (list statement * expr
| C.Ederef r ty =>
do (sl, a) <- transl_expr For_val r;
ret (finish dst sl (Ederef a ty))
- | C.Efield l1 f ty =>
- do (sl, a) <- transl_expr For_val l1;
+ | C.Efield r f ty =>
+ do (sl, a) <- transl_expr For_val r;
ret (finish dst sl (Efield a f ty))
| C.Eval (Vint n) ty =>
ret (finish dst nil (Econst_int n ty))
@@ -189,8 +214,9 @@ Fixpoint transl_expr (dst: destination) (a: C.expr) : mon (list statement * expr
| C.Esizeof ty' ty =>
ret (finish dst nil (Esizeof ty' ty))
| C.Evalof l ty =>
- do (sl, a) <- transl_expr For_val l;
- ret (finish dst sl a)
+ do (sl1, a1) <- transl_expr For_val l;
+ do (sl2, a2) <- transl_valof (C.typeof l) a1;
+ ret (finish dst (sl1 ++ sl2) a2)
| C.Eaddrof l ty =>
do (sl, a) <- transl_expr For_val l;
ret (finish dst sl (Eaddrof a ty))
@@ -234,33 +260,35 @@ Fixpoint transl_expr (dst: destination) (a: C.expr) : mon (list statement * expr
dummy_expr)
end
| C.Eassignop op l1 r2 tyres ty =>
+ let ty1 := C.typeof l1 in
do (sl1, a1) <- transl_expr For_val l1;
do (sl2, a2) <- transl_expr For_val r2;
- let ty1 := C.typeof l1 in
+ do (sl3, a3) <- transl_valof ty1 a1;
match dst with
| For_val | For_test _ _ _ =>
do t <- gensym tyres;
ret (finish dst
- (sl1 ++ sl2 ++
- Sset t (Ebinop op a1 a2 tyres) ::
+ (sl1 ++ sl2 ++ sl3 ++
+ Sset t (Ebinop op a3 a2 tyres) ::
Sassign a1 (Etempvar t tyres) :: nil)
(Ecast (Etempvar t tyres) ty1))
| For_effects =>
- ret (sl1 ++ sl2 ++ Sassign a1 (Ebinop op a1 a2 tyres) :: nil,
+ ret (sl1 ++ sl2 ++ sl3 ++ Sassign a1 (Ebinop op a3 a2 tyres) :: nil,
dummy_expr)
end
| C.Epostincr id l1 ty =>
- do (sl1, a1) <- transl_expr For_val l1;
let ty1 := C.typeof l1 in
+ do (sl1, a1) <- transl_expr For_val l1;
match dst with
| For_val | For_test _ _ _ =>
do t <- gensym ty1;
ret (finish dst
- (sl1 ++ Sset t a1 ::
+ (sl1 ++ make_set t a1 ::
Sassign a1 (transl_incrdecr id (Etempvar t ty1) ty1) :: nil)
(Etempvar t ty1))
| For_effects =>
- ret (sl1 ++ Sassign a1 (transl_incrdecr id a1 ty1) :: nil,
+ do (sl2, a2) <- transl_valof ty1 a1;
+ ret (sl1 ++ sl2 ++ Sassign a1 (transl_incrdecr id a2 ty1) :: nil,
dummy_expr)
end
| C.Ecomma r1 r2 ty =>
@@ -303,7 +331,7 @@ Definition transl_if (r: C.expr) (s1 s2: statement) : mon statement :=
(** Translation of statements *)
-Definition expr_true := Econst_int Int.one (Tint I32 Signed).
+Definition expr_true := Econst_int Int.one type_int32s.
Definition is_Sskip:
forall s, {s = C.Sskip} + {s <> C.Sskip}.
diff --git a/cfrontend/SimplExprproof.v b/cfrontend/SimplExprproof.v
index 2372d02..4df5f03 100644
--- a/cfrontend/SimplExprproof.v
+++ b/cfrontend/SimplExprproof.v
@@ -68,6 +68,12 @@ Lemma varinfo_preserved:
Proof
(Genv.find_var_info_transf_partial transl_fundef _ TRANSL).
+Lemma block_is_volatile_preserved:
+ forall b, block_is_volatile tge b = block_is_volatile ge b.
+Proof.
+ intros. unfold block_is_volatile. rewrite varinfo_preserved. auto.
+Qed.
+
Lemma type_of_fundef_preserved:
forall f tf, transl_fundef f = OK tf ->
type_of_fundef tf = C.type_of_fundef f.
@@ -114,7 +120,7 @@ Proof.
rewrite H0; auto. simpl; auto.
rewrite H0; auto. simpl; auto.
destruct H1; congruence.
- rewrite H0; auto. simpl; auto.
+ destruct H6. inv H1; try congruence. rewrite H0; auto. simpl; auto.
rewrite H0; auto. simpl; auto.
rewrite H0; auto. simpl; auto.
destruct H7. rewrite H0; auto. rewrite H2; auto. simpl; auto.
@@ -134,6 +140,31 @@ Proof (proj2 tr_simple_nil).
(** Evaluation of simple expressions and of their translation *)
+Remark deref_loc_preserved:
+ forall ty m b ofs t v,
+ deref_loc ge ty m b ofs t v -> deref_loc tge ty m b ofs t v.
+Proof.
+ intros. inv H.
+ eapply deref_loc_value; eauto.
+ eapply deref_loc_volatile; eauto.
+ eapply volatile_load_preserved with (ge1 := ge); auto.
+ exact symbols_preserved. exact block_is_volatile_preserved.
+ eapply deref_loc_reference; eauto.
+ eapply deref_loc_copy; eauto.
+Qed.
+
+Remark assign_loc_preserved:
+ forall ty m b ofs v t m',
+ assign_loc ge ty m b ofs v t m' -> assign_loc tge ty m b ofs v t m'.
+Proof.
+ intros. inv H.
+ eapply assign_loc_value; eauto.
+ eapply assign_loc_volatile; eauto.
+ eapply volatile_store_preserved with (ge1 := ge); auto.
+ exact symbols_preserved. exact block_is_volatile_preserved.
+ eapply assign_loc_copy; eauto.
+Qed.
+
Lemma tr_simple:
forall e m,
(forall r v,
@@ -163,9 +194,11 @@ Opaque makeif.
auto.
exists a0; auto.
(* rvalof *)
+ inv H7; try congruence.
exploit H0; eauto. intros [A [B C]].
subst sl1; simpl.
- assert (eval_expr tge e le m a v). eapply eval_Elvalue. eauto. congruence.
+ assert (eval_expr tge e le m a v).
+ eapply eval_Elvalue. eauto. congruence. rewrite <- B. eapply deref_loc_preserved; eauto.
destruct dst; auto.
econstructor. split. simpl; eauto. auto.
(* addrof *)
@@ -346,9 +379,9 @@ Ltac TR :=
Ltac NOTIN :=
match goal with
| [ H1: In ?x ?l, H2: list_disjoint ?l _ |- ~In ?x _ ] =>
- red; intro; elim (H2 x x); auto
+ red; intro; elim (H2 x x); auto; fail
| [ H1: In ?x ?l, H2: list_disjoint _ ?l |- ~In ?x _ ] =>
- red; intro; elim (H2 x x); auto
+ red; intro; elim (H2 x x); auto; fail
end.
Ltac UNCHANGED :=
@@ -377,8 +410,9 @@ Ltac UNCHANGED :=
(* rvalof *)
inv H1.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
- TR. subst sl1; rewrite app_ass; eauto. auto.
- 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.
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
@@ -475,15 +509,16 @@ Ltac UNCHANGED :=
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.
+ 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 [U [R S]]]]]]]]].
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 tr_expr_invariant; eauto. UNCHANGED.
+ eauto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto.
eapply typeof_context; eauto.
(* assignop right *)
inv H2.
@@ -492,25 +527,24 @@ Ltac UNCHANGED :=
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [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.
+ 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. 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 [U [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.
+ 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. 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 [U [R S]]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto. red; auto.
- intros. replace (C.typeof (C e)) with (C.typeof (C e')). rewrite <- app_ass.
- econstructor; eauto.
- eapply typeof_context; eauto.
+ intros. rewrite <- app_ass. econstructor; eauto.
+ symmetry; eapply typeof_context; eauto.
(* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto. red; auto.
@@ -706,7 +740,20 @@ Proof.
apply star_one. eapply step_ifthenelse; eauto.
Qed.
-(** Matching between continuations *)
+Lemma step_make_set:
+ forall id a ty m b ofs t v e le f k,
+ deref_loc ge ty m b ofs t v ->
+ eval_lvalue tge e le m a b ofs ->
+ typeof a = ty ->
+ step tge (State f (make_set id a) k e le m)
+ t (State f Sskip k e (PTree.set id v le) m).
+Proof.
+ intros. exploit deref_loc_preserved; eauto. rewrite <- H1. intros DEREF.
+ unfold make_set. destruct (type_is_volatile (typeof a)) as []_eqn.
+ econstructor; eauto.
+ assert (t = E0). inv H; congruence. subst t.
+ constructor. eapply eval_Elvalue; eauto.
+Qed.
Fixpoint Kseqlist (sl: list statement) (k: cont) :=
match sl with
@@ -721,6 +768,40 @@ Proof.
induction sl1; simpl; congruence.
Qed.
+(*
+Axiom only_scalar_types_volatile:
+ forall ty, type_is_volatile ty = true -> exists chunk, access_mode ty = By_value chunk.
+*)
+
+Lemma step_tr_rvalof:
+ forall ty m b ofs t v e le a sl a' tmp f k,
+ deref_loc ge ty m b ofs t v ->
+ eval_lvalue tge e le m a b ofs ->
+ tr_rvalof ty a sl a' tmp ->
+ typeof a = ty ->
+ exists le',
+ star step tge (State f Sskip (Kseqlist sl k) e le m)
+ t (State f Sskip k e le' m)
+ /\ eval_expr tge e le' m a' v
+ /\ typeof a' = typeof a
+ /\ forall x, ~In x tmp -> le'!x = le!x.
+Proof.
+ intros. inv H1.
+ (* non volatile *)
+ assert (t = E0). inv H; auto. congruence. subst t.
+ exists le; intuition. apply star_refl.
+ eapply eval_Elvalue; eauto. eapply deref_loc_preserved; eauto.
+ (* volatile *)
+ exists (PTree.set t0 v le); intuition.
+ simpl. eapply star_two. econstructor. eapply step_vol_read; eauto.
+ eapply deref_loc_preserved; eauto. traceEq.
+ constructor. apply PTree.gss.
+ apply PTree.gso. congruence.
+Qed.
+
+(** Matching between continuations *)
+
+
Inductive match_cont : Csem.cont -> cont -> Prop :=
| match_Kstop:
match_cont Csem.Kstop Kstop
@@ -853,7 +934,9 @@ Inductive match_states: Csem.state -> state -> Prop :=
| match_returnstates: forall res k m tk,
match_cont k tk ->
match_states (Csem.Returnstate res k m)
- (Returnstate res tk m).
+ (Returnstate res tk m)
+ | match_stuckstate: forall S,
+ match_states Csem.Stuckstate S.
Lemma push_seq:
forall f sl k e le m,
@@ -931,6 +1014,12 @@ Proof.
red; simpl; intros. rewrite H; 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. red; simpl; auto.
+Qed.
+
Definition nolabel_dest (dst: destination) : Prop :=
match dst with
| For_val => True
@@ -970,13 +1059,18 @@ Lemma tr_find_label_expr:
/\(forall le rl sl al tmps, tr_exprlist le rl sl al tmps -> nolabel_list sl).
Proof.
apply tr_expr_exprlist; intros; NoLabelTac.
- destruct H1. apply makeif_nolabel; auto.
+ destruct H1. apply makeif_nolabel; auto.
+ eapply tr_rvalof_nolabel; eauto.
apply makeif_nolabel; NoLabelTac.
rewrite (makeseq_nolabel sl2); auto.
rewrite (makeseq_nolabel sl3); auto.
apply makeif_nolabel; NoLabelTac.
rewrite (makeseq_nolabel sl2). auto. apply H3. apply nolabel_dest_cast; auto.
rewrite (makeseq_nolabel sl3). auto. apply H5. apply nolabel_dest_cast; auto.
+ eapply tr_rvalof_nolabel; eauto.
+ eapply tr_rvalof_nolabel; eauto.
+ eapply tr_rvalof_nolabel; eauto.
+ unfold make_set. destruct (type_is_volatile (typeof a1)); auto.
apply nolabel_dest_cast; auto.
Qed.
@@ -1234,7 +1328,19 @@ Proof.
exploit tr_simple_rvalue; eauto. destruct dest.
intros [A [B C]]. subst sl. apply tr_top_val_val; auto.
intros A. subst sl. apply tr_top_base. constructor.
- intros [b [A [B C]]]. subst sl. apply tr_top_val_test; auto.
+ intros [b [A [B C]]]. subst sl. apply tr_top_val_test; auto.
+(* rval volatile *)
+ exploit tr_top_leftcontext; eauto. clear H11.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
+ 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_vol_read; eauto.
+ rewrite <- TY. eapply deref_loc_preserved; eauto. congruence. auto.
+ econstructor; eauto. change (final dst' (Etempvar t0 (C.typeof l)) ++ sl2) with (nil ++ (final dst' (Etempvar t0 (C.typeof l)) ++ sl2)).
+ 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.
+ auto.
(* condition true *)
exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
@@ -1297,14 +1403,14 @@ Proof.
left. eapply plus_left. constructor.
apply star_one. econstructor; eauto.
rewrite <- TY1; rewrite <- TY2; eauto.
- rewrite <- TY1; eauto.
+ rewrite <- TY1. eapply assign_loc_preserved; 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 t 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.
@@ -1314,7 +1420,7 @@ Proof.
eapply star_left. constructor.
apply star_one. econstructor; eauto. constructor. apply PTree.gss.
simpl. rewrite <- TY1; eauto.
- rewrite <- TY1; eauto.
+ rewrite <- TY1. eapply assign_loc_preserved; eauto.
reflexivity. reflexivity. traceEq.
econstructor. auto. apply S.
apply tr_val_gen. auto. intros. econstructor; eauto. constructor.
@@ -1322,62 +1428,100 @@ Proof.
intros. apply PTree.gso. intuition congruence.
auto. auto.
(* assignop *)
- exploit tr_top_leftcontext; eauto. clear H14.
+ exploit tr_top_leftcontext; eauto. clear H15.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
inv P. inv H6.
(* for effects *)
- exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
- subst; simpl Kseqlist.
+ 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.
+ intros. apply INV. NOTIN. intros [? [? EV1']].
+ 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.
econstructor; split.
- left. eapply plus_left. constructor.
- apply star_one. econstructor; eauto.
- econstructor; eauto. eapply eval_Elvalue; eauto. rewrite <- TY1; eauto.
- rewrite <- TY1; rewrite <- TY2; eauto.
+ left. eapply star_plus_trans. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
+ eapply plus_two. simpl. econstructor. econstructor. eexact EV1'.
+ econstructor. eexact EV3. eexact EV2.
+ rewrite TY3; rewrite <- TY1; rewrite <- TY2; eauto.
rewrite <- TY1; eauto.
- rewrite <- TY1; eauto.
- traceEq.
+ rewrite <- TY1. eapply assign_loc_preserved; eauto.
+ reflexivity. 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. intros [SL1 [TY1 EV1]].
- exploit tr_simple_lvalue. eauto.
- eapply tr_expr_invariant with (le' := PTree.set t v3 le). eauto.
- intros. apply PTree.gso. intuition congruence.
- intros [SL3 [TY3 EV3]].
+ 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.
+ intros. apply INV. NOTIN. intros [? [? EV1']].
+ 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.
+ intros. rewrite PTree.gso. apply INV. NOTIN. intuition congruence.
+ intros [? [? EV1'']].
subst; simpl Kseqlist.
econstructor; split.
- left. eapply plus_left. constructor.
- eapply star_left. constructor.
- econstructor. eapply eval_Elvalue; eauto. rewrite <- TY1; eauto. eauto.
- rewrite <- TY1; rewrite <- TY2. eauto.
- eapply star_left. constructor.
- apply star_one. econstructor. eauto. constructor. apply PTree.gss.
- rewrite <- TY1. eauto. rewrite <- TY1. eauto.
- reflexivity. reflexivity. traceEq.
+ left. rewrite app_ass. rewrite Kseqlist_app.
+ eapply star_plus_trans. eexact EXEC.
+ simpl. eapply plus_four. econstructor. econstructor.
+ econstructor. eexact EV3. eexact EV2.
+ rewrite TY3; rewrite <- TY1; rewrite <- TY2. eauto.
+ econstructor. econstructor.
+ eexact EV1''. constructor. apply PTree.gss.
+ simpl. rewrite <- TY1; eauto.
+ rewrite <- TY1. eapply assign_loc_preserved; eauto.
+ reflexivity. traceEq.
econstructor. auto. apply S.
apply tr_val_gen. auto. intros. econstructor; eauto. constructor.
- rewrite H6; auto. apply PTree.gss.
- intros. apply PTree.gso. intuition congruence.
+ 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.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [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.
+ 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]].
+ exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
+ exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
+ subst; simpl Kseqlist.
+ econstructor; split.
+ right; split. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
+ simpl. omega.
+ constructor.
(* postincr *)
- exploit tr_top_leftcontext; eauto. clear H13.
+ exploit tr_top_leftcontext; eauto. clear H14.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
inv P. inv H5.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
- assert (EV2: eval_expr tge e le m a1 v1). eapply eval_Elvalue; eauto. rewrite <- TY1; auto.
+ 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.
+ intros. apply INV. NOTIN. intros [? [? EV1']].
subst; simpl Kseqlist.
econstructor; split.
- left. eapply plus_two. constructor.
+ left. rewrite app_ass; rewrite Kseqlist_app.
+ eapply star_plus_trans. eexact EXEC.
+ eapply plus_two. simpl. constructor.
econstructor; eauto.
unfold transl_incrdecr. destruct id; simpl in H2.
- econstructor. eauto. constructor. simpl. rewrite <- TY1. eauto.
- econstructor. eauto. constructor. simpl. rewrite <- TY1. eauto.
+ econstructor. eauto. constructor. simpl. rewrite TY3; rewrite <- TY1. eauto.
+ econstructor. eauto. constructor. simpl. rewrite TY3; rewrite <- TY1. eauto.
rewrite <- TY1. instantiate (1 := v3). destruct id; auto.
- rewrite <- TY1. eauto.
- traceEq.
+ rewrite <- TY1. eapply assign_loc_preserved; eauto.
+ reflexivity. traceEq.
econstructor. auto. change sl2 with (nil ++ sl2). apply S.
constructor. auto. auto. auto.
(* for value *)
@@ -1388,21 +1532,40 @@ Proof.
intros [SL2 [TY2 EV2]].
subst; simpl Kseqlist.
econstructor; split.
- left. eapply plus_four. constructor.
- constructor. eapply eval_Elvalue; eauto. rewrite <- TY1; eauto.
+ left. eapply plus_four. constructor.
+ eapply step_make_set; eauto.
constructor.
econstructor. eauto.
unfold transl_incrdecr. destruct id; simpl in H2.
econstructor. constructor. apply PTree.gss. constructor. simpl. eauto.
econstructor. constructor. apply PTree.gss. constructor. simpl. eauto.
rewrite <- TY1. instantiate (1 := v3). destruct id; auto.
- rewrite <- TY1. eauto.
+ rewrite <- TY1. eapply assign_loc_preserved; eauto.
traceEq.
econstructor. auto. apply S.
apply tr_val_gen. auto. intros. econstructor; eauto.
rewrite H5; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
auto. auto.
+(* postincr stuck *)
+ exploit tr_top_leftcontext; eauto. clear H11.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [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.
+ 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.
+ econstructor; split.
+ left. eapply plus_two. constructor. eapply step_make_set; eauto.
+ traceEq.
+ constructor.
(* comma *)
exploit tr_top_leftcontext; eauto. clear H9.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
@@ -1510,6 +1673,14 @@ Proof.
inv H0. exists a0; auto.
Qed.
+Lemma bind_parameters_preserved:
+ forall e m params args m',
+ bind_parameters ge e m params args m' ->
+ bind_parameters tge e m params args m'.
+Proof.
+ induction 1; econstructor; eauto. eapply assign_loc_preserved; eauto.
+Qed.
+
Lemma sstep_simulation:
forall S1 t S2, Csem.sstep ge S1 t S2 ->
forall S1' (MS: match_states S1 S1'),
@@ -1766,7 +1937,7 @@ Proof.
left; apply plus_one. eapply step_internal_function.
rewrite C; rewrite D; auto.
rewrite C; rewrite D; eauto.
- rewrite C; eauto.
+ rewrite C. eapply bind_parameters_preserved; eauto.
constructor; auto.
(* external function *)
diff --git a/cfrontend/SimplExprspec.v b/cfrontend/SimplExprspec.v
index 1224ea9..b3efd7f 100644
--- a/cfrontend/SimplExprspec.v
+++ b/cfrontend/SimplExprspec.v
@@ -48,6 +48,14 @@ Definition final (dst: destination) (a: expr) : list statement :=
| For_test tyl s1 s2 => makeif (fold_left Ecast tyl a) s1 s2 :: nil
end.
+Inductive tr_rvalof: type -> expr -> list statement -> expr -> list ident -> Prop :=
+ | tr_rvalof_nonvol: forall ty a tmp,
+ type_is_volatile ty = false ->
+ tr_rvalof ty a nil a tmp
+ | tr_rvalof_vol: forall ty a t tmp,
+ type_is_volatile ty = true -> In t tmp ->
+ tr_rvalof ty a (Svolread t a :: nil) (Etempvar t ty) tmp.
+
Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -> list ident -> Prop :=
| tr_var: forall le dst id ty tmp,
tr_expr le dst (C.Evar id ty)
@@ -80,11 +88,13 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
tr_expr le dst (C.Esizeof ty' ty)
(final dst (Esizeof ty' ty))
(Esizeof ty' ty) tmp
- | tr_valof: forall le dst e1 ty tmp sl1 a1,
- tr_expr le For_val e1 sl1 a1 tmp ->
+ | tr_valof: forall le dst e1 ty tmp sl1 a1 tmp1 sl2 a2 tmp2,
+ tr_expr le For_val e1 sl1 a1 tmp1 ->
+ tr_rvalof (C.typeof e1) a1 sl2 a2 tmp2 ->
+ list_disjoint tmp1 tmp2 -> incl tmp1 tmp -> incl tmp2 tmp ->
tr_expr le dst (C.Evalof e1 ty)
- (sl1 ++ final dst a1)
- a1 tmp
+ (sl1 ++ sl2 ++ final dst a2)
+ 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 (C.Eaddrof e1 ty)
@@ -153,38 +163,45 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
Sassign a1 (Etempvar t ty2) ::
final dst (Ecast (Etempvar t ty2) ty1))
(Ecast (Etempvar t ty2) ty1) tmp
- | tr_assignop_effects: forall le op e1 e2 tyres ty sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
+ | tr_assignop_effects: forall le op e1 e2 tyres ty ty1 sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 any tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_expr le For_val e2 sl2 a2 tmp2 ->
- list_disjoint tmp1 tmp2 ->
- incl tmp1 tmp -> incl tmp2 tmp ->
+ ty1 = C.typeof e1 ->
+ tr_rvalof ty1 a1 sl3 a3 tmp3 ->
+ list_disjoint tmp1 tmp2 -> list_disjoint tmp1 tmp3 -> list_disjoint tmp2 tmp3 ->
+ incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp ->
tr_expr le For_effects (C.Eassignop op e1 e2 tyres ty)
- (sl1 ++ sl2 ++ Sassign a1 (Ebinop op a1 a2 tyres) :: nil)
+ (sl1 ++ sl2 ++ sl3 ++ Sassign a1 (Ebinop op a3 a2 tyres) :: nil)
any tmp
- | tr_assignop_val: forall le dst op e1 e2 tyres ty sl1 a1 tmp1 sl2 a2 tmp2 t tmp ty1,
+ | tr_assignop_val: forall le dst op e1 e2 tyres ty sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 t tmp ty1,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_expr le For_val e2 sl2 a2 tmp2 ->
- list_disjoint tmp1 tmp2 ->
- incl tmp1 tmp -> incl tmp2 tmp ->
- In t tmp -> ~In t tmp1 -> ~In t tmp2 ->
+ tr_rvalof ty1 a1 sl3 a3 tmp3 ->
+ list_disjoint tmp1 tmp2 -> list_disjoint tmp1 tmp3 -> list_disjoint tmp2 tmp3 ->
+ incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp ->
+ In t tmp -> ~In t tmp1 -> ~In t tmp2 -> ~In t tmp3 ->
ty1 = C.typeof e1 ->
tr_expr le dst (C.Eassignop op e1 e2 tyres ty)
- (sl1 ++ sl2 ++
- Sset t (Ebinop op a1 a2 tyres) ::
+ (sl1 ++ sl2 ++ sl3 ++
+ Sset t (Ebinop op a3 a2 tyres) ::
Sassign a1 (Etempvar t tyres) ::
final dst (Ecast (Etempvar t tyres) ty1))
(Ecast (Etempvar t tyres) ty1) tmp
- | tr_postincr_effects: forall le id e1 ty sl1 a1 tmp any,
- tr_expr le For_val e1 sl1 a1 tmp ->
+ | tr_postincr_effects: forall le id e1 ty ty1 sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
+ tr_expr le For_val e1 sl1 a1 tmp1 ->
+ tr_rvalof ty1 a1 sl2 a2 tmp2 ->
+ ty1 = C.typeof e1 ->
+ incl tmp1 tmp -> incl tmp2 tmp ->
+ list_disjoint tmp1 tmp2 ->
tr_expr le For_effects (C.Epostincr id e1 ty)
- (sl1 ++ Sassign a1 (transl_incrdecr id a1 (C.typeof e1)) :: nil)
+ (sl1 ++ sl2 ++ Sassign a1 (transl_incrdecr id a2 ty1) :: nil)
any tmp
| tr_postincr_val: forall le dst id e1 ty sl1 a1 tmp1 t ty1 tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
incl tmp1 tmp -> In t tmp -> ~In t tmp1 ->
ty1 = C.typeof e1 ->
tr_expr le dst (C.Epostincr id e1 ty)
- (sl1 ++ Sset t a1 ::
+ (sl1 ++ make_set t a1 ::
Sassign a1 (transl_incrdecr id (Etempvar t ty1) ty1) ::
final dst (Etempvar t ty1))
(Etempvar t ty1) tmp
@@ -253,6 +270,13 @@ Proof.
induction 1; intros; econstructor; eauto.
Qed.
+Lemma tr_rvalof_monotone:
+ forall ty a sl b tmps, tr_rvalof ty a sl b tmps ->
+ forall tmps', incl tmps tmps' -> tr_rvalof ty a sl b tmps'.
+Proof.
+ induction 1; intros; econstructor; unfold incl in *; eauto.
+Qed.
+
Lemma tr_expr_monotone:
forall le dst r sl a tmps, tr_expr le dst r sl a tmps ->
forall tmps', incl tmps tmps' -> tr_expr le dst r sl a tmps'
@@ -260,6 +284,7 @@ with tr_exprlist_monotone:
forall le rl sl al tmps, tr_exprlist le rl sl al tmps ->
forall tmps', incl tmps tmps' -> tr_exprlist le rl sl al tmps'.
Proof.
+ specialize tr_rvalof_monotone. intros RVALOF.
induction 1; intros; econstructor; unfold incl in *; eauto.
induction 1; intros; econstructor; unfold incl in *; eauto.
Qed.
@@ -572,6 +597,25 @@ Ltac UseFinish :=
repeat rewrite app_ass
end.
+Lemma transl_valof_meets_spec:
+ forall ty a g sl b g' I,
+ transl_valof ty a g = Res (sl, b) g' I ->
+ exists tmps, tr_rvalof ty a sl b tmps /\ contained tmps g g'.
+Proof.
+ unfold transl_valof; intros.
+ destruct (type_is_volatile ty) as []_eqn; monadInv H.
+ exists (x :: nil); split; eauto with gensym. econstructor; eauto with coqlib.
+ exists (@nil ident); split; eauto with gensym. constructor; auto.
+(*
+ destruct (access_mode ty) as []_eqn.
+ destruct (Csem.type_is_volatile ty) as []_eqn; monadInv H.
+ exists (x :: nil); split; eauto with gensym. econstructor; eauto with coqlib.
+ exists (@nil ident); split; eauto with gensym. constructor; auto.
+ monadInv H. exists (@nil ident); split; eauto with gensym. constructor; auto.
+ monadInv H. exists (@nil ident); split; eauto with gensym. constructor; auto.
+*)
+Qed.
+
Scheme expr_ind2 := Induction for C.expr Sort Prop
with exprlist_ind2 := Induction for C.exprlist Sort Prop.
Combined Scheme expr_exprlist_ind from expr_ind2, exprlist_ind2.
@@ -603,9 +647,11 @@ Opaque makeif.
monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
econstructor; split; eauto. constructor; auto.
(* valof *)
- monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
- econstructor; split.
- econstructor; eauto. eauto with gensym.
+ monadInv H0. exploit H; eauto. intros [tmp1 [A B]].
+ exploit transl_valof_meets_spec; eauto. intros [tmp2 [C D]]. UseFinish.
+ exists (tmp1 ++ tmp2); split.
+ econstructor; eauto with gensym.
+ eauto with gensym.
(* deref *)
monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
econstructor; split; eauto. constructor; auto.
@@ -668,18 +714,19 @@ Opaque makeif.
(* assignop *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [C D]].
- destruct dst; monadInv EQ2.
+ exploit transl_valof_meets_spec; eauto. intros [tmp3 [E F]].
+ destruct dst; monadInv EQ3.
(* for value *)
- exists (x1 :: tmp1 ++ tmp2); split.
+ 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_app; eauto with gensym.
(* for effects *)
- exists (tmp1 ++ tmp2); split.
+ exists (tmp1 ++ tmp2 ++ tmp3); split.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
(* for test *)
- exists (x1 :: tmp1 ++ tmp2); split.
+ exists (x2 :: tmp1 ++ tmp2 ++ tmp3); split.
repeat rewrite app_ass. simpl.
intros. eapply tr_assignop_val with (dst := For_test tyl s1 s2); eauto with gensym.
apply contained_cons. eauto with gensym.
@@ -692,8 +739,10 @@ Opaque makeif.
econstructor; eauto with gensym.
apply contained_cons; eauto with gensym.
(* for effects *)
- exists tmp1; split.
- econstructor; eauto with gensym. auto.
+ exploit transl_valof_meets_spec; eauto. intros [tmp2 [C D]].
+ exists (tmp1 ++ tmp2); split.
+ econstructor; eauto with gensym.
+ eauto with gensym.
(* for test *)
repeat rewrite app_ass; simpl.
exists (x0 :: tmp1); split.
diff --git a/common/Behaviors.v b/common/Behaviors.v
index ccb5749..454ea34 100644
--- a/common/Behaviors.v
+++ b/common/Behaviors.v
@@ -530,6 +530,156 @@ Qed.
End BACKWARD_SIMULATIONS.
+(** * Program behaviors for the "atomic" construction *)
+
+Section ATOMIC.
+
+Variable L: semantics.
+Hypothesis Lwb: well_behaved_traces L.
+
+Remark atomic_finish: forall s t, output_trace t -> Star (atomic L) (t, s) t (E0, s).
+Proof.
+ induction t; intros.
+ apply star_refl.
+ simpl in H; destruct H. eapply star_left; eauto.
+ simpl. apply atomic_step_continue; auto. simpl; auto. auto.
+Qed.
+
+Lemma step_atomic_plus:
+ forall s1 t s2, Step L s1 t s2 -> Plus (atomic L) (E0,s1) t (E0,s2).
+Proof.
+ intros. destruct t.
+ apply plus_one. simpl; apply atomic_step_silent; auto.
+ exploit Lwb; eauto. simpl; intros.
+ eapply plus_left. eapply atomic_step_start; eauto. eapply atomic_finish; eauto. auto.
+Qed.
+
+Lemma star_atomic_star:
+ forall s1 t s2, Star L s1 t s2 -> Star (atomic L) (E0,s1) t (E0,s2).
+Proof.
+ induction 1. apply star_refl. eapply star_trans with (s2 := (E0,s2)).
+ apply plus_star. eapply step_atomic_plus; eauto. eauto. auto.
+Qed.
+
+Lemma atomic_forward_simulation: forward_simulation L (atomic L).
+Proof.
+ set (ms := fun (s: state L) (ts: state (atomic L)) => ts = (E0,s)).
+ apply forward_simulation_plus with ms; intros.
+ auto.
+ exists (E0,s1); split. simpl; auto. red; auto.
+ red in H. subst s2. simpl; auto.
+ red in H0. subst s2. exists (E0,s1'); split.
+ apply step_atomic_plus; auto. red; auto.
+Qed.
+
+Lemma atomic_star_star_gen:
+ forall ts1 t ts2, Star (atomic L) ts1 t ts2 ->
+ exists t', Star L (snd ts1) t' (snd ts2) /\ fst ts1 ** t' = t ** fst ts2.
+Proof.
+ induction 1.
+ exists E0; split. apply star_refl. traceEq.
+ destruct IHstar as [t' [A B]].
+ simpl in H; inv H; simpl in *.
+ exists t'; split. eapply star_left; eauto. auto.
+ exists (ev :: t0 ** t'); split. eapply star_left; eauto. rewrite B; auto.
+ exists t'; split. auto. rewrite B; auto.
+Qed.
+
+Lemma atomic_star_star:
+ forall s1 t s2, Star (atomic L) (E0,s1) t (E0,s2) -> Star L s1 t s2.
+Proof.
+ intros. exploit atomic_star_star_gen; eauto. intros [t' [A B]].
+ simpl in *. replace t with t'. auto. subst; traceEq.
+Qed.
+
+Lemma atomic_forever_silent_forever_silent:
+ forall s, Forever_silent (atomic L) s -> Forever_silent L (snd s).
+Proof.
+ cofix COINDHYP; intros. inv H. inv H0.
+ apply forever_silent_intro with (snd (E0, s')). auto. apply COINDHYP; auto.
+Qed.
+
+Remark star_atomic_output_trace:
+ forall s t t' s',
+ Star (atomic L) (E0, s) t (t', s') -> output_trace t'.
+Proof.
+ assert (forall ts1 t ts2, Star (atomic L) ts1 t ts2 ->
+ output_trace (fst ts1) -> output_trace (fst ts2)).
+ induction 1; intros. auto. inv H; simpl in *.
+ apply IHstar. auto.
+ apply IHstar. exploit Lwb; eauto.
+ destruct H2. apply IHstar. auto.
+ intros. change t' with (fst (t',s')). eapply H; eauto. simpl; auto.
+Qed.
+
+Lemma atomic_forever_reactive_forever_reactive:
+ forall s T, Forever_reactive (atomic L) (E0,s) T -> Forever_reactive L s T.
+Proof.
+ assert (forall t s T, Forever_reactive (atomic L) (t,s) T ->
+ exists T', Forever_reactive (atomic L) (E0,s) T' /\ T = t *** T').
+ induction t; intros. exists T; auto.
+ inv H. inv H0. congruence. simpl in H; inv H.
+ destruct (IHt s (t2***T0)) as [T' [A B]]. eapply star_forever_reactive; eauto.
+ exists T'; split; auto. simpl. congruence.
+
+ cofix COINDHYP; intros. inv H0. destruct s2 as [t2 s2].
+ destruct (H _ _ _ H3) as [T' [A B]].
+ assert (Star (atomic L) (E0, s) (t**t2) (E0, s2)).
+ eapply star_trans. eauto. apply atomic_finish. eapply star_atomic_output_trace; eauto. auto.
+ replace (t *** T0) with ((t ** t2) *** T'). apply forever_reactive_intro with s2.
+ apply atomic_star_star; auto. destruct t; simpl in *; unfold E0 in *; congruence.
+ apply COINDHYP. auto.
+ subst T0; traceEq.
+Qed.
+
+Theorem atomic_behaviors:
+ forall beh, program_behaves L beh <-> program_behaves (atomic L) beh.
+Proof.
+ intros; split; intros.
+ (* L -> atomic L *)
+ exploit forward_simulation_behavior_improves. eapply atomic_forward_simulation. eauto.
+ intros [beh2 [A B]]. red in B. destruct B as [EQ | [t [C D]]].
+ congruence.
+ subst beh. inv H. inv H1.
+ apply program_runs with (E0,s). simpl; auto.
+ apply state_goes_wrong with (E0,s'). apply star_atomic_star; auto.
+ red; intros; red; intros. inv H. eelim H3; eauto. eelim H3; eauto.
+ intros; red; intros. simpl in H. destruct H. eelim H4; eauto.
+ apply program_goes_initially_wrong.
+ intros; red; intros. simpl in H; destruct H. eelim H1; eauto.
+ (* atomic L -> L *)
+ inv H.
+ (* initial state defined *)
+ destruct s as [t s]. simpl in H0. destruct H0; subst t.
+ apply program_runs with s; auto.
+ inv H1.
+ (* termination *)
+ destruct s' as [t' s']. simpl in H2; destruct H2; subst t'.
+ econstructor. eapply atomic_star_star; eauto. auto.
+ (* silent divergence *)
+ destruct s' as [t' s'].
+ assert (t' = E0). inv H2. inv H1; auto. subst t'.
+ econstructor. eapply atomic_star_star; eauto.
+ change s' with (snd (E0,s')). apply atomic_forever_silent_forever_silent. auto.
+ (* reactive divergence *)
+ econstructor. apply atomic_forever_reactive_forever_reactive. auto.
+ (* going wrong *)
+ destruct s' as [t' s'].
+ assert (t' = E0).
+ destruct t'; auto. eelim H2. simpl. apply atomic_step_continue.
+ eapply star_atomic_output_trace; eauto.
+ subst t'. econstructor. apply atomic_star_star; eauto.
+ red; intros; red; intros. destruct t0.
+ elim (H2 E0 (E0,s'0)). constructor; auto.
+ elim (H2 (e::nil) (t0,s'0)). constructor; auto.
+ intros; red; intros. elim (H3 r). simpl; auto.
+ (* initial state undefined *)
+ apply program_goes_initially_wrong.
+ intros; red; intros. elim (H0 (E0,s)); simpl; auto.
+Qed.
+
+End ATOMIC.
+
(** * Additional results about infinite reduction sequences *)
(** We now show that any infinite sequence of reductions is either of
diff --git a/common/Events.v b/common/Events.v
index 018314e..3d082a7 100644
--- a/common/Events.v
+++ b/common/Events.v
@@ -490,6 +490,59 @@ Qed.
End MATCH_TRACES_INV.
+(** An output trace is a trace composed only of output events,
+ that is, events that do not take any result from the outside world. *)
+
+Definition output_event (ev: event) : Prop :=
+ match ev with
+ | Event_syscall _ _ _ => False
+ | Event_vload _ _ _ _ => False
+ | Event_vstore _ _ _ _ => True
+ | Event_annot _ _ => True
+ end.
+
+Fixpoint output_trace (t: trace) : Prop :=
+ match t with
+ | nil => True
+ | ev :: t' => output_event ev /\ output_trace t'
+ end.
+
+(** * Semantics of volatile memory accesses *)
+
+Definition block_is_volatile (F V: Type) (ge: Genv.t F V) (b: block) : bool :=
+ match Genv.find_var_info ge b with
+ | None => false
+ | Some gv => gv.(gvar_volatile)
+ end.
+
+Inductive volatile_load (F V: Type) (ge: Genv.t F V):
+ memory_chunk -> mem -> block -> int -> trace -> val -> Prop :=
+ | volatile_load_vol: forall chunk m b ofs id ev v,
+ block_is_volatile ge b = true ->
+ Genv.find_symbol ge id = Some b ->
+ eventval_match ge ev (type_of_chunk chunk) v ->
+ volatile_load ge chunk m b ofs
+ (Event_vload chunk id ofs ev :: nil)
+ (Val.load_result chunk v)
+ | volatile_load_nonvol: forall chunk m b ofs v,
+ block_is_volatile ge b = false ->
+ Mem.load chunk m b (Int.unsigned ofs) = Some v ->
+ volatile_load ge chunk m b ofs E0 v.
+
+Inductive volatile_store (F V: Type) (ge: Genv.t F V):
+ memory_chunk -> mem -> block -> int -> val -> trace -> mem -> Prop :=
+ | volatile_store_vol: forall chunk m b ofs id ev v,
+ block_is_volatile ge b = true ->
+ Genv.find_symbol ge id = Some b ->
+ eventval_match ge ev (type_of_chunk chunk) v ->
+ volatile_store ge chunk m b ofs v
+ (Event_vstore chunk id ofs ev :: nil)
+ m
+ | volatile_store_nonvol: forall chunk m b ofs v m',
+ block_is_volatile ge b = false ->
+ Mem.store chunk m b (Int.unsigned ofs) v = Some m' ->
+ volatile_store ge chunk m b ofs v E0 m'.
+
(** * Semantics of external functions *)
(** For each external function, its behavior is defined by a predicate relating:
@@ -530,12 +583,6 @@ Definition inject_separated (f f': meminj) (m1 m2: mem): Prop :=
f b1 = None -> f' b1 = Some(b2, delta) ->
~Mem.valid_block m1 b1 /\ ~Mem.valid_block m2 b2.
-Definition block_is_volatile (F V: Type) (ge: Genv.t F V) (b: block) : bool :=
- match Genv.find_var_info ge b with
- | None => false
- | Some gv => gv.(gvar_volatile)
- end.
-
Record extcall_properties (sem: extcall_sem)
(sg: signature) : Prop := mk_extcall_properties {
@@ -624,20 +671,34 @@ Record extcall_properties (sem: extcall_sem)
Inductive volatile_load_sem (chunk: memory_chunk) (F V: Type) (ge: Genv.t F V):
list val -> mem -> trace -> val -> mem -> Prop :=
- | volatile_load_sem_vol: forall b ofs m id ev v,
- Genv.find_symbol ge id = Some b -> block_is_volatile ge b = true ->
- eventval_match ge ev (type_of_chunk chunk) v ->
- volatile_load_sem chunk ge
- (Vptr b ofs :: nil) m
- (Event_vload chunk id ofs ev :: nil)
- (Val.load_result chunk v) m
- | volatile_load_sem_nonvol: forall b ofs m v,
- block_is_volatile ge b = false ->
- Mem.load chunk m b (Int.unsigned ofs) = Some v ->
- volatile_load_sem chunk ge
- (Vptr b ofs :: nil) m
- E0
- v m.
+ | volatile_load_sem_intro: forall b ofs m t v,
+ volatile_load ge chunk m b ofs t v ->
+ volatile_load_sem chunk ge (Vptr b ofs :: nil) m t v m.
+
+Lemma volatile_load_preserved:
+ forall F1 V1 (ge1: Genv.t F1 V1) F2 V2 (ge2: Genv.t F2 V2) chunk m b ofs t v,
+ (forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id) ->
+ (forall b, block_is_volatile ge2 b = block_is_volatile ge1 b) ->
+ volatile_load ge1 chunk m b ofs t v ->
+ volatile_load ge2 chunk m b ofs t v.
+Proof.
+ intros. inv H1; constructor; auto.
+ rewrite H0; auto.
+ rewrite H; auto.
+ eapply eventval_match_preserved; eauto.
+ rewrite H0; auto.
+Qed.
+
+Lemma volatile_load_extends:
+ forall F V (ge: Genv.t F V) chunk m b ofs t v m',
+ volatile_load ge chunk m b ofs t v ->
+ Mem.extends m m' ->
+ exists v', volatile_load ge chunk m' b ofs t v' /\ Val.lessdef v v'.
+Proof.
+ intros. inv H.
+ econstructor; split; eauto. econstructor; eauto.
+ exploit Mem.load_extends; eauto. intros [v' [A B]]. exists v'; split; auto. constructor; auto.
+Qed.
Remark meminj_preserves_block_is_volatile:
forall F V (ge: Genv.t F V) f b1 b2 delta,
@@ -653,6 +714,35 @@ Proof.
auto.
Qed.
+Lemma volatile_load_inject:
+ forall F V (ge: Genv.t F V) f chunk m b ofs t v b' ofs' m',
+ meminj_preserves_globals ge f ->
+ volatile_load ge chunk m b ofs t v ->
+ val_inject f (Vptr b ofs) (Vptr b' ofs') ->
+ Mem.inject f m m' ->
+ exists v', volatile_load ge chunk m' b' ofs' t v' /\ val_inject f v v'.
+Proof.
+ intros. inv H0.
+ inv H1. exploit (proj1 H); eauto. intros EQ; rewrite H8 in EQ; inv EQ.
+ rewrite Int.add_zero. exists (Val.load_result chunk v0); split.
+ constructor; auto.
+ apply val_load_result_inject. eapply eventval_match_inject_2; eauto.
+ exploit Mem.loadv_inject; eauto. simpl; eauto. simpl; intros [v' [A B]]. exists v'; split; auto.
+ constructor; auto. rewrite <- H3. inv H1. eapply meminj_preserves_block_is_volatile; eauto.
+Qed.
+
+Lemma volatile_load_receptive:
+ forall F V (ge: Genv.t F V) chunk m b ofs t1 t2 v1,
+ volatile_load ge chunk m b ofs t1 v1 -> match_traces ge t1 t2 ->
+ exists v2, volatile_load ge chunk m b ofs t2 v2.
+Proof.
+ intros. inv H; inv H0.
+ exploit eventval_match_valid; eauto. intros [A B].
+ exploit eventval_valid_match. eexact H9. rewrite <- H10; eauto.
+ intros [v' EVM]. exists (Val.load_result chunk v'). constructor; auto.
+ exists v1; constructor; auto.
+Qed.
+
Lemma volatile_load_ok:
forall chunk,
extcall_properties (volatile_load_sem chunk)
@@ -660,64 +750,36 @@ Lemma volatile_load_ok:
Proof.
intros; constructor; intros.
(* well typed *)
- unfold proj_sig_res; simpl. destruct H.
+ unfold proj_sig_res; simpl. inv H. inv H0.
destruct chunk; destruct v; simpl; constructor.
eapply Mem.load_type; eauto.
(* arity *)
- destruct H; simpl; auto.
+ inv H; inv H0; auto.
(* symbols *)
- destruct H1.
- econstructor; eauto. rewrite H; auto. eapply eventval_match_preserved; eauto.
- econstructor; eauto.
+ inv H1. constructor. eapply volatile_load_preserved; eauto.
(* valid blocks *)
- destruct H; auto.
+ inv H; auto.
(* bounds *)
- destruct H; auto.
+ inv H; auto.
(* mem extends *)
- destruct H.
- inv H1. inv H8. inv H6.
- exists (Val.load_result chunk v); exists m1'; intuition.
- constructor; auto.
- red; auto.
- inv H1. inv H7. inv H5.
- exploit Mem.load_extends; eauto. intros [v' [A B]].
- exists v'; exists m1'; intuition.
- econstructor; eauto.
- red; auto.
+ inv H. inv H1. inv H6. inv H4.
+ exploit volatile_load_extends; eauto. intros [v' [A B]].
+ exists v'; exists m1'; intuition. constructor; auto. red; auto.
(* mem injects *)
- destruct H0.
- inv H2. inv H9. inv H7.
- generalize H; intros [P [Q R]].
- exploit P; eauto. intro EQ; rewrite H6 in EQ; inv EQ.
- exists f; exists (Val.load_result chunk v); exists m1'; intuition.
- rewrite Int.add_zero. constructor; auto.
- apply val_load_result_inject. eapply eventval_match_inject_2; eauto.
- red; auto.
- red; auto.
- red; intros. congruence.
- inv H2. inv H8.
- exploit Mem.loadv_inject; eauto. simpl. eauto. intros [v1 [A B]].
- inv H6; simpl in *.
- exists f; exists v1; exists m1'; intuition.
- econstructor; eauto.
- rewrite <- H0. eapply meminj_preserves_block_is_volatile; eauto.
- red; auto.
- red; auto.
- red; intros. congruence.
+ inv H0. inv H2. inv H7. inversion H5; subst.
+ exploit volatile_load_inject; eauto. intros [v' [A B]].
+ exists f; exists v'; exists m1'; intuition. constructor; auto.
+ red; auto. red; auto. red; intros. congruence.
(* trace length *)
- inv H; simpl; omega.
+ inv H; inv H0; simpl; omega.
(* receptive *)
- inv H; inv H0.
- exploit eventval_match_valid; eauto. intros [A B].
- exploit eventval_valid_match. eexact H9. rewrite <- H10; eauto.
- intros [v' EVM]. exists (Val.load_result chunk v'); exists m1.
- eapply volatile_load_sem_vol; eauto.
- exists vres1; exists m1; eapply volatile_load_sem_nonvol; eauto.
+ inv H. exploit volatile_load_receptive; eauto. intros [v2 A].
+ exists v2; exists m1; constructor; auto.
(* determ *)
- inv H; inv H0; try congruence.
+ inv H; inv H0. inv H1; inv H7; try congruence.
assert (id = id0) by (eapply Genv.genv_vars_inj; eauto). subst id0.
- exploit eventval_match_valid. eexact H3. intros [V1 T1].
- exploit eventval_match_valid. eexact H11. intros [V2 T2].
+ exploit eventval_match_valid. eexact H2. intros [V1 T1].
+ exploit eventval_match_valid. eexact H4. intros [V2 T2].
split. constructor; auto. congruence.
intros EQ; inv EQ.
assert (v = v0) by (eapply eventval_match_determ_1; eauto). subst v0.
@@ -728,31 +790,19 @@ Qed.
Inductive volatile_load_global_sem (chunk: memory_chunk) (id: ident) (ofs: int)
(F V: Type) (ge: Genv.t F V):
list val -> mem -> trace -> val -> mem -> Prop :=
- | volatile_load_global_sem_vol: forall b m ev v,
- Genv.find_symbol ge id = Some b -> block_is_volatile ge b = true ->
- eventval_match ge ev (type_of_chunk chunk) v ->
- volatile_load_global_sem chunk id ofs ge
- nil m
- (Event_vload chunk id ofs ev :: nil)
- (Val.load_result chunk v) m
- | volatile_load_global_sem_nonvol: forall b m v,
- Genv.find_symbol ge id = Some b -> block_is_volatile ge b = false ->
- Mem.load chunk m b (Int.unsigned ofs) = Some v ->
- volatile_load_global_sem chunk id ofs ge
- nil m
- E0
- v m.
+ | volatile_load_global_sem_intro: forall b t v m,
+ Genv.find_symbol ge id = Some b ->
+ volatile_load ge chunk m b ofs t v ->
+ volatile_load_global_sem chunk id ofs ge nil m t v m.
Remark volatile_load_global_charact:
forall chunk id ofs (F V: Type) (ge: Genv.t F V) vargs m t vres m',
volatile_load_global_sem chunk id ofs ge vargs m t vres m' <->
exists b, Genv.find_symbol ge id = Some b /\ volatile_load_sem chunk ge (Vptr b ofs :: vargs) m t vres m'.
Proof.
- intros; split.
- intros. inv H; exists b; split; auto; constructor; auto.
- intros [b [P Q]]. inv Q.
- assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0. econstructor; eauto.
- econstructor; eauto.
+ intros; split.
+ intros. inv H. exists b; split; auto. constructor; auto.
+ intros [b [P Q]]. inv Q. econstructor; eauto.
Qed.
Lemma volatile_load_global_ok:
@@ -762,117 +812,79 @@ Lemma volatile_load_global_ok:
Proof.
intros; constructor; intros.
(* well typed *)
- unfold proj_sig_res; simpl. destruct H.
+ unfold proj_sig_res; simpl. inv H. inv H1.
destruct chunk; destruct v; simpl; constructor.
eapply Mem.load_type; eauto.
(* arity *)
- destruct H; simpl; auto.
+ inv H; inv H1; auto.
(* symbols *)
- destruct H1.
- econstructor; eauto. rewrite H; auto. eapply eventval_match_preserved; eauto.
- econstructor; eauto. rewrite H; auto.
+ inv H1. econstructor. rewrite H; eauto. eapply volatile_load_preserved; eauto.
(* valid blocks *)
- destruct H; auto.
+ inv H; auto.
(* bounds *)
- destruct H; auto.
-(* mem extends *)
- destruct H.
- inv H1.
- exists (Val.load_result chunk v); exists m1'; intuition.
- econstructor; eauto.
- red; auto.
- inv H1.
- exploit Mem.load_extends; eauto. intros [v' [A B]].
- exists v'; exists m1'; intuition.
- econstructor; eauto.
- red; auto.
-(* mem injects *)
- destruct H0.
- inv H2.
- exists f; exists (Val.load_result chunk v); exists m1'; intuition.
- econstructor; eauto.
- apply val_load_result_inject. eapply eventval_match_inject_2; eauto.
- red; auto.
- red; auto.
- red; intros. congruence.
- inv H2. destruct H as [P [Q R]]. exploit P; eauto. intros EQ.
- exploit Mem.load_inject; eauto. rewrite Zplus_0_r. intros [v1 [A B]].
- exists f; exists v1; exists m1'; intuition.
- econstructor; eauto.
- red; auto.
- red; auto.
- red; intros. congruence.
+ inv H; auto.
+(* extends *)
+ inv H. inv H1. exploit volatile_load_extends; eauto. intros [v' [A B]].
+ exists v'; exists m1'; intuition. econstructor; eauto. red; auto.
+(* inject *)
+ inv H0. inv H2.
+ assert (val_inject f (Vptr b ofs) (Vptr b ofs)).
+ exploit (proj1 H); eauto. intros EQ. econstructor. eauto. rewrite Int.add_zero; auto.
+ exploit volatile_load_inject; eauto. intros [v' [A B]].
+ exists f; exists v'; exists m1'; intuition. econstructor; eauto.
+ red; auto. red; auto. red; intros; congruence.
(* trace length *)
- inv H; simpl; omega.
+ inv H; inv H1; simpl; omega.
(* receptive *)
- inv H; inv H0.
- exploit eventval_match_valid; eauto. intros [A B].
- exploit eventval_valid_match. eexact H9. rewrite <- H10; eauto.
- intros [v' EVM]. exists (Val.load_result chunk v'); exists m1.
- eapply volatile_load_global_sem_vol; eauto.
- exists vres1; exists m1; eapply volatile_load_global_sem_nonvol; eauto.
+ inv H. exploit volatile_load_receptive; eauto. intros [v2 A].
+ exists v2; exists m1; econstructor; eauto.
(* determ *)
- inv H; inv H0; try congruence.
- assert (b = b0) by congruence. subst b0.
- exploit eventval_match_valid. eexact H3. intros [V1 T1].
- exploit eventval_match_valid. eexact H5. intros [V2 T2].
- split. constructor; auto. congruence.
- intros EQ; inv EQ.
- assert (v = v0) by (eapply eventval_match_determ_1; eauto). subst v0.
- auto.
- split. constructor. intuition congruence.
+ rewrite volatile_load_global_charact in *.
+ destruct H as [b1 [A1 B1]]. destruct H0 as [b2 [A2 B2]].
+ rewrite A1 in A2; inv A2.
+ eapply ec_determ. eapply volatile_load_ok. eauto. eauto.
Qed.
(** ** Semantics of volatile stores *)
Inductive volatile_store_sem (chunk: memory_chunk) (F V: Type) (ge: Genv.t F V):
list val -> mem -> trace -> val -> mem -> Prop :=
- | volatile_store_sem_vol: forall b ofs m id ev v,
- Genv.find_symbol ge id = Some b -> block_is_volatile ge b = true ->
- eventval_match ge ev (type_of_chunk chunk) v ->
- volatile_store_sem chunk ge
- (Vptr b ofs :: v :: nil) m
- (Event_vstore chunk id ofs ev :: nil)
- Vundef m
- | volatile_store_sem_nonvol: forall b ofs m v m',
- block_is_volatile ge b = false ->
- Mem.store chunk m b (Int.unsigned ofs) v = Some m' ->
- volatile_store_sem chunk ge
- (Vptr b ofs :: v :: nil) m
- E0
- Vundef m'.
+ | volatile_store_sem_intro: forall b ofs m1 v t m2,
+ volatile_store ge chunk m1 b ofs v t m2 ->
+ volatile_store_sem chunk ge (Vptr b ofs :: v :: nil) m1 t Vundef m2.
-Lemma volatile_store_ok:
- forall chunk,
- extcall_properties (volatile_store_sem chunk)
- (mksignature (Tint :: type_of_chunk chunk :: nil) None).
+Lemma volatile_store_preserved:
+ forall F1 V1 (ge1: Genv.t F1 V1) F2 V2 (ge2: Genv.t F2 V2) chunk m1 b ofs v t m2,
+ (forall id, Genv.find_symbol ge2 id = Genv.find_symbol ge1 id) ->
+ (forall b, block_is_volatile ge2 b = block_is_volatile ge1 b) ->
+ volatile_store ge1 chunk m1 b ofs v t m2 ->
+ volatile_store ge2 chunk m1 b ofs v t m2.
Proof.
- intros; constructor; intros.
-(* well typed *)
- unfold proj_sig_res; simpl. inv H; constructor.
-(* arity *)
- inv H; simpl; auto.
-(* symbols preserved *)
- inv H1.
- constructor. rewrite H; auto. rewrite H0; auto. eapply eventval_match_preserved; eauto.
- constructor; auto. rewrite H0; auto.
-(* valid block *)
- inv H. auto. eauto with mem.
-(* bounds *)
- inv H. auto. eapply Mem.bounds_store; eauto.
-(* mem extends*)
- inv H.
- inv H1. inv H6. inv H8. inv H7.
- exists Vundef; exists m1'; intuition.
- constructor; auto. eapply eventval_match_lessdef; eauto.
- red; auto.
- inv H1. inv H5. inv H7. inv H6.
- exploit Mem.store_within_extends; eauto. intros [m' [A B]].
- exists Vundef; exists m'; intuition.
- constructor; auto.
+ intros. inv H1; constructor; auto.
+ rewrite H0; auto.
+ rewrite H; auto.
+ eapply eventval_match_preserved; eauto.
+ rewrite H0; auto.
+Qed.
+
+Lemma volatile_store_extends:
+ forall F V (ge: Genv.t F V) chunk m1 b ofs v t m2 m1' v',
+ volatile_store ge chunk m1 b ofs v t m2 ->
+ Mem.extends m1 m1' ->
+ Val.lessdef v v' ->
+ exists m2',
+ volatile_store ge chunk m1' b ofs v' t m2'
+ /\ Mem.extends m2 m2'
+ /\ mem_unchanged_on (loc_out_of_bounds m1) m1' m2'.
+Proof.
+ intros. inv H.
+ econstructor; split. econstructor; eauto. eapply eventval_match_lessdef; eauto.
+ split. auto. red; auto.
+ exploit Mem.store_within_extends; eauto. intros [m2' [A B]].
+ exists m2'; intuition. econstructor; eauto.
red; split; intros.
eapply Mem.perm_store_1; eauto.
- rewrite <- H1. eapply Mem.load_store_other; eauto.
+ rewrite <- H4. eapply Mem.load_store_other; eauto.
destruct (eq_block b0 b); auto. subst b0; right.
exploit Mem.valid_access_in_bounds.
eapply Mem.store_valid_access_3. eexact H3.
@@ -883,30 +895,37 @@ Proof.
(Int.unsigned ofs, Int.unsigned ofs + size_chunk chunk)).
red; intros. generalize (H x H5). unfold loc_out_of_bounds, Intv.In; simpl. omega.
simpl; omega. simpl; omega.
-(* mem injects *)
- inv H0.
- inv H2. inv H7. inv H9. inv H10.
- generalize H; intros [P [Q R]].
- exploit P; eauto. intro EQ; rewrite H6 in EQ; inv EQ.
- exists f; exists Vundef; exists m1'; intuition.
- rewrite Int.add_zero. constructor; auto.
- eapply eventval_match_inject; eauto.
- red; auto.
- red; auto.
- red; intros; congruence.
- inv H2. inv H8. inv H9. inv H6.
+Qed.
+
+Lemma volatile_store_inject:
+ forall F V (ge: Genv.t F V) f chunk m1 b ofs v t m2 m1' b' ofs' v',
+ meminj_preserves_globals ge f ->
+ volatile_store ge chunk m1 b ofs v t m2 ->
+ val_inject f (Vptr b ofs) (Vptr b' ofs') ->
+ val_inject f v v' ->
+ Mem.inject f m1 m1' ->
+ exists m2',
+ volatile_store ge chunk m1' b' ofs' v' t m2'
+ /\ Mem.inject f m2 m2'
+ /\ mem_unchanged_on (loc_unmapped f) m1 m2
+ /\ mem_unchanged_on (loc_out_of_reach f m1) m1' m2'.
+Proof.
+ intros. inv H0.
+ inv H1. exploit (proj1 H); eauto. intros EQ; rewrite H9 in EQ; inv EQ.
+ rewrite Int.add_zero. exists m1'.
+ split. constructor; auto. eapply eventval_match_inject; eauto.
+ split. auto. split. red; auto. red; auto.
assert (Mem.storev chunk m1 (Vptr b ofs) v = Some m2). simpl; auto.
exploit Mem.storev_mapped_inject; eauto. intros [m2' [A B]].
- inv H4.
- exists f; exists Vundef; exists m2'; intuition.
- constructor; auto. rewrite <- H3. eapply meminj_preserves_block_is_volatile; eauto.
+ inv H1. exists m2'; intuition.
+ constructor; auto. rewrite <- H4. eapply meminj_preserves_block_is_volatile; eauto.
split; intros. eapply Mem.perm_store_1; eauto.
- rewrite <- H4. eapply Mem.load_store_other; eauto.
- left. exploit (H2 ofs0). generalize (size_chunk_pos chunk0). omega.
+ rewrite <- H6. eapply Mem.load_store_other; eauto.
+ left. exploit (H1 ofs0). generalize (size_chunk_pos chunk0). omega.
unfold loc_unmapped. congruence.
split; intros. eapply Mem.perm_store_1; eauto.
- rewrite <- H4. eapply Mem.load_store_other; eauto.
- destruct (eq_block b0 b2); auto. subst b0; right.
+ rewrite <- H6. eapply Mem.load_store_other; eauto.
+ destruct (eq_block b0 b'); auto. subst b0; right.
assert (EQ: Int.unsigned (Int.add ofs (Int.repr delta)) = Int.unsigned ofs + delta).
eapply Mem.address_inject; eauto with mem.
unfold Mem.storev in A. rewrite EQ in A. rewrite EQ.
@@ -917,16 +936,48 @@ Proof.
generalize (size_chunk_pos chunk). intro G.
apply (Intv.range_disjoint' (ofs0, ofs0 + size_chunk chunk0)
(Int.unsigned ofs + delta, Int.unsigned ofs + delta + size_chunk chunk)).
- red; intros. exploit (H2 x H8). eauto. unfold Intv.In; simpl. omega.
+ red; intros. exploit (H1 x H7). eauto. unfold Intv.In; simpl. omega.
simpl; omega. simpl; omega.
- red; intros; congruence.
+Qed.
+
+Lemma volatile_store_receptive:
+ forall F V (ge: Genv.t F V) chunk m b ofs v t1 m1 t2,
+ volatile_store ge chunk m b ofs v t1 m1 -> match_traces ge t1 t2 -> t1 = t2.
+Proof.
+ intros. inv H; inv H0; auto.
+Qed.
+
+Lemma volatile_store_ok:
+ forall chunk,
+ extcall_properties (volatile_store_sem chunk)
+ (mksignature (Tint :: type_of_chunk chunk :: nil) None).
+Proof.
+ intros; constructor; intros.
+(* well typed *)
+ unfold proj_sig_res; simpl. inv H; constructor.
+(* arity *)
+ inv H; simpl; auto.
+(* symbols preserved *)
+ inv H1. constructor. eapply volatile_store_preserved; eauto.
+(* valid block *)
+ inv H. inv H1. auto. eauto with mem.
+(* bounds *)
+ inv H. inv H1. auto. eapply Mem.bounds_store; eauto.
+(* mem extends*)
+ inv H. inv H1. inv H6. inv H7. inv H4.
+ exploit volatile_store_extends; eauto. intros [m2' [A [B C]]].
+ exists Vundef; exists m2'; intuition. constructor; auto.
+(* mem inject *)
+ inv H0. inv H2. inv H7. inv H8. inversion H5; subst.
+ exploit volatile_store_inject; eauto. intros [m2' [A [B [C D]]]].
+ exists f; exists Vundef; exists m2'; intuition. constructor; auto. red; intros; congruence.
(* trace length *)
- inv H; simpl; omega.
+ inv H; inv H0; simpl; omega.
(* receptive *)
- assert (t1 = t2). inv H; inv H0; auto.
- exists vres1; exists m1; congruence.
+ assert (t1 = t2). inv H. eapply volatile_store_receptive; eauto.
+ subst t2; exists vres1; exists m1; auto.
(* determ *)
- inv H; inv H0; try congruence.
+ inv H; inv H0. inv H1; inv H8; try congruence.
assert (id = id0) by (eapply Genv.genv_vars_inj; eauto). subst id0.
assert (ev = ev0) by (eapply eventval_match_determ_2; eauto). subst ev0.
split. constructor. auto.
@@ -936,20 +987,10 @@ Qed.
Inductive volatile_store_global_sem (chunk: memory_chunk) (id: ident) (ofs: int)
(F V: Type) (ge: Genv.t F V):
list val -> mem -> trace -> val -> mem -> Prop :=
- | volatile_store_global_sem_vol: forall b m ev v,
- Genv.find_symbol ge id = Some b -> block_is_volatile ge b = true ->
- eventval_match ge ev (type_of_chunk chunk) v ->
- volatile_store_global_sem chunk id ofs ge
- (v :: nil) m
- (Event_vstore chunk id ofs ev :: nil)
- Vundef m
- | volatile_store_global_sem_nonvol: forall b m v m',
- Genv.find_symbol ge id = Some b -> block_is_volatile ge b = false ->
- Mem.store chunk m b (Int.unsigned ofs) v = Some m' ->
- volatile_store_global_sem chunk id ofs ge
- (v :: nil) m
- E0
- Vundef m'.
+ | volatile_store_global_sem_intro: forall b m1 v t m2,
+ Genv.find_symbol ge id = Some b ->
+ volatile_store ge chunk m1 b ofs v t m2 ->
+ volatile_store_global_sem chunk id ofs ge (v :: nil) m1 t Vundef m2.
Remark volatile_store_global_charact:
forall chunk id ofs (F V: Type) (ge: Genv.t F V) vargs m t vres m',
@@ -958,9 +999,7 @@ Remark volatile_store_global_charact:
Proof.
intros; split.
intros. inv H; exists b; split; auto; econstructor; eauto.
- intros [b [P Q]]. inv Q.
- assert (id0 = id) by (eapply Genv.genv_vars_inj; eauto). subst id0. econstructor; eauto.
- econstructor; eauto.
+ intros [b [P Q]]. inv Q. econstructor; eauto.
Qed.
Lemma volatile_store_global_ok:
@@ -974,13 +1013,11 @@ Proof.
(* arity *)
inv H; simpl; auto.
(* symbols preserved *)
- inv H1.
- econstructor. rewrite H; eauto. rewrite H0; auto. eapply eventval_match_preserved; eauto.
- econstructor; eauto. rewrite H; auto.
+ inv H1. econstructor. rewrite H; eauto. eapply volatile_store_preserved; eauto.
(* valid block *)
- inv H. auto. eauto with mem.
+ inv H. inv H2. auto. eauto with mem.
(* bounds *)
- inv H. auto. eapply Mem.bounds_store; eauto.
+ inv H. inv H2. auto. eapply Mem.bounds_store; eauto.
(* mem extends*)
rewrite volatile_store_global_charact in H. destruct H as [b [P Q]].
exploit ec_mem_extends. eapply volatile_store_ok. eexact Q. eauto. eauto.
@@ -995,16 +1032,14 @@ Proof.
exists f'; exists vres'; exists m2'; intuition.
rewrite volatile_store_global_charact. exists b; auto.
(* trace length *)
- inv H; simpl; omega.
+ inv H. inv H1; simpl; omega.
(* receptive *)
- assert (t1 = t2). inv H; inv H0; auto.
+ assert (t1 = t2). inv H. eapply volatile_store_receptive; eauto. subst t2.
exists vres1; exists m1; congruence.
(* determ *)
- inv H; inv H0; try congruence.
- assert (b = b0) by congruence. subst b0.
- assert (ev = ev0) by (eapply eventval_match_determ_2; eauto). subst ev0.
- split. constructor. auto.
- split. constructor. intuition congruence.
+ rewrite volatile_store_global_charact in *.
+ destruct H as [b1 [A1 B1]]. destruct H0 as [b2 [A2 B2]]. rewrite A1 in A2; inv A2.
+ eapply ec_determ. eapply volatile_store_ok. eauto. eauto.
Qed.
(** ** Semantics of dynamic memory allocation (malloc) *)
diff --git a/common/Smallstep.v b/common/Smallstep.v
index 63ab5ea..458e8c6 100644
--- a/common/Smallstep.v
+++ b/common/Smallstep.v
@@ -817,12 +817,15 @@ End COMPOSE_SIMULATIONS.
(** * Receptiveness and determinacy *)
+Definition single_events (L: semantics) : Prop :=
+ forall s t s', Step L s t s' -> (length t <= 1)%nat.
+
Record receptive (L: semantics) : Prop :=
Receptive {
sr_receptive: forall s t1 s1 t2,
Step L s t1 s1 -> match_traces (globalenv L) t1 t2 -> exists s2, Step L s t2 s2;
- sr_traces: forall s t s',
- Step L s t s' -> (length t <= 1)%nat
+ sr_traces:
+ single_events L
}.
Record determinate (L: semantics) : Prop :=
@@ -830,8 +833,8 @@ Record determinate (L: semantics) : Prop :=
sd_determ: forall s t1 s1 t2 s2,
Step L s t1 s1 -> Step L s t2 s2 ->
match_traces (globalenv L) t1 t2 /\ (t1 = t2 -> s1 = s2);
- sd_traces: forall s t s',
- Step L s t s' -> (length t <= 1)%nat;
+ sd_traces:
+ single_events L;
sd_initial_determ: forall s1 s2,
initial_state L s1 -> initial_state L s2 -> s1 = s2;
sd_final_nostep: forall s r,
@@ -925,8 +928,6 @@ Record backward_simulation (L1 L2: semantics) : Type :=
exists i', exists s1',
(Plus L1 s1 t s1' \/ (Star L1 s1 t s1' /\ bsim_order i' i))
/\ bsim_match_states i' s1' s2';
- bsim_traces:
- forall s2 t s2', Step L2 s2 t s2' -> (length t <= 1)%nat;
bsim_symbols_preserved:
forall id, Genv.find_symbol (globalenv L2) id = Genv.find_symbol (globalenv L1) id
}.
@@ -960,9 +961,6 @@ Variable L2: semantics.
Hypothesis symbols_preserved:
forall id, Genv.find_symbol (globalenv L2) id = Genv.find_symbol (globalenv L1) id.
-Hypothesis length_traces:
- forall s2 t s2', Step L2 s2 t s2' -> (length t <= 1)%nat.
-
Variable match_states: state L1 -> state L2 -> Prop.
Hypothesis initial_states_exist:
@@ -1085,6 +1083,7 @@ Section COMPOSE_BACKWARD_SIMULATIONS.
Variable L1: semantics.
Variable L2: semantics.
Variable L3: semantics.
+Hypothesis L3_single_events: single_events L3.
Variable S12: backward_simulation L1 L2.
Variable S23: backward_simulation L2 L3.
@@ -1117,7 +1116,7 @@ Proof.
intros [ [i2' [s2' [PLUS2 MATCH2]]] | [i2' [ORD2 [EQ MATCH2]]]].
(* 1 L2 makes one or several transitions *)
assert (EITHER: t = E0 \/ (length t = 1)%nat).
- exploit bsim_traces; eauto.
+ exploit L3_single_events; eauto.
destruct t; auto. destruct t; auto. simpl. intros. omegaContradiction.
destruct EITHER.
(* 1.1 these are silent transitions *)
@@ -1202,8 +1201,6 @@ Proof.
eapply (bsim_progress S23). eauto. eapply star_safe; eauto. eapply bsim_safe; eauto.
(* simulation *)
exact bb_simulation.
-(* trace length *)
- exact (bsim_traces S23).
(* symbols *)
intros. transitivity (Genv.find_symbol (globalenv L2) id); apply bsim_symbols_preserved; auto.
Qed.
@@ -1217,9 +1214,8 @@ Section FORWARD_TO_BACKWARD.
Variable L1: semantics.
Variable L2: semantics.
Variable FS: forward_simulation L1 L2.
-
-Hypothesis receptive: receptive L1.
-Hypothesis determinate: determinate L2.
+Hypothesis L1_receptive: receptive L1.
+Hypothesis L2_determinate: determinate L2.
(** Exploiting forward simulation *)
@@ -1391,7 +1387,7 @@ Proof.
(* 1. At matching states *)
exploit f2b_progress; eauto. intros TRANS; inv TRANS.
(* 1.1 L1 can reach final state and L2 is at final state: impossible! *)
- exploit (sd_final_nostep determinate); eauto. contradiction.
+ exploit (sd_final_nostep L2_determinate); eauto. contradiction.
(* 1.2 L1 can make 0 or several steps; L2 can make 1 or several matching steps. *)
inv H2.
exploit f2b_determinacy_inv. eexact H5. eexact STEP2.
@@ -1409,15 +1405,15 @@ Proof.
right; split. auto. constructor.
econstructor. eauto. auto. apply star_one; eauto. eauto. eauto.
(* 1.2.2 L2 makes a non-silent transition, and so does L1 *)
- exploit not_silent_length. eapply (sr_traces receptive); eauto. intros [EQ | EQ].
+ exploit not_silent_length. eapply (sr_traces L1_receptive); eauto. intros [EQ | EQ].
congruence.
subst t2. rewrite E0_right in H1.
(* Use receptiveness to equate the traces *)
- exploit (sr_receptive receptive); eauto. intros [s1''' STEP1].
+ exploit (sr_receptive L1_receptive); eauto. intros [s1''' STEP1].
exploit fsim_simulation_not_E0. eexact STEP1. auto. eauto.
intros [i''' [s2''' [P Q]]]. inv P.
(* Exploit determinacy *)
- exploit not_silent_length. eapply (sr_traces receptive); eauto. intros [EQ | EQ].
+ exploit not_silent_length. eapply (sr_traces L1_receptive); eauto. intros [EQ | EQ].
subst t0. simpl in *. exploit sd_determ_1. eauto. eexact STEP2. eexact H2.
intros. elim NOT2. inv H8. auto.
subst t2. rewrite E0_right in *.
@@ -1436,17 +1432,17 @@ Proof.
right; split. apply star_refl. constructor. omega.
econstructor; eauto. eapply star_right; eauto.
(* 2.2 L2 make a non-silent transition *)
- exploit not_silent_length. eapply (sr_traces receptive); eauto. intros [EQ | EQ].
+ exploit not_silent_length. eapply (sr_traces L1_receptive); eauto. intros [EQ | EQ].
congruence.
subst. rewrite E0_right in *.
(* Use receptiveness to equate the traces *)
- exploit (sr_receptive receptive); eauto. intros [s1''' STEP1].
+ exploit (sr_receptive L1_receptive); eauto. intros [s1''' STEP1].
exploit fsim_simulation_not_E0. eexact STEP1. auto. eauto.
intros [i''' [s2''' [P Q]]].
(* Exploit determinacy *)
exploit f2b_determinacy_star. eauto. eexact STEP2. auto. apply plus_star; eauto.
intro R. inv R. congruence.
- exploit not_silent_length. eapply (sr_traces receptive); eauto. intros [EQ | EQ].
+ exploit not_silent_length. eapply (sr_traces L1_receptive); eauto. intros [EQ | EQ].
subst. simpl in *. exploit sd_determ_1. eauto. eexact STEP2. eexact H2.
intros. elim NOT2. inv H7; auto.
subst. rewrite E0_right in *.
@@ -1482,11 +1478,11 @@ Proof.
(* final states *)
intros. inv H.
exploit f2b_progress; eauto. intros TRANS; inv TRANS.
- assert (r0 = r) by (eapply (sd_final_determ determinate); eauto). subst r0.
+ assert (r0 = r) by (eapply (sd_final_determ L2_determinate); eauto). subst r0.
exists s1'; auto.
- inv H4. exploit (sd_final_nostep determinate); eauto. contradiction.
- inv H5. congruence. exploit (sd_final_nostep determinate); eauto. contradiction.
- inv H2. exploit (sd_final_nostep determinate); eauto. contradiction.
+ inv H4. exploit (sd_final_nostep L2_determinate); eauto. contradiction.
+ inv H5. congruence. exploit (sd_final_nostep L2_determinate); eauto. contradiction.
+ inv H2. exploit (sd_final_nostep L2_determinate); eauto. contradiction.
(* progress *)
intros. inv H.
exploit f2b_progress; eauto. intros TRANS; inv TRANS.
@@ -1496,14 +1492,265 @@ Proof.
inv H1. right; econstructor; econstructor; eauto.
(* simulation *)
exact f2b_simulation_step.
-(* trace length *)
- exact (sd_traces determinate).
(* symbols preserved *)
exact (fsim_symbols_preserved FS).
Qed.
End FORWARD_TO_BACKWARD.
+(** * Transforming a semantics into a single-event, equivalent semantics *)
+
+Definition well_behaved_traces (L: semantics) : Prop :=
+ forall s t s', Step L s t s' ->
+ match t with nil => True | ev :: t' => output_trace t' end.
+
+Section ATOMIC.
+
+Variable L: semantics.
+
+Hypothesis Lwb: well_behaved_traces L.
+
+Inductive atomic_step (ge: Genv.t (funtype L) (vartype L)): (trace * state L) -> trace -> (trace * state L) -> Prop :=
+ | atomic_step_silent: forall s s',
+ Step L s E0 s' ->
+ atomic_step ge (E0, s) E0 (E0, s')
+ | atomic_step_start: forall s ev t s',
+ Step L s (ev :: t) s' ->
+ atomic_step ge (E0, s) (ev :: nil) (t, s')
+ | atomic_step_continue: forall ev t s,
+ output_trace (ev :: t) ->
+ atomic_step ge (ev :: t, s) (ev :: nil) (t, s).
+
+Definition atomic : semantics := {|
+ state := (trace * state L)%type;
+ funtype := funtype L;
+ vartype := vartype L;
+ step := atomic_step;
+ initial_state := fun s => initial_state L (snd s) /\ fst s = E0;
+ final_state := fun s r => final_state L (snd s) r /\ fst s = E0;
+ globalenv := globalenv L
+|}.
+
+End ATOMIC.
+
+(** A forward simulation from a semantics [L1] to a single-event semantics [L2]
+ can be "factored" into a forward simulation from [atomic L1] to [L2]. *)
+
+Section FACTOR_FORWARD_SIMULATION.
+
+Variable L1: semantics.
+Variable L2: semantics.
+Variable sim: forward_simulation L1 L2.
+Hypothesis L2single: single_events L2.
+
+Inductive ffs_match: fsim_index sim -> (trace * state L1) -> state L2 -> Prop :=
+ | ffs_match_at: forall i s1 s2,
+ sim i s1 s2 ->
+ ffs_match i (E0, s1) s2
+ | ffs_match_buffer: forall i ev t s1 s2 s2',
+ Star L2 s2 (ev :: t) s2' -> sim i s1 s2' ->
+ ffs_match i (ev :: t, s1) s2.
+
+Lemma star_non_E0_split':
+ forall s2 t s2', Star L2 s2 t s2' ->
+ match t with
+ | nil => True
+ | ev :: t' => exists s2x, Plus L2 s2 (ev :: nil) s2x /\ Star L2 s2x t' s2'
+ end.
+Proof.
+ induction 1. simpl. auto.
+ exploit L2single; eauto. intros LEN.
+ destruct t1. simpl in *. subst. destruct t2. auto.
+ destruct IHstar as [s2x [A B]]. exists s2x; split; auto.
+ eapply plus_left. eauto. apply plus_star; eauto. auto.
+ destruct t1. simpl in *. subst t. exists s2; split; auto. apply plus_one; auto.
+ simpl in LEN. omegaContradiction.
+Qed.
+
+Lemma ffs_simulation:
+ forall s1 t s1', Step (atomic L1) s1 t s1' ->
+ forall i s2, ffs_match i s1 s2 ->
+ exists i', exists s2',
+ (Plus L2 s2 t s2' \/ (Star L2 s2 t s2') /\ fsim_order sim i' i)
+ /\ ffs_match i' s1' s2'.
+Proof.
+ induction 1; intros.
+(* silent step *)
+ inv H0.
+ exploit (fsim_simulation sim); eauto.
+ intros [i' [s2' [A B]]].
+ exists i'; exists s2'; split. auto. constructor; auto.
+(* start step *)
+ inv H0.
+ exploit (fsim_simulation sim); eauto.
+ intros [i' [s2' [A B]]].
+ destruct t as [ | ev' t].
+ (* single event *)
+ exists i'; exists s2'; split. auto. constructor; auto.
+ (* multiple events *)
+ assert (C: Star L2 s2 (ev :: ev' :: t) s2'). intuition. apply plus_star; auto.
+ exploit star_non_E0_split'. eauto. simpl. intros [s2x [P Q]].
+ exists i'; exists s2x; split. auto. econstructor; eauto.
+(* continue step *)
+ inv H0.
+ exploit star_non_E0_split'. eauto. simpl. intros [s2x [P Q]].
+ destruct t.
+ exists i; exists s2'; split. left. eapply plus_star_trans; eauto. constructor; auto.
+ exists i; exists s2x; split. auto. econstructor; eauto.
+Qed.
+
+Theorem factor_forward_simulation:
+ forward_simulation (atomic L1) L2.
+Proof.
+ apply Forward_simulation with (fsim_match_states := ffs_match) (fsim_order := fsim_order sim).
+(* wf *)
+ apply fsim_order_wf.
+(* initial states *)
+ intros. destruct s1 as [t1 s1]. simpl in H. destruct H. subst.
+ exploit (fsim_match_initial_states sim); eauto. intros [i [s2 [A B]]].
+ exists i; exists s2; split; auto. constructor; auto.
+(* final states *)
+ intros. destruct s1 as [t1 s1]. simpl in H0; destruct H0; subst. inv H.
+ eapply (fsim_match_final_states sim); eauto.
+(* simulation *)
+ exact ffs_simulation.
+(* symbols preserved *)
+ simpl. exact (fsim_symbols_preserved sim).
+Qed.
+
+End FACTOR_FORWARD_SIMULATION.
+
+(** Likewise, a backward simulation from a single-event semantics [L1] to a semantics [L2]
+ can be "factored" as a backward simulation from [L1] to [atomic L2]. *)
+
+Section FACTOR_BACKWARD_SIMULATION.
+
+Variable L1: semantics.
+Variable L2: semantics.
+Variable sim: backward_simulation L1 L2.
+Hypothesis L1single: single_events L1.
+Hypothesis L2wb: well_behaved_traces L2.
+
+Inductive fbs_match: bsim_index sim -> state L1 -> (trace * state L2) -> Prop :=
+ | fbs_match_intro: forall i s1 t s2 s1',
+ Star L1 s1 t s1' -> sim i s1' s2 ->
+ t = E0 \/ output_trace t ->
+ fbs_match i s1 (t, s2).
+
+Lemma fbs_simulation:
+ forall s2 t s2', Step (atomic L2) s2 t s2' ->
+ forall i s1, fbs_match i s1 s2 -> safe L1 s1 ->
+ exists i', exists s1',
+ (Plus L1 s1 t s1' \/ (Star L1 s1 t s1' /\ bsim_order sim i' i))
+ /\ fbs_match i' s1' s2'.
+Proof.
+ induction 1; intros.
+(* silent step *)
+ inv H0.
+ exploit (bsim_simulation sim); eauto. eapply star_safe; eauto.
+ intros [i' [s1'' [A B]]].
+ exists i'; exists s1''; split.
+ destruct A as [P | [P Q]]. left. eapply star_plus_trans; eauto. right; split; auto. eapply star_trans; eauto.
+ econstructor. apply star_refl. auto. auto.
+(* start step *)
+ inv H0.
+ exploit (bsim_simulation sim); eauto. eapply star_safe; eauto.
+ intros [i' [s1'' [A B]]].
+ assert (C: Star L1 s1 (ev :: t) s1'').
+ eapply star_trans. eauto. destruct A as [P | [P Q]]. apply plus_star; eauto. eauto. auto.
+ exploit star_non_E0_split'; eauto. simpl. intros [s1x [P Q]].
+ exists i'; exists s1x; split.
+ left; auto.
+ econstructor; eauto.
+ exploit L2wb; eauto.
+(* continue step *)
+ inv H0. unfold E0 in H8; destruct H8; try congruence.
+ exploit star_non_E0_split'; eauto. simpl. intros [s1x [P Q]].
+ exists i; exists s1x; split. left; auto. econstructor; eauto. simpl in H0; tauto.
+Qed.
+
+Lemma fbs_progress:
+ forall i s1 s2,
+ fbs_match i s1 s2 -> safe L1 s1 ->
+ (exists r, final_state (atomic L2) s2 r) \/
+ (exists t, exists s2', Step (atomic L2) s2 t s2').
+Proof.
+ intros. inv H. destruct t.
+(* 1. no buffered events *)
+ exploit (bsim_progress sim); eauto. eapply star_safe; eauto.
+ intros [[r A] | [t [s2' A]]].
+(* final state *)
+ left; exists r; simpl; auto.
+(* L2 can step *)
+ destruct t.
+ right; exists E0; exists (nil, s2'). constructor. auto.
+ right; exists (e :: nil); exists (t, s2'). constructor. auto.
+(* 2. some buffered events *)
+ unfold E0 in H3; destruct H3. congruence.
+ right; exists (e :: nil); exists (t, s3). constructor. auto.
+Qed.
+
+Theorem factor_backward_simulation:
+ backward_simulation L1 (atomic L2).
+Proof.
+ apply Backward_simulation with (bsim_match_states := fbs_match) (bsim_order := bsim_order sim).
+(* wf *)
+ apply bsim_order_wf.
+(* initial states exist *)
+ intros. exploit (bsim_initial_states_exist sim); eauto. intros [s2 A].
+ exists (E0, s2). simpl; auto.
+(* initial states match *)
+ intros. destruct s2 as [t s2]; simpl in H0; destruct H0; subst.
+ exploit (bsim_match_initial_states sim); eauto. intros [i [s1' [A B]]].
+ exists i; exists s1'; split. auto. econstructor. apply star_refl. auto. auto.
+(* final states match *)
+ intros. destruct s2 as [t s2]; simpl in H1; destruct H1; subst.
+ inv H. exploit (bsim_match_final_states sim); eauto. eapply star_safe; eauto.
+ intros [s1'' [A B]]. exists s1''; split; auto. eapply star_trans; eauto.
+(* progress *)
+ exact fbs_progress.
+(* simulation *)
+ exact fbs_simulation.
+(* symbols *)
+ simpl. exact (bsim_symbols_preserved sim).
+Qed.
+
+End FACTOR_BACKWARD_SIMULATION.
+
+(** Receptiveness of [atomic L]. *)
+
+Record strongly_receptive (L: semantics) : Prop :=
+ Strongly_receptive {
+ ssr_receptive: forall s ev1 t1 s1 ev2,
+ Step L s (ev1 :: t1) s1 ->
+ match_traces (globalenv L) (ev1 :: nil) (ev2 :: nil) ->
+ exists s2, exists t2, Step L s (ev2 :: t2) s2;
+ ssr_well_behaved:
+ well_behaved_traces L
+ }.
+
+Theorem atomic_receptive:
+ forall L, strongly_receptive L -> receptive (atomic L).
+Proof.
+ intros. constructor; intros.
+(* receptive *)
+ inv H0.
+ (* silent step *)
+ inv H1. exists (E0, s'). constructor; auto.
+ (* start step *)
+ assert (exists ev2, t2 = ev2 :: nil). inv H1; econstructor; eauto.
+ destruct H0 as [ev2 EQ]; subst t2.
+ exploit ssr_receptive; eauto. intros [s2 [t2 P]].
+ exploit ssr_well_behaved. eauto. eexact P. simpl; intros Q.
+ exists (t2, s2). constructor; auto.
+ (* continue step *)
+ simpl in H2; destruct H2.
+ assert (t2 = ev :: nil). inv H1; simpl in H0; tauto.
+ subst t2. exists (t, s0). constructor; auto. simpl; auto.
+(* single-event *)
+ red. intros. inv H0; simpl; omega.
+Qed.
+
(** * Connections with big-step semantics *)
(** The general form of a big-step semantics *)
diff --git a/common/Values.v b/common/Values.v
index 7fae3b7..54eac86 100644
--- a/common/Values.v
+++ b/common/Values.v
@@ -160,6 +160,13 @@ Definition notint (v: val) : val :=
Definition of_bool (b: bool): val := if b then Vtrue else Vfalse.
+Definition boolval (v: val) : val :=
+ match v with
+ | Vint n => of_bool (negb (Int.eq n Int.zero))
+ | Vptr b ofs => Vtrue
+ | _ => Vundef
+ end.
+
Definition notbool (v: val) : val :=
match v with
| Vint n => of_bool (Int.eq n Int.zero)
diff --git a/configure b/configure
index da2e59f..9c2618a 100755
--- a/configure
+++ b/configure
@@ -32,6 +32,8 @@ usage='Usage: ./configure [options] target
Supported targets:
ppc-macosx (PowerPC, MacOS X)
ppc-linux (PowerPC, Linux)
+ ppc-eabi-unix (PowerPC, EABI with Unix tools)
+ ppc-eabi-diab (PowerPC, EABI with Diab tools)
arm-linux (ARM, Linux EABI)
ia32-linux (x86 32 bits, Linux)
ia32-bsd (x86 32 bits, BSD)
@@ -76,7 +78,7 @@ case "$target" in
clinker="gcc -arch ppc"
libmath=""
need_stdlib_wrapper="true";;
- powerpc-linux|ppc-linux)
+ powerpc-linux|ppc-linux|powerpc-eabi-unix|ppc-eabi-unix)
arch="powerpc"
variant="eabi"
system="linux"
@@ -86,6 +88,16 @@ case "$target" in
clinker="gcc"
libmath="-lm"
need_stdlib_wrapper="false";;
+ powerpc-eabi-diab|ppc-eabi-diab)
+ arch="powerpc"
+ variant="eabi"
+ system="diab"
+ cc="dcc"
+ cprepro="dcc -E"
+ casm="das"
+ clinker="dcc"
+ libmath="-lm"
+ need_stdlib_wrapper="false";;
arm-linux)
arch="arm"
variant="linux"
@@ -227,7 +239,7 @@ fi
if test "$target" = "manual"; then
cat <<EOF
-Please finish the configuration by editing file ./Makefile.config
+Please finish the configuration by editing file ./Makefile.config.
EOF
@@ -251,6 +263,8 @@ CompCert configuration:
Binaries installed in......... $bindirexp
Library files installed in.... $libdirexp
+If anything above looks wrong, please edit file ./Makefile.config to correct.
+
EOF
fi
diff --git a/cparser/.depend b/cparser/.depend
index 51f3b5e..4a7f538 100644
--- a/cparser/.depend
+++ b/cparser/.depend
@@ -1,4 +1,3 @@
-AddCasts.cmi: C.cmi
Bitfields.cmi: C.cmi
Builtins.cmi: Env.cmi C.cmi
C.cmi:
@@ -12,18 +11,13 @@ Errors.cmi:
GCC.cmi: Builtins.cmi
Lexer.cmi: Parser.cmi
Machine.cmi:
-PackedStructs.cmi: C.cmi
Parse.cmi: C.cmi
Parse_aux.cmi:
Parser.cmi: Cabs.cmo
Rename.cmi: C.cmi
-SimplExpr.cmi: C.cmi
-StructAssign.cmi: C.cmi
-StructByValue.cmi: C.cmi
+StructReturn.cmi: C.cmi
Transform.cmi: Env.cmi C.cmi
Unblock.cmi: C.cmi
-AddCasts.cmo: Transform.cmi Cutil.cmi C.cmi AddCasts.cmi
-AddCasts.cmx: Transform.cmx Cutil.cmx C.cmi AddCasts.cmi
Bitfields.cmo: Transform.cmi Machine.cmi Cutil.cmi C.cmi Bitfields.cmi
Bitfields.cmx: Transform.cmx Machine.cmx Cutil.cmx C.cmi Bitfields.cmi
Builtins.cmo: Env.cmi Cutil.cmi C.cmi Builtins.cmi
@@ -59,31 +53,21 @@ Machine.cmx: Machine.cmi
Main.cmo: Parse.cmi GCC.cmi Cprint.cmi Builtins.cmi
Main.cmx: Parse.cmx GCC.cmx Cprint.cmx Builtins.cmx
PackedStructs.cmo: Transform.cmi Machine.cmi Errors.cmi Env.cmi Cutil.cmi \
- C.cmi Builtins.cmi PackedStructs.cmi
+ Cprint.cmi C.cmi Builtins.cmi
PackedStructs.cmx: Transform.cmx Machine.cmx Errors.cmx Env.cmx Cutil.cmx \
- C.cmi Builtins.cmx PackedStructs.cmi
-Parse.cmo: Unblock.cmi StructByValue.cmi StructAssign.cmi SimplVolatile.cmo \
- SimplExpr.cmi Rename.cmi PackedStructs.cmi Errors.cmi Elab.cmi \
- Bitfields.cmi AddCasts.cmi Parse.cmi
-Parse.cmx: Unblock.cmx StructByValue.cmx StructAssign.cmx SimplVolatile.cmx \
- SimplExpr.cmx Rename.cmx PackedStructs.cmx Errors.cmx Elab.cmx \
- Bitfields.cmx AddCasts.cmx Parse.cmi
+ Cprint.cmx C.cmi Builtins.cmx
+Parse.cmo: Unblock.cmi StructReturn.cmi Rename.cmi PackedStructs.cmo \
+ Errors.cmi Elab.cmi Bitfields.cmi Parse.cmi
+Parse.cmx: Unblock.cmx StructReturn.cmx Rename.cmx PackedStructs.cmx \
+ Errors.cmx Elab.cmx Bitfields.cmx Parse.cmi
Parse_aux.cmo: Errors.cmi Cabshelper.cmo Parse_aux.cmi
Parse_aux.cmx: Errors.cmx Cabshelper.cmx Parse_aux.cmi
Parser.cmo: Parse_aux.cmi Cabshelper.cmo Cabs.cmo Parser.cmi
Parser.cmx: Parse_aux.cmx Cabshelper.cmx Cabs.cmx Parser.cmi
Rename.cmo: Errors.cmi Cutil.cmi C.cmi Builtins.cmi Rename.cmi
Rename.cmx: Errors.cmx Cutil.cmx C.cmi Builtins.cmx Rename.cmi
-SimplExpr.cmo: Transform.cmi Errors.cmi Cutil.cmi C.cmi SimplExpr.cmi
-SimplExpr.cmx: Transform.cmx Errors.cmx Cutil.cmx C.cmi SimplExpr.cmi
-SimplVolatile.cmo: Transform.cmi Cutil.cmi C.cmi
-SimplVolatile.cmx: Transform.cmx Cutil.cmx C.cmi
-StructAssign.cmo: Transform.cmi Machine.cmi Errors.cmi Env.cmi Cutil.cmi \
- C.cmi StructAssign.cmi
-StructAssign.cmx: Transform.cmx Machine.cmx Errors.cmx Env.cmx Cutil.cmx \
- C.cmi StructAssign.cmi
-StructByValue.cmo: Transform.cmi Env.cmi Cutil.cmi C.cmi StructByValue.cmi
-StructByValue.cmx: Transform.cmx Env.cmx Cutil.cmx C.cmi StructByValue.cmi
+StructReturn.cmo: Transform.cmi Env.cmi Cutil.cmi C.cmi StructReturn.cmi
+StructReturn.cmx: Transform.cmx Env.cmx Cutil.cmx C.cmi StructReturn.cmi
Transform.cmo: Env.cmi Cutil.cmi C.cmi Builtins.cmi Transform.cmi
Transform.cmx: Env.cmx Cutil.cmx C.cmi Builtins.cmx Transform.cmi
Unblock.cmo: Transform.cmi Errors.cmi Cutil.cmi C.cmi Unblock.cmi
diff --git a/cparser/AddCasts.ml b/cparser/AddCasts.ml
deleted file mode 100644
index eb3fa08..0000000
--- a/cparser/AddCasts.ml
+++ /dev/null
@@ -1,243 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Materialize implicit casts *)
-
-(* Assumes: simplified code
- Produces: simplified code
- Preserves: unblocked code *)
-
-open C
-open Cutil
-open Transform
-
-(* We have the option of materializing all casts or leave "widening"
- casts implicit. Widening casts are:
-- from a small integer type to a larger integer type,
- provided both types have the same signedness;
-- from a small float type to a larger float type;
-- from a pointer type to void *.
-*)
-
-let omit_widening_casts = ref false
-
-let widening_cast env tfrom tto =
- begin match unroll env tfrom, unroll env tto with
- | TInt(k1, _), TInt(k2, _) ->
- integer_rank k1 <= integer_rank k2
- && is_signed_ikind k1 = is_signed_ikind k2
- | TFloat(k1, _), TFloat(k2, _) ->
- float_rank k1 <= float_rank k2
- | TPtr(ty1, _), TPtr(ty2, _) -> is_void_type env ty2
- | _, _ -> false
- end
-
-let cast_not_needed env tfrom tto =
- let tfrom = pointer_decay env tfrom
- and tto = pointer_decay env tto in
- compatible_types env tfrom tto
- || (!omit_widening_casts && widening_cast env tfrom tto)
-
-let cast env e tto =
- if cast_not_needed env e.etyp tto
- then e
- else {edesc = ECast(tto, e); etyp = tto}
-
-(* Note: this pass applies only to simplified expressions
- because casts cannot be materialized in op= expressions... *)
-
-let rec add_expr env e =
- match e.edesc with
- | EConst _ -> e
- | EVar _ -> e
- | ESizeof _ -> e
- | EUnop(op, e1) ->
- let e1' = add_expr env e1 in
- let desc =
- match op with
- | Ominus | Oplus | Onot ->
- EUnop(op, cast env e1' e.etyp)
- | Olognot | Oderef | Oaddrof
- | Odot _ | Oarrow _ ->
- EUnop(op, e1')
- | Opreincr | Opredecr | Opostincr | Opostdecr ->
- assert false (* not simplified *)
- in { edesc = desc; etyp = e.etyp }
- | EBinop(op, e1, e2, ty) ->
- let e1' = add_expr env e1 in
- let e2' = add_expr env e2 in
- let desc =
- match op with
- | Oadd ->
- if is_pointer_type env ty
- then EBinop(Oadd, e1', e2', ty)
- else EBinop(Oadd, cast env e1' ty, cast env e2' ty, ty)
- | Osub ->
- if is_pointer_type env ty
- then EBinop(Osub, e1', e2', ty)
- else EBinop(Osub, cast env e1' ty, cast env e2' ty, ty)
- | Omul|Odiv|Omod|Oand|Oor|Oxor|Oeq|One|Olt|Ogt|Ole|Oge ->
- EBinop(op, cast env e1' ty, cast env e2' ty, ty)
- | Oshl|Oshr ->
- EBinop(op, cast env e1' ty, e2', ty)
- | Oindex | Ologand | Ologor | Ocomma ->
- EBinop(op, e1', e2', ty)
- | Oassign
- | Oadd_assign|Osub_assign|Omul_assign|Odiv_assign|Omod_assign
- | Oand_assign|Oor_assign|Oxor_assign|Oshl_assign|Oshr_assign ->
- assert false (* not simplified *)
- in { edesc = desc; etyp = e.etyp }
- | EConditional(e1, e2, e3) ->
- { edesc =
- EConditional(add_expr env e1, add_expr env e2, add_expr env e3);
- etyp = e.etyp }
- | ECast(ty, e1) ->
- { edesc = ECast(ty, add_expr env e1); etyp = e.etyp }
- | ECall(e1, el) ->
- assert false (* not simplified *)
-
-(* Arguments to a prototyped function *)
-
-let rec add_proto env args params =
- match args, params with
- | [], _ -> []
- | _::_, [] -> add_noproto env args
- | arg1 :: argl, (_, ty_p) :: paraml ->
- cast env (add_expr env arg1) ty_p ::
- add_proto env argl paraml
-
-(* Arguments to a non-prototyped function *)
-
-and add_noproto env args =
- match args with
- | [] -> []
- | arg1 :: argl ->
- cast env (add_expr env arg1) (default_argument_conversion env arg1.etyp) ::
- add_noproto env argl
-
-(* Arguments to function calls in general *)
-
-let add_arguments env ty_fun args =
- let ty_args =
- match unroll env ty_fun with
- | TFun(res, args, vararg, a) -> args
- | TPtr(ty, a) ->
- begin match unroll env ty with
- | TFun(res, args, vararg, a) -> args
- | _ -> assert false
- end
- | _ -> assert false in
- match ty_args with
- | None -> add_noproto env args
- | Some targs -> add_proto env args targs
-
-(* Toplevel expressions (appearing in Sdo statements) *)
-
-let add_topexpr env loc e =
- match e.edesc with
- | EBinop(Oassign, lhs, {edesc = ECall(e1, el); etyp = ty}, _) ->
- let ecall =
- {edesc = ECall(add_expr env e1, add_arguments env e1.etyp el);
- etyp = ty} in
- if cast_not_needed env ty lhs.etyp then
- sassign loc (add_expr env lhs) ecall
- else begin
- let tmp = new_temp (erase_attributes_type env ty) in
- sseq loc (sassign loc tmp ecall)
- (sassign loc (add_expr env lhs) (cast env tmp lhs.etyp))
- end
- | EBinop(Oassign, lhs, rhs, _) ->
- sassign loc (add_expr env lhs) (cast env (add_expr env rhs) lhs.etyp)
- | ECall(e1, el) ->
- let ecall =
- {edesc = ECall(add_expr env e1, add_arguments env e1.etyp el);
- etyp = e.etyp} in
- {sdesc = Sdo ecall; sloc = loc}
- | _ ->
- assert false
-
-(* Initializers *)
-
-let rec add_init env tto = function
- | Init_single e ->
- Init_single (cast env (add_expr env e) tto)
- | Init_array il ->
- let ty_elt =
- match unroll env tto with
- | TArray(ty, _, _) -> ty | _ -> assert false in
- Init_array (List.map (add_init env ty_elt) il)
- | Init_struct(id, fil) ->
- Init_struct (id, List.map
- (fun (fld, i) -> (fld, add_init env fld.fld_typ i))
- fil)
- | Init_union(id, fld, i) ->
- Init_union(id, fld, add_init env fld.fld_typ i)
-
-(* Declarations *)
-
-let add_decl env (sto, id, ty, optinit) =
- (sto, id, ty,
- begin match optinit with
- | None -> None
- | Some init -> Some(add_init env ty init)
- end)
-
-(* Statements *)
-
-let rec add_stmt env f s =
- match s.sdesc with
- | Sskip -> s
- | Sdo e -> add_topexpr env s.sloc e
- | Sseq(s1, s2) ->
- {sdesc = Sseq(add_stmt env f s1, add_stmt env f s2); sloc = s.sloc }
- | Sif(e, s1, s2) ->
- {sdesc = Sif(add_expr env e, add_stmt env f s1, add_stmt env f s2);
- sloc = s.sloc}
- | Swhile(e, s1) ->
- {sdesc = Swhile(add_expr env e, add_stmt env f s1);
- sloc = s.sloc}
- | Sdowhile(s1, e) ->
- {sdesc = Sdowhile(add_stmt env f s1, add_expr env e);
- sloc = s.sloc}
- | Sfor(s1, e, s2, s3) ->
- {sdesc = Sfor(add_stmt env f s1, add_expr env e, add_stmt env f s2,
- add_stmt env f s3);
- sloc = s.sloc}
- | Sbreak -> s
- | Scontinue -> s
- | Sswitch(e, s1) ->
- {sdesc = Sswitch(add_expr env e, add_stmt env f s1); sloc = s.sloc}
- | Slabeled(lbl, s) ->
- {sdesc = Slabeled(lbl, add_stmt env f s); sloc = s.sloc}
- | Sgoto lbl -> s
- | Sreturn None -> s
- | Sreturn (Some e) ->
- {sdesc = Sreturn(Some(cast env (add_expr env e) f.fd_ret)); sloc = s.sloc}
- | Sblock sl ->
- {sdesc = Sblock(List.map (add_stmt env f) sl); sloc = s.sloc}
- | Sdecl d ->
- {sdesc = Sdecl(add_decl env d); sloc = s.sloc}
-
-let add_fundef env f =
- reset_temps();
- let body' = add_stmt env f f.fd_body in
- let temps = get_temps () in
- (* fd_locals have no initializers, so no need to transform them *)
- { f with fd_locals = f.fd_locals @ temps; fd_body = body' }
-
-
-let program ?(all = false) p =
- omit_widening_casts := not all;
- Transform.program ~decl:add_decl ~fundef:add_fundef p
diff --git a/cparser/AddCasts.mli b/cparser/AddCasts.mli
deleted file mode 100644
index 318ecc6..0000000
--- a/cparser/AddCasts.mli
+++ /dev/null
@@ -1,16 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-val program: ?all: bool -> C.program -> C.program
diff --git a/cparser/Elab.ml b/cparser/Elab.ml
index 870385d..2da1936 100644
--- a/cparser/Elab.ml
+++ b/cparser/Elab.ml
@@ -1508,13 +1508,13 @@ let rec enter_decdefs local loc env = function
(* check for incomplete type *)
if sto' <> Storage_extern && incomplete_type env ty' then
warning loc "'%s' has incomplete type" s;
- if local && sto <> Storage_extern && sto <> Storage_static then begin
+ if local && sto' <> Storage_extern && sto' <> Storage_static then begin
(* Local definition *)
let (decls, env3) = enter_decdefs local loc env2 rem in
((sto', id, ty', init') :: decls, env3)
end else begin
(* Global definition *)
- emit_elab (elab_loc loc) (Gdecl(sto, id, ty', init'));
+ emit_elab (elab_loc loc) (Gdecl(sto', id, ty', init'));
enter_decdefs local loc env2 rem
end
diff --git a/cparser/Makefile b/cparser/Makefile
index 527cdd3..da2c28b 100644
--- a/cparser/Makefile
+++ b/cparser/Makefile
@@ -15,8 +15,8 @@ SRCS=Errors.ml Cabs.ml Cabshelper.ml Parse_aux.ml Parser.ml Lexer.ml \
Builtins.ml GCC.ml \
Cleanup.ml Elab.ml Rename.ml \
Transform.ml \
- Unblock.ml SimplExpr.ml AddCasts.ml StructByValue.ml StructAssign.ml \
- Bitfields.ml PackedStructs.ml SimplVolatile.ml \
+ Unblock.ml StructReturn.ml \
+ Bitfields.ml PackedStructs.ml \
Parse.ml
COBJS=uint64.o
diff --git a/cparser/Parse.ml b/cparser/Parse.ml
index dcd01e9..2c467a7 100644
--- a/cparser/Parse.ml
+++ b/cparser/Parse.ml
@@ -20,28 +20,18 @@ module CharSet = Set.Make(struct type t = char let compare = compare end)
let transform_program t p =
let run_pass pass flag p = if CharSet.mem flag t then pass p else p in
Rename.program
- (run_pass (AddCasts.program ~all:(CharSet.mem 'C' t)) 'c'
- (run_pass (SimplExpr.program ~volatile:(CharSet.mem 'V' t)) 'e'
- (run_pass SimplVolatile.program 'v'
- (run_pass StructAssign.program 'S'
- (run_pass StructByValue.program 's'
+ (run_pass StructReturn.program 's'
(run_pass PackedStructs.program 'p'
(run_pass Bitfields.program 'f'
(run_pass Unblock.program 'b'
- p))))))))
+ p))))
let parse_transformations s =
let t = ref CharSet.empty in
let set s = String.iter (fun c -> t := CharSet.add c !t) s in
String.iter
(function 'b' -> set "b"
- | 'e' -> set "e"
- | 'c' -> set "ec"
- | 'C' -> set "ecC"
| 's' -> set "s"
- | 'S' -> set "bsS"
- | 'v' -> set "v"
- | 'V' -> set "eV"
| 'f' -> set "bf"
| 'p' -> set "bp"
| _ -> ())
diff --git a/cparser/SimplExpr.ml b/cparser/SimplExpr.ml
deleted file mode 100644
index 4184d95..0000000
--- a/cparser/SimplExpr.ml
+++ /dev/null
@@ -1,568 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Pulling side-effects out of expressions *)
-
-(* Assumes: nothing
- Produces: simplified code *)
-
-open C
-open Cutil
-open Transform
-
-(* Grammar of simplified expressions:
- e ::= EConst cst
- | ESizeof ty
- | EVar id
- | EUnop pure-unop e
- | EBinop pure-binop e e
- | EConditional e e e
- | ECast ty e
-
- Grammar of statements produced to reflect side-effects in expressions:
- s ::= Sskip
- | Sdo (EBinop Oassign e e)
- | Sdo (EBinop Oassign e (ECall e e* ))
- | Sdo (Ecall e el)
- | Sseq s s
- | Sif e s s
-*)
-
-let rec is_simpl_expr e =
- match e.edesc with
- | EConst cst -> true
- | ESizeof ty -> true
- | EVar id -> true
- | EUnop((Ominus|Oplus|Olognot|Onot|Oderef|Oaddrof), e1) ->
- is_simpl_expr e1
- | EBinop((Oadd|Osub|Omul|Odiv|Omod|Oand|Oor|Oxor|Oshl|Oshr|
- Oeq|One|Olt|Ogt|Ole|Oge|Oindex|Ologand|Ologor), e1, e2, _) ->
- is_simpl_expr e1 && is_simpl_expr e2
- | EConditional(e1, e2, e3) ->
- is_simpl_expr e1 && is_simpl_expr e2 && is_simpl_expr e3
- | ECast(ty, e1) ->
- is_simpl_expr e1
- | _ ->
- false
-
-(* "Destination" of a simplified expression *)
-
-type exp_destination =
- | RHS (* evaluate as a r-value *)
- | LHS (* evaluate as a l-value *)
- | Drop (* drop its value, we only need the side-effects *)
- | Set of exp (* assign it to the given simplified l.h.s. *)
-
-let voidconst = { nullconst with etyp = TVoid [] }
-
-(* Reads from volatile lvalues are also considered as side-effects if
- [volatilize] is true. *)
-
-let volatilize = ref false
-
-(* [simpl_expr loc env e act] returns a pair [s, e'] of
- a statement that performs the side-effects present in [e] and
- a simplified, side-effect-free expression [e'].
- If [act] is [RHS], [e'] evaluates to the same value as [e].
- If [act] is [LHS], [e'] evaluates to the same location as [e].
- If [act] is [Drop], [e'] is not meaningful and must be ignored.
- If [act] is [Set lhs], [s] also performs an assignment
- equivalent to [lhs = e]. [e'] is not meaningful. *)
-
-let simpl_expr loc env e act =
-
- (* Temporaries should not be [const] because we assign into them,
- and need not be [volatile] because no one else is writing into them.
- As for [restrict] it doesn't make sense anyway. *)
-
- let new_temp ty =
- Transform.new_temp (erase_attributes_type env ty) in
-
- let eboolvalof e =
- { edesc = EBinop(One, e, intconst 0L IInt, TInt(IInt, []));
- etyp = TInt(IInt, []) } in
-
- let sseq s1 s2 = Cutil.sseq loc s1 s2 in
-
- let sassign e1 e2 =
- { sdesc = Sdo {edesc = EBinop(Oassign, e1, e2, e1.etyp); etyp = e1.etyp};
- sloc = loc } in
-
- let sif e s1 s2 =
- { sdesc = Sif(e, s1, s2); sloc = loc } in
-
- let is_volatile_read e =
- !volatilize
- && List.mem AVolatile (attributes_of_type env e.etyp)
- && is_lvalue e in
-
- let lhs_to_rhs e =
- if is_volatile_read e
- then (let t = new_temp e.etyp in (sassign t e, t))
- else (sskip, e) in
-
- let finish act s e =
- match act with
- | RHS ->
- if is_volatile_read e
- then (let t = new_temp e.etyp in (sseq s (sassign t e), t))
- else (s, e)
- | LHS ->
- (s, e)
- | Drop ->
- if is_volatile_read e
- then (let t = new_temp e.etyp in (sseq s (sassign t e), voidconst))
- else (s, voidconst)
- | Set lhs ->
- if is_volatile_read e
- then (let t = new_temp e.etyp in
- (sseq s (sseq (sassign t e) (sassign lhs t)), voidconst))
- else (sseq s (sassign lhs e), voidconst) in
-
- let rec simpl e act =
- match e.edesc with
-
- | EConst cst ->
- finish act sskip e
-
- | ESizeof ty ->
- finish act sskip e
-
- | EVar id ->
- finish act sskip e
-
- | EUnop(op, e1) ->
-
- begin match op with
-
- | Ominus | Oplus | Olognot | Onot | Oderef | Oarrow _ ->
- let (s1, e1') = simpl e1 RHS in
- finish act s1 {edesc = EUnop(op, e1'); etyp = e.etyp}
-
- | Oaddrof ->
- let (s1, e1') = simpl e1 LHS in
- finish act s1 {edesc = EUnop(op, e1'); etyp = e.etyp}
-
- | Odot _ ->
- let (s1, e1') = simpl e1 (if act = LHS then LHS else RHS) in
- finish act s1 {edesc = EUnop(op, e1'); etyp = e.etyp}
-
- | Opreincr | Opredecr ->
- let (s1, e1') = simpl e1 LHS in
- let (s2, e2') = lhs_to_rhs e1' in
- let op' = match op with Opreincr -> Oadd | _ -> Osub in
- let ty = unary_conversion env e.etyp in
- let e3 =
- {edesc = EBinop(op', e2', intconst 1L IInt, ty); etyp = ty} in
- begin match act with
- | Drop ->
- (sseq s1 (sseq s2 (sassign e1' e3)), voidconst)
- | _ ->
- let tmp = new_temp e.etyp in
- finish act (sseq s1 (sseq s2 (sseq (sassign tmp e3)
- (sassign e1' tmp))))
- tmp
- end
-
- | Opostincr | Opostdecr ->
- let (s1, e1') = simpl e1 LHS in
- let op' = match op with Opostincr -> Oadd | _ -> Osub in
- let ty = unary_conversion env e.etyp in
- begin match act with
- | Drop ->
- let (s2, e2') = lhs_to_rhs e1' in
- let e3 =
- {edesc = EBinop(op', e2', intconst 1L IInt, ty); etyp = ty} in
- (sseq s1 (sseq s2 (sassign e1' e3)), voidconst)
- | _ ->
- let tmp = new_temp e.etyp in
- let e3 =
- {edesc = EBinop(op', tmp, intconst 1L IInt, ty); etyp = ty} in
- finish act (sseq s1 (sseq (sassign tmp e1') (sassign e1' e3)))
- tmp
- end
-
- end
-
- | EBinop(op, e1, e2, ty) ->
-
- begin match op with
-
- | Oadd | Osub | Omul | Odiv | Omod | Oand | Oor | Oxor
- | Oshl | Oshr | Oeq | One | Olt | Ogt | Ole | Oge | Oindex ->
- let (s1, e1') = simpl e1 RHS in
- let (s2, e2') = simpl e2 RHS in
- finish act (sseq s1 s2)
- {edesc = EBinop(op, e1', e2', ty); etyp = e.etyp}
-
- | Oassign ->
- if act = Drop && is_simpl_expr e1 then
- simpl e2 (Set e1)
- else begin
- match act with
- | Drop ->
- let (s1, e1') = simpl e1 LHS in
- let (s2, e2') = simpl e2 RHS in
- (sseq s1 (sseq s2 (sassign e1' e2')), voidconst)
- | _ ->
- let tmp = new_temp e.etyp in
- let (s1, e1') = simpl e1 LHS in
- let (s2, e2') = simpl e2 (Set tmp) in
- finish act (sseq s1 (sseq s2 (sassign e1' tmp)))
- tmp
- end
-
- | Oadd_assign | Osub_assign | Omul_assign | Odiv_assign
- | Omod_assign | Oand_assign | Oor_assign | Oxor_assign
- | Oshl_assign | Oshr_assign ->
- let (s1, e1') = simpl e1 LHS in
- let (s11, e11') = lhs_to_rhs e1' in
- let (s2, e2') = simpl e2 RHS in
- let op' =
- match op with
- | Oadd_assign -> Oadd | Osub_assign -> Osub
- | Omul_assign -> Omul | Odiv_assign -> Odiv
- | Omod_assign -> Omod | Oand_assign -> Oand
- | Oor_assign -> Oor | Oxor_assign -> Oxor
- | Oshl_assign -> Oshl | Oshr_assign -> Oshr
- | _ -> assert false in
- let e3 =
- { edesc = EBinop(op', e11', e2', ty); etyp = ty } in
- begin match act with
- | Drop ->
- (sseq s1 (sseq s11 (sseq s2 (sassign e1' e3))), voidconst)
- | _ ->
- let tmp = new_temp e.etyp in
- finish act (sseq s1 (sseq s11 (sseq s2
- (sseq (sassign tmp e3) (sassign e1' tmp)))))
- tmp
- end
-
- | Ocomma ->
- let (s1, _) = simpl e1 Drop in
- let (s2, e2') = simpl e2 act in
- (sseq s1 s2, e2')
-
- | Ologand ->
- let (s1, e1') = simpl e1 RHS in
- if is_simpl_expr e2 then begin
- finish act s1
- {edesc = EBinop(Ologand, e1', e2, ty); etyp = e.etyp}
- end else begin
- match act with
- | Drop ->
- let (s2, _) = simpl e2 Drop in
- (sseq s1 (sif e1' s2 sskip), voidconst)
- | RHS | LHS -> (* LHS should not happen *)
- let (s2, e2') = simpl e2 RHS in
- let tmp = new_temp e.etyp in
- (sseq s1 (sif e1'
- (sseq s2 (sassign tmp (eboolvalof e2')))
- (sassign tmp (intconst 0L IInt))),
- tmp)
- | Set lv ->
- let (s2, e2') = simpl e2 RHS in
- (sseq s1 (sif e1'
- (sseq s2 (sassign lv (eboolvalof e2')))
- (sassign lv (intconst 0L IInt))),
- voidconst)
- end
-
- | Ologor ->
- let (s1, e1') = simpl e1 RHS in
- if is_simpl_expr e2 then begin
- finish act s1
- {edesc = EBinop(Ologor, e1', e2, ty); etyp = e.etyp}
- end else begin
- match act with
- | Drop ->
- let (s2, _) = simpl e2 Drop in
- (sseq s1 (sif e1' sskip s2), voidconst)
- | RHS | LHS -> (* LHS should not happen *)
- let (s2, e2') = simpl e2 RHS in
- let tmp = new_temp e.etyp in
- (sseq s1 (sif e1'
- (sassign tmp (intconst 1L IInt))
- (sseq s2 (sassign tmp (eboolvalof e2')))),
- tmp)
- | Set lv ->
- let (s2, e2') = simpl e2 RHS in
- (sseq s1 (sif e1'
- (sassign lv (intconst 1L IInt))
- (sseq s2 (sassign lv (eboolvalof e2')))),
- voidconst)
- end
-
- end
-
- | EConditional(e1, e2, e3) ->
- let (s1, e1') = simpl e1 RHS in
- if is_simpl_expr e2 && is_simpl_expr e3 then begin
- finish act s1 {edesc = EConditional(e1', e2, e3); etyp = e.etyp}
- end else begin
- match act with
- | Drop ->
- let (s2, _) = simpl e2 Drop in
- let (s3, _) = simpl e3 Drop in
- (sseq s1 (sif e1' s2 s3), voidconst)
- | RHS | LHS -> (* LHS should not happen *)
- let tmp = new_temp e.etyp in
- let (s2, _) = simpl e2 (Set tmp) in
- let (s3, _) = simpl e3 (Set tmp) in
- (sseq s1 (sif e1' s2 s3), tmp)
- | Set lv ->
- let (s2, _) = simpl e2 (Set lv) in
- let (s3, _) = simpl e3 (Set lv) in
- (sseq s1 (sif e1' s2 s3), voidconst)
- end
-
- | ECast(ty, e1) ->
- if is_void_type env ty then begin
- if act <> Drop then
- Errors.warning "%acast to 'void' in a context expecting a value\n"
- formatloc loc;
- simpl e1 act
- end else begin
- let (s1, e1') = simpl e1 RHS in
- finish act s1 {edesc = ECast(ty, e1'); etyp = e.etyp}
- end
-
- | ECall(e1, el) ->
- let (s1, e1') = simpl e1 RHS in
- let (s2, el') = simpl_list el in
- let e2 = { edesc = ECall(e1', el'); etyp = e.etyp } in
- begin match act with
- | Drop ->
- (sseq s1 (sseq s2 {sdesc = Sdo e2; sloc=loc}), voidconst)
- | Set({edesc = EVar _} as lhs) ->
- (* CompCert wants the destination of a call to be a variable,
- not a more complex lhs. In the latter case, we
- fall through the catch-all case below *)
- (sseq s1 (sseq s2 (sassign lhs e2)), voidconst)
- | _ ->
- let tmp = new_temp e.etyp in
- finish act (sseq s1 (sseq s2 (sassign tmp e2))) tmp
- end
-
- and simpl_list = function
- | [] -> (sskip, [])
- | e1 :: el ->
- let (s1, e1') = simpl e1 RHS in
- let (s2, el') = simpl_list el in
- (sseq s1 s2, e1' :: el')
-
- in simpl e act
-
-(* Simplification of an initializer *)
-
-let simpl_initializer loc env i =
-
- let rec simpl_init = function
- | Init_single e ->
- let (s, e') = simpl_expr loc env e RHS in
- (s, Init_single e)
- | Init_array il ->
- let rec simpl = function
- | [] -> (sskip, [])
- | i1 :: il ->
- let (s1, i1') = simpl_init i1 in
- let (s2, il') = simpl il in
- (sseq loc s1 s2, i1' :: il') in
- let (s, il') = simpl il in
- (s, Init_array il')
- | Init_struct(id, il) ->
- let rec simpl = function
- | [] -> (sskip, [])
- | (f1, i1) :: il ->
- let (s1, i1') = simpl_init i1 in
- let (s2, il') = simpl il in
- (sseq loc s1 s2, (f1, i1') :: il') in
- let (s, il') = simpl il in
- (s, Init_struct(id, il'))
- | Init_union(id, f, i) ->
- let (s, i') = simpl_init i in
- (s, Init_union(id, f, i'))
-
- in simpl_init i
-
-(* Construct a simplified statement equivalent to [if (e) s1; else s2;].
- Optimizes the case where e contains [&&] or [||] or [?].
- [s1] or [s2] can be duplicated, so use only for small [s1] and [s2]
- that do not define any labels. *)
-
-let rec simpl_if loc env e s1 s2 =
- match e.edesc with
- | EUnop(Olognot, e1) ->
- simpl_if loc env e1 s2 s1
- | EBinop(Ologand, e1, e2, _) ->
- simpl_if loc env e1
- (simpl_if loc env e2 s1 s2)
- s2
- | EBinop(Ologor, e1, e2, _) ->
- simpl_if loc env e1
- s1
- (simpl_if loc env e2 s1 s2)
- | EConditional(e1, e2, e3) ->
- simpl_if loc env e1
- (simpl_if loc env e2 s1 s2)
- (simpl_if loc env e3 s1 s2)
- | _ ->
- let (s, e') = simpl_expr loc env e RHS in
- sseq loc s {sdesc = Sif(e', s1, s2); sloc = loc}
-
-(* Trivial statements for which [simpl_if] is applicable *)
-
-let trivial_stmt s =
- match s.sdesc with
- | Sskip | Scontinue | Sbreak | Sgoto _ -> true
- | _ -> false
-
-(* Construct a simplified statement equivalent to [if (!e) exit; ]. *)
-
-let break_if_false loc env e =
- simpl_if loc env e
- {sdesc = Sskip; sloc = loc}
- {sdesc = Sbreak; sloc = loc}
-
-(* Simplification of a statement *)
-
-let simpl_statement env s =
-
- let rec simpl_stmt s =
- match s.sdesc with
-
- | Sskip ->
- s
-
- | Sdo e ->
- let (s', _) = simpl_expr s.sloc env e Drop in
- s'
-
- | Sseq(s1, s2) ->
- {sdesc = Sseq(simpl_stmt s1, simpl_stmt s2); sloc = s.sloc}
-
- | Sif(e, s1, s2) ->
- if trivial_stmt s1 && trivial_stmt s2 then
- simpl_if s.sloc env e (simpl_stmt s1) (simpl_stmt s2)
- else begin
- let (s', e') = simpl_expr s.sloc env e RHS in
- sseq s.sloc s'
- {sdesc = Sif(e', simpl_stmt s1, simpl_stmt s2);
- sloc = s.sloc}
- end
-
- | Swhile(e, s1) ->
- if is_simpl_expr e then
- {sdesc = Swhile(e, simpl_stmt s1); sloc = s.sloc}
- else
- {sdesc =
- Swhile(intconst 1L IInt,
- sseq s.sloc (break_if_false s.sloc env e)
- (simpl_stmt s1));
- sloc = s.sloc}
-
- | Sdowhile(s1, e) ->
- if is_simpl_expr e then
- {sdesc = Sdowhile(simpl_stmt s1, e); sloc = s.sloc}
- else begin
- let tmp = new_temp (TInt(IInt, [])) in
- let (s', _) = simpl_expr s.sloc env e (Set tmp) in
- let s_init =
- {sdesc = Sdo {edesc = EBinop(Oassign, tmp, intconst 1L IInt, tmp.etyp);
- etyp = tmp.etyp};
- sloc = s.sloc} in
- {sdesc = Sfor(s_init, tmp, s', simpl_stmt s1); sloc = s.sloc}
- end
-(*** Alternate translation that unfortunately puts a "break" in the
- "next" part of a "for", something that is not supported
- by Clight semantics, and has unknown behavior in gcc.
- {sdesc =
- Sfor(sskip,
- intconst 1L IInt,
- break_if_false s.sloc env e,
- simpl_stmt s1);
- sloc = s.sloc}
-***)
-
- | Sfor(s1, e, s2, s3) ->
- if is_simpl_expr e then
- {sdesc = Sfor(simpl_stmt s1,
- e,
- simpl_stmt s2,
- simpl_stmt s3);
- sloc = s.sloc}
- else
- let (s', e') = simpl_expr s.sloc env e RHS in
- {sdesc = Sfor(sseq s.sloc (simpl_stmt s1) s',
- e',
- sseq s.sloc (simpl_stmt s2) s',
- simpl_stmt s3);
- sloc = s.sloc}
-
- | Sbreak ->
- s
- | Scontinue ->
- s
-
- | Sswitch(e, s1) ->
- let (s', e') = simpl_expr s.sloc env e RHS in
- sseq s.sloc s' {sdesc = Sswitch(e', simpl_stmt s1); sloc = s.sloc}
-
- | Slabeled(lbl, s1) ->
- {sdesc = Slabeled(lbl, simpl_stmt s1); sloc = s.sloc}
-
- | Sgoto lbl ->
- s
-
- | Sreturn None ->
- s
-
- | Sreturn (Some e) ->
- let (s', e') = simpl_expr s.sloc env e RHS in
- sseq s.sloc s' {sdesc = Sreturn(Some e'); sloc = s.sloc}
-
- | Sblock sl ->
- {sdesc = Sblock(simpl_block sl); sloc = s.sloc}
-
- | Sdecl d -> assert false
-
- and simpl_block = function
- | [] -> []
- | ({sdesc = Sdecl(sto, id, ty, None)} as s) :: sl ->
- s :: simpl_block sl
- | ({sdesc = Sdecl(sto, id, ty, Some i)} as s) :: sl ->
- let (s', i') = simpl_initializer s.sloc env i in
- let sl' =
- {sdesc = Sdecl(sto, id, ty, Some i'); sloc = s.sloc}
- :: simpl_block sl in
- if s'.sdesc = Sskip then sl' else s' :: sl'
- | s :: sl ->
- simpl_stmt s :: simpl_block sl
-
- in simpl_stmt s
-
-(* Simplification of a function definition *)
-
-let simpl_fundef env f =
- reset_temps();
- let body' = simpl_statement env f.fd_body in
- let temps = get_temps() in
- { f with fd_locals = f.fd_locals @ temps; fd_body = body' }
-
-(* Entry point *)
-
-let program ?(volatile = false) p =
- volatilize := volatile;
- Transform.program ~fundef:simpl_fundef p
diff --git a/cparser/SimplExpr.mli b/cparser/SimplExpr.mli
deleted file mode 100644
index cdeb30c..0000000
--- a/cparser/SimplExpr.mli
+++ /dev/null
@@ -1,20 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Pulling side effects out of expressions.
- If [volatile] is [true], treats reads from volatile rvalues
- as side-effects *)
-
-val program: ?volatile: bool -> C.program -> C.program
diff --git a/cparser/SimplVolatile.ml b/cparser/SimplVolatile.ml
deleted file mode 100644
index ef7a3a0..0000000
--- a/cparser/SimplVolatile.ml
+++ /dev/null
@@ -1,81 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Elimination of read-modify-write operators (+=, ++, etc) applied
- to volatile expressions. *)
-
-open Printf
-open C
-open Cutil
-open Transform
-
-(* Rewriting of expressions *)
-
-let transf_expr loc env ctx e =
-
- let is_volatile ty =
- List.mem AVolatile (attributes_of_type env ty) in
-
- let rec texp ctx e =
- match e.edesc with
- | EConst _ -> e
- | ESizeof _ -> e
- | EVar _ -> e
- | EUnop((Opreincr|Opredecr as op), e1) when is_volatile e1.etyp ->
- expand_preincrdecr ~read:(fun e -> e) ~write:eassign
- env ctx op (texp Val e1)
- | EUnop((Opostincr|Opostdecr as op), e1) when is_volatile e1.etyp ->
- expand_postincrdecr ~read:(fun e -> e) ~write:eassign
- env ctx op (texp Val e1)
- | EUnop(op, e1) ->
- {edesc = EUnop(op, texp Val e1); etyp = e.etyp}
- | EBinop(Oassign, e1, e2, ty) when is_volatile e1.etyp ->
- expand_assign ~write:eassign env ctx (texp Val e1) (texp Val e2)
- | EBinop((Oadd_assign | Osub_assign | Omul_assign
- | Odiv_assign | Omod_assign
- | Oand_assign | Oor_assign | Oxor_assign
- | Oshl_assign | Oshr_assign) as op, e1, e2, ty)
- when is_volatile e1.etyp ->
- expand_assignop ~read:(fun e -> e) ~write:eassign
- env ctx op (texp Val e1) (texp Val e2) ty
- | EBinop(Ocomma, e1, e2, ty) ->
- {edesc = EBinop(Ocomma, texp Effects e1, texp ctx e2, ty);
- etyp = e.etyp}
- | EBinop(op, e1, e2, ty) ->
- {edesc = EBinop(op, texp Val e1, texp Val e2, ty); etyp = e.etyp}
- | EConditional(e1, e2, e3) ->
- {edesc = EConditional(texp Val e1, texp ctx e2, texp ctx e3);
- etyp = e.etyp}
- | ECast(ty, e1) ->
- {edesc = ECast(ty, texp Val e1); etyp = e.etyp}
- | ECall(e1, el) ->
- {edesc = ECall(texp Val e1, List.map (texp Val) el); etyp = e.etyp}
-
- in texp ctx e
-
-(* Statements *)
-
-let transf_stmt env s =
- Transform.stmt transf_expr env s
-
-(* Functions *)
-
-let transf_fundef env f =
- Transform.fundef transf_stmt env f
-
-(* Programs *)
-
-let program p =
- Transform.program ~fundef:transf_fundef p
diff --git a/cparser/StructAssign.ml b/cparser/StructAssign.ml
deleted file mode 100644
index d9ad8f5..0000000
--- a/cparser/StructAssign.ml
+++ /dev/null
@@ -1,165 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Expand assignments between structs and between unions *)
-
-(* Assumes: unblocked code.
- Preserves: unblocked code *)
-
-open C
-open Machine
-open Cutil
-open Env
-open Errors
-open Transform
-
-(* Finding appropriate memcpy functions *)
-
-let memcpy_decl = ref (None : ident option)
-
-let memcpy_type =
- TFun(TPtr(TVoid [], []),
- Some [(Env.fresh_ident "", TPtr(TVoid [], []));
- (Env.fresh_ident "", TPtr(TVoid [AConst], []));
- (Env.fresh_ident "", TInt(size_t_ikind, []))],
- false, [])
-
-let lookup_function env name =
- match Env.lookup_ident env name with
- | (id, II_ident(sto, ty)) -> (id, ty)
- | (id, II_enum _) -> raise (Env.Error(Env.Unbound_identifier name))
-
-let default_memcpy () =
- match !memcpy_decl with
- | Some id ->
- (id, memcpy_type)
- | None ->
- let id = Env.fresh_ident "memcpy" in
- memcpy_decl := Some id;
- (id, memcpy_type)
-
-let find_memcpy env =
- try
- (lookup_function env "__builtin_memcpy_aligned", true)
- with Env.Error _ ->
- try
- (lookup_function env "__builtin_memcpy", false)
- with Env.Error _ ->
- try
- (lookup_function env "memcpy", false)
- with Env.Error _ ->
- (default_memcpy(), false)
-
-(* Smart constructor that "bubble up" sequence expressions *)
-
-let rec edot f e ty =
- match e.edesc with
- | EBinop(Ocomma, e1, e2, _) -> ecomma e1 (edot f e2 ty)
- | _ -> { edesc = EUnop(Odot f, e); etyp = ty }
-
-(* Translate an assignment [lhs = rhs] between composite types.
- [lhs] and [rhs] must be l-values. *)
-
-let transf_assign env lhs rhs =
- let al =
- match Cutil.alignof env lhs.etyp with Some al -> al | None -> 1 in
- let ((ident, ty), four_args) = find_memcpy env in
- let memcpy = {edesc = EVar(ident); etyp = ty} in
- let e_lhs = eaddrof lhs
- and e_rhs = eaddrof rhs
- and e_size = {edesc = ESizeof(lhs.etyp); etyp = TInt(size_t_ikind, [])}
- and e_align = intconst (Int64.of_int al) size_t_ikind in
- let args =
- if four_args
- then [e_lhs; e_rhs; e_size; e_align]
- else [e_lhs; e_rhs; e_size] in
- {edesc = ECall(memcpy, args); etyp = TVoid[]}
-
-(* Transformation of expressions. *)
-
-let transf_expr loc env ctx e =
- let rec texp ctx e =
- match e.edesc with
- | EBinop(Oassign, lhs, rhs, _) when is_composite_type env lhs.etyp ->
- let lhs' = texp Val lhs in
- let rhs' = texp Val rhs in
- begin match ctx with
- | Effects ->
- transf_assign env lhs' rhs'
- | Val ->
- bind_lvalue env lhs' (fun l -> ecomma (transf_assign env l rhs') l)
- end
- | EConst c -> e
- | ESizeof ty -> e
- | EVar x ->
- if ctx = Effects && is_composite_type env e.etyp
- then nullconst
- else e
- | EUnop(Oaddrof, e1) ->
- eaddrof (texp Val e1)
- | EUnop(Oderef, e1) ->
- if ctx = Effects && is_composite_type env e.etyp
- then texp Effects e1
- else {edesc = EUnop(Oderef, texp Val e1); etyp = e.etyp}
- | EUnop(Odot f, e1) ->
- if ctx = Effects && is_composite_type env e.etyp
- then texp Effects e1
- else edot f (texp Val e1) e.etyp
- | EUnop(Oarrow f, e1) ->
- if ctx = Effects && is_composite_type env e.etyp
- then texp Effects e1
- else {edesc = EUnop(Oarrow f, texp Val e1); etyp = e.etyp}
- | EUnop(op, e1) ->
- {edesc = EUnop(op, texp Val e1); etyp = e.etyp}
- | EBinop(Oindex, e1, e2, ty) ->
- if ctx = Effects && is_composite_type env e.etyp
- then ecomma (texp Effects e1) (texp Effects e2)
- else {edesc = EBinop(Oindex, texp Val e1, texp Val e2, ty); etyp = e.etyp}
- | EBinop(Ocomma, e1, e2, ty) ->
- ecomma (texp Effects e1) (texp ctx e2)
- | EBinop(op, e1, e2, ty) ->
- {edesc = EBinop(op, texp Val e1, texp Val e2, ty);
- etyp = e.etyp}
- | EConditional(e1, e2, e3) ->
- {edesc = EConditional(texp Val e1, texp ctx e2, texp ctx e3);
- etyp = e.etyp}
- | ECast(ty, e1) ->
- {edesc = ECast(ty, texp Val e1); etyp = e.etyp}
- | ECall(e1, el) ->
- {edesc = ECall(texp Val e1,
- List.map (texp Val) el);
- etyp = e.etyp}
- in texp ctx e
-
-(* Transformation of statements *)
-
-let transf_stmt env s =
- Transform.stmt transf_expr env s
-
-(* Transformation of function definitions *)
-
-let transf_fundef env f =
- Transform.fundef transf_stmt env f
-
-(* Transformation of programs *)
-
-let program p =
- memcpy_decl := None;
- let p' = Transform.program ~fundef:transf_fundef p in
- match !memcpy_decl with
- | None -> p'
- | Some id ->
- {gdesc = Gdecl(Storage_extern, id, memcpy_type, None); gloc = no_loc}
- :: p'
diff --git a/cparser/StructAssign.mli b/cparser/StructAssign.mli
deleted file mode 100644
index 5549282..0000000
--- a/cparser/StructAssign.mli
+++ /dev/null
@@ -1,18 +0,0 @@
-(* *********************************************************************)
-(* *)
-(* The Compcert verified compiler *)
-(* *)
-(* Xavier Leroy, INRIA Paris-Rocquencourt *)
-(* *)
-(* Copyright Institut National de Recherche en Informatique et en *)
-(* Automatique. All rights reserved. This file is distributed *)
-(* under the terms of the GNU General Public License as published by *)
-(* the Free Software Foundation, either version 2 of the License, or *)
-(* (at your option) any later version. This file is also distributed *)
-(* under the terms of the INRIA Non-Commercial License Agreement. *)
-(* *)
-(* *********************************************************************)
-
-(* Expand assignments between structs and between unions *)
-
-val program: C.program -> C.program
diff --git a/cparser/StructByValue.ml b/cparser/StructReturn.ml
index 1b74ec5..2a4bbc1 100644
--- a/cparser/StructByValue.ml
+++ b/cparser/StructReturn.ml
@@ -13,17 +13,14 @@
(* *)
(* *********************************************************************)
-(* Eliminate by-value passing of structs and unions. *)
-
-(* Assumes: nothing.
- Preserves: unblocked code *)
+(* Eliminate structs and unions being returned by value as function results *)
+(* This is a simpler special case of [StructByValue]. *)
open C
open Cutil
open Transform
-(* In function argument types, struct s -> const struct s *
- In function result types, struct s -> void + add 1st parameter struct s *
+(* In function result types, struct s -> void + add 1st parameter struct s *
Try to preserve original typedef names when no change.
*)
@@ -49,11 +46,7 @@ let rec transf_type env t =
if t1' = t1 then t else TArray(transf_type env t1, sz, attr)
| _ -> t
-and transf_funarg env (id, t) =
- let t = transf_type env t in
- if is_composite_type env t
- then (id, TPtr(add_attributes_type [AConst] t, []))
- else (id, t)
+and transf_funarg env (id, t) = (id, transf_type env t)
(* Expressions: transform calls + rewrite the types *)
@@ -92,26 +85,21 @@ let rec transf_expr env ctx e =
if is_composite_type env e.etyp then
transf_composite_call env ctx None fn args e.etyp
else
- {edesc = ECall(transf_expr env Val fn, List.map (transf_arg env) args);
+ {edesc = ECall(transf_expr env Val fn,
+ List.map (transf_expr env Val) args);
etyp = newty}
-(* Function arguments: pass by reference those having composite type *)
-
-and transf_arg env e =
- let e' = transf_expr env Val e in
- if is_composite_type env e'.etyp then eaddrof e' else e'
-
(* Function calls returning a composite: add first argument.
- ctx = Effects: lv = f(...) -> f(&lv, ...)
+ ctx = Effects: lv = f(...) -> f(&lv, ...) [copy optimization]
f(...) -> f(&newtemp, ...)
- ctx = Val: lv = f(...) -> f(&newtemp, ...), lv = newtemp, newtemp
+ ctx = Val: lv = f(...) -> f(&newtemp, ...), lv = newtemp
f(...) -> f(&newtemp, ...), newtemp
*)
and transf_composite_call env ctx opt_lhs fn args ty =
let ty = transf_type env ty in
let fn = transf_expr env Val fn in
- let args = List.map (transf_arg env) args in
+ let args = List.map (transf_expr env Val) args in
match ctx, opt_lhs with
| Effects, None ->
let tmp = new_temp ~name:"_res" ty in
@@ -125,54 +113,14 @@ and transf_composite_call env ctx opt_lhs fn args ty =
| Val, Some lhs ->
let lhs = transf_expr env Val lhs in
let tmp = new_temp ~name:"_res" ty in
- ecomma (ecomma {edesc = ECall(fn, eaddrof tmp :: args); etyp = TVoid []}
- (eassign lhs tmp))
- tmp
-
-(* The transformation above can create ill-formed lhs containing ",", as in
- f().x = y ---> (f(&tmp), tmp).x = y
- f(g(x)); ---> f(&(g(&tmp),tmp))
- We fix this by floating the "," above the lhs, up to the nearest enclosing
- rhs:
- f().x = y ---> (f(&tmp), tmp).x = y --> f(&tmp), tmp.x = y
- f(g(x)); ---> f(&(g(&tmp),tmp)) --> f((g(&tmp), &tmp))
-*)
-
-let rec float_comma e =
- match e.edesc with
- | EConst c -> e
- | ESizeof ty -> e
- | EVar x -> e
- (* lvalue-consuming unops *)
- | EUnop((Oaddrof|Opreincr|Opredecr|Opostincr|Opostdecr|Odot _) as op,
- {edesc = EBinop(Ocomma, e1, e2, _)}) ->
- ecomma (float_comma e1)
- (float_comma {edesc = EUnop(op, e2); etyp = e.etyp})
- (* lvalue-consuming binops *)
- | EBinop((Oassign|Oadd_assign|Osub_assign|Omul_assign|Odiv_assign
- |Omod_assign|Oand_assign|Oor_assign|Oxor_assign
- |Oshl_assign|Oshr_assign) as op,
- {edesc = EBinop(Ocomma, e1, e2, _)}, e3, tyres) ->
- ecomma (float_comma e1)
- (float_comma {edesc = EBinop(op, e2, e3, tyres); etyp = e.etyp})
- (* other expressions *)
- | EUnop(op, e1) ->
- {edesc = EUnop(op, float_comma e1); etyp = e.etyp}
- | EBinop(op, e1, e2, tyres) ->
- {edesc = EBinop(op, float_comma e1, float_comma e2, tyres); etyp = e.etyp}
- | EConditional(e1, e2, e3) ->
- {edesc = EConditional(float_comma e1, float_comma e2, float_comma e3);
- etyp = e.etyp}
- | ECast(ty, e1) ->
- {edesc = ECast(ty, float_comma e1); etyp = e.etyp}
- | ECall(e1, el) ->
- {edesc = ECall(float_comma e1, List.map float_comma el); etyp = e.etyp}
+ ecomma {edesc = ECall(fn, eaddrof tmp :: args); etyp = TVoid []}
+ (eassign lhs tmp)
(* Initializers *)
let rec transf_init env = function
| Init_single e ->
- Init_single (float_comma(transf_expr env Val e))
+ Init_single (transf_expr env Val e)
| Init_array il ->
Init_array (List.map (transf_init env) il)
| Init_struct(id, fil) ->
@@ -190,7 +138,7 @@ let transf_decl env (sto, id, ty, init) =
let transf_funbody env body optres =
-let transf_expr ctx e = float_comma(transf_expr env ctx e) in
+let transf_expr ctx e = transf_expr env ctx e in
(* Function returns: if return type is struct or union,
return x -> _res = x; return
@@ -239,33 +187,11 @@ let rec transf_stmt s =
in
transf_stmt body
-let transf_params loc env params =
- let rec transf_prm = function
- | [] ->
- ([], [], sskip)
- | (id, ty) :: params ->
- let ty = transf_type env ty in
- if is_composite_type env ty then begin
- let id' = Env.fresh_ident id.name in
- let ty' = TPtr(add_attributes_type [AConst] ty, []) in
- let (params', decls, init) = transf_prm params in
- ((id', ty') :: params',
- (Storage_default, id, ty, None) :: decls,
- sseq loc
- (sassign loc {edesc = EVar id; etyp = ty}
- {edesc = EUnop(Oderef, {edesc = EVar id'; etyp = ty'});
- etyp = ty})
- init)
- end else begin
- let (params', decls, init) = transf_prm params in
- ((id, ty) :: params', decls, init)
- end
- in transf_prm params
-
let transf_fundef env f =
reset_temps();
let ret = transf_type env f.fd_ret in
- let (params, newdecls, init) = transf_params f.fd_body.sloc env f.fd_params in
+ let params =
+ List.map (fun (id, ty) -> (id, transf_type env ty)) f.fd_params in
let (ret1, params1, body1) =
if is_composite_type env ret then begin
let vres = Env.fresh_ident "_res" in
@@ -277,10 +203,9 @@ let transf_fundef env f =
transf_funbody env f.fd_body (Some eeres))
end else
(ret, params, transf_funbody env f.fd_body None) in
- let body2 = sseq body1.sloc init body1 in
let temps = get_temps() in
{f with fd_ret = ret1; fd_params = params1;
- fd_locals = newdecls @ f.fd_locals @ temps; fd_body = body2}
+ fd_locals = f.fd_locals @ temps; fd_body = body1}
(* Composites *)
diff --git a/cparser/StructByValue.mli b/cparser/StructReturn.mli
index 45899a4..45899a4 100644
--- a/cparser/StructByValue.mli
+++ b/cparser/StructReturn.mli
diff --git a/driver/Clflags.ml b/driver/Clflags.ml
index cb26b72..95f4209 100644
--- a/driver/Clflags.ml
+++ b/driver/Clflags.ml
@@ -17,12 +17,10 @@ let linker_options = ref ([]: string list)
let exe_name = ref "a.out"
let option_flonglong = ref true
let option_flongdouble = ref false
-let option_fstruct_passing = ref false
-let option_fstruct_assign = ref false
+let option_fstruct_return = ref false
let option_fbitfields = ref false
let option_fvararg_calls = ref true
let option_fpacked_structs = ref false
-let option_fvolatile_rmw = ref true
let option_fmadd = ref false
let option_fsse = ref true
let option_dparse = ref false
diff --git a/driver/Compiler.v b/driver/Compiler.v
index ce9db20..6779aaf 100644
--- a/driver/Compiler.v
+++ b/driver/Compiler.v
@@ -400,7 +400,7 @@ Theorem transf_cstrategy_program_correct:
forall p tp,
transf_c_program p = OK tp ->
forward_simulation (Cstrategy.semantics p) (Asm.semantics tp)
- * backward_simulation (Cstrategy.semantics p) (Asm.semantics tp).
+ * backward_simulation (atomic (Cstrategy.semantics p)) (Asm.semantics tp).
Proof.
intros.
assert (F: forward_simulation (Cstrategy.semantics p) (Asm.semantics tp)).
@@ -411,8 +411,9 @@ Proof.
exact (fst (transf_clight_program_correct _ _ EQ1)).
split. auto.
- apply forward_to_backward_simulation. auto.
- apply Cstrategy.semantics_receptive.
+ apply forward_to_backward_simulation.
+ apply factor_forward_simulation. auto. eapply sd_traces. eapply Asm.semantics_determinate.
+ apply atomic_receptive. apply Cstrategy.semantics_strongly_receptive.
apply Asm.semantics_determinate.
Qed.
@@ -422,7 +423,11 @@ Theorem transf_c_program_correct:
backward_simulation (Csem.semantics p) (Asm.semantics tp).
Proof.
intros.
- eapply compose_backward_simulation.
+ apply compose_backward_simulation with (atomic (Cstrategy.semantics p)).
+ eapply sd_traces; eapply Asm.semantics_determinate.
+ apply factor_backward_simulation.
apply Cstrategy.strategy_simulation.
+ apply Csem.semantics_single_events.
+ eapply ssr_well_behaved; eapply Cstrategy.semantics_strongly_receptive.
exact (snd (transf_cstrategy_program_correct _ _ H)).
Qed.
diff --git a/driver/Complements.v b/driver/Complements.v
index 1b7e974..57351a2 100644
--- a/driver/Complements.v
+++ b/driver/Complements.v
@@ -78,15 +78,21 @@ Theorem transf_cstrategy_program_preservation:
program_behaves (Asm.semantics tp) beh ->
program_behaves (Cstrategy.semantics p) beh).
Proof.
+ assert (WBT: forall p, well_behaved_traces (Cstrategy.semantics p)).
+ intros. eapply ssr_well_behaved. apply Cstrategy.semantics_strongly_receptive.
intros. intuition.
eapply forward_simulation_behavior_improves; eauto.
apply (fst (transf_cstrategy_program_correct _ _ H)).
- eapply backward_simulation_behavior_improves; eauto.
+ exploit backward_simulation_behavior_improves.
apply (snd (transf_cstrategy_program_correct _ _ H)).
+ eauto.
+ intros [beh1 [A B]]. exists beh1; split; auto. rewrite atomic_behaviors; auto.
eapply forward_simulation_same_safe_behavior; eauto.
apply (fst (transf_cstrategy_program_correct _ _ H)).
- eapply backward_simulation_same_safe_behavior; eauto.
+ exploit backward_simulation_same_safe_behavior.
apply (snd (transf_cstrategy_program_correct _ _ H)).
+ intros. rewrite <- atomic_behaviors in H2; eauto. eauto.
+ intros. rewrite atomic_behaviors; auto.
Qed.
(** We can also use the alternate big-step semantics for [Cstrategy]
diff --git a/driver/Driver.ml b/driver/Driver.ml
index cee6250..9813939 100644
--- a/driver/Driver.ml
+++ b/driver/Driver.ml
@@ -69,11 +69,9 @@ let parse_c_file sourcename ifile =
(* Simplification options *)
let simplifs =
"b" (* blocks: mandatory *)
- ^ (if !option_fstruct_passing then "s" else "")
- ^ (if !option_fstruct_assign then "S" else "")
+ ^ (if !option_fstruct_return then "s" else "")
^ (if !option_fbitfields then "f" else "")
^ (if !option_fpacked_structs then "p" else "")
- ^ (if !option_fvolatile_rmw then "v" else "")
in
(* Parsing and production of a simplified C AST *)
let ast =
@@ -333,10 +331,8 @@ Language support options (use -fno-<opt> to turn off -f<opt>) :
-fbitfields Emulate bit fields in structs [off]
-flonglong Partial emulation of 'long long' types [on]
-flongdouble Treat 'long double' as 'double' [off]
- -fstruct-passing Emulate passing structs and unions by value [off]
- -fstruct-assign Emulate assignment between structs or unions [off]
+ -fstruct-return Emulate returning structs and unions by value [off]
-fvararg-calls Emulate calls to variable-argument functions [on]
- -fvolatile-rmw Emulate ++, -- and op= on volatile l-values [on]
-fpacked-structs Emulate packed structs [off]
-fall Activate all language support options above
-fnone Turn off all language support options above
@@ -374,9 +370,8 @@ Interpreter mode:
"
let language_support_options = [
- option_fbitfields; option_flonglong; option_flongdouble; option_fstruct_passing;
- option_fstruct_assign; option_fvararg_calls; option_fpacked_structs;
- option_fvolatile_rmw
+ option_fbitfields; option_flonglong; option_flongdouble;
+ option_fstruct_return; option_fvararg_calls; option_fpacked_structs
]
let cmdline_actions =
@@ -434,14 +429,12 @@ let cmdline_actions =
]
@ f_opt "longlong" option_flonglong
@ f_opt "longdouble" option_flongdouble
- @ f_opt "struct-passing" option_fstruct_passing
- @ f_opt "struct-assign" option_fstruct_assign
+ @ f_opt "struct-return" option_fstruct_return
@ f_opt "bitfields" option_fbitfields
@ f_opt "vararg-calls" option_fvararg_calls
@ f_opt "madd" option_fmadd
@ f_opt "packed-structs" option_fpacked_structs
@ f_opt "sse" option_fsse
- @ f_opt "volatile-rmw" option_fvolatile_rmw
let _ =
Gc.set { (Gc.get()) with Gc.minor_heap_size = 524288 };
diff --git a/driver/Interp.ml b/driver/Interp.ml
index a1e01f0..50e512b 100644
--- a/driver/Interp.ml
+++ b/driver/Interp.ml
@@ -113,12 +113,15 @@ let print_state prog p = function
| Returnstate(res, k, m) ->
fprintf p "returning@ %a"
print_val res
+ | Stuckstate ->
+ fprintf p "stuck after an undefined expression"
let mem_of_state = function
| State(f, s, k, e, m) -> m
| ExprState(f, r, k, e, m) -> m
| Callstate(fd, args, k, m) -> m
| Returnstate(res, k, m) -> m
+ | Stuckstate -> assert false
(* Comparing memory states *)
@@ -211,6 +214,7 @@ let rank_state = function
| ExprState _ -> 1
| Callstate _ -> 2
| Returnstate _ -> 3
+ | Stuckstate -> 4
let compare_state s1 s2 =
if s1 == s2 then 0 else
@@ -382,16 +386,17 @@ let do_step p prog ge time s =
exit (Int32.to_int (camlint_of_coqint r))
end
| None ->
- match Cexec.do_step ge world s with
- | [] ->
- if !trace = 1 then
- fprintf p "@[<hov 2>Time %d: %a@]@." time (print_state prog) s;
- fprintf p "ERROR: Undefined behavior@.";
- fprintf p "@].";
- exit 126
- | l ->
- List.iter (fun (t, s') -> do_events p ge time s t) l;
- List.map snd l
+ let l = Cexec.do_step ge world s in
+ if l = [] || List.exists (fun (t,s) -> s = Stuckstate) l then begin
+ if !trace = 1 then
+ fprintf p "@[<hov 2>Time %d: %a@]@." time (print_state prog) s;
+ fprintf p "ERROR: Undefined behavior@.";
+ fprintf p "@].";
+ exit 126
+ end else begin
+ List.iter (fun (t, s') -> do_events p ge time s t) l;
+ List.map snd l
+ end
let rec explore p prog ge time ss =
let succs =
@@ -415,14 +420,13 @@ let unvolatile prog =
{prog with prog_vars = List.map unvolatile_globvar prog.prog_vars}
let change_main_function p old_main old_main_ty =
- let tint = Tint(I32, Signed) in
let old_main = Evalof(Evar(old_main, old_main_ty), old_main_ty) in
- let arg1 = Eval(Vint(coqint_of_camlint 0l), tint) in
+ let arg1 = Eval(Vint(coqint_of_camlint 0l), type_int32s) in
let arg2 = arg1 in
let body =
- Sreturn(Some(Ecall(old_main, Econs(arg1, Econs(arg2, Enil)), tint))) in
+ Sreturn(Some(Ecall(old_main, Econs(arg1, Econs(arg2, Enil)), type_int32s))) in
let new_main_fn =
- { fn_return = tint; fn_params = []; fn_vars = []; fn_body = body } in
+ { fn_return = type_int32s; fn_params = []; fn_vars = []; fn_body = body } in
let new_main_id = intern_string "___main" in
{ p with
prog_main = new_main_id;
@@ -431,9 +435,9 @@ let change_main_function p old_main old_main_ty =
let fixup_main p =
try
match type_of_fundef (List.assoc p.prog_main p.prog_funct) with
- | Tfunction(Tnil, Tint(I32, Signed)) ->
+ | Tfunction(Tnil, Tint(I32, Signed, _)) ->
Some p
- | Tfunction(Tcons(Tint _, Tcons(Tpointer(Tpointer(Tint(I8,_))), Tnil)),
+ | Tfunction(Tcons(Tint _, Tcons(Tpointer(Tpointer(Tint(I8,_,_),_),_), Tnil)),
Tint _) as ty ->
Some (change_main_function p p.prog_main ty)
| _ ->
diff --git a/ia32/Asm.v b/ia32/Asm.v
index 63149aa..e2176fd 100644
--- a/ia32/Asm.v
+++ b/ia32/Asm.v
@@ -848,7 +848,7 @@ Ltac Equalities :=
exploit external_call_determ. eexact H3. eexact H8. intros [A B].
split. auto. intros. destruct B; auto. subst. auto.
(* trace length *)
- inv H; simpl.
+ red; intros; inv H; simpl.
omega.
eapply external_call_trace_length; eauto.
eapply external_call_trace_length; eauto.
diff --git a/ia32/SelectOp.vp b/ia32/SelectOp.vp
index 71dc83b..98db388 100644
--- a/ia32/SelectOp.vp
+++ b/ia32/SelectOp.vp
@@ -62,6 +62,21 @@ Definition addrstack (ofs: int) :=
Definition notint (e: expr) := Eop (Oxorimm Int.mone) (e ::: Enil).
+(** ** Boolean value and boolean negation *)
+
+Fixpoint boolval (e: expr) {struct e} : expr :=
+ let default := Eop (Ocmp (Ccompuimm Cne Int.zero)) (e ::: Enil) in
+ match e with
+ | Eop (Ointconst n) Enil =>
+ Eop (Ointconst (if Int.eq n Int.zero then Int.zero else Int.one)) Enil
+ | Eop (Ocmp cond) args =>
+ Eop (Ocmp cond) args
+ | Econdition e1 e2 e3 =>
+ Econdition e1 (boolval e2) (boolval e3)
+ | _ =>
+ default
+ end.
+
(** ** Boolean negation *)
Fixpoint notbool (e: expr) {struct e} : expr :=
diff --git a/ia32/SelectOpproof.v b/ia32/SelectOpproof.v
index f14b6a9..e6fd809 100644
--- a/ia32/SelectOpproof.v
+++ b/ia32/SelectOpproof.v
@@ -138,6 +138,31 @@ Proof.
unfold notint; red; intros. TrivialExists.
Qed.
+Theorem eval_boolval: unary_constructor_sound boolval Val.boolval.
+Proof.
+ assert (DFL:
+ forall le a x,
+ eval_expr ge sp e m le a x ->
+ exists v, eval_expr ge sp e m le (Eop (Ocmp (Ccompuimm Cne Int.zero)) (a ::: Enil)) v
+ /\ Val.lessdef (Val.boolval x) v).
+ intros. TrivialExists. simpl. destruct x; simpl; auto.
+
+ red. induction a; simpl; intros; eauto. destruct o; eauto.
+(* intconst *)
+ destruct e0; eauto. InvEval. TrivialExists. simpl. destruct (Int.eq i Int.zero); auto.
+(* cmp *)
+ inv H. simpl in H5.
+ destruct (eval_condition c vl m) as []_eqn.
+ TrivialExists. simpl. inv H5. rewrite Heqo. destruct b; auto.
+ simpl in H5. inv H5.
+ exists Vundef; split; auto. EvalOp; simpl. rewrite Heqo; auto.
+
+(* condition *)
+ inv H. destruct v1.
+ exploit IHa1; eauto. intros [v [A B]]. exists v; split; auto. eapply eval_Econdition; eauto.
+ exploit IHa2; eauto. intros [v [A B]]. exists v; split; auto. eapply eval_Econdition; eauto.
+Qed.
+
Theorem eval_notbool: unary_constructor_sound notbool Val.notbool.
Proof.
assert (DFL:
diff --git a/powerpc/Asm.v b/powerpc/Asm.v
index 7174f79..2d71ca9 100644
--- a/powerpc/Asm.v
+++ b/powerpc/Asm.v
@@ -508,12 +508,6 @@ Definition compare_float (rs: regset) (v1 v2: val) :=
#CR0_2 <- (Val.cmpf Ceq v1 v2)
#CR0_3 <- Vundef.
-Definition val_cond_reg (rs: regset) :=
- Val.or (Val.shl rs#CR0_0 (Vint (Int.repr 31)))
- (Val.or (Val.shl rs#CR0_1 (Vint (Int.repr 30)))
- (Val.or (Val.shl rs#CR0_2 (Vint (Int.repr 29)))
- (Val.shl rs#CR0_3 (Vint (Int.repr 28))))).
-
(** Execution of a single instruction [i] in initial state
[rs] and [m]. Return updated state. For instructions
that correspond to actual PowerPC instructions, the cases are
@@ -968,7 +962,7 @@ Ltac Equalities :=
exploit external_call_determ. eexact H3. eexact H8. intros [A B].
split. auto. intros. destruct B; auto. subst. auto.
(* trace length *)
- inv H; simpl.
+ red; intros. inv H; simpl.
omega.
eapply external_call_trace_length; eauto.
eapply external_call_trace_length; eauto.
diff --git a/powerpc/PrintAsm.ml b/powerpc/PrintAsm.ml
index e566e3c..f6c1c49 100644
--- a/powerpc/PrintAsm.ml
+++ b/powerpc/PrintAsm.ml
@@ -202,8 +202,8 @@ let name_of_section_MacOS = function
let name_of_section_Linux = function
| Section_text -> ".text"
- | Section_data i -> ".data" (*if i then ".data" else ".bss"*)
- | Section_small_data i -> ".sdata" (*if i then ".sdata" else ".sbss"*)
+ | Section_data i -> if i then ".data" else "COMM"
+ | Section_small_data i -> if i then ".sdata" else "COMM"
| Section_const -> ".rodata"
| Section_small_const -> ".sdata2"
| Section_string -> ".rodata"
@@ -238,7 +238,9 @@ let name_of_section =
| Diab -> name_of_section_Diab
let section oc sec =
- fprintf oc " %s\n" (name_of_section sec)
+ let name = name_of_section sec in
+ assert (name <> "COMM");
+ fprintf oc " %s\n" name
(* Encoding masks for rlwinm instructions *)
@@ -663,8 +665,8 @@ let print_instruction oc = function
| Plwzx(r1, r2, r3) ->
fprintf oc " lwzx %a, %a, %a\n" ireg r1 ireg r2 ireg r3
| Pmfcrbit(r1, bit) ->
- fprintf oc " mfcr %a\n" ireg GPR12;
- fprintf oc " rlwinm %a, %a, %d, 31, 31\n" ireg r1 ireg GPR12 (1 + num_crbit bit)
+ fprintf oc " mfcr %a\n" ireg r1;
+ fprintf oc " rlwinm %a, %a, %d, 31, 31\n" ireg r1 ireg r1 (1 + num_crbit bit)
| Pmflr(r1) ->
fprintf oc " mflr %a\n" ireg r1
| Pmr(r1, r2) ->
@@ -1012,21 +1014,32 @@ let print_var oc (name, v) =
let init =
match v.gvar_init with [Init_space _] -> false | _ -> true in
let sec =
- Sections.section_for_variable name init
- and align =
+ Sections.section_for_variable name init in
+ let align =
match C2C.atom_alignof name with
| Some a -> log2 a
- | None -> 3 (* 8-alignment is a safe default *)
- in
- section oc sec;
- fprintf oc " .align %d\n" align;
- if not (C2C.atom_is_static name) then
- fprintf oc " .globl %a\n" symbol name;
- fprintf oc "%a:\n" symbol name;
- print_init_data oc name v.gvar_init;
- if target <> MacOS then begin
- fprintf oc " .type %a, @object\n" symbol name;
- fprintf oc " .size %a, . - %a\n" symbol name symbol name
+ | None -> 3 in (* 8-alignment is a safe default *)
+ let name_sec =
+ name_of_section sec in
+ if name_sec <> "COMM" then begin
+ fprintf oc " %s\n" name_sec;
+ fprintf oc " .align %d\n" align;
+ if not (C2C.atom_is_static name) then
+ fprintf oc " .globl %a\n" symbol name;
+ fprintf oc "%a:\n" symbol name;
+ print_init_data oc name v.gvar_init;
+ if target <> MacOS then begin
+ fprintf oc " .type %a, @object\n" symbol name;
+ fprintf oc " .size %a, . - %a\n" symbol name symbol name
+ end
+ end else begin
+ let sz =
+ match v.gvar_init with [Init_space sz] -> sz | _ -> assert false in
+ fprintf oc " %s %a, %ld, %d\n"
+ (if C2C.atom_is_static name then ".lcomm" else ".comm")
+ symbol name
+ (camlint_of_coqint sz)
+ (1 lsl align)
end
let print_program oc p =
diff --git a/powerpc/SelectOp.vp b/powerpc/SelectOp.vp
index 3bb5544..08968f7 100644
--- a/powerpc/SelectOp.vp
+++ b/powerpc/SelectOp.vp
@@ -69,7 +69,20 @@ Nondetfunction notint (e: expr) :=
| _ => Elet e (Eop Onor (Eletvar O ::: Eletvar O ::: Enil))
end.
-(** ** Boolean negation *)
+(** ** Boolean value and boolean negation *)
+
+Fixpoint boolval (e: expr) {struct e} : expr :=
+ let default := Eop (Ocmp (Ccompuimm Cne Int.zero)) (e ::: Enil) in
+ match e with
+ | Eop (Ointconst n) Enil =>
+ Eop (Ointconst (if Int.eq n Int.zero then Int.zero else Int.one)) Enil
+ | Eop (Ocmp cond) args =>
+ Eop (Ocmp cond) args
+ | Econdition e1 e2 e3 =>
+ Econdition e1 (boolval e2) (boolval e3)
+ | _ =>
+ default
+ end.
Fixpoint notbool (e: expr) {struct e} : expr :=
let default := Eop (Ocmp (Ccompuimm Ceq Int.zero)) (e ::: Enil) in
diff --git a/powerpc/SelectOpproof.v b/powerpc/SelectOpproof.v
index cc14d33..59f2a41 100644
--- a/powerpc/SelectOpproof.v
+++ b/powerpc/SelectOpproof.v
@@ -148,6 +148,31 @@ Proof.
simpl. destruct x; simpl; auto. rewrite Int.or_idem. auto.
Qed.
+Theorem eval_boolval: unary_constructor_sound boolval Val.boolval.
+Proof.
+ assert (DFL:
+ forall le a x,
+ eval_expr ge sp e m le a x ->
+ exists v, eval_expr ge sp e m le (Eop (Ocmp (Ccompuimm Cne Int.zero)) (a ::: Enil)) v
+ /\ Val.lessdef (Val.boolval x) v).
+ intros. TrivialExists. simpl. destruct x; simpl; auto.
+
+ red. induction a; simpl; intros; eauto. destruct o; eauto.
+(* intconst *)
+ destruct e0; eauto. InvEval. TrivialExists. simpl. destruct (Int.eq i Int.zero); auto.
+(* cmp *)
+ inv H. simpl in H5.
+ destruct (eval_condition c vl m) as []_eqn.
+ TrivialExists. simpl. inv H5. rewrite Heqo. destruct b; auto.
+ simpl in H5. inv H5.
+ exists Vundef; split; auto. EvalOp; simpl. rewrite Heqo; auto.
+
+(* condition *)
+ inv H. destruct v1.
+ exploit IHa1; eauto. intros [v [A B]]. exists v; split; auto. eapply eval_Econdition; eauto.
+ exploit IHa2; eauto. intros [v [A B]]. exists v; split; auto. eapply eval_Econdition; eauto.
+Qed.
+
Theorem eval_notbool: unary_constructor_sound notbool Val.notbool.
Proof.
assert (DFL:
diff --git a/test/raytracer/Makefile b/test/raytracer/Makefile
index db359b1..a4b8894 100644
--- a/test/raytracer/Makefile
+++ b/test/raytracer/Makefile
@@ -1,7 +1,7 @@
include ../../Makefile.config
CC=../../ccomp
-CFLAGS=-stdlib ../../runtime -dparse -dclight -dasm -fstruct-passing -fstruct-assign
+CFLAGS=-stdlib ../../runtime -dparse -dclight -dasm -fstruct-return
LIBS=$(LIBMATH)
TIME=xtime -mintime 2.0
diff --git a/test/regression/Makefile b/test/regression/Makefile
index f3dcf4d..e9d0318 100644
--- a/test/regression/Makefile
+++ b/test/regression/Makefile
@@ -12,7 +12,7 @@ TESTS=attribs1 bitfields1 bitfields2 bitfields3 bitfields4 \
expr1 expr6 initializers volatile1 volatile2 volatile3 \
funct3 expr5 struct7 struct8 struct11 casts1 casts2 char1 \
sizeof1 sizeof2 packedstruct1 packedstruct2 \
- instrsel
+ instrsel bool
# Other tests: should compile to .s without errors (but expect warnings)
EXTRAS=annot1 commaprec expr2 expr3 expr4 extern1 funct2 funptr1 init1 \
diff --git a/test/regression/Results/bool b/test/regression/Results/bool
new file mode 100644
index 0000000..923c6e2
--- /dev/null
+++ b/test/regression/Results/bool
@@ -0,0 +1,9 @@
+a = 1
+b = 0
+c = 1
+d = 0
+e = 1
+f = 0
+g = 1
+h = 1
+i = 0
diff --git a/test/regression/bool.c b/test/regression/bool.c
new file mode 100644
index 0000000..d2b3857
--- /dev/null
+++ b/test/regression/bool.c
@@ -0,0 +1,29 @@
+/* Testing _Bool type support */
+
+#include <stdio.h>
+
+int x = 42;
+
+int main()
+{
+ _Bool a, b, c, d, e, f, g, h, i;
+ a = x;
+ b = x >= 100;
+ c = &x;
+ d = a && b;
+ e = a || b;
+ f = a & b;
+ g = a | b;
+ h = 3.14;
+ i = 0.0;
+ printf("a = %d\n", a);
+ printf("b = %d\n", b);
+ printf("c = %d\n", c);
+ printf("d = %d\n", d);
+ printf("e = %d\n", e);
+ printf("f = %d\n", f);
+ printf("g = %d\n", g);
+ printf("h = %d\n", h);
+ printf("i = %d\n", i);
+ return 0;
+}
diff --git a/test/spass/Makefile b/test/spass/Makefile
index 536488b..6797475 100644
--- a/test/spass/Makefile
+++ b/test/spass/Makefile
@@ -1,7 +1,7 @@
include ../../Makefile.config
CC=../../ccomp
-CFLAGS=-stdlib ../../runtime -dparse -dclight -dasm -fstruct-passing -fstruct-assign
+CFLAGS=-stdlib ../../runtime -dparse -dclight -dasm -fstruct-return
SRCS=analyze.c clause.c clock.c closure.c cnf.c component.c \
condensing.c context.c defs.c dfgparser.c dfgscanner.c doc-proof.c \