summaryrefslogtreecommitdiff
path: root/cfrontend
diff options
context:
space:
mode:
authorGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-10-06 15:46:47 +0000
committerGravatar xleroy <xleroy@fca1b0fc-160b-0410-b1d3-a4f43f01ea2e>2012-10-06 15:46:47 +0000
commitf7693b3d897b90fd3bc2533be002dc0bdcd9f6c2 (patch)
tree93ea9491693324d2d690c4236a2c88c3b461e225 /cfrontend
parent261ef24f7fd2ef443f73c468b9b1fa496371f3bf (diff)
Merge of branch seq-and-or. See Changelog for details.
git-svn-id: https://yquem.inria.fr/compcert/svn/compcert/trunk@2059 fca1b0fc-160b-0410-b1d3-a4f43f01ea2e
Diffstat (limited to 'cfrontend')
-rw-r--r--cfrontend/C2C.ml65
-rw-r--r--cfrontend/Cexec.v115
-rw-r--r--cfrontend/Clight.v762
-rw-r--r--cfrontend/Cminorgen.v42
-rw-r--r--cfrontend/Cminorgenproof.v73
-rw-r--r--cfrontend/Csem.v134
-rw-r--r--cfrontend/Csharpminor.v21
-rw-r--r--cfrontend/Cshmgen.v166
-rw-r--r--cfrontend/Cshmgenproof.v577
-rw-r--r--cfrontend/Cstrategy.v203
-rw-r--r--cfrontend/Csyntax.v36
-rw-r--r--cfrontend/Initializers.v16
-rw-r--r--cfrontend/Initializersproof.v48
-rw-r--r--cfrontend/PrintClight.ml53
-rw-r--r--cfrontend/PrintCsyntax.ml23
-rw-r--r--cfrontend/SimplExpr.v192
-rw-r--r--cfrontend/SimplExprproof.v1035
-rw-r--r--cfrontend/SimplExprspec.v410
18 files changed, 2036 insertions, 1935 deletions
diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml
index a24d552..c9beaf7 100644
--- a/cfrontend/C2C.ml
+++ b/cfrontend/C2C.ml
@@ -173,62 +173,25 @@ let register_stub_function name tres targs =
register_special_external stub_name ef targs tres;
(stub_name, Tfunction (targs, tres))
-(** ** Handling of annotations *)
-
-let annot_function_next = ref 0
-
-let register_annotation_stmt txt targs =
- let tres = Tvoid in
- incr annot_function_next;
- let fun_name =
- Printf.sprintf "__builtin_annot_%d" !annot_function_next
- and ef =
- EF_annot (intern_string txt, typlist_of_typelist targs) in
- register_special_external fun_name ef targs tres;
- Evalof(Evar(intern_string fun_name, Tfunction(targs, tres)),
- Tfunction(targs, tres))
-
-let register_annotation_val txt targ =
- let targs = Tcons(targ, Tnil)
- and tres = targ in
- incr annot_function_next;
- let fun_name =
- Printf.sprintf "__builtin_annot_val_%d" !annot_function_next
- and ef =
- EF_annot_val (intern_string txt, typ_of_type targ) in
- register_special_external fun_name ef targs tres;
- Evalof(Evar(intern_string fun_name, Tfunction(targs, tres)),
- Tfunction(targs, tres))
-
(** ** Handling of inlined memcpy functions *)
-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, 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;
- Evalof(Evar(intern_string name, Tfunction(targs, tres)),
- Tfunction(targs, tres))
-
let make_builtin_memcpy args =
match args with
| Econs(dst, Econs(src, Econs(sz, Econs(al, Enil)))) ->
let sz1 =
match Initializers.constval sz with
- | Errors.OK(Vint n) -> camlint_of_coqint n
+ | Errors.OK(Vint n) -> n
| _ -> error "ill-formed __builtin_memcpy_aligned (3rd argument must be
-a constant)"; 0l in
+a constant)"; Integers.Int.zero in
let al1 =
match Initializers.constval al with
- | Errors.OK(Vint n) -> camlint_of_coqint n
+ | Errors.OK(Vint n) -> n
| _ -> error "ill-formed __builtin_memcpy_aligned (4th argument must be
-a constant)"; 0l in
- let fn = register_inlined_memcpy sz1 al1 in
- Ecall(fn, Econs(dst, Econs(src, Enil)), Tvoid)
+a constant)"; Integers.Int.one in
+ (* to check: sz1 > 0, al1 divides sz1, al1 = 1|2|4|8 *)
+ Ebuiltin(EF_memcpy(sz1, al1),
+ Tcons(typeof dst, Tcons(typeof src, Tnil)),
+ Econs(dst, Econs(src, Enil)), Tvoid)
| _ ->
assert false
@@ -481,9 +444,9 @@ let rec convertExpr env e =
| C.EBinop(C.Ocomma, e1, e2, _) ->
Ecomma(convertExpr env e1, convertExpr env e2, ty)
| C.EBinop(C.Ologand, e1, e2, _) ->
- coq_Eseqand (convertExpr env e1) (convertExpr env e2) ty
+ Eseqand(convertExpr env e1, convertExpr env e2, ty)
| C.EBinop(C.Ologor, e1, e2, _) ->
- coq_Eseqor (convertExpr env e1) (convertExpr env e2) ty
+ Eseqor(convertExpr env e1, convertExpr env e2, ty)
| C.EConditional(e1, e2, e3) ->
Econdition(convertExpr env e1, convertExpr env e2, convertExpr env e3, ty)
@@ -494,8 +457,8 @@ let rec convertExpr env e =
begin match args with
| {edesc = C.EConst(CStr txt)} :: args1 ->
let targs1 = convertTypList env (List.map (fun e -> e.etyp) args1) in
- let fn' = register_annotation_stmt txt targs1 in
- Ecall(fn', convertExprList env args1, ty)
+ Ebuiltin(EF_annot(intern_string txt, typlist_of_typelist targs1),
+ targs1, convertExprList env args1, ty)
| _ ->
error "ill-formed __builtin_annot (first argument must be string literal)";
ezero
@@ -505,8 +468,8 @@ let rec convertExpr env e =
begin match args with
| [ {edesc = C.EConst(CStr txt)}; arg ] ->
let targ = convertTyp env arg.etyp in
- let fn' = register_annotation_val txt targ in
- Ecall(fn', convertExprList env [arg], ty)
+ Ebuiltin(EF_annot_val(intern_string txt, typ_of_type targ),
+ Tcons(targ, Tnil), convertExprList env [arg], ty)
| _ ->
error "ill-formed __builtin_annot_intval (first argument must be string literal)";
ezero
diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v
index 5427ac6..9391d8e 100644
--- a/cfrontend/Cexec.v
+++ b/cfrontend/Cexec.v
@@ -778,6 +778,24 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
| None =>
incontext (fun x => Ecast x ty) (step_expr RV r1 m)
end
+ | RV, Eseqand r1 r2 ty =>
+ match is_val r1 with
+ | Some(v1, ty1) =>
+ do b <- bool_val v1 ty1;
+ if b then topred (Rred (Eparen (Eparen r2 type_bool) ty) m E0)
+ else topred (Rred (Eval (Vint Int.zero) ty) m E0)
+ | None =>
+ incontext (fun x => Eseqand x r2 ty) (step_expr RV r1 m)
+ end
+ | RV, Eseqor r1 r2 ty =>
+ match is_val r1 with
+ | Some(v1, ty1) =>
+ do b <- bool_val v1 ty1;
+ if b then topred (Rred (Eval (Vint Int.one) ty) m E0)
+ else topred (Rred (Eparen (Eparen r2 type_bool) ty) m E0)
+ | None =>
+ incontext (fun x => Eseqor x r2 ty) (step_expr RV r1 m)
+ end
| RV, Econdition r1 r2 r3 ty =>
match is_val r1 with
| Some(v1, ty1) =>
@@ -859,6 +877,17 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
incontext2 (fun x => Ecall x rargs ty) (step_expr RV r1 m)
(fun x => Ecall r1 x ty) (step_exprlist rargs m)
end
+ | RV, Ebuiltin ef tyargs rargs ty =>
+ match is_val_list rargs with
+ | Some vtl =>
+ do vargs <- sem_cast_arguments vtl tyargs;
+ match do_external ef w vargs m with
+ | None => stuck
+ | Some(w',t,v,m') => topred (Rred (Eval v ty) m' t)
+ end
+ | _ =>
+ incontext (fun x => Ebuiltin ef tyargs x ty) (step_exprlist rargs m)
+ end
| _, _ => stuck
end
@@ -949,6 +978,10 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
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
+ | Eseqand (Eval v1 ty1) r2 ty =>
+ exists b, bool_val v1 ty1 = Some b
+ | Eseqor (Eval v1 ty1) r2 ty =>
+ exists b, bool_val v1 ty1 = Some b
| Econdition (Eval v1 ty1) r1 r2 ty =>
exists b, bool_val v1 ty1 = Some b
| Eassign (Eloc b ofs ty1) (Eval v2 ty2) ty =>
@@ -971,6 +1004,12 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
/\ Genv.find_funct ge vf = Some fd
/\ cast_arguments rargs tyargs vl
/\ type_of_fundef fd = Tfunction tyargs tyres
+ | Ebuiltin ef tyargs rargs ty =>
+ exprlist_all_values rargs ->
+ exists vargs, exists t, exists vres, exists m', exists w',
+ cast_arguments rargs tyargs vargs
+ /\ external_call ef ge vargs m t vres m'
+ /\ possible_trace w t w'
| _ => True
end.
@@ -993,11 +1032,14 @@ Proof.
exists v; auto.
exists v; auto.
exists v; auto.
+ exists true; auto. exists false; auto.
+ exists true; auto. exists false; 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.
+ intros; exists vargs; exists t; exists vres; exists m'; exists w'; auto.
Qed.
Lemma callred_invert:
@@ -1035,12 +1077,15 @@ Proof.
destruct (C a); auto; contradiction.
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.
+ intros. elim (H0 a m); auto.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
red; intros. destruct (C a); auto.
@@ -1224,6 +1269,20 @@ Proof.
destruct H5. destruct k1; intuition congruence. destruct k2; intuition congruence.
Qed.
+Lemma incontext_list_ok:
+ forall ef tyargs al ty m res,
+ list_reducts_ok al m res ->
+ is_val_list al = None ->
+ reducts_ok RV (Ebuiltin ef tyargs al ty) m
+ (incontext (fun x => Ebuiltin ef tyargs x ty) res).
+Proof.
+ 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 [a'' [k'' [U [V W]]]].
+ exists a''; exists k''. split. eauto. rewrite V; auto.
+ destruct res; simpl in H2. elim H1; auto. congruence.
+Qed.
+
Lemma incontext2_list_ok:
forall a1 a2 ty m res1 res2,
reducts_ok RV a1 m res1 ->
@@ -1357,6 +1416,22 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
apply topred_ok; auto. split. apply red_cast; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
+(* seqand *)
+ destruct (is_val a1) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
+ (* top *)
+ destruct (bool_val v ty') as [v'|]_eqn... destruct v'.
+ apply topred_ok; auto. split. eapply red_seqand_true; eauto. exists w; constructor.
+ apply topred_ok; auto. split. eapply red_seqand_false; eauto. exists w; constructor.
+ (* depth *)
+ eapply incontext_ok; eauto.
+(* seqor *)
+ destruct (is_val a1) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
+ (* top *)
+ destruct (bool_val v ty') as [v'|]_eqn... destruct v'.
+ apply topred_ok; auto. split. eapply red_seqor_true; eauto. exists w; constructor.
+ apply topred_ok; auto. split. eapply red_seqor_false; eauto. exists w; constructor.
+ (* depth *)
+ eapply incontext_ok; eauto.
(* condition *)
destruct (is_val a1) as [[v ty'] | ]_eqn. rewrite (is_val_inv _ _ _ Heqo).
(* top *)
@@ -1430,7 +1505,26 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv. congruence.
(* depth *)
eapply incontext2_list_ok; eauto.
- eapply incontext2_list_ok; eauto.
+ eapply incontext2_list_ok; eauto.
+(* builtin *)
+ destruct (is_val_list rargs) as [vtl | ]_eqn.
+ exploit is_val_list_all_values; eauto. intros ALLVAL.
+ (* top *)
+ destruct (sem_cast_arguments vtl tyargs) as [vargs|]_eqn...
+ destruct (do_external ef w vargs m) as [[[[? ?] v] m'] | ]_eqn...
+ exploit do_ef_external_sound; eauto. intros [EC PT].
+ apply topred_ok; auto. red. split; auto. eapply red_builtin; eauto.
+ eapply sem_cast_arguments_sound; eauto.
+ exists w0; auto.
+ apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv.
+ assert (x = vargs).
+ exploit sem_cast_arguments_complete; eauto. intros [vtl' [A B]]. congruence.
+ subst x. exploit do_ef_external_complete; eauto. congruence.
+ apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv.
+ exploit sem_cast_arguments_complete; eauto. intros [vtl' [A B]]. congruence.
+ (* depth *)
+ eapply incontext_list_ok; eauto.
+
(* loc *)
split; intros. tauto. simpl; congruence.
(* paren *)
@@ -1495,6 +1589,12 @@ Proof.
inv H0. rewrite H; auto.
(* cast *)
inv H0. rewrite H; auto.
+(* seqand *)
+ inv H0. rewrite H; auto.
+ inv H0. rewrite H; auto.
+(* seqor *)
+ inv H0. rewrite H; auto.
+ inv H0. rewrite H; auto.
(* condition *)
inv H0. rewrite H; auto.
(* sizeof *)
@@ -1511,6 +1611,10 @@ Proof.
inv H0. rewrite dec_eq_true; auto.
(* paren *)
inv H0. rewrite H; auto.
+(* builtin *)
+ exploit sem_cast_arguments_complete; eauto. intros [vtl [A B]].
+ exploit do_ef_external_complete; eauto. intros C.
+ rewrite A. rewrite B. rewrite C. auto.
Qed.
Lemma callred_topred:
@@ -1631,6 +1735,12 @@ Proof.
(* cast *)
eapply reducts_incl_trans with (C' := fun x => Ecast x ty); eauto.
destruct (is_val (C a)) as [[v ty']|]_eqn; eauto.
+(* seqand *)
+ eapply reducts_incl_trans with (C' := fun x => Eseqand x r2 ty); eauto.
+ destruct (is_val (C a)) as [[v ty']|]_eqn; eauto.
+(* seqor *)
+ eapply reducts_incl_trans with (C' := fun x => Eseqor x r2 ty); eauto.
+ destruct (is_val (C a)) as [[v ty']|]_eqn; eauto.
(* condition *)
eapply reducts_incl_trans with (C' := fun x => Econdition x r2 r3 ty); eauto.
destruct (is_val (C a)) as [[v ty']|]_eqn; eauto.
@@ -1658,6 +1768,9 @@ Proof.
eapply reducts_incl_trans with (C' := fun x => Ecall e1 x ty). apply step_exprlist_context. auto.
destruct (is_val e1) as [[v1 ty1]|]_eqn; eauto.
destruct (is_val_list (C a)) as [vl|]_eqn; eauto.
+(* builtin *)
+ eapply reducts_incl_trans with (C' := fun x => Ebuiltin ef tyargs x ty). apply step_exprlist_context. auto.
+ destruct (is_val_list (C a)) as [vl|]_eqn; eauto.
(* comma *)
eapply reducts_incl_trans with (C' := fun x => Ecomma x e2 ty); eauto.
destruct (is_val (C a)) as [[v ty']|]_eqn; eauto.
diff --git a/cfrontend/Clight.v b/cfrontend/Clight.v
index e4c451b..a8624e9 100644
--- a/cfrontend/Clight.v
+++ b/cfrontend/Clight.v
@@ -56,7 +56,6 @@ Inductive expr : Type :=
| Eunop: unary_operation -> expr -> type -> expr (**r unary operation *)
| Ebinop: binary_operation -> expr -> expr -> type -> expr (**r binary operation *)
| Ecast: expr -> type -> expr (**r type cast ([(ty) e]) *)
- | Econdition: expr -> expr -> expr -> type -> expr (**r conditional ([e1 ? e2 : e3]) *)
| Efield: expr -> ident -> type -> expr. (**r access to a member of a struct or union *)
(** [sizeof] and [alignof] are derived forms. *)
@@ -77,7 +76,6 @@ Definition typeof (e: expr) : type :=
| Eunop _ _ ty => ty
| Ebinop _ _ _ ty => ty
| Ecast _ ty => ty
- | Econdition _ _ _ ty => ty
| Efield _ _ ty => ty
end.
@@ -86,7 +84,8 @@ Definition typeof (e: expr) : type :=
(** Clight statements are similar to those of Compcert C, with the addition
of assigment (of a rvalue to a lvalue), assignment to a temporary,
and function call (with assignment of the result to a temporary).
- The [for] loop is slightly simplified: there is no initial statement. *)
+ The three C loops are replaced by a single infinite loop [Sloop s1 s2]
+ that executes [s1] then [s2] repeatedly. A [continue] in [s1] branches to [s2]. *)
Definition label := ident.
@@ -94,13 +93,11 @@ 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 *)
+ | Sbuiltin: option ident -> external_function -> typelist -> list expr -> statement (**r builtin invocation *)
| Ssequence : statement -> statement -> statement (**r sequence *)
| Sifthenelse : expr -> statement -> statement -> statement (**r conditional *)
- | Swhile : expr -> statement -> statement (**r [while] loop *)
- | Sdowhile : expr -> statement -> statement (**r [do] loop *)
- | Sfor': expr -> statement -> statement -> statement (**r [for] loop *)
+ | Sloop: statement -> statement -> statement (**r infinite loop *)
| Sbreak : statement (**r [break] statement *)
| Scontinue : statement (**r [continue] statement *)
| Sreturn : option expr -> statement (**r [return] statement *)
@@ -112,10 +109,16 @@ with labeled_statements : Type := (**r cases of a [switch] *)
| LSdefault: statement -> labeled_statements
| LScase: int -> statement -> labeled_statements -> labeled_statements.
-(** The full [for] loop is a derived form. *)
+(** The C loops are derived forms. *)
-Definition Sfor (s1: statement) (e2: expr) (s3 s4: statement) :=
- Ssequence s1 (Sfor' e2 s3 s4).
+Definition Swhile (e: expr) (s: statement) :=
+ Sloop (Ssequence (Sifthenelse e Sskip Sbreak) s) Sskip.
+
+Definition Sdowhile (s: statement) (e: expr) :=
+ Sloop s (Sifthenelse e Sskip Sbreak).
+
+Definition Sfor (s1: statement) (e2: expr) (s3: statement) (s4: statement) :=
+ Ssequence s1 (Sloop (Ssequence (Sifthenelse e2 Sskip Sbreak) s3) s4).
(** ** Functions *)
@@ -183,6 +186,80 @@ Definition empty_env: env := (PTree.empty (block * type)).
Definition temp_env := PTree.t val.
+(** [deref_loc ty m b ofs 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 or by copy, the pointer [Vptr b ofs] is returned. *)
+
+Inductive deref_loc (ty: type) (m: mem) (b: block) (ofs: int) : val -> Prop :=
+ | deref_loc_value: forall chunk v,
+ access_mode ty = By_value chunk ->
+ Mem.loadv chunk m (Vptr b ofs) = Some v ->
+ deref_loc ty m b ofs v
+ | deref_loc_reference:
+ access_mode ty = By_reference ->
+ deref_loc ty m b ofs (Vptr b ofs)
+ | deref_loc_copy:
+ access_mode ty = By_copy ->
+ deref_loc ty m b ofs (Vptr b ofs).
+
+(** Symmetrically, [assign_loc ty m b ofs v 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 or by copy.
+ [m'] is the updated memory state. *)
+
+Inductive assign_loc (ty: type) (m: mem) (b: block) (ofs: int):
+ val -> mem -> Prop :=
+ | assign_loc_value: forall v chunk m',
+ access_mode ty = By_value chunk ->
+ Mem.storev chunk m (Vptr b ofs) v = Some m' ->
+ assign_loc ty m b ofs v 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 ty m b ofs (Vptr b' ofs') m'.
+
+(** Initialization of local variables that are parameters to a function.
+ [bind_parameters e m1 params args m2] stores the values [args]
+ in the memory blocks corresponding to the variables [params].
+ [m1] is the initial memory state and [m2] the final memory state. *)
+
+Inductive bind_parameters (e: env):
+ mem -> list (ident * type) -> list val ->
+ mem -> Prop :=
+ | bind_parameters_nil:
+ forall m,
+ bind_parameters e m nil nil m
+ | bind_parameters_cons:
+ forall m id ty params v1 vl b m1 m2,
+ PTree.get id e = Some(b, ty) ->
+ assign_loc ty m b Int.zero v1 m1 ->
+ bind_parameters e m1 params vl m2 ->
+ bind_parameters e m ((id, ty) :: params) (v1 :: vl) m2.
+
+(** Initialization of temporary variables *)
+
+Fixpoint create_undef_temps (temps: list (ident * type)) : temp_env :=
+ match temps with
+ | nil => PTree.empty val
+ | (id, t) :: temps' => PTree.set id Vundef (create_undef_temps temps')
+ end.
+
+(** Optional assignment to a temporary *)
+
+Definition set_opttemp (optid: option ident) (v: val) (le: temp_env) :=
+ match optid with
+ | None => le
+ | Some id => PTree.set id v le
+ end.
+
(** Selection of the appropriate case of a [switch], given the value [n]
of the selector expression. *)
@@ -250,19 +327,13 @@ Inductive eval_expr: expr -> val -> Prop :=
eval_expr a2 v2 ->
sem_binary_operation op v1 (typeof a1) v2 (typeof a2) m = Some v ->
eval_expr (Ebinop op a1 a2 ty) v
- | eval_Econdition: forall a1 a2 a3 ty v1 b v' v,
- eval_expr a1 v1 ->
- bool_val v1 (typeof a1) = Some b ->
- eval_expr (if b then a2 else a3) v' ->
- sem_cast v' (typeof (if b then a2 else a3)) ty = Some v ->
- eval_expr (Econdition a1 a2 a3 ty) v
| eval_Ecast: forall a ty v1 v,
eval_expr a v1 ->
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 -> type_is_volatile (typeof a) = false ->
- deref_loc ge (typeof a) m loc ofs E0 v ->
+ eval_lvalue a loc ofs ->
+ deref_loc (typeof a) m loc ofs v ->
eval_expr a v
(** [eval_lvalue ge e m a b ofs] defines the evaluation of expression [a]
@@ -318,10 +389,8 @@ End EXPR.
Inductive cont: Type :=
| Kstop: cont
| Kseq: statement -> cont -> cont (**r [Kseq s2 k] = after [s1] in [s1;s2] *)
- | Kwhile: expr -> statement -> cont -> cont (**r [Kwhile e s k] = after [s] in [while (e) s] *)
- | Kdowhile: expr -> statement -> cont -> cont (**r [Kdowhile e s k] = after [s] in [do s while (e)] *)
- | Kfor2: expr -> statement -> statement -> cont -> cont (**r [Kfor2 e2 e3 s k] = after [s] in [for'(e2;e3) s] *)
- | Kfor3: expr -> statement -> statement -> cont -> cont (**r [Kfor3 e2 e3 s k] = after [e3] in [for'(e2;e3) s] *)
+ | Kloop1: statement -> statement -> cont -> cont (**r [Kloop1 s1 s2 k] = after [s1] in [Sloop s1 s2] *)
+ | Kloop2: statement -> statement -> cont -> cont (**r [Kloop1 s1 s2 k] = after [s2] in [Sloop s1 s2] *)
| Kswitch: cont -> cont (**r catches [break] statements arising out of [switch] *)
| Kcall: option ident -> (**r where to store result *)
function -> (**r calling function *)
@@ -334,10 +403,8 @@ Inductive cont: Type :=
Fixpoint call_cont (k: cont) : cont :=
match k with
| Kseq s k => call_cont k
- | Kwhile e s k => call_cont k
- | Kdowhile e s k => call_cont k
- | Kfor2 e2 e3 s k => call_cont k
- | Kfor3 e2 e3 s k => call_cont k
+ | Kloop1 s1 s2 k => call_cont k
+ | Kloop2 s1 s2 k => call_cont k
| Kswitch k => call_cont k
| _ => k
end.
@@ -385,14 +452,10 @@ Fixpoint find_label (lbl: label) (s: statement) (k: cont)
| Some sk => Some sk
| None => find_label lbl s2 k
end
- | Swhile a s1 =>
- find_label lbl s1 (Kwhile a s1 k)
- | Sdowhile a s1 =>
- find_label lbl s1 (Kdowhile a s1 k)
- | Sfor' a2 a3 s1 =>
- match find_label lbl s1 (Kfor2 a2 a3 s1 k) with
+ | Sloop s1 s2 =>
+ match find_label lbl s1 (Kloop1 s1 s2 k) with
| Some sk => Some sk
- | None => find_label lbl a3 (Kfor3 a2 a3 s1 k)
+ | None => find_label lbl s2 (Kloop2 s1 s2 k)
end
| Sswitch e sl =>
find_label_ls lbl sl (Kswitch k)
@@ -416,26 +479,19 @@ 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 t m',
+ | step_assign: forall f a1 a2 k e le m loc ofs v2 v 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 ->
- assign_loc ge (typeof a1) m loc ofs v t m' ->
+ assign_loc (typeof a1) m loc ofs v m' ->
step (State f (Sassign a1 a2) k e le m)
- t (State f Sskip k e le m')
+ E0 (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 ->
@@ -445,6 +501,12 @@ Inductive step: state -> trace -> state -> Prop :=
step (State f (Scall optid a al) k e le m)
E0 (Callstate fd vargs (Kcall optid f e le k) m)
+ | step_builtin: forall f optid ef tyargs al k e le m vargs t vres m',
+ eval_exprlist e le m al tyargs vargs ->
+ external_call ef ge vargs m t vres m' ->
+ step (State f (Sbuiltin optid ef tyargs al) k e le m)
+ t (State f Sskip k e (set_opttemp optid vres le) m')
+
| step_seq: forall f s1 s2 k e le m,
step (State f (Ssequence s1 s2) k e le m)
E0 (State f s1 (Kseq s2 k) e le m)
@@ -464,65 +526,21 @@ Inductive step: state -> trace -> state -> Prop :=
step (State f (Sifthenelse a s1 s2) k e le m)
E0 (State f (if b then s1 else s2) k e le m)
- | step_while_false: forall f a s k e le m v,
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some false ->
- step (State f (Swhile a s) k e le m)
- E0 (State f Sskip k e le m)
- | step_while_true: forall f a s k e le m v,
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some true ->
- step (State f (Swhile a s) k e le m)
- E0 (State f s (Kwhile a s k) e le m)
- | step_skip_or_continue_while: forall f x a s k e le m,
- x = Sskip \/ x = Scontinue ->
- step (State f x (Kwhile a s k) e le m)
- E0 (State f (Swhile a s) k e le m)
- | step_break_while: forall f a s k e le m,
- step (State f Sbreak (Kwhile a s k) e le m)
- E0 (State f Sskip k e le m)
-
- | step_dowhile: forall f a s k e le m,
- step (State f (Sdowhile a s) k e le m)
- E0 (State f s (Kdowhile a s k) e le m)
- | step_skip_or_continue_dowhile_false: forall f x a s k e le m v,
+ | step_loop: forall f s1 s2 k e le m,
+ step (State f (Sloop s1 s2) k e le m)
+ E0 (State f s1 (Kloop1 s1 s2 k) e le m)
+ | step_skip_or_continue_loop1: forall f s1 s2 k e le m x,
x = Sskip \/ x = Scontinue ->
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some false ->
- step (State f x (Kdowhile a s k) e le m)
- E0 (State f Sskip k e le m)
- | step_skip_or_continue_dowhile_true: forall f x a s k e le m v,
- x = Sskip \/ x = Scontinue ->
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some true ->
- step (State f x (Kdowhile a s k) e le m)
- E0 (State f (Sdowhile a s) k e le m)
- | step_break_dowhile: forall f a s k e le m,
- step (State f Sbreak (Kdowhile a s k) e le m)
- E0 (State f Sskip k e le m)
-
- | step_for_false: forall f a2 a3 s k e le m v,
- eval_expr e le m a2 v ->
- bool_val v (typeof a2) = Some false ->
- step (State f (Sfor' a2 a3 s) k e le m)
- E0 (State f Sskip k e le m)
- | step_for_true: forall f a2 a3 s k e le m v,
- eval_expr e le m a2 v ->
- bool_val v (typeof a2) = Some true ->
- step (State f (Sfor' a2 a3 s) k e le m)
- E0 (State f s (Kfor2 a2 a3 s k) e le m)
- | step_skip_or_continue_for2: forall f x a2 a3 s k e le m,
- x = Sskip \/ x = Scontinue ->
- step (State f x (Kfor2 a2 a3 s k) e le m)
- E0 (State f a3 (Kfor3 a2 a3 s k) e le m)
- | step_break_for2: forall f a2 a3 s k e le m,
- step (State f Sbreak (Kfor2 a2 a3 s k) e le m)
+ step (State f x (Kloop1 s1 s2 k) e le m)
+ E0 (State f s2 (Kloop2 s1 s2 k) e le m)
+ | step_break_loop1: forall f s1 s2 k e le m,
+ step (State f Sbreak (Kloop1 s1 s2 k) e le m)
E0 (State f Sskip k e le m)
- | step_skip_for3: forall f a2 a3 s k e le m,
- step (State f Sskip (Kfor3 a2 a3 s k) e le m)
- E0 (State f (Sfor' a2 a3 s) k e le m)
- | step_break_for3: forall f a2 a3 s k e le m,
- step (State f Sbreak (Kfor3 a2 a3 s k) e le m)
+ | step_skip_loop2: forall f s1 s2 k e le m,
+ step (State f Sskip (Kloop2 s1 s2 k) e le m)
+ E0 (State f (Sloop s1 s2) k e le m)
+ | step_break_loop2: forall f s1 s2 k e le m,
+ step (State f Sbreak (Kloop2 s1 s2 k) e le m)
E0 (State f Sskip k e le m)
| step_return_0: forall f k e le m m',
@@ -566,25 +584,76 @@ 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 ge e m1 f.(fn_params) vargs m2 ->
+ bind_parameters 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)
+ E0 (State f f.(fn_body) k e (create_undef_temps f.(fn_temps)) m2)
| step_external_function: forall ef targs tres vargs k m vres t m',
- external_call ef ge vargs m t vres m' ->
+ external_call ef ge vargs m t vres m' ->
step (Callstate (External ef targs tres) vargs k m)
t (Returnstate vres k m')
- | step_returnstate_none: forall v f e le k m,
- step (Returnstate v (Kcall None f e le k) m)
- E0 (State f Sskip k e le m)
- | step_returnstate_some: forall v id f e le k m,
- step (Returnstate v (Kcall (Some id) f e le k) m)
- E0 (State f Sskip k e (PTree.set id v le) m).
+ | step_returnstate: forall v optid f e le k m,
+ step (Returnstate v (Kcall optid f e le k) m)
+ E0 (State f Sskip k e (set_opttemp optid v le) m).
+(** ** Whole-program semantics *)
+
+(** Execution of whole programs are described as sequences of transitions
+ from an initial state to a final state. An initial state is a [Callstate]
+ corresponding to the invocation of the ``main'' function of the program
+ without arguments and with an empty continuation. *)
+
+Inductive initial_state (p: program): state -> Prop :=
+ | initial_state_intro: forall b f m0,
+ let ge := Genv.globalenv p in
+ Genv.init_mem p = Some m0 ->
+ Genv.find_symbol ge p.(prog_main) = Some b ->
+ Genv.find_funct_ptr ge b = Some f ->
+ 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. *)
+
+Inductive final_state: state -> int -> Prop :=
+ | final_state_intro: forall r m,
+ final_state (Returnstate (Vint r) Kstop m) r.
+
+End SEMANTICS.
+
+(** Wrapping up these definitions in a small-step semantics. *)
+
+Definition semantics (p: program) :=
+ Semantics step (initial_state p) final_state (Genv.globalenv p).
+
+(** This semantics is receptive to changes in events. *)
+
+Lemma semantics_receptive:
+ forall (p: program), 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; auto.
+ (* builtin *)
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ econstructor; econstructor; eauto.
+ (* external *)
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ exists (Returnstate vres2 k m2). econstructor; eauto.
+(* trace length *)
+ red; intros. inv H; simpl; try omega.
+ eapply external_call_trace_length; eauto.
+ eapply external_call_trace_length; eauto.
+Qed.
(** * Alternate big-step semantics *)
+Section BIGSTEP.
+
+Variable ge: genv.
+
(** ** Big-step semantics for terminating statements and functions *)
(** The execution of a statement produces an ``outcome'', indicating
@@ -630,41 +699,31 @@ 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 t m',
- eval_lvalue e le m a1 loc ofs ->
- eval_expr e le m a2 v2 ->
+ | exec_Sassign: forall e le m a1 a2 loc ofs v2 v m',
+ eval_lvalue ge e le m a1 loc ofs ->
+ eval_expr ge e le m a2 v2 ->
sem_cast v2 (typeof a2) (typeof a1) = Some v ->
- assign_loc ge (typeof a1) m loc ofs v t m' ->
+ assign_loc (typeof a1) m loc ofs v m' ->
exec_stmt e le m (Sassign a1 a2)
- t le m' Out_normal
+ E0 le m' Out_normal
| exec_Sset: forall e le m id a v,
- eval_expr e le m a v ->
+ eval_expr ge e le m a v ->
exec_stmt e le m (Sset id a)
E0 (PTree.set id v le) m Out_normal
- | 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 ->
- eval_exprlist e le m al tyargs vargs ->
- Genv.find_funct ge vf = Some f ->
- type_of_fundef f = Tfunction tyargs tyres ->
- eval_funcall m f vargs t m' vres ->
- exec_stmt e le m (Scall None a al)
- t le m' Out_normal
- | exec_Scall_some: forall e le m id a al tyargs tyres vf vargs f t m' vres,
+ | exec_Scall: forall e le m optid 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 ->
- eval_exprlist e le m al tyargs vargs ->
+ eval_expr ge e le m a vf ->
+ eval_exprlist ge e le m al tyargs vargs ->
Genv.find_funct ge vf = Some f ->
type_of_fundef f = Tfunction tyargs tyres ->
eval_funcall m f vargs t m' vres ->
- exec_stmt e le m (Scall (Some id) a al)
- t (PTree.set id vres le) m' Out_normal
+ exec_stmt e le m (Scall optid a al)
+ t (set_opttemp optid vres le) m' Out_normal
+ | exec_Sbuiltin: forall e le m optid ef al tyargs vargs t m' vres,
+ eval_exprlist ge e le m al tyargs vargs ->
+ external_call ef ge vargs m t vres m' ->
+ exec_stmt e le m (Sbuiltin optid ef tyargs al)
+ t (set_opttemp optid vres le) m' Out_normal
| exec_Sseq_1: forall e le m s1 s2 t1 le1 m1 t2 le2 m2 out,
exec_stmt e le m s1 t1 le1 m1 Out_normal ->
exec_stmt e le1 m1 s2 t2 le2 m2 out ->
@@ -676,7 +735,7 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env ->
exec_stmt e le m (Ssequence s1 s2)
t1 le1 m1 out
| exec_Sifthenelse: forall e le m a s1 s2 v1 b t le' m' out,
- eval_expr e le m a v1 ->
+ eval_expr ge e le m a v1 ->
bool_val v1 (typeof a) = Some b ->
exec_stmt e le m (if b then s1 else s2) t le' m' out ->
exec_stmt e le m (Sifthenelse a s1 s2)
@@ -685,7 +744,7 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env ->
exec_stmt e le m (Sreturn None)
E0 le m (Out_return None)
| exec_Sreturn_some: forall e le m a v,
- eval_expr e le m a v ->
+ eval_expr ge e le m a v ->
exec_stmt e le m (Sreturn (Some a))
E0 le m (Out_return (Some (v, typeof a)))
| exec_Sbreak: forall e le m,
@@ -694,70 +753,27 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env ->
| exec_Scontinue: forall e le m,
exec_stmt e le m Scontinue
E0 le m Out_continue
- | exec_Swhile_false: forall e le m a s v,
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some false ->
- exec_stmt e le m (Swhile a s)
- E0 le m Out_normal
- | exec_Swhile_stop: forall e le m a v s t le' m' out' out,
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some true ->
- exec_stmt e le m s t le' m' out' ->
+ | exec_Sloop_stop1: forall e le m s1 s2 t le' m' out' out,
+ exec_stmt e le m s1 t le' m' out' ->
out_break_or_return out' out ->
- exec_stmt e le m (Swhile a s)
+ exec_stmt e le m (Sloop s1 s2)
t le' m' out
- | exec_Swhile_loop: forall e le m a s v t1 le1 m1 out1 t2 le2 m2 out,
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some true ->
- exec_stmt e le m s t1 le1 m1 out1 ->
+ | exec_Sloop_stop2: forall e le m s1 s2 t1 le1 m1 out1 t2 le2 m2 out2 out,
+ exec_stmt e le m s1 t1 le1 m1 out1 ->
out_normal_or_continue out1 ->
- exec_stmt e le1 m1 (Swhile a s) t2 le2 m2 out ->
- exec_stmt e le m (Swhile a s)
- (t1 ** t2) le2 m2 out
- | exec_Sdowhile_false: forall e le m s a t le1 m1 out1 v,
- exec_stmt e le m s t le1 m1 out1 ->
+ exec_stmt e le1 m1 s2 t2 le2 m2 out2 ->
+ out_break_or_return out2 out ->
+ exec_stmt e le m (Sloop s1 s2)
+ (t1**t2) le2 m2 out
+ | exec_Sloop_loop: forall e le m s1 s2 t1 le1 m1 out1 t2 le2 m2 t3 le3 m3 out,
+ exec_stmt e le m s1 t1 le1 m1 out1 ->
out_normal_or_continue out1 ->
- eval_expr e le1 m1 a v ->
- bool_val v (typeof a) = Some false ->
- exec_stmt e le m (Sdowhile a s)
- t le1 m1 Out_normal
- | exec_Sdowhile_stop: forall e le m s a t le1 m1 out1 out,
- exec_stmt e le m s t le1 m1 out1 ->
- out_break_or_return out1 out ->
- exec_stmt e le m (Sdowhile a s)
- t le1 m1 out
- | exec_Sdowhile_loop: forall e le m s a le1 m1 le2 m2 t1 t2 out out1 v,
- exec_stmt e le m s t1 le1 m1 out1 ->
- out_normal_or_continue out1 ->
- eval_expr e le1 m1 a v ->
- bool_val v (typeof a) = Some true ->
- exec_stmt e le1 m1 (Sdowhile a s) t2 le2 m2 out ->
- exec_stmt e le m (Sdowhile a s)
- (t1 ** t2) le2 m2 out
-
- | exec_Sfor_false: forall e le m s a2 a3 v,
- eval_expr e le m a2 v ->
- bool_val v (typeof a2) = Some false ->
- exec_stmt e le m (Sfor' a2 a3 s)
- E0 le m Out_normal
- | exec_Sfor_stop: forall e le m s a2 a3 v le1 m1 t out1 out,
- eval_expr e le m a2 v ->
- bool_val v (typeof a2) = Some true ->
- exec_stmt e le m s t le1 m1 out1 ->
- out_break_or_return out1 out ->
- exec_stmt e le m (Sfor' a2 a3 s)
- t le1 m1 out
- | exec_Sfor_loop: forall e le m s a2 a3 v le1 m1 le2 m2 le3 m3 t1 t2 t3 out1 out,
- eval_expr e le m a2 v ->
- bool_val v (typeof a2) = Some true ->
- exec_stmt e le m s t1 le1 m1 out1 ->
- out_normal_or_continue out1 ->
- exec_stmt e le1 m1 a3 t2 le2 m2 Out_normal ->
- exec_stmt e le2 m2 (Sfor' a2 a3 s) t3 le3 m3 out ->
- exec_stmt e le m (Sfor' a2 a3 s)
- (t1 ** t2 ** t3) le3 m3 out
+ exec_stmt e le1 m1 s2 t2 le2 m2 Out_normal ->
+ exec_stmt e le2 m2 (Sloop s1 s2) t3 le3 m3 out ->
+ exec_stmt e le m (Sloop s1 s2)
+ (t1**t2**t3) le3 m3 out
| exec_Sswitch: forall e le m a t n sl le1 m1 out,
- eval_expr e le m a (Vint n) ->
+ eval_expr ge e le m a (Vint n) ->
exec_stmt e le m (seq_of_labeled_statement (select_switch n sl)) t le1 m1 out ->
exec_stmt e le m (Sswitch a sl)
t le1 m1 (outcome_switch out)
@@ -770,8 +786,8 @@ 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 ge e m1 f.(fn_params) vargs m2 ->
- exec_stmt e (PTree.empty val) m2 f.(fn_body) t le m3 out ->
+ bind_parameters e m1 f.(fn_params) vargs m2 ->
+ exec_stmt e (create_undef_temps f.(fn_temps)) 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 ->
eval_funcall m (Internal f) vargs t m4 vres
@@ -781,6 +797,7 @@ with eval_funcall: mem -> fundef -> list val -> trace -> mem -> val -> Prop :=
Scheme exec_stmt_ind2 := Minimality for exec_stmt Sort Prop
with eval_funcall_ind2 := Minimality for eval_funcall Sort Prop.
+Combined Scheme exec_stmt_funcall_ind from exec_stmt_ind2, eval_funcall_ind2.
(** ** Big-step semantics for diverging statements and functions *)
@@ -790,22 +807,14 @@ Scheme exec_stmt_ind2 := Minimality for exec_stmt Sort Prop
trace of observable events performed during the execution. *)
CoInductive execinf_stmt: env -> temp_env -> mem -> statement -> traceinf -> Prop :=
- | execinf_Scall_none: forall e le m a al vf tyargs tyres vargs f t,
+ | execinf_Scall: forall e le m optid a al vf tyargs tyres vargs f t,
classify_fun (typeof a) = fun_case_f tyargs tyres ->
- eval_expr e le m a vf ->
- eval_exprlist e le m al tyargs vargs ->
- Genv.find_funct ge vf = Some f ->
- type_of_fundef f = Tfunction tyargs tyres ->
- evalinf_funcall m f vargs t ->
- execinf_stmt e le m (Scall None a al) t
- | execinf_Scall_some: forall e le m id a al vf tyargs tyres vargs f t,
- classify_fun (typeof a) = fun_case_f tyargs tyres ->
- eval_expr e le m a vf ->
- eval_exprlist e le m al tyargs vargs ->
+ eval_expr ge e le m a vf ->
+ eval_exprlist ge e le m al tyargs vargs ->
Genv.find_funct ge vf = Some f ->
type_of_fundef f = Tfunction tyargs tyres ->
evalinf_funcall m f vargs t ->
- execinf_stmt e le m (Scall (Some id) a al) t
+ execinf_stmt e le m (Scall optid a al) t
| execinf_Sseq_1: forall e le m s1 s2 t,
execinf_stmt e le m s1 t ->
execinf_stmt e le m (Ssequence s1 s2) t
@@ -814,54 +823,26 @@ CoInductive execinf_stmt: env -> temp_env -> mem -> statement -> traceinf -> Pro
execinf_stmt e le1 m1 s2 t2 ->
execinf_stmt e le m (Ssequence s1 s2) (t1 *** t2)
| execinf_Sifthenelse: forall e le m a s1 s2 v1 b t,
- eval_expr e le m a v1 ->
+ eval_expr ge e le m a v1 ->
bool_val v1 (typeof a) = Some b ->
execinf_stmt e le m (if b then s1 else s2) t ->
execinf_stmt e le m (Sifthenelse a s1 s2) t
- | execinf_Swhile_body: forall e le m a v s t,
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some true ->
- execinf_stmt e le m s t ->
- execinf_stmt e le m (Swhile a s) t
- | execinf_Swhile_loop: forall e le m a s v t1 le1 m1 out1 t2,
- eval_expr e le m a v ->
- bool_val v (typeof a) = Some true ->
- exec_stmt e le m s t1 le1 m1 out1 ->
- out_normal_or_continue out1 ->
- execinf_stmt e le1 m1 (Swhile a s) t2 ->
- execinf_stmt e le m (Swhile a s) (t1 *** t2)
- | execinf_Sdowhile_body: forall e le m s a t,
- execinf_stmt e le m s t ->
- execinf_stmt e le m (Sdowhile a s) t
- | execinf_Sdowhile_loop: forall e le m s a le1 m1 t1 t2 out1 v,
- exec_stmt e le m s t1 le1 m1 out1 ->
- out_normal_or_continue out1 ->
- eval_expr e le1 m1 a v ->
- bool_val v (typeof a) = Some true ->
- execinf_stmt e le1 m1 (Sdowhile a s) t2 ->
- execinf_stmt e le m (Sdowhile a s) (t1 *** t2)
- | execinf_Sfor_body: forall e le m s a2 a3 v t,
- eval_expr e le m a2 v ->
- bool_val v (typeof a2) = Some true ->
- execinf_stmt e le m s t ->
- execinf_stmt e le m (Sfor' a2 a3 s) t
- | execinf_Sfor_next: forall e le m s a2 a3 v le1 m1 t1 t2 out1,
- eval_expr e le m a2 v ->
- bool_val v (typeof a2) = Some true ->
- exec_stmt e le m s t1 le1 m1 out1 ->
+ | execinf_Sloop_body1: forall e le m s1 s2 t,
+ execinf_stmt e le m s1 t ->
+ execinf_stmt e le m (Sloop s1 s2) t
+ | execinf_Sloop_body2: forall e le m s1 s2 t1 le1 m1 out1 t2,
+ exec_stmt e le m s1 t1 le1 m1 out1 ->
out_normal_or_continue out1 ->
- execinf_stmt e le1 m1 a3 t2 ->
- execinf_stmt e le m (Sfor' a2 a3 s) (t1 *** t2)
- | execinf_Sfor_loop: forall e le m s a2 a3 v le1 m1 le2 m2 t1 t2 t3 out1,
- eval_expr e le m a2 v ->
- bool_val v (typeof a2) = Some true ->
- exec_stmt e le m s t1 le1 m1 out1 ->
+ execinf_stmt e le1 m1 s2 t2 ->
+ execinf_stmt e le m (Sloop s1 s2) (t1***t2)
+ | execinf_Sloop_loop: forall e le m s1 s2 t1 le1 m1 out1 t2 le2 m2 t3,
+ exec_stmt e le m s1 t1 le1 m1 out1 ->
out_normal_or_continue out1 ->
- exec_stmt e le1 m1 a3 t2 le2 m2 Out_normal ->
- execinf_stmt e le2 m2 (Sfor' a2 a3 s) t3 ->
- execinf_stmt e le m (Sfor' a2 a3 s) (t1 *** t2 *** t3)
+ exec_stmt e le1 m1 s2 t2 le2 m2 Out_normal ->
+ execinf_stmt e le2 m2 (Sloop s1 s2) t3 ->
+ execinf_stmt e le m (Sloop s1 s2) (t1***t2***t3)
| execinf_Sswitch: forall e le m a t n sl,
- eval_expr e le m a (Vint n) ->
+ eval_expr ge e le m a (Vint n) ->
execinf_stmt e le m (seq_of_labeled_statement (select_switch n sl)) t ->
execinf_stmt e le m (Sswitch a sl) t
@@ -872,63 +853,11 @@ 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 ge e m1 f.(fn_params) vargs m2 ->
- execinf_stmt e (PTree.empty val) m2 f.(fn_body) t ->
+ bind_parameters e m1 f.(fn_params) vargs m2 ->
+ execinf_stmt e (create_undef_temps f.(fn_temps)) m2 f.(fn_body) t ->
evalinf_funcall m (Internal f) vargs t.
-End SEMANTICS.
-
-(** * Whole-program semantics *)
-
-(** Execution of whole programs are described as sequences of transitions
- from an initial state to a final state. An initial state is a [Callstate]
- corresponding to the invocation of the ``main'' function of the program
- without arguments and with an empty continuation. *)
-
-Inductive initial_state (p: program): state -> Prop :=
- | initial_state_intro: forall b f m0,
- let ge := Genv.globalenv p in
- Genv.init_mem p = Some m0 ->
- Genv.find_symbol ge p.(prog_main) = Some b ->
- Genv.find_funct_ptr ge b = Some f ->
- 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. *)
-
-Inductive final_state: state -> int -> Prop :=
- | final_state_intro: forall r m,
- final_state (Returnstate (Vint r) Kstop m) r.
-
-(** Wrapping up these definitions in a small-step semantics. *)
-
-Definition semantics (p: program) :=
- Semantics step (initial_state p) final_state (Genv.globalenv p).
-
-(** This semantics is receptive to changes in events. *)
-
-Lemma semantics_receptive:
- forall (p: program), 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; 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 *)
- 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.
+End BIGSTEP.
(** Big-step execution of a whole program. *)
@@ -962,13 +891,6 @@ Section BIGSTEP_TO_TRANSITIONS.
Variable prog: program.
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 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 :=
| osm_normal:
@@ -1006,7 +928,7 @@ Lemma exec_stmt_eval_funcall_steps:
is_call_cont k ->
star step ge (Callstate fd args k m) t (Returnstate res k m')).
Proof.
- apply exec_stmt_eval_funcall_ind; intros.
+ apply exec_stmt_funcall_ind; intros.
(* skip *)
econstructor; split. apply star_refl. constructor.
@@ -1017,20 +939,14 @@ Proof.
(* set *)
econstructor; split. apply star_one. econstructor; eauto. constructor.
-(* set volatile *)
- econstructor; split. apply star_one. econstructor; eauto. constructor.
-
-(* call none *)
+(* call *)
econstructor; split.
eapply star_left. econstructor; eauto.
eapply star_right. apply H5. simpl; auto. econstructor. reflexivity. traceEq.
constructor.
-(* call some *)
- econstructor; split.
- eapply star_left. econstructor; eauto.
- eapply star_right. apply H5. simpl; auto. econstructor; eauto. reflexivity. traceEq.
- constructor.
+(* builtin *)
+ econstructor; split. apply star_one. econstructor; eauto. constructor.
(* sequence 2 *)
destruct (H0 f (Kseq s2 k)) as [S1 [A1 B1]]. inv B1.
@@ -1080,112 +996,55 @@ Proof.
(* continue *)
econstructor; split. apply star_refl. constructor.
-(* while false *)
- econstructor; split.
- apply star_one. eapply step_while_false; eauto.
- constructor.
-
-(* while stop *)
- destruct (H2 f (Kwhile a s k)) as [S1 [A1 B1]].
+(* loop stop 1 *)
+ destruct (H0 f (Kloop1 s1 s2 k)) as [S1 [A1 B1]].
set (S2 :=
match out' with
| Out_break => State f Sskip k e le' m'
| _ => S1
end).
exists S2; split.
- eapply star_left. eapply step_while_true; eauto.
- eapply star_trans. eexact A1.
- unfold S2. inversion H3; subst.
- inv B1. apply star_one. constructor.
- apply star_refl.
- reflexivity. traceEq.
- unfold S2. inversion H3; subst. constructor. inv B1; econstructor; eauto.
-
-(* while loop *)
- destruct (H2 f (Kwhile a s k)) as [S1 [A1 B1]].
- destruct (H5 f k) as [S2 [A2 B2]].
- exists S2; split.
- eapply star_left. eapply step_while_true; eauto.
- eapply star_trans. eexact A1.
- eapply star_left.
- inv H3; inv B1; apply step_skip_or_continue_while; auto.
- eexact A2.
- reflexivity. reflexivity. traceEq.
- auto.
-
-(* dowhile false *)
- destruct (H0 f (Kdowhile a s k)) as [S1 [A1 B1]].
- exists (State f Sskip k e le1 m1); split.
- eapply star_left. constructor.
- eapply star_right. eexact A1.
- inv H1; inv B1; eapply step_skip_or_continue_dowhile_false; eauto.
- reflexivity. traceEq.
- constructor.
-
-(* dowhile stop *)
- destruct (H0 f (Kdowhile a s k)) as [S1 [A1 B1]].
- set (S2 :=
- match out1 with
- | Out_break => State f Sskip k e le1 m1
- | _ => S1
- end).
- exists S2; split.
- eapply star_left. apply step_dowhile.
+ eapply star_left. eapply step_loop.
eapply star_trans. eexact A1.
unfold S2. inversion H1; subst.
- inv B1. apply star_one. constructor.
+ inv B1. apply star_one. constructor.
apply star_refl.
reflexivity. traceEq.
unfold S2. inversion H1; subst. constructor. inv B1; econstructor; eauto.
-(* dowhile loop *)
- destruct (H0 f (Kdowhile a s k)) as [S1 [A1 B1]].
- destruct (H5 f k) as [S2 [A2 B2]].
- exists S2; split.
- eapply star_left. apply step_dowhile.
- eapply star_trans. eexact A1.
- eapply star_left.
- inv H1; inv B1; eapply step_skip_or_continue_dowhile_true; eauto.
- eexact A2.
- reflexivity. reflexivity. traceEq.
- auto.
-
-(* for false *)
- econstructor; split.
- eapply star_one. eapply step_for_false; eauto.
- constructor.
-
-(* for stop *)
- destruct (H2 f (Kfor2 a2 a3 s k)) as [S1 [A1 B1]].
- set (S2 :=
- match out1 with
- | Out_break => State f Sskip k e le1 m1
- | _ => S1
+(* loop stop 2 *)
+ destruct (H0 f (Kloop1 s1 s2 k)) as [S1 [A1 B1]].
+ destruct (H3 f (Kloop2 s1 s2 k)) as [S2 [A2 B2]].
+ set (S3 :=
+ match out2 with
+ | Out_break => State f Sskip k e le2 m2
+ | _ => S2
end).
- exists S2; split.
- eapply star_left. eapply step_for_true; eauto.
+ exists S3; split.
+ eapply star_left. eapply step_loop.
eapply star_trans. eexact A1.
- unfold S2. inversion H3; subst.
- inv B1. apply star_one. constructor.
- apply star_refl.
- reflexivity. traceEq.
- unfold S2. inversion H3; subst. constructor. inv B1; econstructor; eauto.
-
-(* for loop *)
- destruct (H2 f (Kfor2 a2 a3 s k)) as [S1 [A1 B1]].
- destruct (H5 f (Kfor3 a2 a3 s k)) as [S2 [A2 B2]]. inv B2.
- destruct (H7 f k) as [S3 [A3 B3]].
+ eapply star_left with (s2 := State f s2 (Kloop2 s1 s2 k) e le1 m1).
+ inv H1; inv B1; constructor; auto.
+ eapply star_trans. eexact A2.
+ unfold S3. inversion H4; subst.
+ inv B2. apply star_one. constructor. apply star_refl.
+ reflexivity. reflexivity. reflexivity. traceEq.
+ unfold S3. inversion H4; subst. constructor. inv B2; econstructor; eauto.
+
+(* loop loop *)
+ destruct (H0 f (Kloop1 s1 s2 k)) as [S1 [A1 B1]].
+ destruct (H3 f (Kloop2 s1 s2 k)) as [S2 [A2 B2]].
+ destruct (H5 f k) as [S3 [A3 B3]].
exists S3; split.
- eapply star_left. eapply step_for_true; eauto.
+ eapply star_left. eapply step_loop.
eapply star_trans. eexact A1.
- eapply star_trans with (s2 := State f a3 (Kfor3 a2 a3 s k) e le1 m1).
- inv H3; inv B1.
- apply star_one. constructor. auto.
- apply star_one. constructor. auto.
- eapply star_trans. eexact A2.
- eapply star_left. constructor.
+ eapply star_left with (s2 := State f s2 (Kloop2 s1 s2 k) e le1 m1).
+ inv H1; inv B1; constructor; auto.
+ eapply star_trans. eexact A2.
+ eapply star_left with (s2 := State f (Sloop s1 s2) k e le2 m2).
+ inversion H4; subst; inv B2; constructor; auto.
eexact A3.
- reflexivity. reflexivity. reflexivity. reflexivity. traceEq.
+ reflexivity. reflexivity. reflexivity. reflexivity. traceEq.
auto.
(* switch *)
@@ -1264,11 +1123,7 @@ Proof.
cofix CIH_STMT.
intros. inv H.
-(* call none *)
- eapply forever_N_plus.
- apply plus_one. eapply step_call; eauto.
- eapply CIH_FUN. eauto. traceEq.
-(* call some *)
+(* call *)
eapply forever_N_plus.
apply plus_one. eapply step_call; eauto.
eapply CIH_FUN. eauto. traceEq.
@@ -1290,65 +1145,29 @@ Proof.
apply plus_one. eapply step_ifthenelse with (b := b); eauto.
apply CIH_STMT; eauto. traceEq.
-(* while body *)
+(* loop body 1 *)
eapply forever_N_plus.
- eapply plus_one. eapply step_while_true; eauto.
+ eapply plus_one. constructor.
apply CIH_STMT; eauto. traceEq.
-(* while loop *)
- destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H2 f (Kwhile a s0 k)) as [S1 [A1 B1]].
- eapply forever_N_plus with (s2 := State f (Swhile a s0) k e le1 m1).
- eapply plus_left. eapply step_while_true; eauto.
+(* loop body 2 *)
+ destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H0 f (Kloop1 s1 s2 k)) as [S1 [A1 B1]].
+ eapply forever_N_plus with (s2 := State f s2 (Kloop2 s1 s2 k) e le1 m1).
+ eapply plus_left. constructor.
eapply star_right. eexact A1.
- inv H3; inv B1; apply step_skip_or_continue_while; auto.
+ inv H1; inv B1; constructor; auto.
reflexivity. reflexivity.
apply CIH_STMT; eauto. traceEq.
-
-(* dowhile body *)
- eapply forever_N_plus.
- eapply plus_one. eapply step_dowhile.
- apply CIH_STMT; eauto.
- traceEq.
-
-(* dowhile loop *)
- destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H0 f (Kdowhile a s0 k)) as [S1 [A1 B1]].
- eapply forever_N_plus with (s2 := State f (Sdowhile a s0) k e le1 m1).
- eapply plus_left. eapply step_dowhile.
- eapply star_right. eexact A1.
- inv H1; inv B1; eapply step_skip_or_continue_dowhile_true; eauto.
- reflexivity. reflexivity.
- apply CIH_STMT. eauto.
- traceEq.
-
-(* for body *)
- eapply forever_N_plus.
- apply plus_one. eapply step_for_true; eauto.
- apply CIH_STMT; eauto.
- traceEq.
-
-(* for next *)
- destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H2 f (Kfor2 a2 a3 s0 k)) as [S1 [A1 B1]].
- eapply forever_N_plus.
- eapply plus_left. eapply step_for_true; eauto.
- eapply star_trans. eexact A1.
- apply star_one.
- inv H3; inv B1; apply step_skip_or_continue_for2; auto.
- reflexivity. reflexivity.
- apply CIH_STMT; eauto.
- traceEq.
-
-(* for body *)
- destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H2 f (Kfor2 a2 a3 s0 k)) as [S1 [A1 B1]].
- destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H4 f (Kfor3 a2 a3 s0 k)) as [S2 [A2 B2]].
- inv B2.
- eapply forever_N_plus.
- eapply plus_left. eapply step_for_true; eauto.
+(* loop loop *)
+ destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H0 f (Kloop1 s1 s2 k)) as [S1 [A1 B1]].
+ destruct (exec_stmt_steps _ _ _ _ _ _ _ _ H2 f (Kloop2 s1 s2 k)) as [S2 [A2 B2]].
+ eapply forever_N_plus with (s2 := State f (Sloop s1 s2) k e le2 m2).
+ eapply plus_left. constructor.
eapply star_trans. eexact A1.
- eapply star_left. inv H3; inv B1; apply step_skip_or_continue_for2; auto.
- eapply star_right. eexact A2.
- constructor.
- reflexivity. reflexivity. reflexivity. reflexivity.
- apply CIH_STMT; eauto.
- traceEq.
+ eapply star_left. inv H1; inv B1; constructor; auto.
+ eapply star_right. eexact A2.
+ inv B2. constructor.
+ reflexivity. reflexivity. reflexivity. reflexivity.
+ apply CIH_STMT; eauto. traceEq.
(* switch *)
eapply forever_N_plus.
@@ -1382,3 +1201,4 @@ Proof.
Qed.
End BIGSTEP_TO_TRANSITIONS.
+
diff --git a/cfrontend/Cminorgen.v b/cfrontend/Cminorgen.v
index c6d6fd5..2c02813 100644
--- a/cfrontend/Cminorgen.v
+++ b/cfrontend/Cminorgen.v
@@ -166,37 +166,6 @@ Definition bge (x y: approx) : bool :=
| _, _ => 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, (Int1|Int7|Int8u) => Int8u
- | Int8u, Int15 => Int15
- | Int8u, Int16u => Int16u
- | Int8u, Int16s => Int16s
- | Int8s, (Int1|Int7|Int8s) => Int8s
- | Int8s, (Int15|Int16s) => Int16s
- | Int15, (Int1|Int7|Int8u|Int15) => Int15
- | Int15, Int16u => Int16u
- | Int15, (Int8s|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 || Int.eq_dec n Int.one then Int1
else if Int.eq_dec n (Int.zero_ext 7 n) then Int7
@@ -229,8 +198,6 @@ Definition unop (op: unary_operation) (a: approx) :=
| Ocast16unsigned => Int16u
| Ocast16signed => Int16s
| Osingleoffloat => Float32
- | Oboolval => Int1
- | Onotbool => Int1
| _ => Any
end.
@@ -240,7 +207,6 @@ 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.
@@ -366,11 +332,6 @@ Fixpoint transl_expr (cenv: compilenv) (e: Csharpminor.expr)
| Csharpminor.Eload chunk e =>
do (te, a) <- transl_expr cenv e;
OK (Eload chunk te, Approx.of_chunk chunk)
- | Csharpminor.Econdition e1 e2 e3 =>
- do (te1, a1) <- transl_expr cenv e1;
- do (te2, a2) <- transl_expr cenv e2;
- do (te3, a3) <- transl_expr cenv e3;
- OK (Econdition te1 te2 te3, Approx.lub a2 a3)
end.
Fixpoint transl_exprlist (cenv: compilenv) (el: list Csharpminor.expr)
@@ -515,9 +476,6 @@ Fixpoint addr_taken_expr (e: Csharpminor.expr): Identset.t :=
| Csharpminor.Ebinop op e1 e2 =>
Identset.union (addr_taken_expr e1) (addr_taken_expr e2)
| Csharpminor.Eload chunk e => addr_taken_expr e
- | Csharpminor.Econdition e1 e2 e3 =>
- Identset.union (addr_taken_expr e1)
- (Identset.union (addr_taken_expr e2) (addr_taken_expr e3))
end.
Fixpoint addr_taken_exprlist (e: list Csharpminor.expr): Identset.t :=
diff --git a/cfrontend/Cminorgenproof.v b/cfrontend/Cminorgenproof.v
index ea5d68e..0fa9ac0 100644
--- a/cfrontend/Cminorgenproof.v
+++ b/cfrontend/Cminorgenproof.v
@@ -987,12 +987,13 @@ Proof.
Qed.
Lemma match_callstack_alloc_right:
- forall f m tm cs tf sp tm' te,
+ forall f le m tm cs tf sp tm' te,
match_callstack f m tm cs (Mem.nextblock m) (Mem.nextblock tm) ->
Mem.alloc tm 0 tf.(fn_stackspace) = (tm', sp) ->
Mem.inject f m tm ->
+ (forall id v, le!id = Some v -> exists v', te!(for_temp id) = Some v' /\ val_inject f v v') ->
match_callstack f m tm'
- (Frame gce tf empty_env empty_temp_env te sp (Mem.nextblock m) (Mem.nextblock m) :: cs)
+ (Frame gce tf empty_env le te sp (Mem.nextblock m) (Mem.nextblock m) :: cs)
(Mem.nextblock m) (Mem.nextblock tm').
Proof.
intros.
@@ -1007,13 +1008,13 @@ Proof.
constructor. apply PTree.gempty. auto.
constructor. apply PTree.gempty.
(* temps *)
- intros. rewrite PTree.gempty in H2. congruence.
+ assumption.
(* low high *)
omega.
(* bounded *)
- intros. rewrite PTree.gempty in H2. congruence.
+ intros. rewrite PTree.gempty in H3. congruence.
(* inj *)
- intros. rewrite PTree.gempty in H2. congruence.
+ intros. rewrite PTree.gempty in H3. congruence.
(* inv *)
intros.
assert (sp <> sp). apply Mem.valid_not_valid_diff with tm.
@@ -1023,7 +1024,7 @@ Proof.
intros. rewrite RES. change (Mem.valid_block tm tb).
eapply Mem.valid_block_inject_2; eauto.
(* bounds *)
- unfold empty_env; intros. rewrite PTree.gempty in H2. congruence.
+ unfold empty_env; intros. rewrite PTree.gempty in H3. congruence.
(* padding freeable *)
red; intros. left. eapply Mem.perm_alloc_2; eauto.
(* previous call stack *)
@@ -1183,18 +1184,6 @@ Proof.
unfold Approx.bge in H; destruct a1; try discriminate; destruct a2; simpl in *; try discriminate; intuition.
Qed.
-Lemma approx_lub_ge_left:
- forall x y, Approx.bge (Approx.lub x y) x = true.
-Proof.
- destruct x; destruct y; auto.
-Qed.
-
-Lemma approx_lub_ge_right:
- forall x y, Approx.bge (Approx.lub x y) y = true.
-Proof.
- destruct x; destruct y; auto.
-Qed.
-
Lemma approx_of_int_sound:
forall n, val_match_approx (Approx.of_int n) (Vint n).
Proof.
@@ -1243,8 +1232,6 @@ 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); simpl; auto.
- destruct v1; simpl; auto. destruct (Int.eq i Int.zero); simpl; auto.
destruct v1; simpl; auto. rewrite Float.singleoffloat_idem; auto.
Qed.
@@ -1364,20 +1351,6 @@ 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.
@@ -1423,9 +1396,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.
inv H; inv H0; simpl; TrivialExists.
inv H; inv H0; simpl; TrivialExists.
@@ -2383,6 +2354,16 @@ Proof.
left; auto.
Qed.
+Lemma create_undef_temps_val:
+ forall id v temps, (create_undef_temps temps)!id = Some v -> In id temps /\ v = Vundef.
+Proof.
+ induction temps; simpl; intros.
+ rewrite PTree.gempty in H. congruence.
+ rewrite PTree.gsspec in H. destruct (peq id a).
+ split. auto. congruence.
+ exploit IHtemps; eauto. tauto.
+Qed.
+
(** Preservation of [match_callstack] by simultaneous allocation
of Csharpminor local variables and of the Cminor stack data block. *)
@@ -2403,7 +2384,7 @@ Lemma match_callstack_alloc_variables:
inject_incr f f'
/\ Mem.inject f' m' tm'
/\ match_callstack f' m' tm'
- (Frame cenv tf e empty_temp_env te sp (Mem.nextblock m) (Mem.nextblock m') :: cs)
+ (Frame cenv tf e (create_undef_temps (Csharpminor.fn_temps fn)) te sp (Mem.nextblock m) (Mem.nextblock m') :: cs)
(Mem.nextblock m') (Mem.nextblock tm').
Proof.
intros.
@@ -2411,6 +2392,11 @@ Proof.
eapply match_callstack_alloc_variables_rec; eauto with mem.
red; intros. eapply Mem.perm_alloc_2; eauto.
eapply match_callstack_alloc_right; eauto.
+ intros. exploit create_undef_temps_val; eauto. intros [A B]. subst v.
+ assert (te!(for_temp id) <> None).
+ unfold te. apply set_locals_defined. left.
+ apply in_or_app. right. apply in_map. auto.
+ destruct (te!(for_temp id)). exists v; auto. congruence.
eapply Mem.alloc_right_inject; eauto. omega.
intros. elim (Mem.valid_not_valid_diff tm sp sp); eauto with mem.
eapply Mem.valid_block_inject_2; eauto.
@@ -2624,7 +2610,7 @@ Lemma function_entry_ok:
/\ Mem.inject f2 m2 tm2
/\ inject_incr f f2
/\ match_callstack f2 m2 tm2
- (Frame cenv tf e empty_temp_env te sp (Mem.nextblock m) (Mem.nextblock m1) :: cs)
+ (Frame cenv tf e (create_undef_temps (Csharpminor.fn_temps fn)) te sp (Mem.nextblock m) (Mem.nextblock m1) :: cs)
(Mem.nextblock m2) (Mem.nextblock tm2).
Proof.
intros.
@@ -2741,17 +2727,6 @@ Proof.
exploit Mem.loadv_inject; eauto. intros [tv [LOAD INJ]].
exists tv; split. econstructor; eauto. split. auto.
destruct v1; simpl in H0; try discriminate. eapply approx_of_chunk_sound; eauto.
- (* Econdition *)
- exploit IHeval_expr1; eauto. intros [tv1 [EVAL1 [INJ1 APP1]]].
- assert (transl_expr cenv (if vb1 then b else c) =
- OK ((if vb1 then x1 else x3), (if vb1 then x2 else x4))).
- destruct vb1; auto.
- exploit IHeval_expr2; eauto. intros [tv2 [EVAL2 [INJ2 APP2]]].
- exists tv2; split. eapply eval_Econdition; eauto.
- eapply bool_of_val_inject; eauto.
- split. auto.
- apply val_match_approx_increasing with (if vb1 then x2 else x4); auto.
- destruct vb1. apply approx_lub_ge_left. apply approx_lub_ge_right.
Qed.
Lemma transl_exprlist_correct:
diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v
index ac7a58f..ddfbeaf 100644
--- a/cfrontend/Csem.v
+++ b/cfrontend/Csem.v
@@ -93,18 +93,18 @@ 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_p2bool =>
+ match v with
+ | Vint i => Some (Vint (cast_int_int IBool Signed i))
+ | Vptr _ _ => Some (Vint 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 =>
@@ -120,12 +120,25 @@ Function sem_cast (v: val) (t1 t2: type) : option val :=
considered as true. The integer zero (which also represents
the null pointer) and the float 0.0 are false. *)
-Function bool_val (v: val) (t: type) : option bool :=
- match v, t with
- | 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
+Definition bool_val (v: val) (t: type) : option bool :=
+ match classify_bool t with
+ | bool_case_i =>
+ match v with
+ | Vint n => Some (negb (Int.eq n Int.zero))
+ | _ => None
+ end
+ | bool_case_f =>
+ match v with
+ | Vfloat f => Some (negb (Float.cmp Ceq f Float.zero))
+ | _ => None
+ end
+ | bool_case_p =>
+ match v with
+ | Vint n => Some (negb (Int.eq n Int.zero))
+ | Vptr b ofs => Some true
+ | _ => None
+ end
+ | bool_default => None
end.
(** The following [sem_] functions compute the result of an operator
@@ -172,10 +185,9 @@ Function sem_notint (v: val) (ty: type): option val :=
Function sem_notbool (v: val) (ty: type) : option val :=
match classify_bool ty with
- | bool_case_ip =>
+ | bool_case_i =>
match v with
| Vint n => Some (Val.of_bool (Int.eq n Int.zero))
- | Vptr _ _ => Some Vfalse
| _ => None
end
| bool_case_f =>
@@ -183,6 +195,12 @@ Function sem_notbool (v: val) (ty: type) : option val :=
| Vfloat f => Some (Val.of_bool (Float.cmp Ceq f Float.zero))
| _ => None
end
+ | bool_case_p =>
+ match v with
+ | Vint n => Some (Val.of_bool (Int.eq n Int.zero))
+ | Vptr _ _ => Some Vfalse
+ | _ => None
+ end
| bool_default => None
end.
@@ -502,7 +520,17 @@ Lemma cast_bool_bool_val:
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.
+ intros.
+ assert (A: classify_bool t =
+ match t with
+ | Tint _ _ _ => bool_case_i
+ | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ => bool_case_p
+ | Tfloat _ _ => bool_case_f
+ | _ => bool_default
+ end).
+ unfold classify_bool; destruct t; simpl; auto. destruct i; auto. destruct s; auto.
+
+ unfold bool_val. rewrite A. unfold sem_cast. destruct t; simpl; auto; 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.
@@ -515,14 +543,8 @@ Lemma notbool_bool_val:
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.
+ intros. unfold sem_notbool, bool_val.
+ destruct (classify_bool t); auto; destruct v; auto; rewrite negb_involutive; auto.
Qed.
(** * Operational semantics *)
@@ -657,6 +679,15 @@ Fixpoint seq_of_labeled_statement (sl: labeled_statements) : statement :=
| LScase c s sl' => Ssequence s (seq_of_labeled_statement sl')
end.
+(** Extract the values from a list of function arguments *)
+
+Inductive cast_arguments: exprlist -> typelist -> list val -> Prop :=
+ | cast_args_nil:
+ cast_arguments Enil Tnil nil
+ | cast_args_cons: forall v ty el targ1 targs v1 vl,
+ sem_cast v ty targ1 = Some v1 -> cast_arguments el targs vl ->
+ cast_arguments (Econs (Eval v ty) el) (Tcons targ1 targs) (v1 :: vl).
+
Section SEMANTICS.
Variable ge: genv.
@@ -731,6 +762,22 @@ Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop :=
sem_cast v1 ty1 ty = Some v ->
rred (Ecast (Eval v1 ty1) ty) m
E0 (Eval v ty) m
+ | red_seqand_true: forall v1 ty1 r2 ty m,
+ bool_val v1 ty1 = Some true ->
+ rred (Eseqand (Eval v1 ty1) r2 ty) m
+ E0 (Eparen (Eparen r2 type_bool) ty) m
+ | red_seqand_false: forall v1 ty1 r2 ty m,
+ bool_val v1 ty1 = Some false ->
+ rred (Eseqand (Eval v1 ty1) r2 ty) m
+ E0 (Eval (Vint Int.zero) ty) m
+ | red_seqor_true: forall v1 ty1 r2 ty m,
+ bool_val v1 ty1 = Some true ->
+ rred (Eseqor (Eval v1 ty1) r2 ty) m
+ E0 (Eval (Vint Int.one) ty) m
+ | red_seqor_false: forall v1 ty1 r2 ty m,
+ bool_val v1 ty1 = Some false ->
+ rred (Eseqor (Eval v1 ty1) r2 ty) m
+ E0 (Eparen (Eparen r2 type_bool) 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
@@ -766,18 +813,17 @@ Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop :=
| red_paren: forall v1 ty1 ty m v,
sem_cast v1 ty1 ty = Some v ->
rred (Eparen (Eval v1 ty1) ty) m
- E0 (Eval v ty) m.
+ E0 (Eval v ty) m
+ | red_builtin: forall ef tyargs el ty m vargs t vres m',
+ cast_arguments el tyargs vargs ->
+ external_call ef ge vargs m t vres m' ->
+ rred (Ebuiltin ef tyargs el ty) m
+ t (Eval vres ty) m'.
+
(** Head reduction for function calls.
(More exactly, identification of function calls that can reduce.) *)
-Inductive cast_arguments: exprlist -> typelist -> list val -> Prop :=
- | cast_args_nil:
- cast_arguments Enil Tnil nil
- | cast_args_cons: forall v ty el targ1 targs v1 vl,
- sem_cast v ty targ1 = Some v1 -> cast_arguments el targs vl ->
- cast_arguments (Econs (Eval v ty) el) (Tcons targ1 targs) (v1 :: vl).
-
Inductive callred: expr -> fundef -> list val -> type -> Prop :=
| red_Ecall: forall vf tyf tyargs tyres el ty fd vargs,
Genv.find_funct ge vf = Some fd ->
@@ -791,13 +837,16 @@ Inductive callred: expr -> fundef -> list val -> type -> Prop :=
we allow reduction both to the left and to the right of a binary operator.
To enforce C's notion of sequence point, reductions within a conditional
[a ? b : c] can only take place in [a], not in [b] nor [c];
- and reductions within a sequence [a, b] can only take place in [a], not in [b].
-
- Reduction contexts are represented by functions [C] from expressions to expressions,
- suitably constrained by the [context from to C] predicate below.
- Contexts are "kinded" with respect to l-values and r-values:
- [from] is the kind of the hole in the context and [to] is the kind of
- the term resulting from filling the hole.
+ reductions within a sequential "or" / "and" [a && b] or [a || b] can
+ only take place in [a], not in [b];
+ and reductions within a sequence [a, b] can only take place in [a],
+ not in [b].
+
+ Reduction contexts are represented by functions [C] from expressions
+ to expressions, suitably constrained by the [context from to C]
+ predicate below. Contexts are "kinded" with respect to l-values and
+ r-values: [from] is the kind of the hole in the context and [to] is
+ the kind of the term resulting from filling the hole.
*)
Inductive kind : Type := LV | RV.
@@ -821,6 +870,10 @@ Inductive context: kind -> kind -> (expr -> expr) -> Prop :=
context k RV C -> context k RV (fun x => Ebinop op e1 (C x) ty)
| ctx_cast: forall k C ty,
context k RV C -> context k RV (fun x => Ecast (C x) ty)
+ | ctx_seqand: forall k C r2 ty,
+ context k RV C -> context k RV (fun x => Eseqand (C x) r2 ty)
+ | ctx_seqor: forall k C r2 ty,
+ context k RV C -> context k RV (fun x => Eseqor (C x) r2 ty)
| ctx_condition: forall k C r2 r3 ty,
context k RV C -> context k RV (fun x => Econdition (C x) r2 r3 ty)
| ctx_assign_left: forall k C e2 ty,
@@ -837,6 +890,8 @@ Inductive context: kind -> kind -> (expr -> expr) -> Prop :=
context k RV C -> context k RV (fun x => Ecall (C x) el ty)
| ctx_call_right: forall k C e1 ty,
contextlist k C -> context k RV (fun x => Ecall e1 (C x) ty)
+ | ctx_builtin: forall k C ef tyargs ty,
+ contextlist k C -> context k RV (fun x => Ebuiltin ef tyargs (C x) ty)
| ctx_comma: forall k C e2 ty,
context k RV C -> context k RV (fun x => Ecomma (C x) e2 ty)
| ctx_paren: forall k C ty,
@@ -912,7 +967,7 @@ Inductive cont: Type :=
| Kdowhile2: expr -> statement -> cont -> cont (**r [Kdowhile2 x s k] = after [x] in [do s while (x)] *)
| Kfor2: expr -> statement -> statement -> cont -> cont (**r [Kfor2 e2 e3 s k] = after [e2] in [for(e1;e2;e3) s] *)
| Kfor3: expr -> statement -> statement -> cont -> cont (**r [Kfor3 e2 e3 s k] = after [s] in [for(e1;e2;e3) s] *)
- | Kfor4: expr -> statement -> statement -> cont -> cont (**r [Kfor3 e2 e3 s k] = after [e3] in [for(e1;e2;e3) s] *)
+ | Kfor4: expr -> statement -> statement -> cont -> cont (**r [Kfor4 e2 e3 s k] = after [e3] in [for(e1;e2;e3) s] *)
| Kswitch1: labeled_statements -> cont -> cont (**r [Kswitch1 ls k] = after [e] in [switch(e) { ls }] *)
| Kswitch2: cont -> cont (**r catches [break] statements arising out of [switch] *)
| Kreturn: cont -> cont (**r [Kreturn k] = after [e] in [return e;] *)
@@ -1256,5 +1311,6 @@ Proof.
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.
+ eapply external_call_trace_length; eauto.
inv H; simpl; try omega. eapply external_call_trace_length; eauto.
Qed.
diff --git a/cfrontend/Csharpminor.v b/cfrontend/Csharpminor.v
index b4993e7..b267891 100644
--- a/cfrontend/Csharpminor.v
+++ b/cfrontend/Csharpminor.v
@@ -51,8 +51,7 @@ Inductive expr : Type :=
| Econst : constant -> expr (**r constants *)
| Eunop : unary_operation -> expr -> expr (**r unary operation *)
| Ebinop : binary_operation -> expr -> expr -> expr (**r binary operation *)
- | Eload : memory_chunk -> expr -> expr (**r memory read *)
- | Econdition : expr -> expr -> expr -> expr. (**r conditional expression *)
+ | Eload : memory_chunk -> expr -> expr. (**r memory read *)
(** Statements include expression evaluation, variable assignment,
memory stores, function calls, an if/then/else conditional,
@@ -155,7 +154,14 @@ Definition env := PTree.t (block * var_kind).
Definition temp_env := PTree.t val.
Definition empty_env : env := PTree.empty (block * var_kind).
-Definition empty_temp_env : temp_env := PTree.empty val.
+
+(** Initialization of temporary variables *)
+
+Fixpoint create_undef_temps (temps: list ident) : temp_env :=
+ match temps with
+ | nil => PTree.empty val
+ | id :: temps' => PTree.set id Vundef (create_undef_temps temps')
+ end.
(** Continuations *)
@@ -395,12 +401,7 @@ Inductive eval_expr: expr -> val -> Prop :=
| eval_Eload: forall chunk a v1 v,
eval_expr a v1 ->
Mem.loadv chunk m v1 = Some v ->
- eval_expr (Eload chunk a) v
- | eval_Econdition: forall a b c v1 vb1 v2,
- eval_expr a v1 ->
- Val.bool_of_val v1 vb1 ->
- eval_expr (if vb1 then b else c) v2 ->
- eval_expr (Econdition a b c) v2.
+ eval_expr (Eload chunk a) v.
(** Evaluation of a list of expressions:
[eval_exprlist prg e m al vl] states that the list [al] of
@@ -530,7 +531,7 @@ Inductive step: state -> trace -> state -> Prop :=
alloc_variables empty_env m (fn_variables f) e m1 ->
bind_parameters e m1 f.(fn_params) vargs m2 ->
step (Callstate (Internal f) vargs k m)
- E0 (State f f.(fn_body) k e empty_temp_env m2)
+ E0 (State f f.(fn_body) k e (create_undef_temps f.(fn_temps)) m2)
| step_external_function: forall ef vargs k m t vres m',
external_call ef ge vargs m t vres m' ->
diff --git a/cfrontend/Cshmgen.v b/cfrontend/Cshmgen.v
index b635b7d..f611902 100644
--- a/cfrontend/Cshmgen.v
+++ b/cfrontend/Cshmgen.v
@@ -94,32 +94,39 @@ Definition make_intoffloat (e: expr) (sg: signedness) :=
end.
(** [make_boolean e ty] returns a Csharpminor expression that evaluates
- to the boolean value of [e]. Recall that:
-- in Csharpminor, [false] is the integer 0,
- [true] any non-zero integer or any pointer
-- in C, [false] is the integer 0, the null pointer, the float 0.0
- [true] is any non-zero integer, non-null pointer, non-null float.
-*)
-Definition make_boolean (e: expr) (ty: type) :=
- match ty with
- | Tfloat _ _ => Ebinop (Ocmpf Cne) e (make_floatconst Float.zero)
- | _ => e
+ to the boolean value of [e]. *)
+
+Definition make_cmp_ne_zero (e: expr) :=
+ match e with
+ | Ebinop (Ocmp c) e1 e2 => e
+ | Ebinop (Ocmpu c) e1 e2 => e
+ | Ebinop (Ocmpf c) e1 e2 => e
+ | _ => Ebinop (Ocmp Cne) e (make_intconst Int.zero)
end.
-Definition make_neg (e: expr) (ty: type) :=
- match classify_neg ty with
- | neg_case_i _ => OK (Eunop Onegint e)
- | neg_case_f => OK (Eunop Onegf e)
- | _ => Error (msg "Cshmgen.make_neg")
+Definition make_boolean (e: expr) (ty: type) :=
+ match classify_bool ty with
+ | bool_case_i => make_cmp_ne_zero e
+ | bool_case_f => Ebinop (Ocmpf Cne) e (make_floatconst Float.zero)
+ | bool_case_p => Ebinop (Ocmpu Cne) e (make_intconst Int.zero)
+ | bool_default => e (**r should not happen *)
end.
Definition make_notbool (e: expr) (ty: type) :=
match classify_bool ty with
- | bool_case_ip => OK (Eunop Onotbool e)
+ | bool_case_i => OK (Ebinop (Ocmp Ceq) e (make_intconst Int.zero))
| bool_case_f => OK (Ebinop (Ocmpf Ceq) e (make_floatconst Float.zero))
+ | bool_case_p => OK (Ebinop (Ocmpu Ceq) e (make_intconst Int.zero))
| _ => Error (msg "Cshmgen.make_notbool")
end.
+Definition make_neg (e: expr) (ty: type) :=
+ match classify_neg ty with
+ | neg_case_i _ => OK (Eunop Onegint e)
+ | neg_case_f => OK (Eunop Onegf e)
+ | _ => Error (msg "Cshmgen.make_neg")
+ end.
+
Definition make_notint (e: expr) (ty: type) :=
Eunop Onotint e.
@@ -219,7 +226,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
+ | IBool, _ => make_cmp_ne_zero e
end.
Definition make_cast_float (e: expr) (sz: floatsize) :=
@@ -235,8 +242,8 @@ 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_p2bool => Ebinop (Ocmpu Cne) e (make_intconst Int.zero)
| cast_case_struct id1 fld1 id2 fld2 => e
| cast_case_union id1 fld1 id2 fld2 => e
| cast_case_void => e
@@ -257,18 +264,6 @@ Definition make_load (addr: expr) (ty_res: type) :=
| By_nothing => Error (msg "Cshmgen.make_load")
end.
-(** [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]. *)
@@ -288,16 +283,6 @@ Definition make_store (addr: expr) (ty: type) (rhs: expr) :=
| _ => 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.
-
(** * Reading and writing variables *)
(** Determine if a C expression is a variable *)
@@ -396,13 +381,6 @@ Fixpoint transl_expr (a: Clight.expr) {struct a} : res expr :=
| Clight.Ecast b ty =>
do tb <- transl_expr b;
OK (make_cast (typeof b) ty tb)
- | Clight.Econdition b c d ty =>
- do tb <- transl_expr b;
- do tc <- transl_expr c;
- do td <- transl_expr d;
- OK(Econdition (make_boolean tb (typeof b))
- (make_cast (typeof c) ty tc)
- (make_cast (typeof d) ty td))
| Clight.Efield b i ty =>
match typeof b with
| Tstruct _ fld _ =>
@@ -445,52 +423,25 @@ with transl_lvalue (a: Clight.expr) {struct a} : res expr :=
Error(msg "Cshmgen.transl_lvalue")
end.
-(** [transl_exprlist al tyl] returns a list of Csharpminor expressions
+(** [transl_arglist al tyl] returns a list of Csharpminor expressions
that compute the values of the list [al] of Csyntax expressions,
casted to the corresponding types in [tyl].
Used for function applications. *)
-Fixpoint transl_exprlist (al: list Clight.expr) (tyl: typelist)
+Fixpoint transl_arglist (al: list Clight.expr) (tyl: typelist)
{struct al}: res (list expr) :=
match al, tyl with
| nil, Tnil => OK nil
| a1 :: a2, Tcons ty1 ty2 =>
do ta1 <- transl_expr a1;
- do ta2 <- transl_exprlist a2 ty2;
+ do ta2 <- transl_arglist a2 ty2;
OK (make_cast (typeof a1) ty1 ta1 :: ta2)
| _, _ =>
- Error(msg "Cshmgen.transl_exprlist: arity mismatch")
+ Error(msg "Cshmgen.transl_arglist: arity mismatch")
end.
(** * Translation of statements *)
-(** [exit_if_false e] return the statement that tests the boolean
- value of the Clight expression [e]. If [e] evaluates to false,
- an [exit 0] is performed. If [e] evaluates to true, the generated
- statement continues in sequence.
-
- The Clight code generated by [SimplExpr] contains many [while(1)]
- and [for(;1;...)] loops, so we optimize the case where [e] is a constant.
- *)
-
-Definition is_constant_bool (e: expr) : option bool :=
- match e with
- | Econst (Ointconst n) => Some (negb (Int.eq n Int.zero))
- | _ => None
- end.
-
-Definition exit_if_false (e: Clight.expr) : res stmt :=
- do te <- transl_expr e;
- match is_constant_bool te with
- | Some true => OK(Sskip)
- | Some false => OK(Sexit 0%nat)
- | None =>
- OK(Sifthenelse
- (make_boolean te (typeof e))
- Sskip
- (Sexit 0%nat))
- end.
-
(** [transl_statement nbrk ncnt s] returns a Csharpminor statement
that performs the same computations as the CabsCoq statement [s].
@@ -504,34 +455,14 @@ Definition exit_if_false (e: Clight.expr) : res stmt :=
The general translation for loops is as follows:
<<
-while (e1) s; ---> block {
- loop {
- if (!e1) exit 0;
- block { s }
- // continue in s branches here
- }
- }
- // break in s branches here
-
-do s; while (e1); ---> block {
- loop {
- block { s }
- // continue in s branches here
- if (!e1) exit 0;
- }
- }
- // break in s branches here
-
-for (;e2;e3) s; ---> block {
+loop s1 s2 ---> block {
loop {
- if (!e2) exit 0;
- block { s }
- // continue in s branches here
- e3;
+ block { s1 };
+ // continue in s1 branches here
+ s2;
}
}
- // break in s branches here
->>
+ // break in s1 and s2 branches here
*)
Fixpoint transl_statement (tyret: type) (nbrk ncnt: nat)
@@ -540,11 +471,7 @@ Fixpoint transl_statement (tyret: type) (nbrk ncnt: nat)
| Clight.Sskip =>
OK Sskip
| Clight.Sassign b c =>
- 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
+ match is_variable b with
| Some id =>
do tc <- transl_expr c;
var_set id (typeof b) (make_cast (typeof c) (typeof b) tc)
@@ -556,17 +483,17 @@ 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 =>
do tb <- transl_expr b;
- do tcl <- transl_exprlist cl args;
+ do tcl <- transl_arglist cl args;
OK(Scall x (signature_of_type args res) tb tcl)
| _ => Error(msg "Cshmgen.transl_stmt(call)")
end
+ | Clight.Sbuiltin x ef tyargs bl =>
+ do tbl <- transl_arglist bl tyargs;
+ OK(Sbuiltin x ef tbl)
| Clight.Ssequence s1 s2 =>
do ts1 <- transl_statement tyret nbrk ncnt s1;
do ts2 <- transl_statement tyret nbrk ncnt s2;
@@ -576,19 +503,10 @@ Fixpoint transl_statement (tyret: type) (nbrk ncnt: nat)
do ts1 <- transl_statement tyret nbrk ncnt s1;
do ts2 <- transl_statement tyret nbrk ncnt s2;
OK (Sifthenelse (make_boolean te (typeof e)) ts1 ts2)
- | Clight.Swhile e s1 =>
- do te <- exit_if_false e;
- do ts1 <- transl_statement tyret 1%nat 0%nat s1;
- OK (Sblock (Sloop (Sseq te (Sblock ts1))))
- | Clight.Sdowhile e s1 =>
- do te <- exit_if_false e;
- do ts1 <- transl_statement tyret 1%nat 0%nat s1;
- OK (Sblock (Sloop (Sseq (Sblock ts1) te)))
- | Clight.Sfor' e2 e3 s1 =>
- do te2 <- exit_if_false e2;
- do te3 <- transl_statement tyret 0%nat (S ncnt) e3;
+ | Clight.Sloop s1 s2 =>
do ts1 <- transl_statement tyret 1%nat 0%nat s1;
- OK (Sblock (Sloop (Sseq te2 (Sseq (Sblock ts1) te3))))
+ do ts2 <- transl_statement tyret 0%nat (S ncnt) s2;
+ OK (Sblock (Sloop (Sseq (Sblock ts1) ts2)))
| Clight.Sbreak =>
OK (Sexit nbrk)
| Clight.Scontinue =>
diff --git a/cfrontend/Cshmgenproof.v b/cfrontend/Cshmgenproof.v
index 6a02e1d..51511b9 100644
--- a/cfrontend/Cshmgenproof.v
+++ b/cfrontend/Cshmgenproof.v
@@ -150,6 +150,19 @@ Proof.
intros. functional inversion H; subst; simpl in H0; congruence.
Qed.
+Remark cast_to_bool_normalized:
+ forall ty1 ty2 chunk,
+ classify_cast ty1 ty2 = cast_case_f2bool \/
+ classify_cast ty1 ty2 = cast_case_p2bool ->
+ access_mode ty2 = By_value chunk ->
+ chunk = Mint8unsigned.
+Proof.
+ intros. destruct ty2; simpl in *; try discriminate.
+ destruct i; destruct ty1; intuition congruence.
+ destruct ty1; intuition discriminate.
+ destruct ty1; intuition discriminate.
+Qed.
+
Lemma cast_result_normalized:
forall chunk v1 ty1 ty2 v2,
sem_cast v1 ty1 ty2 = Some v2 ->
@@ -163,14 +176,10 @@ 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.
+ rewrite (cast_to_bool_normalized ty1 ty2 chunk); auto. red; auto.
+ rewrite (cast_to_bool_normalized ty1 ty2 chunk); auto. red; auto.
+ rewrite (cast_to_bool_normalized ty1 ty2 chunk); auto. destruct (Int.eq i Int.zero); red; auto.
+ rewrite (cast_to_bool_normalized ty1 ty2 chunk); auto. 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.
@@ -376,6 +385,31 @@ Hint Resolve make_intconst_correct make_floatconst_correct
eval_Eunop eval_Ebinop: cshm.
Hint Extern 2 (@eq trace _ _) => traceEq: cshm.
+Lemma make_cmp_ne_zero_correct:
+ forall e le m a n,
+ eval_expr ge e le m a (Vint n) ->
+ eval_expr ge e le m (make_cmp_ne_zero a) (Vint (if Int.eq n Int.zero then Int.zero else Int.one)).
+Proof.
+ intros.
+ assert (DEFAULT: eval_expr ge e le m (Ebinop (Ocmp Cne) a (make_intconst Int.zero))
+ (Vint (if Int.eq n Int.zero then Int.zero else Int.one))).
+ econstructor; eauto with cshm. simpl. unfold Val.cmp, Val.cmp_bool.
+ unfold Int.cmp. destruct (Int.eq n Int.zero); auto.
+ assert (CMP: forall ob,
+ Val.of_optbool ob = Vint n ->
+ n = (if Int.eq n Int.zero then Int.zero else Int.one)).
+ intros. destruct ob; simpl in H0; inv H0. destruct b; inv H2.
+ rewrite Int.eq_false. auto. apply Int.one_not_zero.
+ rewrite Int.eq_true. auto.
+ destruct a; simpl; auto. destruct b; auto.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ simpl in H6. inv H6. unfold Val.cmp in H0. eauto.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ simpl in H6. inv H6. unfold Val.cmp in H0. eauto.
+ inv H. econstructor; eauto. rewrite H6. decEq. decEq.
+ simpl in H6. inv H6. unfold Val.cmp in H0. eauto.
+Qed.
+
Lemma make_cast_int_correct:
forall e le m a n sz si,
eval_expr ge e le m a (Vint n) ->
@@ -386,7 +420,7 @@ Proof.
destruct si; eauto with cshm.
destruct si; eauto with cshm.
auto.
- econstructor. eauto. simpl. destruct (Int.eq n Int.zero); auto.
+ apply make_cmp_ne_zero_correct; auto.
Qed.
Lemma make_cast_float_correct:
@@ -418,14 +452,15 @@ Proof.
rewrite H2. auto with cshm.
(* float -> int *)
rewrite H2. eauto with cshm.
- (* 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.
+ (* pointer -> bool *)
+ rewrite H2. econstructor; eauto with cshm.
+ simpl. unfold Val.cmpu, Val.cmpu_bool, Int.cmpu. destruct (Int.eq i Int.zero); reflexivity.
+ rewrite H2. econstructor; eauto with cshm.
(* struct -> struct *)
rewrite H2. auto.
(* union -> union *)
@@ -442,21 +477,20 @@ Lemma make_boolean_correct:
eval_expr ge e le m (make_boolean a ty) vb
/\ Val.bool_of_val vb b.
Proof.
- assert (VBI: forall n, Val.bool_of_val (Vint n) (negb (Int.eq n Int.zero))).
- constructor.
- 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.
- exists (Val.of_bool (Float.cmp Cne f Float.zero)); split.
- econstructor; eauto with cshm.
- destruct (Float.cmp Cne f Float.zero); simpl; constructor.
+ intros. unfold make_boolean. unfold bool_val in H0.
+ destruct (classify_bool ty); destruct v; inv H0.
+(* int *)
+ econstructor; split. apply make_cmp_ne_zero_correct with (n := i); auto.
+ destruct (Int.eq i Int.zero); simpl; constructor.
+(* float *)
+ econstructor; split. econstructor; eauto with cshm. simpl. eauto.
+ unfold Val.cmpf, Val.cmpf_bool. simpl. rewrite <- Float.cmp_ne_eq.
+ destruct (Float.cmp Cne f Float.zero); constructor.
+(* pointer *)
+ econstructor; split. econstructor; eauto with cshm. simpl. eauto.
+ unfold Val.cmpu, Val.cmpu_bool. simpl.
+ destruct (Int.eq i Int.zero); simpl; constructor.
+ exists Vtrue; split. econstructor; eauto with cshm. constructor.
Qed.
Lemma make_neg_correct:
@@ -682,75 +716,26 @@ 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) ->
- deref_loc ge ty m b ofs E0 v ->
- type_is_volatile ty = false ->
+ deref_loc ty m b ofs v ->
eval_expr ge e le m code v.
Proof.
- unfold make_load; intros until m; intros MKLOAD EVEXP DEREF NONVOL.
+ unfold make_load; intros until m; intros MKLOAD EVEXP DEREF.
inv DEREF.
- (* nonvolatile scalar *)
+ (* 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',
+ forall f dst src ty k e le m b ofs v 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' ->
+ assign_loc ty m b ofs v 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').
+ step ge (State f (make_memcpy dst src ty) k e le m) E0 (State f Sskip k e le m').
Proof.
intros. inv H1; try congruence.
unfold make_memcpy. change le with (set_optvar None Vundef le) at 2.
@@ -763,44 +748,18 @@ Proof.
Qed.
Lemma make_store_correct:
- forall addr ty rhs code e le m b ofs v t m' f k,
+ forall addr ty rhs code e le m b ofs v 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 ->
- 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').
+ assign_loc ty m b ofs v m' ->
+ step ge (State f code k e le m) E0 (State f Sskip k e le m').
Proof.
- unfold make_store. intros until k; intros MKSTORE EV1 EV2 ASSIGN NONVOL.
+ unfold make_store. intros until k; intros MKSTORE EV1 EV2 ASSIGN.
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_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.
@@ -859,31 +818,6 @@ Proof.
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
@@ -1013,7 +947,7 @@ Qed.
Lemma bind_parameters_match:
forall e m1 vars vals m2,
- Csem.bind_parameters ge e m1 vars vals m2 ->
+ Clight.bind_parameters e m1 vars vals m2 ->
forall te tvars,
val_casted_list vals (type_of_params vars) ->
match_env e te ->
@@ -1026,10 +960,7 @@ Proof.
(* inductive case *)
simpl in H2. destruct H2.
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].
+ inv H0.
(* scalar case *)
assert (vk = Vscalar chunk). exploit var_kind_by_value; eauto. congruence.
subst vk.
@@ -1038,8 +969,7 @@ Proof.
eapply val_casted_normalized; eauto.
assumption.
apply IHbind_parameters; auto.
- (* array case *)
- inv H0; try congruence.
+ (* struct case *)
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.
@@ -1050,6 +980,14 @@ Proof.
apply IHbind_parameters; auto.
Qed.
+Lemma create_undef_temps_match:
+ forall temps,
+ create_undef_temps (List.map (@fst ident type) temps) = Clight.create_undef_temps temps.
+Proof.
+ induction temps; simpl. auto.
+ destruct a as [id ty]. simpl. decEq. auto.
+Qed.
+
(* ** Correctness of variable accessors *)
(** Correctness of the code generated by [var_get]. *)
@@ -1057,21 +995,15 @@ 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 ->
- deref_loc ge ty m loc ofs E0 v ->
+ deref_loc ty m loc ofs v ->
var_get id ty = OK code ->
match_env e te ->
eval_expr tge te le m code v.
Proof.
unfold var_get; intros.
- destruct (access_mode ty) as [chunk| | | ]_eqn.
+ inv H0.
(* access mode By_value *)
- assert (Mem.loadv chunk m (Vptr loc ofs) = Some v).
- inv H0.
- congruence.
- inv H5. simpl. congruence.
- congruence.
- congruence.
- inv H1. inv H.
+ rewrite H3 in H1. inv H1. inv H.
(* local variable *)
exploit me_local; eauto. intros [vk [A B]].
assert (vk = Vscalar chunk).
@@ -1088,8 +1020,7 @@ Proof.
eauto. eauto.
assumption.
(* access mode By_reference *)
- assert (v = Vptr loc ofs). inv H0; congruence.
- inv H1. inv H.
+ rewrite H3 in H1. inv H1. inv H.
(* local variable *)
exploit me_local; eauto. intros [vk [A B]].
eapply eval_Eaddrof.
@@ -1100,8 +1031,7 @@ Proof.
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.
+ rewrite H3 in H1. inv H1. inv H.
(* local variable *)
exploit me_local; eauto. intros [vk [A B]].
eapply eval_Eaddrof.
@@ -1111,25 +1041,22 @@ Proof.
eapply eval_Eaddrof.
eapply eval_var_addr_global. auto.
rewrite symbols_preserved. eauto.
- (* access mode By_nothing *)
- congruence.
Qed.
(** Correctness of the code generated by [var_set]. *)
Lemma var_set_correct:
- forall e le m id ty loc ofs v t m' code te rhs f k,
+ forall e le m id ty loc ofs v m' code te rhs f k,
Clight.eval_lvalue ge e le m (Clight.Evar id ty) loc ofs ->
val_casted v ty ->
- assign_loc ge ty m loc ofs v t m' ->
- type_is_volatile ty = false ->
+ assign_loc ty m loc ofs v m' ->
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) t (State f Sskip k te le m').
+ step tge (State f code k te le m) E0 (State f Sskip k te le m').
Proof.
- intros. unfold var_set in H3.
- inversion H1; subst; rewrite H6 in H3; inv H3.
+ intros. unfold var_set in H2.
+ inversion H1; subst; rewrite H5 in H2; inv H2.
(* scalar, non volatile *)
inv H.
(* local variable *)
@@ -1148,9 +1075,7 @@ Proof.
rewrite symbols_preserved. eauto.
eauto. eauto.
eapply val_casted_normalized; eauto. assumption.
- (* scalar, volatile *)
- congruence.
- (* array *)
+ (* struct *)
assert (eval_expr tge te le m (Eaddrof id) (Vptr loc ofs)).
inv H.
exploit me_local; eauto. intros [vk [A B]].
@@ -1158,7 +1083,7 @@ Proof.
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.
+ eapply make_memcpy_correct; eauto.
Qed.
(** * Proof of semantic preservation *)
@@ -1216,11 +1141,6 @@ Proof.
eapply transl_unop_correct; eauto.
(* binop *)
eapply transl_binop_correct; eauto.
-(* condition *)
- exploit make_boolean_correct. eapply H0; eauto. eauto.
- intros [vb [EVAL BVAL]].
- eapply eval_Econdition; eauto.
- destruct b; eapply make_cast_correct; eauto.
(* cast *)
eapply make_cast_correct; eauto.
(* rvalue out of lvalue *)
@@ -1229,7 +1149,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 deref_loc_preserved; eauto.
+ eapply make_load_correct; eauto.
(* var local *)
exploit (me_local _ _ MENV); eauto.
intros [vk [A B]].
@@ -1263,10 +1183,10 @@ Lemma transl_lvalue_correct:
Csharpminor.eval_expr tge te le m ta (Vptr b ofs).
Proof (proj2 transl_expr_lvalue_correct).
-Lemma transl_exprlist_correct:
+Lemma transl_arglist_correct:
forall al tyl vl,
Clight.eval_exprlist ge e le m al tyl vl ->
- forall tal, transl_exprlist al tyl = OK tal ->
+ forall tal, transl_arglist al tyl = OK tal ->
Csharpminor.eval_exprlist tge te le m tal vl.
Proof.
induction 1; intros.
@@ -1278,48 +1198,6 @@ Qed.
End EXPR.
-Lemma is_constant_bool_sound:
- forall te le m a v ty b,
- Csharpminor.eval_expr tge te le m a v ->
- bool_val v ty = Some b ->
- is_constant_bool a = Some b \/ is_constant_bool a = None.
-Proof.
- intros. unfold is_constant_bool. destruct a; auto. destruct c; auto.
- left. decEq. inv H. simpl in H2. inv H2. functional inversion H0; auto.
-Qed.
-
-Lemma exit_if_false_true:
- forall a ts e le m v te f tk,
- exit_if_false a = OK ts ->
- Clight.eval_expr ge e le m a v ->
- bool_val v (typeof a) = Some true ->
- match_env e te ->
- star step tge (State f ts tk te le m) E0 (State f Sskip tk te le m).
-Proof.
- intros. monadInv H.
- exploit transl_expr_correct; eauto. intros EV.
- exploit is_constant_bool_sound; eauto. intros [P | P]; rewrite P in EQ0; inv EQ0.
- constructor.
- exploit make_boolean_correct; eauto. intros [vb [EV' VB]].
- apply star_one. apply step_ifthenelse with (v := vb) (b := true); auto.
-Qed.
-
-Lemma exit_if_false_false:
- forall a ts e le m v te f tk,
- exit_if_false a = OK ts ->
- Clight.eval_expr ge e le m a v ->
- bool_val v (typeof a) = Some false ->
- match_env e te ->
- star step tge (State f ts tk te le m) E0 (State f (Sexit 0) tk te le m).
-Proof.
- intros. monadInv H.
- exploit transl_expr_correct; eauto. intros EV.
- exploit is_constant_bool_sound; eauto. intros [P | P]; rewrite P in EQ0; inv EQ0.
- constructor.
- exploit make_boolean_correct; eauto. intros [vb [EV' VB]].
- apply star_one. apply step_ifthenelse with (v := vb) (b := false); auto.
-Qed.
-
(** ** Semantic preservation for statements *)
(** The simulation diagram for the translation of statements and functions
@@ -1362,36 +1240,20 @@ Inductive match_cont: type -> nat -> nat -> Clight.cont -> Csharpminor.cont -> P
match_cont tyret nbrk ncnt
(Clight.Kseq s k)
(Kseq ts tk)
- | match_Kwhile: forall tyret nbrk ncnt a s k ta ts tk,
- exit_if_false a = OK ta ->
- transl_statement tyret 1%nat 0%nat s = OK ts ->
- match_cont tyret nbrk ncnt k tk ->
- match_cont tyret 1%nat 0%nat
- (Clight.Kwhile a s k)
- (Kblock (Kseq (Sloop (Sseq ta (Sblock ts))) (Kblock tk)))
- | match_Kdowhile: forall tyret nbrk ncnt a s k ta ts tk,
- exit_if_false a = OK ta ->
- transl_statement tyret 1%nat 0%nat s = OK ts ->
+ | match_Kloop1: forall tyret s1 s2 k ts1 ts2 nbrk ncnt tk,
+ transl_statement tyret 1%nat 0%nat s1 = OK ts1 ->
+ transl_statement tyret 0%nat (S ncnt) s2 = OK ts2 ->
match_cont tyret nbrk ncnt k tk ->
match_cont tyret 1%nat 0%nat
- (Clight.Kdowhile a s k)
- (Kblock (Kseq ta (Kseq (Sloop (Sseq (Sblock ts) ta)) (Kblock tk))))
- | match_Kfor2: forall tyret nbrk ncnt a2 a3 s k ta2 ta3 ts tk,
- exit_if_false a2 = OK ta2 ->
- transl_statement tyret 0%nat (S ncnt) a3 = OK ta3 ->
- transl_statement tyret 1%nat 0%nat s = OK ts ->
- match_cont tyret nbrk ncnt k tk ->
- match_cont tyret 1%nat 0%nat
- (Clight.Kfor2 a2 a3 s k)
- (Kblock (Kseq ta3 (Kseq (Sloop (Sseq ta2 (Sseq (Sblock ts) ta3))) (Kblock tk))))
- | match_Kfor3: forall tyret nbrk ncnt a2 a3 s k ta2 ta3 ts tk,
- exit_if_false a2 = OK ta2 ->
- transl_statement tyret 0%nat (S ncnt) a3 = OK ta3 ->
- transl_statement tyret 1%nat 0%nat s = OK ts ->
+ (Clight.Kloop1 s1 s2 k)
+ (Kblock (Kseq ts2 (Kseq (Sloop (Sseq (Sblock ts1) ts2)) (Kblock tk))))
+ | match_Kloop2: forall tyret s1 s2 k ts1 ts2 nbrk ncnt tk,
+ transl_statement tyret 1%nat 0%nat s1 = OK ts1 ->
+ transl_statement tyret 0%nat (S ncnt) s2 = OK ts2 ->
match_cont tyret nbrk ncnt k tk ->
match_cont tyret 0%nat (S ncnt)
- (Clight.Kfor3 a2 a3 s k)
- (Kseq (Sloop (Sseq ta2 (Sseq (Sblock ts) ta3))) (Kblock tk))
+ (Clight.Kloop2 s1 s2 k)
+ (Kseq (Sloop (Sseq (Sblock ts1) ts2)) (Kblock tk))
| match_Kswitch: forall tyret nbrk ncnt k tk,
match_cont tyret nbrk ncnt k tk ->
match_cont tyret 0%nat (S ncnt)
@@ -1446,13 +1308,6 @@ Section FIND_LABEL.
Variable lbl: label.
Variable tyret: type.
-Remark exit_if_false_no_label:
- forall a s, exit_if_false a = OK s -> forall k, find_label lbl s k = None.
-Proof.
- intros. unfold exit_if_false in H. monadInv H.
- destruct (is_constant_bool x). destruct b; inv EQ0; auto. inv EQ0; auto.
-Qed.
-
Lemma transl_find_label:
forall s nbrk ncnt k ts tk
(TR: transl_statement tyret nbrk ncnt s = OK ts)
@@ -1484,19 +1339,18 @@ Proof.
(* skip *)
auto.
(* assign *)
- 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.
+ simpl in TR.
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.
+ 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.
+(* builtin *)
+ auto.
(* seq *)
exploit (transl_find_label s0 nbrk ncnt (Clight.Kseq s1 k)); eauto. constructor; eauto.
destruct (Clight.find_label lbl s0 (Clight.Kseq s1 k)) as [[s' k'] | ].
@@ -1509,23 +1363,13 @@ Proof.
intros [ts' [tk' [nbrk' [ncnt' [A [B C]]]]]].
rewrite A. exists ts'; exists tk'; exists nbrk'; exists ncnt'; auto.
intro. rewrite H. eapply transl_find_label; eauto.
-(* while *)
- rewrite (exit_if_false_no_label _ _ EQ).
- eapply transl_find_label; eauto. econstructor; eauto.
-(* dowhile *)
- exploit (transl_find_label s0 1%nat 0%nat (Clight.Kdowhile e s0 k)); eauto. econstructor; eauto.
- destruct (Clight.find_label lbl s0 (Kdowhile e s0 k)) as [[s' k'] | ].
- intros [ts' [tk' [nbrk' [ncnt' [A [B C]]]]]].
- rewrite A. exists ts'; exists tk'; exists nbrk'; exists ncnt'; auto.
- intro. rewrite H. eapply exit_if_false_no_label; eauto.
-(* for *)
- rewrite (exit_if_false_no_label _ _ EQ).
- exploit (transl_find_label s1 1%nat 0%nat (Kfor2 e s0 s1 k)); eauto. econstructor; eauto.
- destruct (Clight.find_label lbl s1 (Kfor2 e s0 s1 k)) as [[s' k'] | ].
+(* loop *)
+ exploit (transl_find_label s0 1%nat 0%nat (Kloop1 s0 s1 k)); eauto. econstructor; eauto.
+ destruct (Clight.find_label lbl s0 (Kloop1 s0 s1 k)) as [[s' k'] | ].
intros [ts' [tk' [nbrk' [ncnt' [A [B C]]]]]].
rewrite A. exists ts'; exists tk'; exists nbrk'; exists ncnt'; auto.
intro. rewrite H.
- eapply transl_find_label; eauto. econstructor; eauto.
+ eapply transl_find_label; eauto. econstructor; eauto.
(* break *)
auto.
(* continue *)
@@ -1589,21 +1433,9 @@ Proof.
induction 1; intros T1 MST; inv MST.
(* assign *)
- 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 *)
+ simpl in TR.
destruct (is_variable a1) as []_eqn; monadInv TR.
+ (* a variable *)
assert (SAME: ts' = ts /\ tk' = tk).
inversion MTR. auto.
subst ts. unfold var_set, make_memcpy in EQ0. destruct (access_mode (typeof a1)); congruence.
@@ -1613,7 +1445,7 @@ Proof.
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.
- (* Case 3: everything else *)
+ (* not a variable *)
assert (SAME: ts' = ts /\ tk' = tk).
inversion MTR. auto.
subst ts. unfold make_store, make_memcpy in EQ2. destruct (access_mode (typeof a1)); congruence.
@@ -1621,7 +1453,7 @@ Proof.
econstructor; split.
apply plus_one. eapply make_store_correct; eauto.
eapply transl_lvalue_correct; eauto. eapply make_cast_correct; eauto.
- eapply transl_expr_correct; eauto. eapply assign_loc_preserved; eauto.
+ eapply transl_expr_correct; eauto.
eapply match_states_skip; eauto.
(* set *)
@@ -1629,17 +1461,6 @@ 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.
@@ -1648,13 +1469,24 @@ Proof.
econstructor; split.
apply plus_one. econstructor; eauto.
exploit transl_expr_correct; eauto.
- exploit transl_exprlist_correct; eauto.
+ exploit transl_arglist_correct; eauto.
eapply transl_fundef_sig1; eauto.
rewrite H3. auto.
econstructor; eauto.
econstructor; eauto.
simpl. auto.
- eapply eval_exprlist_casted; eauto.
+ eapply eval_exprlist_casted; eauto.
+
+(* builtin *)
+ monadInv TR. inv MTR.
+ econstructor; split.
+ apply plus_one. econstructor.
+ eapply transl_arglist_correct; eauto.
+ eapply external_call_symbols_preserved_2; eauto.
+ exact symbols_preserved.
+ eexact (Genv.find_var_info_transf_partial2 transl_fundef transl_globvar _ TRANSL).
+ eexact (Genv.find_var_info_rev_transf_partial2 transl_fundef transl_globvar _ TRANSL).
+ eapply match_states_skip; eauto.
(* seq *)
monadInv TR. inv MTR.
@@ -1690,33 +1522,17 @@ Proof.
apply plus_one. apply step_ifthenelse with (v := v) (b := b); auto.
destruct b; econstructor; eauto; constructor.
-(* while false *)
+(* loop *)
monadInv TR.
econstructor; split.
eapply star_plus_trans. eapply match_transl_step; eauto.
eapply plus_left. constructor.
eapply star_left. constructor.
- eapply star_trans. eapply exit_if_false_false; eauto.
- eapply star_left. constructor.
- eapply star_left. constructor.
- apply star_one. constructor.
- reflexivity. reflexivity. reflexivity. reflexivity. reflexivity. traceEq.
- eapply match_states_skip; eauto.
-
-(* while true *)
- monadInv TR.
- econstructor; split.
- eapply star_plus_trans. eapply match_transl_step; eauto.
- eapply plus_left. constructor.
- eapply star_left. constructor.
- eapply star_trans. eapply exit_if_false_true; eauto.
- eapply star_left. constructor.
apply star_one. constructor.
- reflexivity. reflexivity. reflexivity. reflexivity. traceEq.
- econstructor; eauto. constructor.
- econstructor; eauto.
+ reflexivity. reflexivity. traceEq.
+ econstructor; eauto. constructor. econstructor; eauto.
-(* skip or continue while *)
+(* skip-or-continue loop *)
assert ((ts' = Sskip \/ ts' = Sexit ncnt) /\ tk' = tk).
destruct H; subst x; monadInv TR; inv MTR; auto.
destruct H0. inv MK.
@@ -1724,126 +1540,32 @@ Proof.
eapply plus_left.
destruct H0; subst ts'; constructor.
apply star_one. constructor. traceEq.
- econstructor; eauto.
- simpl. rewrite H8; simpl. rewrite H10; simpl. reflexivity.
- constructor.
+ econstructor; eauto. constructor. econstructor; eauto.
-(* break while *)
+(* break loop1 *)
monadInv TR. inv MTR. inv MK.
econstructor; split.
eapply plus_left. constructor.
eapply star_left. constructor.
- apply star_one. constructor.
- reflexivity. traceEq.
- eapply match_states_skip; eauto.
-
-(* dowhile *)
- monadInv TR.
- econstructor; split.
- eapply star_plus_trans. eapply match_transl_step; eauto.
- eapply plus_left. constructor.
eapply star_left. constructor.
apply star_one. constructor.
reflexivity. reflexivity. traceEq.
- econstructor; eauto. constructor.
- econstructor; eauto.
-
-(* skip or continue dowhile false *)
- assert ((ts' = Sskip \/ ts' = Sexit ncnt) /\ tk' = tk).
- destruct H; subst x; monadInv TR; inv MTR; auto.
- destruct H2. inv MK.
- econstructor; split.
- eapply plus_left. destruct H2; subst ts'; constructor.
- eapply star_left. constructor.
- eapply star_trans. eapply exit_if_false_false; eauto.
- eapply star_left. constructor.
- apply star_one. constructor.
- reflexivity. reflexivity. reflexivity. traceEq.
eapply match_states_skip; eauto.
-(* skip or continue dowhile true *)
- assert ((ts' = Sskip \/ ts' = Sexit ncnt) /\ tk' = tk).
- destruct H; subst x; monadInv TR; inv MTR; auto.
- destruct H2. inv MK.
- econstructor; split.
- eapply plus_left. destruct H2; subst ts'; constructor.
- eapply star_left. constructor.
- eapply star_trans. eapply exit_if_false_true; eauto.
- apply star_one. constructor.
- reflexivity. reflexivity. traceEq.
- econstructor; eauto.
- simpl. rewrite H10; simpl. rewrite H12; simpl. reflexivity. constructor.
-
-(* break dowhile *)
+(* skip loop2 *)
monadInv TR. inv MTR. inv MK.
econstructor; split.
- eapply plus_left. constructor.
- eapply star_left. constructor.
- eapply star_left. constructor.
- apply star_one. constructor.
- reflexivity. reflexivity. traceEq.
- eapply match_states_skip; eauto.
-
-(* for false *)
- monadInv TR.
- econstructor; split.
- eapply star_plus_trans. eapply match_transl_step; eauto.
- eapply plus_left. constructor.
- eapply star_left. constructor.
- eapply star_trans. eapply exit_if_false_false; eauto.
- eapply star_left. constructor.
- eapply star_left. constructor.
- apply star_one. constructor.
- reflexivity. reflexivity. reflexivity. reflexivity. reflexivity. reflexivity.
- eapply match_states_skip; eauto.
-
-(* for true *)
- monadInv TR.
- econstructor; split.
- eapply star_plus_trans. eapply match_transl_step; eauto.
- eapply plus_left. constructor.
- eapply star_left. constructor.
- eapply star_trans. eapply exit_if_false_true; eauto.
- eapply star_left. constructor.
- eapply star_left. constructor.
- apply star_one. constructor.
- reflexivity. reflexivity. reflexivity. reflexivity. reflexivity. reflexivity.
- econstructor; eauto. constructor.
+ apply plus_one. constructor.
econstructor; eauto.
+ simpl. rewrite H5; simpl. rewrite H7; simpl. eauto.
+ constructor.
-(* skip or continue for2 *)
- assert ((ts' = Sskip \/ ts' = Sexit ncnt) /\ tk' = tk).
- destruct H; subst x; monadInv TR; inv MTR; auto.
- destruct H0. inv MK.
- econstructor; split.
- eapply plus_left. destruct H0; subst ts'; constructor.
- apply star_one. constructor. reflexivity.
- econstructor; eauto. constructor.
- econstructor; eauto.
-
-(* break for2 *)
+(* break loop2 *)
monadInv TR. inv MTR. inv MK.
econstructor; split.
eapply plus_left. constructor.
- eapply star_left. constructor.
- eapply star_left. constructor.
apply star_one. constructor.
- reflexivity. reflexivity. traceEq.
- eapply match_states_skip; eauto.
-
-(* skip for3 *)
- monadInv TR. inv MTR. inv MK.
- econstructor; split.
- apply plus_one. constructor.
- econstructor; eauto.
- simpl. rewrite H6; simpl. rewrite H8; simpl. rewrite H9; simpl. reflexivity.
- constructor.
-
-(* break for3 *)
- monadInv TR. inv MTR. inv MK.
- econstructor; split.
- eapply plus_left. constructor. apply star_one. constructor.
- econstructor; eauto.
+ traceEq.
eapply match_states_skip; eauto.
(* return none *)
@@ -1858,10 +1580,6 @@ Proof.
monadInv TR. inv MTR.
econstructor; split.
apply plus_one. constructor.
-(* monadInv TRF. simpl.
- unfold opttyp_of_type. destruct (Clight.fn_return f); try congruence.
- inv H0. inv H3. inv H3.
-*)
eapply make_cast_correct. eapply transl_expr_correct; eauto. eauto.
eapply match_env_free_blocks; eauto.
econstructor; eauto.
@@ -1930,6 +1648,7 @@ Proof.
eapply transl_names_norepet; eauto.
eexact C. eapply bind_parameters_match; eauto.
simpl in TY. unfold type_of_function in TY. congruence.
+ simpl. rewrite (create_undef_temps_match (Clight.fn_temps f)).
econstructor; eauto.
unfold transl_function. rewrite EQ0; simpl. rewrite EQ; simpl. rewrite EQ1; auto.
constructor.
@@ -1953,10 +1672,6 @@ Proof.
econstructor; split.
apply plus_one. constructor.
econstructor; eauto. simpl; reflexivity. constructor.
- inv MK.
- econstructor; split.
- apply plus_one. constructor.
- econstructor; eauto. simpl; reflexivity. constructor.
Qed.
Lemma transl_initial_states:
diff --git a/cfrontend/Cstrategy.v b/cfrontend/Cstrategy.v
index 13cffb5..5be17ed 100644
--- a/cfrontend/Cstrategy.v
+++ b/cfrontend/Cstrategy.v
@@ -56,7 +56,9 @@ Fixpoint simple (a: expr) : bool :=
| Eunop _ r1 _ => simple r1
| Ebinop _ r1 r2 _ => simple r1 && simple r2
| Ecast r1 _ => simple r1
- | Econdition r1 r2 r3 _ => simple r1 && simple r2 && simple r3
+ | Eseqand _ _ _ => false
+ | Eseqor _ _ _ => false
+ | Econdition _ _ _ _ => false
| Esizeof _ _ => true
| Ealignof _ _ => true
| Eassign _ _ _ => false
@@ -64,6 +66,7 @@ Fixpoint simple (a: expr) : bool :=
| Epostincr _ _ _ => false
| Ecomma _ _ _ => false
| Ecall _ _ _ => false
+ | Ebuiltin _ _ _ _ => false
| Eparen _ _ => false
end.
@@ -126,13 +129,6 @@ with eval_simple_rvalue: expr -> val -> Prop :=
eval_simple_rvalue r1 v1 ->
sem_cast v1 (typeof r1) ty = Some v ->
eval_simple_rvalue (Ecast r1 ty) v
- | esr_condition: forall r1 r2 r3 ty v1 b v2 v,
- simple r2 = true -> simple r3 = true ->
- eval_simple_rvalue r1 v1 ->
- bool_val v1 (typeof r1) = Some b ->
- eval_simple_rvalue (if b then r2 else r3) v2 ->
- sem_cast v2 (typeof (if b then r2 else r3)) ty = Some v ->
- eval_simple_rvalue (Econdition r1 r2 r3 ty) v
| esr_sizeof: forall ty1 ty,
eval_simple_rvalue (Esizeof ty1 ty) (Vint (Int.repr (sizeof ty1)))
| esr_alignof: forall ty1 ty,
@@ -175,6 +171,10 @@ Inductive leftcontext: kind -> kind -> (expr -> expr) -> Prop :=
leftcontext k RV (fun x => Ebinop op e1 (C x) ty)
| lctx_cast: forall k C ty,
leftcontext k RV C -> leftcontext k RV (fun x => Ecast (C x) ty)
+ | lctx_seqand: forall k C r2 ty,
+ leftcontext k RV C -> leftcontext k RV (fun x => Eseqand (C x) r2 ty)
+ | lctx_seqor: forall k C r2 ty,
+ leftcontext k RV C -> leftcontext k RV (fun x => Eseqor (C x) r2 ty)
| lctx_condition: forall k C r2 r3 ty,
leftcontext k RV C -> leftcontext k RV (fun x => Econdition (C x) r2 r3 ty)
| lctx_assign_left: forall k C e2 ty,
@@ -194,6 +194,9 @@ Inductive leftcontext: kind -> kind -> (expr -> expr) -> Prop :=
| lctx_call_right: forall k C e1 ty,
simple e1 = true -> leftcontextlist k C ->
leftcontext k RV (fun x => Ecall e1 (C x) ty)
+ | lctx_builtin: forall k C ef tyargs ty,
+ leftcontextlist k C ->
+ leftcontext k RV (fun x => Ebuiltin ef tyargs (C x) ty)
| lctx_comma: forall k C e2 ty,
leftcontext k RV C -> leftcontext k RV (fun x => Ecomma (C x) e2 ty)
| lctx_paren: forall k C ty,
@@ -240,8 +243,33 @@ Inductive estep: state -> trace -> state -> Prop :=
estep (ExprState f (C (Evalof l ty)) k e m)
t (ExprState f (C (Eval v ty)) k e m)
+ | step_seqand_true: forall f C r1 r2 ty k e m v,
+ leftcontext RV RV C ->
+ eval_simple_rvalue e m r1 v ->
+ bool_val v (typeof r1) = Some true ->
+ estep (ExprState f (C (Eseqand r1 r2 ty)) k e m)
+ E0 (ExprState f (C (Eparen (Eparen r2 type_bool) ty)) k e m)
+ | step_seqand_false: forall f C r1 r2 ty k e m v,
+ leftcontext RV RV C ->
+ eval_simple_rvalue e m r1 v ->
+ bool_val v (typeof r1) = Some false ->
+ estep (ExprState f (C (Eseqand r1 r2 ty)) k e m)
+ E0 (ExprState f (C (Eval (Vint Int.zero) ty)) k e m)
+
+ | step_seqor_true: forall f C r1 r2 ty k e m v,
+ leftcontext RV RV C ->
+ eval_simple_rvalue e m r1 v ->
+ bool_val v (typeof r1) = Some true ->
+ estep (ExprState f (C (Eseqor r1 r2 ty)) k e m)
+ E0 (ExprState f (C (Eval (Vint Int.one) ty)) k e m)
+ | step_seqor_false: forall f C r1 r2 ty k e m v,
+ leftcontext RV RV C ->
+ eval_simple_rvalue e m r1 v ->
+ bool_val v (typeof r1) = Some false ->
+ estep (ExprState f (C (Eseqor r1 r2 ty)) k e m)
+ E0 (ExprState f (C (Eparen (Eparen r2 type_bool) ty)) k e m)
+
| step_condition: forall f C r1 r2 r3 ty k e m v b,
- simple r2 = false \/ simple r3 = false ->
leftcontext RV RV C ->
eval_simple_rvalue e m r1 v ->
bool_val v (typeof r1) = Some b ->
@@ -338,7 +366,14 @@ Inductive estep: state -> trace -> state -> Prop :=
Genv.find_funct ge vf = Some fd ->
type_of_fundef fd = Tfunction targs tres ->
estep (ExprState f (C (Ecall rf rargs ty)) 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_builtin: forall f C ef tyargs rargs ty k e m vargs t vres m',
+ leftcontext RV RV C ->
+ eval_simple_list e m rargs tyargs vargs ->
+ external_call ef ge vargs m t vres m' ->
+ estep (ExprState f (C (Ebuiltin ef tyargs rargs ty)) k e m)
+ t (ExprState f (C (Eval vres ty)) k e m').
Definition step (S: state) (t: trace) (S': state) : Prop :=
estep S t S' \/ sstep ge S t S'.
@@ -503,6 +538,10 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
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
+ | Eseqand (Eval v1 ty1) r2 ty =>
+ exists b, bool_val v1 ty1 = Some b
+ | Eseqor (Eval v1 ty1) r2 ty =>
+ exists b, bool_val v1 ty1 = Some b
| Econdition (Eval v1 ty1) r1 r2 ty =>
exists b, bool_val v1 ty1 = Some b
| Eassign (Eloc b ofs ty1) (Eval v2 ty2) ty =>
@@ -527,6 +566,11 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
/\ Genv.find_funct ge vf = Some fd
/\ cast_arguments rargs tyargs vl
/\ type_of_fundef fd = Tfunction tyargs tyres
+ | Ebuiltin ef tyargs rargs ty =>
+ exprlist_all_values rargs ->
+ exists vargs, exists t, exists vres, exists m',
+ cast_arguments rargs tyargs vargs
+ /\ external_call ef ge vargs m t vres m'
| _ => True
end.
@@ -549,11 +593,14 @@ Proof.
exists v; auto.
exists v; auto.
exists v; auto.
+ exists true; auto. exists false; auto.
+ exists true; auto. exists false; auto.
exists b; auto.
exists v; exists m'; exists t; auto.
exists t; exists v1; auto.
exists t; exists v1; auto.
exists v; auto.
+ intros. exists vargs; exists t; exists vres; exists m'; auto.
Qed.
Lemma callred_invert:
@@ -591,12 +638,15 @@ Proof.
destruct (C a); auto; contradiction.
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.
+ intros. elim (H0 a m); auto.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
red; intros. destruct (C a); auto.
@@ -681,12 +731,6 @@ Ltac FinishL := apply star_one; left; apply step_lred; eauto; simpl; try (econst
FinishR.
(* cast *)
Steps H0 (fun x => C(Ecast x ty)). FinishR.
-(* condition *)
- Steps H2 (fun x => C(Econdition x r2 r3 ty)).
- eapply star_left with (s2 := ExprState f (C (Eparen (if b then r2 else r3) ty)) k e m).
- left; apply step_rred; eauto. constructor; auto.
- Steps H5 (fun x => C(Eparen x ty)).
- FinishR. auto.
(* sizeof *)
FinishR.
(* alignof *)
@@ -801,23 +845,6 @@ Ltac StepR REC C' a :=
StepR IHa (fun x => C(Ecast x ty)) a.
exploit safe_inv. eexact SAFE0. eauto. simpl. intros [v' CAST].
exists v'; econstructor; eauto.
-(* condition *)
- destruct (andb_prop _ _ S) as [S' S3]. destruct (andb_prop _ _ S') as [S1 S2]. clear S S'.
- StepR IHa1 (fun x => C(Econdition x a2 a3 ty)) a1.
- exploit safe_inv. eexact SAFE0. eauto. simpl. intros [b BV].
- set (a' := if b then a2 else a3).
- assert (SAFE1: safe (ExprState f (C (Eparen a' ty)) k e m)).
- eapply safe_steps. eexact SAFE0. apply star_one. left; apply step_rred; auto.
- unfold a'; constructor; auto.
- assert (EV': exists v, eval_simple_rvalue e m a' v).
- unfold a'; destruct b.
- eapply (IHa2 RV (fun x => C(Eparen x ty))); eauto.
- eapply (IHa3 RV (fun x => C(Eparen x ty))); eauto.
- destruct EV' as [v1 E1].
- exploit safe_inv. eapply (eval_simple_rvalue_safe (fun x => C(Eparen x ty))).
- eexact E1. eauto. auto. eauto.
- simpl. intros [v2 CAST].
- exists v2. econstructor; eauto.
(* sizeof *)
econstructor; econstructor.
(* alignof *)
@@ -898,9 +925,12 @@ Proof.
Qed.
Inductive contextlist' : (exprlist -> expr) -> Prop :=
- | contextlist'_intro: forall r1 rl0 ty C,
+ | contextlist'_call: forall r1 rl0 ty C,
+ context RV RV C ->
+ contextlist' (fun rl => C (Ecall r1 (exprlist_app rl0 rl) ty))
+ | contextlist'_builtin: forall ef tyargs rl0 ty C,
context RV RV C ->
- contextlist' (fun rl => C (Ecall r1 (exprlist_app rl0 rl) ty)).
+ contextlist' (fun rl => C (Ebuiltin ef tyargs (exprlist_app rl0 rl) ty)).
Lemma exprlist_app_context:
forall rl1 rl2,
@@ -921,6 +951,10 @@ Proof.
assert (context RV RV C'). constructor. apply exprlist_app_context.
change (context RV RV (fun x => C0 (C' x))).
eapply context_compose; eauto.
+ set (C' := fun x => Ebuiltin ef tyargs (exprlist_app rl0 (Econs x rl)) ty).
+ assert (context RV RV C'). constructor. apply exprlist_app_context.
+ change (context RV RV (fun x => C0 (C' x))).
+ eapply context_compose; eauto.
Qed.
Lemma contextlist'_tail:
@@ -933,6 +967,10 @@ Proof.
with (fun x => C0 (Ecall r0 (exprlist_app (exprlist_app rl0 (Econs r1 Enil)) x) ty)).
constructor. auto.
apply extensionality; intros. f_equal. f_equal. apply exprlist_app_assoc.
+ replace (fun x => C0 (Ebuiltin ef tyargs (exprlist_app rl0 (Econs r1 x)) ty))
+ with (fun x => C0 (Ebuiltin ef tyargs (exprlist_app (exprlist_app rl0 (Econs r1 Enil)) x) ty)).
+ constructor. auto.
+ apply extensionality; intros. f_equal. f_equal. apply exprlist_app_assoc.
Qed.
Hint Resolve contextlist'_head contextlist'_tail.
@@ -992,12 +1030,15 @@ Variable m: mem.
Definition simple_side_effect (r: expr) : Prop :=
match r with
| Evalof l _ => simple l = true /\ type_is_volatile (typeof l) = true
- | Econdition r1 r2 r3 _ => simple r1 = true /\ (simple r2 = false \/ simple r3 = false)
+ | Eseqand r1 r2 _ => simple r1 = true
+ | Eseqor r1 r2 _ => simple r1 = true
+ | Econdition r1 r2 r3 _ => simple r1 = true
| Eassign l1 r2 _ => simple l1 = true /\ simple r2 = true
| Eassignop _ l1 r2 _ _ => simple l1 = true /\ simple r2 = true
| Epostincr _ l1 _ => simple l1 = true
| Ecomma r1 r2 _ => simple r1 = true
| Ecall r1 rl _ => simple r1 = true /\ simplelist rl = true
+ | Ebuiltin ef tyargs rl _ => simplelist rl = true
| Eparen r1 _ => simple r1 = true
| _ => False
end.
@@ -1046,10 +1087,12 @@ Ltac Base :=
Rec H0 RV C (fun x => Ebinop op r1 x ty).
(* cast *)
Kind. Rec H RV C (fun x => Ecast x ty).
+(* seqand *)
+ Kind. Rec H RV C (fun x => Eseqand x r2 ty). Base.
+(* seqor *)
+ Kind. Rec H RV C (fun x => Eseqor x r2 ty). Base.
(* condition *)
- Kind. Rec H RV C (fun x => Econdition x r2 r3 ty). rewrite H4; simpl.
- destruct (simple r2 && simple r3) as []_eqn. auto.
- rewrite andb_false_iff in Heqb. Base.
+ Kind. Rec H RV C (fun x => Econdition x r2 r3 ty). Base.
(* assign *)
Kind. Rec H LV C (fun x => Eassign x r ty). Rec H0 RV C (fun x => Eassign l x ty). Base.
(* assignop *)
@@ -1061,9 +1104,15 @@ Ltac Base :=
(* call *)
Kind. Rec H RV C (fun x => Ecall x rargs ty).
destruct (H0 (fun x => C (Ecall r1 x ty))) as [A | [C' [a' [D [A B]]]]].
- eapply contextlist'_intro with (C := C) (rl0 := Enil). auto. auto.
+ eapply contextlist'_call with (C := C) (rl0 := Enil). auto. auto.
Base.
right; exists (fun x => Ecall r1 (C' x) ty); exists a'. rewrite D; simpl; auto.
+(* builtin *)
+ Kind.
+ destruct (H (fun x => C (Ebuiltin ef tyargs x ty))) as [A | [C' [a' [D [A B]]]]].
+ eapply contextlist'_builtin with (C := C) (rl0 := Enil). auto. auto.
+ Base.
+ right; exists (fun x => Ebuiltin ef tyargs (C' x) ty); exists a'. rewrite D; simpl; auto.
(* rparen *)
Kind. Rec H RV C (fun x => (Eparen x ty)). Base.
(* cons *)
@@ -1102,6 +1151,22 @@ Proof.
eapply plus_right.
eapply eval_simple_lvalue_steps with (C := fun x => C(Evalof x (typeof l))); eauto.
left. apply step_rred; eauto. econstructor; eauto. auto.
+(* seqand true *)
+ eapply plus_right.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqand x r2 ty)); eauto.
+ left. apply step_rred; eauto. apply red_seqand_true; auto. traceEq.
+(* seqand false *)
+ eapply plus_right.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqand x r2 ty)); eauto.
+ left. apply step_rred; eauto. apply red_seqand_false; auto. traceEq.
+(* seqor true *)
+ eapply plus_right.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqor x r2 ty)); eauto.
+ left. apply step_rred; eauto. apply red_seqor_true; auto. traceEq.
+(* seqor false *)
+ eapply plus_right.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eseqor x r2 ty)); eauto.
+ left. apply step_rred; eauto. apply red_seqor_false; auto. traceEq.
(* condition *)
eapply plus_right.
eapply eval_simple_rvalue_steps with (C := fun x => C(Econdition x r2 r3 ty)); eauto.
@@ -1196,9 +1261,16 @@ Proof.
eapply eval_simple_rvalue_steps with (C := fun x => C(Ecall x rargs ty)); eauto.
eapply plus_right.
eapply eval_simple_list_steps with (C := fun x => C(Ecall (Eval vf (typeof rf)) x ty)); eauto.
- eapply contextlist'_intro with (rl0 := Enil); auto.
+ eapply contextlist'_call with (rl0 := Enil); auto.
left; apply Csem.step_call; eauto. econstructor; eauto.
traceEq. auto.
+(* builtin *)
+ exploit eval_simple_list_implies; eauto. intros [vl' [A B]].
+ eapply plus_right.
+ eapply eval_simple_list_steps with (C := fun x => C(Ebuiltin ef tyargs x ty)); eauto.
+ eapply contextlist'_builtin with (rl0 := Enil); auto.
+ left; apply Csem.step_rred; eauto. econstructor; eauto.
+ traceEq.
Qed.
Lemma can_estep:
@@ -1219,8 +1291,21 @@ Proof.
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.
+(* seqand *)
+ exploit (simple_can_eval_rval f k e m b1 (fun x => C(Eseqand x b2 ty))); eauto.
+ intros [v1 [E1 S1]].
+ exploit safe_inv. eexact S1. eauto. simpl. intros [b BV].
+ destruct b.
+ econstructor; econstructor; eapply step_seqand_true; eauto.
+ econstructor; econstructor; eapply step_seqand_false; eauto.
+(* seqor *)
+ exploit (simple_can_eval_rval f k e m b1 (fun x => C(Eseqor x b2 ty))); eauto.
+ intros [v1 [E1 S1]].
+ exploit safe_inv. eexact S1. eauto. simpl. intros [b BV].
+ destruct b.
+ econstructor; econstructor; eapply step_seqor_true; eauto.
+ econstructor; econstructor; eapply step_seqor_false; eauto.
(* condition *)
- destruct Q.
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].
@@ -1276,7 +1361,7 @@ Proof.
exploit (simple_can_eval_rval f k e m b (fun x => C(Ecall x rargs ty))); eauto.
intros [vf [E1 S1]].
pose (C' := fun x => C(Ecall (Eval vf (typeof b)) x ty)).
- assert (contextlist' C'). unfold C'; eapply contextlist'_intro with (rl0 := Enil); auto.
+ assert (contextlist' C'). unfold C'; eapply contextlist'_call with (rl0 := Enil); auto.
exploit (simple_list_can_eval f k e m rargs C'); eauto.
intros [vl E2].
exploit safe_inv. 2: eapply leftcontext_context; eexact R.
@@ -1285,6 +1370,18 @@ Proof.
simpl. intros X. exploit X. eapply rval_list_all_values.
intros [tyargs [tyres [fd [vargs [P [Q [U V]]]]]]].
econstructor; econstructor; eapply step_call; eauto. eapply can_eval_simple_list; eauto.
+(* builtin *)
+ pose (C' := fun x => C(Ebuiltin ef tyargs x ty)).
+ assert (contextlist' C'). unfold C'; eapply contextlist'_builtin with (rl0 := Enil); auto.
+ exploit (simple_list_can_eval f k e m rargs C'); eauto.
+ intros [vl E].
+ exploit safe_inv. 2: eapply leftcontext_context; eexact R.
+ eapply safe_steps. eexact H.
+ apply (eval_simple_list_steps f k e m rargs vl E C'); auto.
+ simpl. intros X. exploit X. eapply rval_list_all_values.
+ intros [vargs [t [vres [m' [U V]]]]].
+ econstructor; econstructor; eapply step_builtin; eauto.
+ eapply can_eval_simple_list; eauto.
(* paren *)
exploit (simple_can_eval_rval f k e m b (fun x => C(Eparen x ty))); eauto.
intros [v1 [E1 S1]].
@@ -1448,6 +1545,11 @@ Proof.
rewrite Heqo; rewrite Heqo0; auto.
econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; auto.
+ (* builtin *)
+ exploit external_call_trace_length; eauto. destruct t1; simpl; intros.
+ exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]].
+ econstructor; econstructor. left; eapply step_builtin; eauto.
+ omegaContradiction.
(* external calls *)
inv H1.
exploit external_call_trace_length; eauto. destruct t1; simpl; intros.
@@ -1474,6 +1576,9 @@ Proof.
tauto.
(* postincr stuck *)
exploit deref_loc_trace; eauto. destruct t; auto. destruct t; tauto.
+ (* builtins *)
+ exploit external_call_trace_length; eauto.
+ destruct t; simpl; auto. destruct t; simpl; auto. intros; omegaContradiction.
(* external calls *)
exploit external_call_trace_length; eauto.
destruct t; simpl; auto. destruct t; simpl; auto. intros; omegaContradiction.
@@ -1502,6 +1607,8 @@ Qed.
(** * A big-step semantics implementing the reduction strategy. *)
+(** TODO: update with seqand / seqor *)
+
Section BIGSTEP.
Variable ge: genv.
@@ -2488,6 +2595,8 @@ Fixpoint esize (a: expr) : nat :=
| Eunop _ r1 _ => S(esize r1)
| Ebinop _ r1 r2 _ => S(esize r1 + esize r2)%nat
| Ecast r1 _ => S(esize r1)
+ | Eseqand r1 r2 _ => S(esize r1)
+ | Eseqor r1 r2 _ => S(esize r1)
| Econdition r1 _ _ _ => S(esize r1)
| Esizeof _ _ => 1%nat
| Ealignof _ _ => 1%nat
@@ -2496,6 +2605,7 @@ Fixpoint esize (a: expr) : nat :=
| Epostincr _ l1 _ => S(esize l1)
| Ecomma r1 r2 _ => S(esize r1 + esize r2)%nat
| Ecall r1 rl2 _ => S(esize r1 + esizelist rl2)%nat
+ | Ebuiltin ef tyargs rl _ => S(esizelist rl)
| Eparen r1 _ => S(esize r1)
end
@@ -2518,8 +2628,11 @@ with leftcontextlist_size:
(esize e1 < esize e2)%nat ->
(esizelist (C e1) < esizelist (C e2))%nat.
Proof.
- induction 1; intros; simpl; auto with arith. exploit leftcontextlist_size; eauto. auto with arith.
- induction 1; intros; simpl; auto with arith. exploit leftcontext_size; eauto. auto with arith.
+ induction 1; intros; simpl; auto with arith.
+ exploit leftcontextlist_size; eauto. auto with arith.
+ exploit leftcontextlist_size; eauto. auto with arith.
+ induction 1; intros; simpl; auto with arith.
+ exploit leftcontext_size; eauto. auto with arith.
Qed.
Lemma evalinf_funcall_steps:
diff --git a/cfrontend/Csyntax.v b/cfrontend/Csyntax.v
index e9e260b..86bba85 100644
--- a/cfrontend/Csyntax.v
+++ b/cfrontend/Csyntax.v
@@ -208,6 +208,8 @@ Inductive expr : Type :=
| Ebinop (op: binary_operation) (r1 r2: expr) (ty: type)
(**r binary arithmetic operation *)
| Ecast (r: expr) (ty: type) (**r type cast [(ty)r] *)
+ | Eseqand (r1 r2: expr) (ty: type) (**r sequential "and" [r1 && r2] *)
+ | Eseqor (r1 r2: expr) (ty: type) (**r sequential "or" [r1 && r2] *)
| Econdition (r1 r2 r3: expr) (ty: type) (**r conditional [r1 ? r2 : r3] *)
| Esizeof (ty': type) (ty: type) (**r size of a type *)
| Ealignof (ty': type) (ty: type) (**r natural alignment of a type *)
@@ -219,6 +221,8 @@ Inductive expr : Type :=
| Ecomma (r1 r2: expr) (ty: type) (**r sequence expression [r1, r2] *)
| Ecall (r1: expr) (rargs: exprlist) (ty: type)
(**r function call [r1(rargs)] *)
+ | Ebuiltin (ef: external_function) (tyargs: typelist) (rargs: exprlist) (ty: type)
+ (**r builtin function call *)
| Eloc (b: block) (ofs: int) (ty: type)
(**r memory location, result of evaluating a l-value *)
| Eparen (r: expr) (ty: type) (**r marked subexpression *)
@@ -263,24 +267,6 @@ Definition Epreincr (id: incr_or_decr) (l: expr) (ty: type) :=
Eassignop (match id with Incr => Oadd | Decr => Osub end)
l (Eval (Vint Int.one) type_int32s) (typeconv ty) ty.
-(** 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
- (Ecast r2 type_bool)
- (Eval (Vint Int.zero) type_int32s)
- ty.
-
-(** 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) type_int32s)
- (Ecast r2 type_bool)
- ty.
-
(** Extract the type part of a type-annotated expression. *)
Definition typeof (a: expr) : type :=
@@ -296,6 +282,8 @@ Definition typeof (a: expr) : type :=
| Ebinop _ _ _ ty => ty
| Ecast _ ty => ty
| Econdition _ _ _ ty => ty
+ | Eseqand _ _ ty => ty
+ | Eseqor _ _ ty => ty
| Esizeof _ ty => ty
| Ealignof _ ty => ty
| Eassign _ _ ty => ty
@@ -303,6 +291,7 @@ Definition typeof (a: expr) : type :=
| Epostincr _ _ ty => ty
| Ecomma _ _ ty => ty
| Ecall _ _ ty => ty
+ | Ebuiltin _ _ _ ty => ty
| Eparen _ ty => ty
end.
@@ -790,14 +779,15 @@ Definition classify_notint (ty: type) : classify_notint_cases :=
the [!] and [?] operators, as well as the [if], [while], [for] statements. *)
Inductive classify_bool_cases : Type :=
- | bool_case_ip (**r integer or pointer *)
+ | bool_case_i (**r integer *)
| bool_case_f (**r float *)
+ | bool_case_p (**r pointer *)
| bool_default.
Definition classify_bool (ty: type) : classify_bool_cases :=
match typeconv ty with
- | Tint _ _ _ => bool_case_ip
- | Tpointer _ _ => bool_case_ip
+ | Tint _ _ _ => bool_case_i
+ | Tpointer _ _ => bool_case_p
| Tfloat _ _ => bool_case_f
| _ => bool_default
end.
@@ -949,8 +939,8 @@ 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_p2bool (**r pointer -> 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 *)
@@ -959,8 +949,8 @@ Inductive classify_cast_cases : Type :=
Function classify_cast (tfrom tto: type) : classify_cast_cases :=
match tto, tfrom with
| 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 IBool _ _, (Tpointer _ _ | Tarray _ _ _ | Tfunction _ _) => cast_case_p2bool
| 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
diff --git a/cfrontend/Initializers.v b/cfrontend/Initializers.v
index b4e3984..41dbe3f 100644
--- a/cfrontend/Initializers.v
+++ b/cfrontend/Initializers.v
@@ -82,6 +82,22 @@ Fixpoint constval (a: expr) : res val :=
OK (Vint (Int.repr (sizeof ty1)))
| Ealignof ty1 ty =>
OK (Vint (Int.repr (alignof ty1)))
+ | Eseqand r1 r2 ty =>
+ do v1 <- constval r1;
+ do v2 <- constval r2;
+ match bool_val v1 (typeof r1) with
+ | Some true => do v3 <- do_cast v2 (typeof r2) type_bool; do_cast v3 type_bool ty
+ | Some false => OK (Vint Int.zero)
+ | None => Error(msg "undefined && operation")
+ end
+ | Eseqor r1 r2 ty =>
+ do v1 <- constval r1;
+ do v2 <- constval r2;
+ match bool_val v1 (typeof r1) with
+ | Some false => do v3 <- do_cast v2 (typeof r2) type_bool; do_cast v3 type_bool ty
+ | Some true => OK (Vint Int.one)
+ | None => Error(msg "undefined || operation")
+ end
| Econdition r1 r2 r3 ty =>
do v1 <- constval r1;
do v2 <- constval r2;
diff --git a/cfrontend/Initializersproof.v b/cfrontend/Initializersproof.v
index 76f08f3..a68013e 100644
--- a/cfrontend/Initializersproof.v
+++ b/cfrontend/Initializersproof.v
@@ -51,6 +51,8 @@ Fixpoint simple (a: expr) : Prop :=
| Eunop _ r1 _ => simple r1
| Ebinop _ r1 r2 _ => simple r1 /\ simple r2
| Ecast r1 _ => simple r1
+ | Eseqand r1 r2 _ => simple r1 /\ simple r2
+ | Eseqor r1 r2 _ => simple r1 /\ simple r2
| Econdition r1 r2 r3 _ => simple r1 /\ simple r2 /\ simple r3
| Esizeof _ _ => True
| Ealignof _ _ => True
@@ -59,6 +61,7 @@ Fixpoint simple (a: expr) : Prop :=
| Epostincr _ _ _ => False
| Ecomma r1 r2 _ => simple r1 /\ simple r2
| Ecall _ _ _ => False
+ | Ebuiltin _ _ _ _ => False
| Eparen r1 _ => simple r1
end.
@@ -123,6 +126,24 @@ with eval_simple_rvalue: expr -> val -> Prop :=
eval_simple_rvalue (Esizeof ty1 ty) (Vint (Int.repr (sizeof ty1)))
| esr_alignof: forall ty1 ty,
eval_simple_rvalue (Ealignof ty1 ty) (Vint (Int.repr (alignof ty1)))
+ | esr_seqand_true: forall r1 r2 ty v1 v2 v3 v4,
+ eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) = Some true ->
+ eval_simple_rvalue r2 v2 ->
+ sem_cast v2 (typeof r2) type_bool = Some v3 ->
+ sem_cast v3 type_bool ty = Some v4 ->
+ eval_simple_rvalue (Eseqand r1 r2 ty) v4
+ | esr_seqand_false: forall r1 r2 ty v1,
+ eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) = Some false ->
+ eval_simple_rvalue (Eseqand r1 r2 ty) (Vint Int.zero)
+ | esr_seqor_false: forall r1 r2 ty v1 v2 v3 v4,
+ eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) = Some false ->
+ eval_simple_rvalue r2 v2 ->
+ sem_cast v2 (typeof r2) type_bool = Some v3 ->
+ sem_cast v3 type_bool ty = Some v4 ->
+ eval_simple_rvalue (Eseqor r1 r2 ty) v4
+ | esr_seqor_true: forall r1 r2 ty v1,
+ eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) = Some true ->
+ eval_simple_rvalue (Eseqor r1 r2 ty) (Vint Int.one)
| esr_condition: forall r1 r2 r3 ty v v1 b v',
eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) = Some b ->
eval_simple_rvalue (if b then r2 else r3) v' ->
@@ -185,6 +206,10 @@ Proof.
inv EV. econstructor; eauto. constructor.
inv EV. econstructor; eauto. constructor. constructor.
inv EV. econstructor; eauto. constructor.
+ inv EV. inv H2. eapply esr_seqand_true; eauto. constructor.
+ inv EV. eapply esr_seqand_false; eauto. constructor.
+ inv EV. eapply esr_seqor_true; eauto. constructor.
+ inv EV. inv H2. eapply esr_seqor_false; eauto. constructor.
inv EV. eapply esr_condition; eauto. constructor.
inv EV. constructor.
inv EV. constructor.
@@ -210,6 +235,12 @@ Proof.
inv H0. econstructor; eauto. congruence.
inv H0. econstructor; eauto. congruence.
inv H0. econstructor; eauto. congruence.
+ inv H0.
+ eapply esr_seqand_true; eauto. rewrite TY; auto.
+ eapply esr_seqand_false; eauto. rewrite TY; auto.
+ inv H0.
+ eapply esr_seqor_false; eauto. rewrite TY; auto.
+ eapply esr_seqor_true; eauto. rewrite TY; auto.
inv H0. eapply esr_condition; eauto. congruence.
inv H0.
inv H0.
@@ -218,6 +249,7 @@ Proof.
inv H0.
inv H0.
red; split; intros. auto. inv H0.
+ red; split; intros. auto. inv H0.
inv H0. econstructor; eauto.
inv H0. econstructor; eauto. congruence.
Qed.
@@ -412,10 +444,10 @@ 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. inv H0. inv H. constructor.
+ rewrite H2 in H. inv H0. inv H. 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.
@@ -464,6 +496,16 @@ Proof.
constructor.
(* alignof *)
constructor.
+ (* seqand *)
+ rewrite (bool_val_match x v1 (typeof r1)) in EQ2; auto. rewrite H0 in EQ2.
+ monadInv EQ2. eapply sem_cast_match; eauto. eapply sem_cast_match; eauto.
+ rewrite (bool_val_match x v1 (typeof r1)) in EQ2; auto. rewrite H0 in EQ2.
+ monadInv EQ2. constructor.
+ (* seqor *)
+ rewrite (bool_val_match x v1 (typeof r1)) in EQ2; auto. rewrite H0 in EQ2.
+ monadInv EQ2. eapply sem_cast_match; eauto. eapply sem_cast_match; eauto.
+ rewrite (bool_val_match x v1 (typeof r1)) in EQ2; auto. rewrite H0 in EQ2.
+ monadInv EQ2. constructor.
(* conditional *)
rewrite (bool_val_match x v1 (typeof r1)) in EQ3; auto.
rewrite H0 in EQ3. destruct b; eapply sem_cast_match; eauto.
@@ -496,7 +538,7 @@ Proof.
destruct (access_mode ty); discriminate || eauto.
intuition eauto.
Qed.
-
+
(** Soundness of [constval] with respect to the reduction semantics. *)
Theorem constval_steps:
diff --git a/cfrontend/PrintClight.ml b/cfrontend/PrintClight.ml
index ced5717..ae8e31d 100644
--- a/cfrontend/PrintClight.ml
+++ b/cfrontend/PrintClight.ml
@@ -20,6 +20,7 @@ open Camlcoq
open Datatypes
open Values
open AST
+open PrintAST
open Csyntax
open PrintCsyntax
open Clight
@@ -65,7 +66,6 @@ let rec precedence = function
| Ebinop(Oand, _, _, _) -> (8, LtoR)
| Ebinop(Oxor, _, _, _) -> (7, LtoR)
| Ebinop(Oor, _, _, _) -> (6, LtoR)
- | Econdition _ -> (3, RtoL)
(* Expressions *)
@@ -100,8 +100,6 @@ let rec expr p (prec, e) =
expr (prec1, a1) (name_binop op) expr (prec2, a2)
| Ecast(a1, ty) ->
fprintf p "(%s) %a" (name_type ty) expr (prec', a1)
- | Econdition(a1, a2, a3, _) ->
- fprintf p "%a@ ? %a@ : %a" expr (4, a1) expr (4, a2) expr (4, a3)
end;
if prec' < prec then fprintf p ")@]" else fprintf p "@]"
@@ -125,8 +123,6 @@ 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
@@ -136,6 +132,15 @@ let rec print_stmt p s =
(temp_name id)
print_expr e1
print_expr_list (true, el)
+ | Sbuiltin(None, ef, tyargs, el) ->
+ fprintf p "@[<hv 2>builtin %s@,(@[<hov 0>%a@]);@]"
+ (name_of_external ef)
+ print_expr_list (true, el)
+ | Sbuiltin(Some id, ef, tyargs, el) ->
+ fprintf p "@[<hv 2>%s =@ builtin %s@,(@[<hov 0>%a@]);@]"
+ (temp_name id)
+ (name_of_external ef)
+ print_expr_list (true, el)
| Ssequence(Sskip, s2) ->
print_stmt p s2
| Ssequence(s1, Sskip) ->
@@ -155,19 +160,13 @@ let rec print_stmt p s =
print_expr e
print_stmt s1
print_stmt s2
- | Swhile(e, s) ->
- fprintf p "@[<v 2>while (%a) {@ %a@;<0 -2>}@]"
- print_expr e
- print_stmt s
- | Sdowhile(e, s) ->
- fprintf p "@[<v 2>do {@ %a@;<0 -2>} while(%a);@]"
- print_stmt s
- print_expr e
- | Sfor'(e, s_iter, s_body) ->
- fprintf p "@[<v 2>for (@[<hv 0>;@ %a;@ %a) {@]@ %a@;<0 -2>}@]"
- print_expr e
- print_stmt_for s_iter
- print_stmt s_body
+ | Sloop(s1, Sskip) ->
+ fprintf p "@[<v 2>while (1) {@ %a@;<0 -2>}@]"
+ print_stmt s1
+ | Sloop(s1, s2) ->
+ fprintf p "@[<v 2>for (@[<hv 0>;@ 1;@ %a) {@]@ %a@;<0 -2>}@]"
+ print_stmt_for s2
+ print_stmt s1
| Sbreak ->
fprintf p "break;"
| Scontinue ->
@@ -220,6 +219,15 @@ and print_stmt_for p s =
(temp_name id)
print_expr e1
print_expr_list (true, el)
+ | Sbuiltin(None, ef, tyargs, el) ->
+ fprintf p "@[<hv 2>builtin %s@,(@[<hov 0>%a@]);@]"
+ (name_of_external ef)
+ print_expr_list (true, el)
+ | Sbuiltin(Some id, ef, tyargs, el) ->
+ fprintf p "@[<hv 2>%s =@ builtin %s@,(@[<hov 0>%a@]);@]"
+ (temp_name id)
+ (name_of_external ef)
+ print_expr_list (true, el)
| _ ->
fprintf p "({ %a })" print_stmt s
@@ -265,8 +273,6 @@ let rec collect_expr e =
| Eunop(_, r, _) -> collect_expr r
| Ebinop(_, r1, r2, _) -> collect_expr r1; collect_expr r2
| Ecast(r, _) -> collect_expr r
- | Econdition(r1, r2, r3, _) ->
- collect_expr r1; collect_expr r2; collect_expr r3
let rec collect_exprlist = function
| [] -> ()
@@ -276,14 +282,11 @@ 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
+ | Sbuiltin(optid, ef, tyargs, el) -> 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
- | Swhile(e, s) -> collect_expr e; collect_stmt s
- | Sdowhile(e, s) -> collect_stmt s; collect_expr e
- | Sfor'(e, s_iter, s_body) ->
- collect_expr e; collect_stmt s_iter; collect_stmt s_body
+ | Sloop(s1, s2) -> collect_stmt s1; collect_stmt s2
| Sbreak -> ()
| Scontinue -> ()
| Sswitch(e, cases) -> collect_expr e; collect_cases cases
diff --git a/cfrontend/PrintCsyntax.ml b/cfrontend/PrintCsyntax.ml
index f63c150..2c803bb 100644
--- a/cfrontend/PrintCsyntax.ml
+++ b/cfrontend/PrintCsyntax.ml
@@ -147,7 +147,7 @@ let rec precedence = function
| Evalof(l, _) -> precedence l
| Esizeof _ -> (15, RtoL)
| Ealignof _ -> (15, RtoL)
- | Ecall _ -> (16, LtoR)
+ | Ecall _ | Ebuiltin _ -> (16, LtoR)
| Epostincr _ -> (16, LtoR)
| Eunop _ -> (15, RtoL)
| Eaddrof _ -> (15, RtoL)
@@ -160,6 +160,8 @@ let rec precedence = function
| Ebinop(Oand, _, _, _) -> (8, LtoR)
| Ebinop(Oxor, _, _, _) -> (7, LtoR)
| Ebinop(Oor, _, _, _) -> (6, LtoR)
+ | Eseqand _ -> (5, LtoR)
+ | Eseqor _ -> (4, LtoR)
| Econdition _ -> (3, RtoL)
| Eassign _ -> (2, RtoL)
| Eassignop _ -> (2, RtoL)
@@ -221,12 +223,28 @@ let rec expr p (prec, e) =
fprintf p "%a =@ %a" expr (prec1, a1) expr (prec2, a2)
| Eassignop(op, a1, a2, _, _) ->
fprintf p "%a %s=@ %a" expr (prec1, a1) (name_binop op) expr (prec2, a2)
+ | Eseqand(a1, a2, _) ->
+ fprintf p "%a@ && %a" expr (prec1, a1) expr (prec2, a2)
+ | Eseqor(a1, a2, _) ->
+ fprintf p "%a@ || %a" expr (prec1, a1) expr (prec2, a2)
| Econdition(a1, a2, a3, _) ->
fprintf p "%a@ ? %a@ : %a" expr (4, a1) expr (4, a2) expr (4, a3)
| Ecomma(a1, a2, _) ->
fprintf p "%a,@ %a" expr (prec1, a1) expr (prec2, a2)
| Ecall(a1, al, _) ->
fprintf p "%a@[<hov 1>(%a)@]" expr (prec', a1) exprlist (true, al)
+ | Ebuiltin(EF_memcpy(sz, al), _, args, _) ->
+ fprintf p "__builtin_memcpy_aligned@[<hov 1>(%ld,@ %ld,@ %a)@]"
+ (camlint_of_coqint sz) (camlint_of_coqint al)
+ exprlist (true, args)
+ | Ebuiltin(EF_annot(txt, _), _, args, _) ->
+ fprintf p "__builtin_annot@[<hov 1>(%S,@ %a)@]"
+ (extern_atom txt) exprlist (true, args)
+ | Ebuiltin(EF_annot_val(txt, _), _, args, _) ->
+ fprintf p "__builtin_annot_val@[<hov 1>(%S,@ %a)@]"
+ (extern_atom txt) exprlist (true, args)
+ | Ebuiltin(_, _, args, _) ->
+ fprintf p "<unknown builtin>@[<hov 1>(%a)@]" exprlist (true, args)
| Eparen(a1, ty) ->
fprintf p "(%s) %a" (name_type ty) expr (prec', a1)
end;
@@ -453,6 +471,8 @@ let rec collect_expr e =
| Eunop(_, r, _) -> collect_expr r
| Ebinop(_, r1, r2, _) -> collect_expr r1; collect_expr r2
| Ecast(r, _) -> collect_expr r
+ | Eseqand(r1, r2, _) -> collect_expr r1; collect_expr r2
+ | Eseqor(r1, r2, _) -> collect_expr r1; collect_expr r2
| Econdition(r1, r2, r3, _) ->
collect_expr r1; collect_expr r2; collect_expr r3
| Esizeof(ty, _) -> collect_type ty
@@ -462,6 +482,7 @@ let rec collect_expr e =
| Epostincr(_, l, _) -> collect_expr l
| Ecomma(r1, r2, _) -> collect_expr r1; collect_expr r2
| Ecall(r1, rl, _) -> collect_expr r1; collect_exprlist rl
+ | Ebuiltin(_, _, rl, _) -> collect_exprlist rl
| Eparen _ -> assert false
and collect_exprlist = function
diff --git a/cfrontend/SimplExpr.v b/cfrontend/SimplExpr.v
index 3144b65..6886d81 100644
--- a/cfrontend/SimplExpr.v
+++ b/cfrontend/SimplExpr.v
@@ -134,30 +134,42 @@ Definition transl_incrdecr (id: incr_or_decr) (a: expr) (ty: type) : expr :=
| Decr => Ebinop Osub a (Econst_int Int.one type_int32s) (typeconv ty)
end.
-(** Generate a [Sset] or [Svolread] operation as appropriate
+(** Generate a [Sset] or [Sbuiltin] operation as appropriate
to dereference a l-value [l] and store its result in temporary variable [id]. *)
+Definition chunk_for_volatile_type (ty: type) : option memory_chunk :=
+ if type_is_volatile ty
+ then match access_mode ty with By_value chunk => Some chunk | _ => None end
+ else None.
+
Definition make_set (id: ident) (l: expr) : statement :=
- if type_is_volatile (typeof l)
- then Svolread id l
- else Sset id l.
+ match chunk_for_volatile_type (typeof l) with
+ | None => Sset id l
+ | Some chunk =>
+ let typtr := Tpointer (typeof l) noattr in
+ Sbuiltin (Some id) (EF_vload chunk) (Tcons typtr Tnil) ((Eaddrof l typtr):: nil)
+ end.
(** 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))
+ then do t <- gensym ty; ret (make_set 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)
+
+(** Translation of an assignment. *)
+
+Definition make_assign (l r: expr) : statement :=
+ match chunk_for_volatile_type (typeof l) with
+ | None =>
+ Sassign l r
+ | Some chunk =>
+ let ty := typeof l in
+ let typtr := Tpointer ty noattr in
+ Sbuiltin None (EF_vstore chunk) (Tcons typtr (Tcons ty Tnil))
+ (Eaddrof l typtr :: r :: nil)
end.
-*)
(** Translation of expressions. Return a pair [(sl, a)] of
a list of statements [sl] and a pure expression [a].
@@ -167,31 +179,31 @@ Definition transl_valof (ty: type) (l: expr) : mon (list statement * expr) :=
- If the [dst] argument is [For_effects], the statements [sl]
perform the side effects of the original expression,
and [a] is meaningless.
-- If the [dst] argument is [For_test s1 s2], the statements [sl]
- perform the side effects of the original expression, followed
- by an [if (v) { s1 } else { s2 }] test, where [v] is the value
- of the original expression. [a] is meaningless.
+- If the [dst] argument is [For_set tyl tvar], the statements [sl]
+ perform the side effects of the original expression, then
+ assign the value of the original expression to the temporary [tvar].
+ The value is casted according to the list of types [tyl] before
+ assignment. In this case, [a] is meaningless.
*)
Inductive destination : Type :=
| For_val
| For_effects
- | For_test (tyl: list type) (s1 s2: statement).
+ | For_set (tyl: list type) (tmp: ident).
Definition dummy_expr := Econst_int Int.zero type_int32s.
+Fixpoint do_set (tmp: ident) (tyl: list type) (a: expr) : list statement :=
+ match tyl with
+ | nil => nil
+ | ty1 :: tys => Sset tmp (Ecast a ty1) :: do_set tmp tys (Etempvar tmp ty1)
+ end.
+
Definition finish (dst: destination) (sl: list statement) (a: expr) :=
match dst with
| For_val => (sl, a)
| For_effects => (sl, a)
- | For_test tyl s1 s2 => (sl ++ makeif (fold_left Ecast tyl a) s1 s2 :: nil, a)
- end.
-
-Definition cast_destination (ty: type) (dst: destination) :=
- match dst with
- | For_val => For_val
- | For_effects => For_effects
- | For_test tyl s1 s2 => For_test (ty :: tyl) s1 s2
+ | For_set tyl tmp => (sl ++ do_set tmp tyl a, a)
end.
Fixpoint transl_expr (dst: destination) (a: C.expr) : mon (list statement * expr) :=
@@ -233,39 +245,75 @@ Fixpoint transl_expr (dst: destination) (a: C.expr) : mon (list statement * expr
| C.Ecast r1 ty =>
do (sl1, a1) <- transl_expr For_val r1;
ret (finish dst sl1 (Ecast a1 ty))
+ | C.Eseqand r1 r2 ty =>
+ do (sl1, a1) <- transl_expr For_val r1;
+ match dst with
+ | For_val =>
+ do t <- gensym ty;
+ do (sl2, a2) <- transl_expr (For_set (type_bool :: ty :: nil) t) r2;
+ ret (sl1 ++
+ makeif a1 (makeseq sl2) (Sset t (Econst_int Int.zero ty)) :: nil,
+ Etempvar t ty)
+ | For_effects =>
+ do (sl2, a2) <- transl_expr For_effects r2;
+ ret (sl1 ++ makeif a1 (makeseq sl2) Sskip :: nil, dummy_expr)
+ | For_set tyl t =>
+ do (sl2, a2) <- transl_expr (For_set (type_bool :: ty :: tyl) t) r2;
+ ret (sl1 ++
+ makeif a1 (makeseq sl2) (makeseq (do_set t tyl (Econst_int Int.zero ty))) :: nil,
+ dummy_expr)
+ end
+ | C.Eseqor r1 r2 ty =>
+ do (sl1, a1) <- transl_expr For_val r1;
+ match dst with
+ | For_val =>
+ do t <- gensym ty;
+ do (sl2, a2) <- transl_expr (For_set (type_bool :: ty :: nil) t) r2;
+ ret (sl1 ++
+ makeif a1 (Sset t (Econst_int Int.one ty)) (makeseq sl2) :: nil,
+ Etempvar t ty)
+ | For_effects =>
+ do (sl2, a2) <- transl_expr For_effects r2;
+ ret (sl1 ++ makeif a1 Sskip (makeseq sl2) :: nil, dummy_expr)
+ | For_set tyl t =>
+ do (sl2, a2) <- transl_expr (For_set (type_bool :: ty :: tyl) t) r2;
+ ret (sl1 ++
+ makeif a1 (makeseq (do_set t tyl (Econst_int Int.one ty))) (makeseq sl2) :: nil,
+ dummy_expr)
+ end
| C.Econdition r1 r2 r3 ty =>
- if Cstrategy.simple r2 && Cstrategy.simple r3 then (
- do (sl1, a1) <- transl_expr For_val r1;
- do (sl2, a2) <- transl_expr For_val r2;
- do (sl3, a3) <- transl_expr For_val r3;
- ret (finish dst sl1 (Econdition a1 a2 a3 ty))
- ) else (
- do (sl1, a1) <- transl_expr For_val r1;
- do (sl2, a2) <- transl_expr (cast_destination ty dst) r2;
- do (sl3, a3) <- transl_expr (cast_destination ty dst) r3;
- match dst with
- | For_val =>
- do t <- gensym ty;
- ret (sl1 ++ makeif a1 (Ssequence (makeseq sl2) (Sset t (Ecast a2 ty)))
- (Ssequence (makeseq sl3) (Sset t (Ecast a3 ty))) :: nil,
- Etempvar t ty)
- | For_effects | For_test _ _ _ =>
- ret (sl1 ++ makeif a1 (makeseq sl2) (makeseq sl3) :: nil,
- dummy_expr)
- end)
+ do (sl1, a1) <- transl_expr For_val r1;
+ match dst with
+ | For_val =>
+ do t <- gensym ty;
+ do (sl2, a2) <- transl_expr (For_set (ty::nil) t) r2;
+ do (sl3, a3) <- transl_expr (For_set (ty::nil) t) r3;
+ ret (sl1 ++ makeif a1 (makeseq sl2) (makeseq sl3) :: nil,
+ Etempvar t ty)
+ | For_effects =>
+ do (sl2, a2) <- transl_expr For_effects r2;
+ do (sl3, a3) <- transl_expr For_effects r3;
+ ret (sl1 ++ makeif a1 (makeseq sl2) (makeseq sl3) :: nil,
+ dummy_expr)
+ | For_set tyl t =>
+ do (sl2, a2) <- transl_expr (For_set (ty::tyl) t) r2;
+ do (sl3, a3) <- transl_expr (For_set (ty::tyl) t) r3;
+ ret (sl1 ++ makeif a1 (makeseq sl2) (makeseq sl3) :: nil,
+ dummy_expr)
+ end
| C.Eassign l1 r2 ty =>
do (sl1, a1) <- transl_expr For_val l1;
do (sl2, a2) <- transl_expr For_val r2;
let ty1 := C.typeof l1 in
let ty2 := C.typeof r2 in
match dst with
- | For_val | For_test _ _ _ =>
+ | For_val | For_set _ _ =>
do t <- gensym ty2;
ret (finish dst
- (sl1 ++ sl2 ++ Sset t a2 :: Sassign a1 (Etempvar t ty2) :: nil)
+ (sl1 ++ sl2 ++ Sset t a2 :: make_assign a1 (Etempvar t ty2) :: nil)
(Ecast (Etempvar t ty2) ty1))
| For_effects =>
- ret (sl1 ++ sl2 ++ Sassign a1 a2 :: nil,
+ ret (sl1 ++ sl2 ++ make_assign a1 a2 :: nil,
dummy_expr)
end
| C.Eassignop op l1 r2 tyres ty =>
@@ -274,30 +322,30 @@ Fixpoint transl_expr (dst: destination) (a: C.expr) : mon (list statement * expr
do (sl2, a2) <- transl_expr For_val r2;
do (sl3, a3) <- transl_valof ty1 a1;
match dst with
- | For_val | For_test _ _ _ =>
+ | For_val | For_set _ _ =>
do t <- gensym tyres;
ret (finish dst
(sl1 ++ sl2 ++ sl3 ++
Sset t (Ebinop op a3 a2 tyres) ::
- Sassign a1 (Etempvar t tyres) :: nil)
+ make_assign a1 (Etempvar t tyres) :: nil)
(Ecast (Etempvar t tyres) ty1))
| For_effects =>
- ret (sl1 ++ sl2 ++ sl3 ++ Sassign a1 (Ebinop op a3 a2 tyres) :: nil,
+ ret (sl1 ++ sl2 ++ sl3 ++ make_assign a1 (Ebinop op a3 a2 tyres) :: nil,
dummy_expr)
end
| C.Epostincr id l1 ty =>
let ty1 := C.typeof l1 in
do (sl1, a1) <- transl_expr For_val l1;
match dst with
- | For_val | For_test _ _ _ =>
+ | For_val | For_set _ _ =>
do t <- gensym ty1;
ret (finish dst
(sl1 ++ make_set t a1 ::
- Sassign a1 (transl_incrdecr id (Etempvar t ty1) ty1) :: nil)
+ make_assign a1 (transl_incrdecr id (Etempvar t ty1) ty1) :: nil)
(Etempvar t ty1))
| For_effects =>
do (sl2, a2) <- transl_valof ty1 a1;
- ret (sl1 ++ sl2 ++ Sassign a1 (transl_incrdecr id a2 ty1) :: nil,
+ ret (sl1 ++ sl2 ++ make_assign a1 (transl_incrdecr id a2 ty1) :: nil,
dummy_expr)
end
| C.Ecomma r1 r2 ty =>
@@ -308,13 +356,23 @@ Fixpoint transl_expr (dst: destination) (a: C.expr) : mon (list statement * expr
do (sl1, a1) <- transl_expr For_val r1;
do (sl2, al2) <- transl_exprlist rl2;
match dst with
- | For_val | For_test _ _ _ =>
+ | For_val | For_set _ _ =>
do t <- gensym ty;
ret (finish dst (sl1 ++ sl2 ++ Scall (Some t) a1 al2 :: nil)
(Etempvar t ty))
| For_effects =>
ret (sl1 ++ sl2 ++ Scall None a1 al2 :: nil, dummy_expr)
end
+ | C.Ebuiltin ef tyargs rl ty =>
+ do (sl, al) <- transl_exprlist rl;
+ match dst with
+ | For_val | For_set _ _ =>
+ do t <- gensym ty;
+ ret (finish dst (sl ++ Sbuiltin (Some t) ef tyargs al :: nil)
+ (Etempvar t ty))
+ | For_effects =>
+ ret (sl ++ Sbuiltin None ef tyargs al :: nil, dummy_expr)
+ end
| C.Eparen r1 ty =>
error (msg "SimplExpr.transl_expr: paren")
end
@@ -335,8 +393,14 @@ Definition transl_expression (r: C.expr) : mon (statement * expr) :=
Definition transl_expr_stmt (r: C.expr) : mon statement :=
do (sl, a) <- transl_expr For_effects r; ret (makeseq sl).
+(*
Definition transl_if (r: C.expr) (s1 s2: statement) : mon statement :=
do (sl, a) <- transl_expr (For_test nil s1 s2) r; ret (makeseq sl).
+*)
+
+Definition transl_if (r: C.expr) (s1 s2: statement) : mon statement :=
+ do (sl, a) <- transl_expr For_val r;
+ ret (makeseq (sl ++ makeif a s1 s2 :: nil)).
(** Translation of statements *)
@@ -353,7 +417,7 @@ Defined.
but can duplicate the "then" and "else" branches.
The other produces no code duplication. We choose between the
two based on the shape of the "then" and "else" branches. *)
-
+(*
Fixpoint small_stmt (s: statement) : bool :=
match s with
| Sskip => true
@@ -364,6 +428,7 @@ Fixpoint small_stmt (s: statement) : bool :=
| Ssequence s1 s2 => small_stmt s1 && small_stmt s2
| _ => false
end.
+*)
Fixpoint transl_stmt (s: C.statement) : mon statement :=
match s with
@@ -376,28 +441,25 @@ Fixpoint transl_stmt (s: C.statement) : mon statement :=
| C.Sifthenelse e s1 s2 =>
do ts1 <- transl_stmt s1;
do ts2 <- transl_stmt s2;
- if small_stmt ts1 && small_stmt ts2 then
- transl_if e ts1 ts2
- else
- (do (s', a) <- transl_expression e;
- ret (Ssequence s' (Sifthenelse a ts1 ts2)))
+ do (s', a) <- transl_expression e;
+ ret (Ssequence s' (Sifthenelse a ts1 ts2))
| C.Swhile e s1 =>
do s' <- transl_if e Sskip Sbreak;
do ts1 <- transl_stmt s1;
- ret (Swhile expr_true (Ssequence s' ts1))
+ ret (Sloop (Ssequence s' ts1) Sskip)
| C.Sdowhile e s1 =>
do s' <- transl_if e Sskip Sbreak;
do ts1 <- transl_stmt s1;
- ret (Sfor' expr_true s' ts1)
+ ret (Sloop ts1 s')
| C.Sfor s1 e2 s3 s4 =>
do ts1 <- transl_stmt s1;
do s' <- transl_if e2 Sskip Sbreak;
do ts3 <- transl_stmt s3;
do ts4 <- transl_stmt s4;
if is_Sskip s1 then
- ret (Sfor' expr_true ts3 (Ssequence s' ts4))
+ ret (Sloop (Ssequence s' ts4) ts3)
else
- ret (Ssequence ts1 (Sfor' expr_true ts3 (Ssequence s' ts4)))
+ ret (Ssequence ts1 (Sloop (Ssequence s' ts4) ts3))
| C.Sbreak =>
ret Sbreak
| C.Scontinue =>
diff --git a/cfrontend/SimplExprproof.v b/cfrontend/SimplExprproof.v
index 868f7ab..8a50fcb 100644
--- a/cfrontend/SimplExprproof.v
+++ b/cfrontend/SimplExprproof.v
@@ -121,16 +121,13 @@ Proof.
rewrite H0; auto. simpl; auto.
destruct H1; congruence.
destruct (andb_prop _ _ H6). inv H1.
- rewrite H0; auto. simpl; auto.
- rewrite H9 in H8; discriminate.
+ rewrite H0; eauto. simpl; auto.
+ unfold chunk_for_volatile_type in H9.
+ destruct (type_is_volatile (Csyntax.typeof e1)); simpl in H8; congruence.
rewrite H0; auto. simpl; auto.
rewrite H0; auto. simpl; auto.
destruct (andb_prop _ _ H7). rewrite H0; auto. rewrite H2; auto. simpl; auto.
rewrite H0; auto. simpl; auto.
- destruct (andb_prop _ _ H13). destruct (andb_prop _ _ H14).
- rewrite H2; auto. simpl; auto.
- destruct (andb_prop _ _ H14). destruct (andb_prop _ _ H15). intuition congruence.
- destruct (andb_prop _ _ H13). destruct (andb_prop _ _ H14). intuition congruence.
destruct (andb_prop _ _ H6). rewrite H0; auto.
Qed.
@@ -144,33 +141,48 @@ Lemma tr_simple_exprlist_nil:
simplelist rl = true -> sl = nil.
Proof (proj2 tr_simple_nil).
-(** Evaluation of simple expressions and of their translation *)
+(** Translation of [deref_loc] and [assign_loc] operations. *)
-Remark deref_loc_preserved:
+Remark deref_loc_translated:
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.
+ Csem.deref_loc ge ty m b ofs t v ->
+ match chunk_for_volatile_type ty with
+ | None => t = E0 /\ Clight.deref_loc ty m b ofs v
+ | Some chunk => volatile_load tge chunk m b ofs t v
+ end.
Proof.
- intros. inv H.
- eapply deref_loc_value; eauto.
- eapply deref_loc_volatile; eauto.
- eapply volatile_load_preserved with (ge1 := ge); auto.
+ intros. unfold chunk_for_volatile_type. inv H.
+ (* By_value, not volatile *)
+ rewrite H1. split; auto. eapply deref_loc_value; eauto.
+ (* By_value, volatile *)
+ rewrite H0; rewrite H1. 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.
+ (* By reference *)
+ rewrite H0. destruct (type_is_volatile ty); split; auto; eapply deref_loc_reference; eauto.
+ (* By copy *)
+ rewrite H0. destruct (type_is_volatile ty); split; auto; eapply deref_loc_copy; eauto.
Qed.
-Remark assign_loc_preserved:
+Remark assign_loc_translated:
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'.
+ Csem.assign_loc ge ty m b ofs v t m' ->
+ match chunk_for_volatile_type ty with
+ | None => t = E0 /\ Clight.assign_loc ty m b ofs v m'
+ | Some chunk => volatile_store tge chunk m b ofs v t m'
+ end.
Proof.
- intros. inv H.
- eapply assign_loc_value; eauto.
- eapply assign_loc_volatile; eauto.
- eapply volatile_store_preserved with (ge1 := ge); auto.
+ intros. unfold chunk_for_volatile_type. inv H.
+ (* By_value, not volatile *)
+ rewrite H1. split; auto. eapply assign_loc_value; eauto.
+ (* By_value, volatile *)
+ rewrite H0; rewrite H1. eapply volatile_store_preserved with (ge1 := ge); auto.
exact symbols_preserved. exact block_is_volatile_preserved.
- eapply assign_loc_copy; eauto.
+ (* By copy *)
+ rewrite H0. destruct (type_is_volatile ty); split; auto; eapply assign_loc_copy; eauto.
Qed.
+(** Evaluation of simple expressions and of their translation *)
+
Lemma tr_simple:
forall e m,
(forall r v,
@@ -180,8 +192,8 @@ Lemma tr_simple:
match dst with
| For_val => sl = nil /\ C.typeof r = typeof a /\ eval_expr tge e le m a v
| For_effects => sl = nil
- | For_test tyl s1 s2 =>
- exists b, sl = makeif (fold_left Ecast tyl b) s1 s2 :: nil
+ | For_set tyl t =>
+ exists b, sl = do_set t tyl b
/\ C.typeof r = typeof b
/\ eval_expr tge e le m b v
end)
@@ -200,11 +212,13 @@ Opaque makeif.
auto.
exists a0; auto.
(* rvalof *)
- inv H7; try congruence.
+ 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. rewrite <- B. eapply deref_loc_preserved; eauto.
+ eapply eval_Elvalue. eauto.
+ rewrite <- B.
+ exploit deref_loc_translated; eauto. unfold chunk_for_volatile_type; rewrite H2. tauto.
destruct dst; auto.
econstructor. split. simpl; eauto. auto.
(* addrof *)
@@ -228,16 +242,6 @@ Opaque makeif.
subst sl1; simpl.
assert (eval_expr tge e le m (Ecast a1 ty) v). econstructor; eauto. congruence.
destruct dst; auto. simpl; econstructor; eauto.
-(* condition *)
- exploit H2; eauto. intros [A [B C]]. subst sl1.
- assert (tr_expr le For_val (if b then r2 else r3) (if b then sl2 else sl3) (if b then a2 else a3) (if b then tmp2 else tmp3)).
- destruct b; auto.
- exploit H5; eauto. intros [D [E F]].
- assert (eval_expr tge e le m (Econdition a1 a2 a3 ty) v).
- econstructor; eauto. congruence. rewrite <- E. auto.
- destruct dst; auto. simpl; econstructor; eauto.
- intuition congruence.
- intuition congruence.
(* sizeof *)
destruct dst.
split; auto. split; auto. constructor.
@@ -273,8 +277,10 @@ Lemma tr_simple_rvalue:
match dst with
| For_val => sl = nil /\ C.typeof r = typeof a /\ eval_expr tge e le m a v
| For_effects => sl = nil
- | For_test tyl s1 s2 =>
- exists b, sl = makeif (fold_left Ecast tyl b) s1 s2 :: nil /\ C.typeof r = typeof b /\ eval_expr tge e le m b v
+ | For_set tyl t =>
+ exists b, sl = do_set t tyl b
+ /\ C.typeof r = typeof b
+ /\ eval_expr tge e le m b v
end.
Proof.
intros e m. exact (proj1 (tr_simple e m)).
@@ -315,48 +321,6 @@ Proof.
induction 1; intros; auto.
Qed.
-Inductive compat_dest: (C.expr -> C.expr) -> destination -> destination -> list statement -> Prop :=
- | compat_dest_base: forall dst,
- compat_dest (fun x => x) dst dst nil
- | compat_dest_val: forall C dst sl,
- compat_dest C For_val dst sl
- | compat_dest_effects: forall C dst sl,
- compat_dest C For_effects dst sl
- | compat_dest_paren: forall C ty dst' dst sl,
- compat_dest C dst' (cast_destination ty dst) sl ->
- compat_dest (fun x => C.Eparen (C x) ty) dst' dst sl.
-
-Lemma compat_dest_not_test:
- forall C dst' dst sl,
- compat_dest C dst' dst sl ->
- dst = For_val \/ dst = For_effects ->
- dst' = For_val \/ dst' = For_effects.
-Proof.
- induction 1; intros; auto.
- apply IHcompat_dest. destruct H0; subst; auto.
-Qed.
-
-Lemma compat_dest_change:
- forall C1 dst' dst1 sl1 C2 dst2 sl2,
- compat_dest C1 dst' dst1 sl1 ->
- dst1 = For_val \/ dst1 = For_effects ->
- compat_dest C2 dst' dst2 sl2.
-Proof.
- intros. exploit compat_dest_not_test; eauto. intros [A | A]; subst dst'; constructor.
-Qed.
-
-Lemma compat_dest_test:
- forall C tyl s1 s2 dst sl,
- compat_dest C (For_test tyl s1 s2) dst sl ->
- exists tyl', exists tyl'', dst = For_test tyl'' s1 s2 /\ tyl = tyl' ++ tyl''.
-Proof.
- intros. dependent induction H.
- exists (@nil type); exists tyl; auto.
- exploit IHcompat_dest; eauto. intros [l1 [l2 [A B]]].
- destruct dst; simpl in A; inv A.
- exists (l1 ++ ty :: nil); exists tyl0; split; auto. rewrite app_ass. auto.
-Qed.
-
Scheme leftcontext_ind2 := Minimality for leftcontext Sort Prop
with leftcontextlist_ind2 := Minimality for leftcontextlist Sort Prop.
Combined Scheme leftcontext_leftcontextlist_ind from leftcontext_ind2, leftcontextlist_ind2.
@@ -369,7 +333,6 @@ Lemma tr_expr_leftcontext_rec:
exists dst', exists sl1, exists sl2, exists a', exists tmp',
tr_expr le dst' e sl1 a' tmp'
/\ sl = sl1 ++ sl2
- /\ compat_dest C dst' dst sl2
/\ incl tmp' tmps
/\ (forall le' e' sl3,
tr_expr le' dst' e' sl3 a' tmp' ->
@@ -383,7 +346,6 @@ Lemma tr_expr_leftcontext_rec:
exists dst', exists sl1, exists sl2, exists a', exists tmp',
tr_expr le dst' e sl1 a' tmp'
/\ sl = sl1 ++ sl2
- /\ match dst' with For_test _ _ _ => False | _ => True end
/\ incl tmp' tmps
/\ (forall le' e' sl3,
tr_expr le' dst' e' sl3 a' tmp' ->
@@ -395,7 +357,7 @@ Proof.
Ltac TR :=
econstructor; econstructor; econstructor; econstructor; econstructor;
- split; [eauto | split; [idtac | split; [eauto | split]]].
+ split; [eauto | split; [idtac | split]].
Ltac NOTIN :=
match goal with
@@ -412,41 +374,41 @@ Ltac UNCHANGED :=
intros; apply H; NOTIN
end.
- generalize compat_dest_change; intro CDC.
+ (*generalize compat_dest_change; intro CDC.*)
apply leftcontext_leftcontextlist_ind; intros.
(* base *)
- TR. rewrite <- app_nil_end; auto. constructor. red; auto.
+ TR. rewrite <- app_nil_end; auto. red; auto.
intros. rewrite <- app_nil_end; auto.
(* deref *)
inv H1.
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
- TR. subst sl1; rewrite app_ass; eauto. auto.
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR. subst sl1; rewrite app_ass; eauto. auto.
intros. rewrite <- app_ass. econstructor; eauto.
(* field *)
inv H1.
- exploit H0. eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0. eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1; rewrite app_ass; eauto. auto.
intros. rewrite <- app_ass. econstructor; eauto.
(* rvalof *)
inv H1.
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1; rewrite app_ass; eauto. red; eauto.
intros. rewrite <- app_ass; econstructor; eauto.
exploit typeof_context; eauto. congruence.
(* addrof *)
inv H1.
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1; rewrite app_ass; eauto. auto.
intros. rewrite <- app_ass. econstructor; eauto.
(* unop *)
inv H1.
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1; rewrite app_ass; eauto. auto.
intros. rewrite <- app_ass. econstructor; eauto.
(* binop left *)
inv H1.
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. econstructor; eauto.
@@ -454,55 +416,108 @@ Ltac UNCHANGED :=
(* binop right *)
inv H2.
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
- exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl2. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. change (sl3 ++ sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor; eauto.
eapply tr_expr_invariant; eauto. UNCHANGED.
(* cast *)
inv H1.
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1; rewrite app_ass; eauto. auto.
intros. rewrite <- app_ass. econstructor; eauto.
-(* condition *)
+(* seqand *)
inv H1.
- (* simple *)
- 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. auto. auto. eauto.
+ (* for val *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
+ auto. auto. auto. auto.
+ (* for effects *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
+ auto. auto. auto. auto.
+ (* for set *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
+ auto. auto. auto. auto.
+(* seqor *)
+ inv H1.
+ (* for val *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
- auto. auto. auto. auto. auto.
+ auto. auto. auto. auto.
+ (* for effects *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
+ auto. auto. auto. auto.
+ (* for set *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
+ auto. auto. auto. auto.
+(* condition *)
+ inv H1.
(* for val *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR.
rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. auto. apply S; auto.
+ intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
eapply tr_expr_invariant; eauto. UNCHANGED.
- auto. auto. auto. auto. auto. auto. auto.
+ auto. auto. auto. auto. auto. auto.
(* for effects *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR.
rewrite Q. rewrite app_ass. eauto.
red; auto.
- intros. rewrite <- app_ass. econstructor. auto. auto. apply S; auto.
+ intros. rewrite <- app_ass. eapply tr_condition_effects. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto. auto.
+ (* for set *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR.
+ rewrite Q. rewrite app_ass. eauto.
+ red; auto.
+ intros. rewrite <- app_ass. eapply tr_condition_set. apply S; auto.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
+ eapply tr_expr_invariant; eauto. UNCHANGED.
+ auto. auto. auto. auto. auto. auto.
(* assign left *)
inv H1.
(* for effects *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [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.
(* for val *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
@@ -514,7 +529,7 @@ Ltac UNCHANGED :=
inv H2.
(* for effects *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
- exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl2. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. change (sl3 ++ sl2') with (nil ++ (sl3 ++ sl2')). rewrite app_ass.
@@ -523,7 +538,7 @@ Ltac UNCHANGED :=
apply S; auto. auto. auto. auto.
(* for val *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
- exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl2. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. change (sl3 ++ sl2') with (nil ++ (sl3 ++ sl2')). rewrite app_ass.
@@ -534,7 +549,7 @@ Ltac UNCHANGED :=
(* assignop left *)
inv H1.
(* for effects *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
@@ -542,7 +557,7 @@ Ltac 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]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
@@ -553,7 +568,7 @@ Ltac UNCHANGED :=
inv H2.
(* for effects *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
- exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl2. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. change (sl0 ++ sl2') with (nil ++ sl0 ++ sl2'). rewrite app_ass. econstructor.
@@ -561,7 +576,7 @@ Ltac UNCHANGED :=
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]]]]]]]]].
+ exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl2. rewrite app_ass. eauto.
red; auto.
intros. rewrite <- app_ass. change (sl0 ++ sl2') with (nil ++ sl0 ++ sl2'). rewrite app_ass. econstructor.
@@ -570,25 +585,25 @@ Ltac UNCHANGED :=
(* postincr *)
inv H1.
(* for effects *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto. red; auto.
intros. rewrite <- app_ass. econstructor; eauto.
symmetry; eapply typeof_context; eauto.
(* for val *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto. red; auto.
intros. rewrite <- app_ass. econstructor; eauto.
eapply typeof_context; eauto.
(* call left *)
inv H1.
(* for effects *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto. red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_exprlist_invariant; eauto. UNCHANGED.
auto. auto. auto.
(* for val *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto. red; auto.
intros. rewrite <- app_ass. econstructor. auto. apply S; auto.
eapply tr_exprlist_invariant; eauto. UNCHANGED.
@@ -597,24 +612,40 @@ Ltac UNCHANGED :=
inv H2.
(* for effects *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
- exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
- TR. rewrite Q; rewrite app_ass; eauto. destruct dst'; contradiction || constructor.
+ exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR. rewrite Q; rewrite app_ass; eauto.
+ (*destruct dst'; constructor||contradiction.*)
red; auto.
intros. rewrite <- app_ass. change (sl3++sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
apply S; auto. auto. auto. auto.
(* for val *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
- exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
- TR. rewrite Q; rewrite app_ass; eauto. destruct dst'; contradiction || constructor.
+ exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR. rewrite Q; rewrite app_ass; eauto.
+ (*destruct dst'; constructor||contradiction.*)
red; auto.
intros. rewrite <- app_ass. change (sl3++sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor.
auto. eapply tr_expr_invariant; eauto. UNCHANGED.
apply S; auto.
auto. auto. auto. auto.
+(* builtin *)
+ inv H1.
+ (* for effects *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR. rewrite Q; rewrite app_ass; eauto.
+ red; auto.
+ intros. rewrite <- app_ass. change (sl3++sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor.
+ apply S; auto. auto.
+ (* for val *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR. rewrite Q; rewrite app_ass; eauto.
+ red; auto.
+ intros. rewrite <- app_ass. change (sl3++sl2') with (nil ++ sl3 ++ sl2'). rewrite app_ass. econstructor.
+ auto. apply S; auto. auto. auto.
(* comma *)
inv H1.
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. rewrite Q; rewrite app_ass; eauto. red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
@@ -622,18 +653,21 @@ Ltac UNCHANGED :=
(* paren *)
inv H1.
(* 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.
- intros. rewrite <- app_ass. econstructor; eauto.
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR. rewrite Q. eauto. red; auto.
+ intros. econstructor; eauto.
(* for effects *)
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
- TR. rewrite Q. eauto. constructor; auto. auto.
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR. rewrite Q. eauto. auto.
+ intros. econstructor; eauto.
+ (* for set *)
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
+ TR. rewrite Q. eauto. auto.
intros. econstructor; eauto.
(* cons left *)
inv H1.
- exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1. rewrite app_ass. eauto.
- exploit compat_dest_not_test; eauto. intros [A|A]; subst dst'; auto.
red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_exprlist_invariant; eauto. UNCHANGED.
@@ -641,7 +675,7 @@ Ltac UNCHANGED :=
(* cons right *)
inv H2.
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
- exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl2. eauto.
red; auto.
intros. change sl3 with (nil ++ sl3). rewrite app_ass. econstructor.
@@ -657,7 +691,6 @@ Theorem tr_expr_leftcontext:
exists dst', exists sl1, exists sl2, exists a', exists tmp',
tr_expr le dst' r sl1 a' tmp'
/\ sl = sl1 ++ sl2
- /\ compat_dest C dst' dst sl2
/\ incl tmp' tmps
/\ (forall le' r' sl3,
tr_expr le' dst' r' sl3 a' tmp' ->
@@ -677,7 +710,6 @@ Theorem tr_top_leftcontext:
exists dst', exists sl1, exists sl2, exists a', exists tmp',
tr_top tge e le m dst' r sl1 a' tmp'
/\ sl = sl1 ++ sl2
- /\ compat_dest C dst' dst sl2
/\ incl tmp' tmps
/\ (forall le' m' r' sl3,
tr_expr le' dst' r' sl3 a' tmp' ->
@@ -691,55 +723,15 @@ Proof.
exists For_val; econstructor; econstructor; econstructor; econstructor.
split. apply tr_top_val_val; eauto.
split. instantiate (1 := nil); auto.
- split. constructor.
split. apply incl_refl.
intros. rewrite <- app_nil_end. constructor; auto.
-(* val for test *)
- inv H2; inv H1.
- exists (For_test tyl s1 s2); econstructor; econstructor; econstructor; econstructor.
- split. apply tr_top_val_test; eauto.
- split. instantiate (1 := nil); auto.
- split. constructor.
- split. apply incl_refl.
- intros. rewrite <- app_nil_end. constructor; eauto.
(* base *)
subst r. exploit tr_expr_leftcontext; eauto.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R [S T]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
exists dst'; exists sl1; exists sl2; exists a'; exists tmp'.
split. apply tr_top_base; auto.
- split. auto. split. auto. split. auto.
- intros. apply tr_top_base. apply T; auto.
-(* paren *)
- inv H1; inv H0.
- (* at top *)
- exists (For_test tyl s1 s2); econstructor; econstructor; econstructor; econstructor.
- split. apply tr_top_paren_test; eauto.
- split. instantiate (1 := nil). rewrite <- app_nil_end; auto.
- split. constructor.
- split. apply incl_refl.
- intros. rewrite <- app_nil_end. constructor; eauto.
- (* below *)
- exploit (IHtr_top r0 C0); auto.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
- exists dst'; exists sl1; exists sl2; exists a'; exists tmp'.
- split. auto.
- split. auto.
- split. constructor; auto.
- split. auto.
- intros. apply tr_top_paren_test. apply S; auto.
-Qed.
-
-Theorem tr_top_testcontext:
- forall C tyl s1 s2 dst sl2 r sl1 a tmps e le m,
- compat_dest C (For_test tyl s1 s2) dst sl2 ->
- tr_top tge e le m (For_test tyl s1 s2) r sl1 a tmps ->
- tr_top tge e le m dst (C r) (sl1 ++ sl2) a tmps.
-Proof.
- intros. dependent induction H.
- rewrite <- app_nil_end. auto.
- exploit compat_dest_test; eauto. intros [l1 [l2 [A B]]].
- destruct dst; simpl in A; inv A.
- apply tr_top_paren_test. eapply IHcompat_dest; eauto.
+ split. auto. split. auto.
+ intros. apply tr_top_base. apply S; auto.
Qed.
(** Semantics of smart constructors *)
@@ -771,18 +763,44 @@ Qed.
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 ->
+ Csem.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.
+ intros. exploit deref_loc_translated; eauto. rewrite <- H1.
+ unfold make_set. destruct (chunk_for_volatile_type (typeof a)) as [chunk|].
+(* volatile case *)
+ intros. change (PTree.set id v le) with (set_opttemp (Some id) v le). econstructor.
+ econstructor. constructor. eauto.
+ simpl. unfold sem_cast. simpl. eauto. constructor.
+ simpl. econstructor; eauto.
+(* nonvolatile case *)
+ intros [A B]. subst t. constructor. eapply eval_Elvalue; eauto.
+Qed.
+
+Lemma step_make_assign:
+ forall a1 a2 ty m b ofs t v m' v2 e le f k,
+ Csem.assign_loc ge ty m b ofs v t m' ->
+ eval_lvalue tge e le m a1 b ofs ->
+ eval_expr tge e le m a2 v2 ->
+ sem_cast v2 (typeof a2) ty = Some v ->
+ typeof a1 = ty ->
+ step tge (State f (make_assign a1 a2) k e le m)
+ t (State f Sskip k e le m').
+Proof.
+ intros. exploit assign_loc_translated; eauto. rewrite <- H3.
+ unfold make_assign. destruct (chunk_for_volatile_type (typeof a1)) as [chunk|].
+(* volatile case *)
+ intros. change le with (set_opttemp None Vundef le) at 2. econstructor.
+ econstructor. constructor. eauto.
+ simpl. unfold sem_cast. simpl. eauto.
+ econstructor; eauto. rewrite H3; eauto. constructor.
+ simpl. econstructor; eauto.
+(* nonvolatile case *)
+ intros [A B]. subst t. econstructor; eauto. congruence.
+Qed.
Fixpoint Kseqlist (sl: list statement) (k: cont) :=
match sl with
@@ -797,14 +815,20 @@ 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 push_seq:
+ forall f sl k e le m,
+ star step tge (State f (makeseq sl) k e le m)
+ E0 (State f Sskip (Kseqlist sl k) e le m).
+Proof.
+ intros. unfold makeseq. generalize Sskip. revert sl k.
+ induction sl; simpl; intros.
+ apply star_refl.
+ eapply star_right. apply IHsl. constructor. traceEq.
+Qed.
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 ->
+ Csem.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 ->
@@ -815,22 +839,23 @@ Lemma step_tr_rvalof:
/\ 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.
+ intros. inv H1.
+ (* not volatile *)
+ exploit deref_loc_translated; eauto. unfold chunk_for_volatile_type; rewrite H3.
+ intros [A B]. subst t.
+ exists le; split. apply star_refl.
+ split. eapply eval_Elvalue; eauto.
+ auto.
(* 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.
+ intros. exists (PTree.set t0 v le); split.
+ simpl. eapply star_two. econstructor. eapply step_make_set; eauto. traceEq.
+ split. constructor. apply PTree.gss.
+ split. auto.
+ intros. apply PTree.gso. congruence.
Qed.
(** Matching between continuations *)
-
Inductive match_cont : Csem.cont -> cont -> Prop :=
| match_Kstop:
match_cont Csem.Kstop Kstop
@@ -843,37 +868,38 @@ Inductive match_cont : Csem.cont -> cont -> Prop :=
tr_stmt s ts ->
match_cont k tk ->
match_cont (Csem.Kwhile2 r s k)
- (Kwhile expr_true (Ssequence s' ts) tk)
+ (Kloop1 (Ssequence s' ts) Sskip tk)
| match_Kdowhile1: forall r s k s' ts tk,
tr_if r Sskip Sbreak s' ->
tr_stmt s ts ->
match_cont k tk ->
match_cont (Csem.Kdowhile1 r s k)
- (Kfor2 expr_true s' ts tk)
+ (Kloop1 ts s' tk)
| match_Kfor3: forall r s3 s k ts3 s' ts tk,
tr_if r Sskip Sbreak s' ->
tr_stmt s3 ts3 ->
tr_stmt s ts ->
match_cont k tk ->
match_cont (Csem.Kfor3 r s3 s k)
- (Kfor2 expr_true ts3 (Ssequence s' ts) tk)
+ (Kloop1 (Ssequence s' ts) ts3 tk)
| match_Kfor4: forall r s3 s k ts3 s' ts tk,
tr_if r Sskip Sbreak s' ->
tr_stmt s3 ts3 ->
tr_stmt s ts ->
match_cont k tk ->
match_cont (Csem.Kfor4 r s3 s k)
- (Kfor3 expr_true ts3 (Ssequence s' ts) tk)
+ (Kloop2 (Ssequence s' ts) ts3 tk)
| match_Kswitch2: forall k tk,
match_cont k tk ->
match_cont (Csem.Kswitch2 k) (Kswitch tk)
- | match_Kcall_none: forall f e C ty k tf le sl tk a dest tmps,
+ | match_Kcall: forall f e C ty k optid tf le sl tk a dest tmps,
transl_function f = Errors.OK tf ->
leftcontext RV RV C ->
- (forall v m, tr_top tge e le m dest (C (C.Eval v ty)) sl a tmps) ->
+ (forall v m, tr_top tge e (set_opttemp optid v le) m dest (C (C.Eval v ty)) sl a tmps) ->
match_cont_exp dest a k tk ->
match_cont (Csem.Kcall f e C ty k)
- (Kcall None tf e le (Kseqlist sl tk))
+ (Kcall optid tf e le (Kseqlist sl tk))
+(*
| match_Kcall_some: forall f e C ty k dst tf le sl tk a dest tmps,
transl_function f = Errors.OK tf ->
leftcontext RV RV C ->
@@ -881,6 +907,7 @@ Inductive match_cont : Csem.cont -> cont -> Prop :=
match_cont_exp dest a k tk ->
match_cont (Csem.Kcall f e C ty k)
(Kcall (Some dst) tf e le (Kseqlist sl tk))
+*)
with match_cont_exp : destination -> expr -> Csem.cont -> cont -> Prop :=
| match_Kdo: forall k a tk,
@@ -890,32 +917,30 @@ with match_cont_exp : destination -> expr -> Csem.cont -> cont -> Prop :=
tr_stmt s1 ts1 -> tr_stmt s2 ts2 ->
match_cont k tk ->
match_cont_exp For_val a (Csem.Kifthenelse s1 s2 k) (Kseq (Sifthenelse a ts1 ts2) tk)
- | match_Kifthenelse_2: forall a s1 s2 k ts1 ts2 tk,
- tr_stmt s1 ts1 -> tr_stmt s2 ts2 ->
- match_cont k tk ->
- match_cont_exp (For_test nil ts1 ts2) a (Csem.Kifthenelse s1 s2 k) tk
| match_Kwhile1: forall r s k s' a ts tk,
tr_if r Sskip Sbreak s' ->
tr_stmt s ts ->
match_cont k tk ->
- match_cont_exp (For_test nil Sskip Sbreak) a
+ match_cont_exp For_val a
(Csem.Kwhile1 r s k)
- (Kseq ts (Kwhile expr_true (Ssequence s' ts) tk))
+ (Kseq (makeif a Sskip Sbreak)
+ (Kseq ts (Kloop1 (Ssequence s' ts) Sskip tk)))
| match_Kdowhile2: forall r s k s' a ts tk,
tr_if r Sskip Sbreak s' ->
tr_stmt s ts ->
match_cont k tk ->
- match_cont_exp (For_test nil Sskip Sbreak) a
+ match_cont_exp For_val a
(Csem.Kdowhile2 r s k)
- (Kfor3 expr_true s' ts tk)
+ (Kseq (makeif a Sskip Sbreak) (Kloop2 ts s' tk))
| match_Kfor2: forall r s3 s k s' a ts3 ts tk,
tr_if r Sskip Sbreak s' ->
tr_stmt s3 ts3 ->
tr_stmt s ts ->
match_cont k tk ->
- match_cont_exp (For_test nil Sskip Sbreak) a
+ match_cont_exp For_val a
(Csem.Kfor2 r s3 s k)
- (Kseq ts (Kfor2 expr_true ts3 (Ssequence s' ts) tk))
+ (Kseq (makeif a Sskip Sbreak)
+ (Kseq ts (Kloop1 (Ssequence s' ts) ts3 tk)))
| match_Kswitch1: forall ls k a tls tk,
tr_lblstmts ls tls ->
match_cont k tk ->
@@ -929,15 +954,7 @@ Lemma match_cont_call:
match_cont k tk ->
match_cont (Csem.call_cont k) (call_cont tk).
Proof.
- induction 1; simpl; auto. constructor. econstructor; eauto. econstructor; eauto.
-Qed.
-
-Lemma match_cont_exp_for_test_inv:
- forall tyl s1 s2 a a' k tk,
- match_cont_exp (For_test tyl s1 s2) a k tk ->
- match_cont_exp (For_test tyl s1 s2) a' k tk.
-Proof.
- intros. inv H; econstructor; eauto.
+ induction 1; simpl; auto. constructor. econstructor; eauto.
Qed.
(** Matching between states *)
@@ -967,17 +984,6 @@ Inductive match_states: Csem.state -> state -> Prop :=
| match_stuckstate: forall S,
match_states Csem.Stuckstate S.
-Lemma push_seq:
- forall f sl k e le m,
- star step tge (State f (makeseq sl) k e le m)
- E0 (State f Sskip (Kseqlist sl k) e le m).
-Proof.
- intros. unfold makeseq. generalize Sskip. revert sl k.
- induction sl; simpl; intros.
- apply star_refl.
- eapply star_right. apply IHsl. constructor. traceEq.
-Qed.
-
(** Additional results on translation of statements *)
Lemma tr_select_switch:
@@ -1028,13 +1034,6 @@ Proof.
intros. unfold makeseq. apply H; auto. red. auto.
Qed.
-Lemma small_stmt_nolabel:
- forall s, small_stmt s = true -> nolabel s.
-Proof.
- induction s; simpl; intros; congruence || (red; auto).
- destruct (andb_prop _ _ H). intros; simpl. rewrite IHs1; auto. apply IHs2; auto.
-Qed.
-
Lemma makeif_nolabel:
forall a s1 s2, nolabel s1 -> nolabel s2 -> nolabel (makeif a s1 s2).
Proof.
@@ -1043,72 +1042,81 @@ Proof.
red; simpl; intros. rewrite H; auto.
Qed.
+Lemma make_set_nolabel:
+ forall t a, nolabel (make_set t a).
+Proof.
+ unfold make_set; intros; red; intros.
+ destruct (chunk_for_volatile_type (typeof a)); auto.
+Qed.
+
+Lemma make_assign_nolabel:
+ forall l r, nolabel (make_assign l r).
+Proof.
+ unfold make_assign; intros; red; intros.
+ destruct (chunk_for_volatile_type (typeof l)); auto.
+Qed.
+
Lemma tr_rvalof_nolabel:
forall ty a sl a' tmp, tr_rvalof ty a sl a' tmp -> nolabel_list sl.
Proof.
- destruct 1; simpl; intuition. red; simpl; auto.
+ destruct 1; simpl; intuition. apply make_set_nolabel.
Qed.
-Definition nolabel_dest (dst: destination) : Prop :=
- match dst with
- | For_val => True
- | For_effects => True
- | For_test tyl s1 s2 => nolabel s1 /\ nolabel s2
- end.
+Lemma nolabel_do_set:
+ forall tmp tyl a, nolabel_list (do_set tmp tyl a).
+Proof.
+ induction tyl; intros; simpl; split; auto; red; auto.
+Qed.
Lemma nolabel_final:
- forall dst a, nolabel_dest dst -> nolabel_list (final dst a).
+ forall dst a, nolabel_list (final dst a).
Proof.
- destruct dst; simpl; intros. auto. auto.
- split; auto. destruct H. apply makeif_nolabel; auto.
+ destruct dst; simpl; intros. auto. auto. apply nolabel_do_set.
Qed.
+(*
Lemma nolabel_dest_cast:
forall ty dst, nolabel_dest dst -> nolabel_dest (cast_destination ty dst).
Proof.
unfold nolabel_dest; intros; destruct dst; auto.
Qed.
+*)
Ltac NoLabelTac :=
match goal with
| [ |- nolabel_list nil ] => exact I
- | [ |- nolabel_list (final _ _) ] => apply nolabel_final; NoLabelTac
+ | [ |- nolabel_list (final _ _) ] => apply nolabel_final (*; NoLabelTac*)
| [ |- nolabel_list (_ :: _) ] => simpl; split; NoLabelTac
| [ |- nolabel_list (_ ++ _) ] => apply nolabel_list_app; NoLabelTac
- | [ |- nolabel_dest For_val ] => exact I
- | [ |- nolabel_dest For_effects ] => exact I
| [ H: _ -> nolabel_list ?x |- nolabel_list ?x ] => apply H; NoLabelTac
+ | [ |- nolabel (makeseq _) ] => apply makeseq_nolabel; NoLabelTac
+ | [ |- nolabel (makeif _ _ _) ] => apply makeif_nolabel; NoLabelTac
+ | [ |- nolabel (make_set _ _) ] => apply make_set_nolabel
+ | [ |- nolabel (make_assign _ _) ] => apply make_assign_nolabel
| [ |- nolabel _ ] => red; intros; simpl; auto
| [ |- _ /\ _ ] => split; NoLabelTac
| _ => auto
end.
Lemma tr_find_label_expr:
- (forall le dst r sl a tmps, tr_expr le dst r sl a tmps -> nolabel_dest dst -> nolabel_list sl)
+ (forall le dst r sl a tmps, tr_expr le dst r sl a tmps -> nolabel_list sl)
/\(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.
+ apply nolabel_do_set.
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 H4. apply nolabel_dest_cast; auto.
- rewrite (makeseq_nolabel sl3). auto. apply H6. apply nolabel_dest_cast; auto.
+ apply nolabel_do_set.
+ apply nolabel_do_set.
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.
Lemma tr_find_label_top:
forall e le m dst r sl a tmps,
- tr_top tge e le m dst r sl a tmps -> nolabel_dest dst -> nolabel_list sl.
+ tr_top tge e le m dst r sl a tmps -> nolabel_list sl.
Proof.
induction 1; intros; NoLabelTac.
- destruct H1. apply makeif_nolabel; auto.
eapply (proj1 tr_find_label_expr); eauto.
Qed.
@@ -1118,8 +1126,7 @@ Proof.
intros. inv H.
assert (nolabel (makeseq sl)). apply makeseq_nolabel.
eapply tr_find_label_top with (e := empty_env) (le := PTree.empty val) (m := Mem.empty).
- eauto. exact I.
- apply H.
+ eauto. apply H.
Qed.
Lemma tr_find_label_expr_stmt:
@@ -1128,20 +1135,21 @@ Proof.
intros. inv H.
assert (nolabel (makeseq sl)). apply makeseq_nolabel.
eapply tr_find_label_top with (e := empty_env) (le := PTree.empty val) (m := Mem.empty).
- eauto. exact I.
- apply H.
+ eauto. apply H.
Qed.
Lemma tr_find_label_if:
- forall r s1 s2 s,
- tr_if r s1 s2 s ->
- small_stmt s1 = true -> small_stmt s2 = true ->
+ forall r s,
+ tr_if r Sskip Sbreak s ->
forall k, find_label lbl s k = None.
Proof.
intros. inv H.
- assert (nolabel (makeseq sl)). apply makeseq_nolabel.
+ assert (nolabel (makeseq (sl ++ makeif a Sskip Sbreak :: nil))).
+ apply makeseq_nolabel.
+ apply nolabel_list_app.
eapply tr_find_label_top with (e := empty_env) (le := PTree.empty val) (m := Mem.empty).
- eauto. split; apply small_stmt_nolabel; auto.
+ eauto.
+ simpl; split; auto. apply makeif_nolabel. red; simpl; auto. red; simpl; auto.
apply H.
Qed.
@@ -1180,39 +1188,30 @@ Proof.
destruct (Csem.find_label lbl s1 (Csem.Kseq s2 k)) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; auto.
intro EQ. rewrite EQ. eapply IHs2; eauto.
-(* if no-opt *)
+(* if *)
rename s' into sr.
rewrite (tr_find_label_expression _ _ _ H2).
exploit (IHs1 k); eauto.
destruct (Csem.find_label lbl s1 k) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
intro EQ. rewrite EQ. eapply IHs2; eauto.
-(* if opt *)
- rewrite (tr_find_label_if _ _ _ _ H7); auto.
- exploit (IHs1 k); eauto.
- destruct (Csem.find_label lbl s1 k) as [[s' k'] | ].
- intros [ts' [tk' [A [B C]]]].
- exploit small_stmt_nolabel. eexact H4. instantiate (1 := tk). congruence.
- intros.
- exploit (IHs2 k); eauto.
- destruct (Csem.find_label lbl s2 k) as [[s' k'] | ].
- intros [ts' [tk' [A [B C]]]].
- exploit small_stmt_nolabel. eexact H6. instantiate (1 := tk). congruence.
- auto.
(* while *)
rename s' into sr.
- rewrite (tr_find_label_if _ _ _ _ H1); auto.
- eapply IHs; eauto. econstructor; eauto.
+ rewrite (tr_find_label_if _ _ H1); auto.
+ exploit (IHs (Kwhile2 e s k)); eauto. econstructor; eauto.
+ destruct (Csem.find_label lbl s (Kwhile2 e s k)) as [[s' k'] | ].
+ intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
+ intro EQ. rewrite EQ. auto.
(* dowhile *)
rename s' into sr.
- rewrite (tr_find_label_if _ _ _ _ H1); auto.
+ rewrite (tr_find_label_if _ _ H1); auto.
exploit (IHs (Kdowhile1 e s k)); eauto. econstructor; eauto.
destruct (Csem.find_label lbl s (Kdowhile1 e s k)) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
intro EQ. rewrite EQ. auto.
(* for skip *)
rename s' into sr.
- rewrite (tr_find_label_if _ _ _ _ H4); auto.
+ rewrite (tr_find_label_if _ _ H4); auto.
exploit (IHs3 (Csem.Kfor3 e s2 s3 k)); eauto. econstructor; eauto.
destruct (Csem.find_label lbl s3 (Csem.Kfor3 e s2 s3 k)) as [[s' k'] | ].
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; intuition.
@@ -1220,7 +1219,7 @@ Proof.
exploit (IHs2 (Csem.Kfor4 e s2 s3 k)); eauto. econstructor; eauto.
(* for not skip *)
rename s' into sr.
- rewrite (tr_find_label_if _ _ _ _ H3); auto.
+ rewrite (tr_find_label_if _ _ H3); auto.
exploit (IHs1 (Csem.Kseq (C.Sfor C.Sskip e s2 s3) k)); eauto.
econstructor; eauto. econstructor; eauto.
destruct (Csem.find_label lbl s1
@@ -1254,7 +1253,7 @@ Proof.
intros [ts' [tk' [A [B C]]]]. rewrite A. exists ts'; exists tk'; auto.
intro EQ. rewrite EQ. eapply IHs; eauto.
Qed.
-
+
End FIND_LABEL.
(** Anti-stuttering measure *)
@@ -1283,6 +1282,8 @@ Fixpoint esize (a: C.expr) : nat :=
| C.Eunop _ r1 _ => S(esize r1)
| C.Ebinop _ r1 r2 _ => S(esize r1 + esize r2)%nat
| C.Ecast r1 _ => S(esize r1)
+ | C.Eseqand r1 _ _ => S(esize r1)
+ | C.Eseqor r1 _ _ => S(esize r1)
| C.Econdition r1 _ _ _ => S(esize r1)
| C.Esizeof _ _ => 1%nat
| C.Ealignof _ _ => 1%nat
@@ -1291,6 +1292,7 @@ Fixpoint esize (a: C.expr) : nat :=
| C.Epostincr _ l1 _ => S(esize l1)
| C.Ecomma r1 r2 _ => S(esize r1 + esize r2)%nat
| C.Ecall r1 rl2 _ => S(esize r1 + esizelist rl2)%nat
+ | C.Ebuiltin ef _ rl _ => S(esizelist rl)%nat
| C.Eparen r1 _ => S(esize r1)
end
@@ -1322,7 +1324,9 @@ with leftcontextlist_size:
(esize e1 < esize e2)%nat ->
(esizelist (C e1) < esizelist (C e2))%nat.
Proof.
- induction 1; intros; simpl; auto with arith. exploit leftcontextlist_size; eauto. auto with arith.
+ induction 1; intros; simpl; auto with arith.
+ exploit leftcontextlist_size; eauto. auto with arith.
+ exploit leftcontextlist_size; eauto. auto with arith.
induction 1; intros; simpl; auto with arith. exploit leftcontext_size; eauto. auto with arith.
Qed.
@@ -1350,57 +1354,226 @@ Proof.
induction 1; intros; inv MS.
(* expr *)
assert (tr_expr le dest r sl a tmps).
- inv H9. contradiction. contradiction. auto. inv H.
+ inv H9. contradiction. auto.
+ exploit tr_simple_rvalue; eauto. destruct dest.
+ (* for val *)
+ intros [SL1 [TY1 EV1]]. subst sl.
econstructor; split.
right; split. apply star_refl. destruct r; simpl; (contradiction || omega).
econstructor; eauto.
- instantiate (1 := tmps).
- 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.
+ instantiate (1 := tmps). apply tr_top_val_val; auto.
+ (* for effects *)
+ intros SL1. subst sl.
+ econstructor; split.
+ right; split. apply star_refl. destruct r; simpl; (contradiction || omega).
+ econstructor; eauto.
+ instantiate (1 := tmps). apply tr_top_base. constructor.
+ (* for set *)
+ inv H10.
(* 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.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [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.
+ left. eapply plus_two. constructor. eapply step_make_set; eauto. traceEq.
+ 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.
+(* seqand true *)
+ exploit tr_top_leftcontext; eauto. clear H9.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
+ inv P. inv H2.
+ (* for val *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. apply tr_paren_val with (a1 := dummy_expr); auto. econstructor; eauto.
+ apply tr_expr_monotone with tmp2; eauto. auto. auto.
+ (* for effects *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. apply tr_paren_effects with (a1 := dummy_expr); auto. econstructor; eauto.
+ apply tr_expr_monotone with tmp2; eauto. auto. auto.
+ (* for set *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. apply tr_paren_set with (a1 := dummy_expr); auto. econstructor; eauto.
+ apply tr_expr_monotone with tmp2; eauto. auto. auto.
+(* seqand false *)
+ exploit tr_top_leftcontext; eauto. clear H9.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
+ inv P. inv H2.
+ (* for val *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
+ apply star_one. constructor. constructor. reflexivity. reflexivity.
+ eapply match_exprstates; eauto.
+ change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
+ intros. constructor. rewrite H2. apply PTree.gss. auto.
+ intros. apply PTree.gso. congruence.
+ auto.
+ (* for effects *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ apply step_makeif with (b := false) (v1 := v); auto. congruence.
+ reflexivity.
+ eapply match_exprstates; eauto.
+ change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
+ auto. auto.
+ (* for set *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. econstructor; eauto. intros. constructor. auto. auto.
+(* seqor true *)
+ exploit tr_top_leftcontext; eauto. clear H9.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
+ inv P. inv H2.
+ (* for val *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
+ apply star_one. constructor. constructor. reflexivity. reflexivity.
+ eapply match_exprstates; eauto.
+ change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
+ intros. constructor. rewrite H2. apply PTree.gss. auto.
+ intros. apply PTree.gso. congruence.
+ auto.
+ (* for effects *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ apply step_makeif with (b := true) (v1 := v); auto. congruence.
+ reflexivity.
+ eapply match_exprstates; eauto.
+ change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
+ auto. auto.
+ (* for set *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. econstructor; eauto. intros. constructor. auto. auto.
+(* seqand false *)
+ exploit tr_top_leftcontext; eauto. clear H9.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
+ inv P. inv H2.
+ (* for val *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. apply tr_paren_val with (a1 := dummy_expr); auto. econstructor; eauto.
+ apply tr_expr_monotone with tmp2; eauto. auto. auto.
+ (* for effects *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. apply tr_paren_effects with (a1 := dummy_expr); auto. econstructor; eauto.
+ apply tr_expr_monotone with tmp2; eauto. auto. auto.
+ (* for set *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. apply tr_paren_set with (a1 := dummy_expr); auto. econstructor; eauto.
+ apply tr_expr_monotone with tmp2; eauto. auto. auto.
(* condition *)
- exploit tr_top_leftcontext; eauto. clear H10.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
- inv P. inv H3.
- (* simple *)
- intuition congruence.
+ exploit tr_top_leftcontext; eauto. clear H9.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
+ inv P. inv H2.
(* for value *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist. destruct b.
econstructor; split.
left. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
- eapply star_left. constructor. apply push_seq. reflexivity. reflexivity. traceEq.
- replace (Kseqlist sl3 (Kseq (Sset t (Ecast a2 ty)) (Kseqlist sl2 tk)))
- with (Kseqlist (sl3 ++ Sset t (Ecast a2 ty) :: sl2) tk).
- eapply match_exprstates; eauto.
- change (Sset t (Ecast a2 ty) :: sl2) with ((Sset t (Ecast a2 ty) :: nil) ++ sl2). rewrite <- app_ass.
- apply S. econstructor; eauto. auto. auto.
- rewrite Kseqlist_app. auto.
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. econstructor; eauto. apply tr_expr_monotone with tmp2; eauto. auto. auto.
econstructor; split.
left. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
- eapply star_left. constructor. apply push_seq. reflexivity. reflexivity. traceEq.
- replace (Kseqlist sl4 (Kseq (Sset t (Ecast a3 ty)) (Kseqlist sl2 tk)))
- with (Kseqlist (sl4 ++ Sset t (Ecast a3 ty) :: sl2) tk).
- eapply match_exprstates; eauto.
- change (Sset t (Ecast a3 ty) :: sl2) with ((Sset t (Ecast a3 ty) :: nil) ++ sl2). rewrite <- app_ass.
- apply S. econstructor; eauto. auto. auto.
- rewrite Kseqlist_app. auto.
- (* for effects *)
+ apply push_seq. reflexivity. reflexivity.
+ rewrite <- Kseqlist_app.
+ eapply match_exprstates; eauto.
+ apply S. econstructor; eauto. apply tr_expr_monotone with tmp3; eauto. auto. auto.
+ (* for effects *)
+ exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
+ subst sl0; simpl Kseqlist. destruct b.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := true) (v1 := v); auto. congruence.
+ apply push_seq.
+ reflexivity. traceEq.
+ rewrite <- Kseqlist_app.
+ econstructor. eauto. apply S.
+ econstructor; eauto. apply tr_expr_monotone with tmp2; eauto.
+ econstructor; eauto.
+ auto. auto.
+ econstructor; split.
+ left. eapply plus_left. constructor.
+ eapply star_trans. apply step_makeif with (b := false) (v1 := v); auto. congruence.
+ apply push_seq.
+ reflexivity. traceEq.
+ rewrite <- Kseqlist_app.
+ econstructor. eauto. apply S.
+ econstructor; eauto. apply tr_expr_monotone with tmp3; eauto.
+ econstructor; eauto.
+ auto. auto.
+ (* for set *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist. destruct b.
econstructor; split.
@@ -1425,7 +1598,7 @@ Proof.
auto. auto.
(* assign *)
exploit tr_top_leftcontext; eauto. clear H12.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H4.
(* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
@@ -1433,10 +1606,8 @@ Proof.
subst; simpl Kseqlist.
econstructor; split.
left. eapply plus_left. constructor.
- apply star_one. econstructor; eauto.
- rewrite <- TY1; rewrite <- TY2; eauto.
- rewrite <- TY1. eapply assign_loc_preserved; eauto.
- traceEq.
+ apply star_one. eapply step_make_assign; eauto.
+ rewrite <- TY2; eauto. traceEq.
econstructor. auto. change sl2 with (nil ++ sl2). apply S.
constructor. auto. auto. auto.
(* for value *)
@@ -1450,10 +1621,8 @@ Proof.
left. eapply plus_left. constructor.
eapply star_left. constructor. eauto.
eapply star_left. constructor.
- apply star_one. econstructor; eauto. constructor. apply PTree.gss.
- simpl. rewrite <- TY1; eauto.
- rewrite <- TY1. eapply assign_loc_preserved; eauto.
- reflexivity. reflexivity. traceEq.
+ apply star_one. eapply step_make_assign; eauto.
+ constructor. apply PTree.gss. reflexivity. reflexivity. traceEq.
econstructor. auto. apply S.
apply tr_val_gen. auto. intros. econstructor; eauto. constructor.
rewrite H4; auto. apply PTree.gss.
@@ -1461,7 +1630,7 @@ Proof.
auto. auto.
(* assignop *)
exploit tr_top_leftcontext; eauto. clear H15.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H6.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
@@ -1473,11 +1642,9 @@ Proof.
subst; simpl Kseqlist.
econstructor; split.
left. eapply star_plus_trans. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
- eapply plus_two. simpl. econstructor. econstructor. eexact EV1'.
+ eapply plus_two. simpl. econstructor. eapply step_make_assign; eauto.
econstructor. eexact EV3. eexact EV2.
rewrite TY3; rewrite <- TY1; rewrite <- TY2; eauto.
- rewrite <- TY1; eauto.
- rewrite <- TY1. eapply assign_loc_preserved; eauto.
reflexivity. traceEq.
econstructor. auto. change sl2 with (nil ++ sl2). apply S.
constructor. auto. auto. auto.
@@ -1499,11 +1666,9 @@ Proof.
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. eapply step_make_assign; eauto.
+ constructor. apply PTree.gss.
+ reflexivity. traceEq.
econstructor. auto. apply S.
apply tr_val_gen. auto. intros. econstructor; eauto. constructor.
rewrite H10; auto. apply PTree.gss.
@@ -1513,7 +1678,7 @@ Proof.
auto. auto.
(* assignop stuck *)
exploit tr_top_leftcontext; eauto. clear H12.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H4.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
@@ -1535,7 +1700,7 @@ Proof.
constructor.
(* postincr *)
exploit tr_top_leftcontext; eauto. clear H14.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H5.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
@@ -1547,12 +1712,11 @@ Proof.
left. rewrite app_ass; rewrite Kseqlist_app.
eapply star_plus_trans. eexact EXEC.
eapply plus_two. simpl. constructor.
- econstructor; eauto.
+ eapply step_make_assign; eauto.
unfold transl_incrdecr. destruct id; simpl in H2.
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. eapply assign_loc_preserved; eauto.
+ destruct id; auto.
reflexivity. traceEq.
econstructor. auto. change sl2 with (nil ++ sl2). apply S.
constructor. auto. auto. auto.
@@ -1567,12 +1731,11 @@ Proof.
left. eapply plus_four. constructor.
eapply step_make_set; eauto.
constructor.
- econstructor. eauto.
+ eapply step_make_assign; 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. eapply assign_loc_preserved; eauto.
+ destruct id; auto.
traceEq.
econstructor. auto. apply S.
apply tr_val_gen. auto. intros. econstructor; eauto.
@@ -1581,7 +1744,7 @@ Proof.
auto. auto.
(* postincr stuck *)
exploit tr_top_leftcontext; eauto. clear H11.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H3.
(* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
@@ -1600,7 +1763,7 @@ Proof.
constructor.
(* comma *)
exploit tr_top_leftcontext; eauto. clear H9.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H1.
exploit tr_simple_rvalue; eauto. simpl; intro SL1.
subst sl0; simpl Kseqlist.
@@ -1612,11 +1775,11 @@ Proof.
auto. auto.
(* paren *)
exploit tr_top_leftcontext; eauto. clear H9.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [T [R S]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
(* for value *)
- exploit tr_simple_rvalue; eauto. intros [SL1 [TY1 EV1]].
- subst sl0; simpl Kseqlist.
+ exploit tr_simple_rvalue; eauto. intros [b [SL1 [TY1 EV1]]].
+ subst sl1; simpl Kseqlist.
econstructor; split.
left. eapply plus_left. constructor. apply star_one.
econstructor. econstructor; eauto. rewrite <- TY1; eauto. traceEq.
@@ -1630,29 +1793,23 @@ Proof.
right; split. apply star_refl. simpl. apply plus_lt_compat_r.
apply (leftcontext_size _ _ _ H). simpl. omega.
econstructor; eauto.
- exploit tr_simple_rvalue; eauto. destruct dst'; simpl.
- (* dst' = For_val: impossible *)
- congruence.
- (* dst' = For_effects: easy *)
- intros A. subst sl1. apply S. constructor; auto. auto. auto.
- (* dst' = For_test: then dest is For_test as well and C is a string of C.Eparen,
- so we can apply tr_top_paren. *)
- intros [b [A [B D]]].
- eapply tr_top_testcontext; eauto.
- subst sl1. apply tr_top_val_test; auto. econstructor; eauto. rewrite <- B; auto.
- (* already reduced *)
+ exploit tr_simple_rvalue; eauto. simpl. intros A. subst sl1.
+ apply S. constructor; auto. auto. auto.
+ (* for set *)
+ exploit tr_simple_rvalue; eauto. simpl. intros [b [SL1 [TY1 EV1]]]. subst sl1.
+ simpl Kseqlist.
econstructor; split.
- right; split. apply star_refl. simpl. apply plus_lt_compat_r.
- apply (leftcontext_size _ _ _ H). simpl. omega.
- econstructor; eauto. instantiate (1 := @nil ident).
- inv H9.
- inv H0. eapply tr_top_testcontext; eauto. simpl. constructor. auto. econstructor; eauto.
- exploit tr_simple_rvalue; eauto. simpl. intros [b [A [B D]]].
- eapply tr_top_testcontext; eauto. subst sl1. apply tr_top_val_test. auto. econstructor; eauto. congruence.
- inv H0.
+ left. eapply plus_left. constructor. apply star_one. econstructor. econstructor; eauto.
+ rewrite <- TY1; eauto. traceEq.
+ econstructor; eauto.
+ apply S. constructor; auto.
+ intros. constructor. rewrite H2. apply PTree.gss. auto.
+ intros. apply PTree.gso. congruence.
+ auto.
+
(* call *)
exploit tr_top_leftcontext; eauto. clear H12.
- intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [U [R S]]]]]]]]].
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H5.
(* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL1 [TY1 EV1]].
@@ -1684,6 +1841,36 @@ Proof.
auto. intros. constructor. rewrite H5; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
auto.
+
+(* builtin *)
+ exploit tr_top_leftcontext; eauto. clear H9.
+ intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
+ inv P. inv H2.
+ (* for effects *)
+ exploit tr_simple_exprlist; eauto. intros [SL EV].
+ subst. simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor. apply star_one.
+ econstructor; eauto.
+ eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+ traceEq.
+ econstructor; eauto.
+ change sl2 with (nil ++ sl2). apply S. constructor. simpl; auto. auto.
+ (* for value *)
+ exploit tr_simple_exprlist; eauto. intros [SL EV].
+ subst. simpl Kseqlist.
+ econstructor; split.
+ left. eapply plus_left. constructor. apply star_one.
+ econstructor; eauto.
+ eapply external_call_symbols_preserved; eauto.
+ exact symbols_preserved. exact varinfo_preserved.
+ traceEq.
+ econstructor; eauto.
+ change sl2 with (nil ++ sl2). apply S.
+ apply tr_val_gen. auto. intros. constructor. rewrite H2; auto. simpl. apply PTree.gss.
+ intros; simpl. apply PTree.gso. intuition congruence.
+ auto.
Qed.
(** Forward simulation for statements. *)
@@ -1696,21 +1883,16 @@ Proof.
intros. inv H. auto. inv H0. auto.
Qed.
-Lemma tr_top_val_for_test_inv:
- forall s1 s2 e le m v ty sl a tmps,
- tr_top tge e le m (For_test nil s1 s2) (C.Eval v ty) sl a tmps ->
- exists b, sl = makeif b s1 s2 :: nil /\ typeof b = ty /\ eval_expr tge e le m b v.
-Proof.
- intros. inv H. exists a0; auto.
- 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'.
+ Csem.bind_parameters ge e m params args m' ->
+ bind_parameters e m params args m'.
Proof.
- induction 1; econstructor; eauto. eapply assign_loc_preserved; eauto.
+ induction 1; econstructor; eauto.
+ inv H0.
+ eapply assign_loc_value; eauto.
+ inv H4. eapply assign_loc_value; eauto.
+ eapply assign_loc_copy; eauto.
Qed.
Lemma sstep_simulation:
@@ -1752,49 +1934,37 @@ Proof.
(* ifthenelse *)
inv H6.
- (* not optimized *)
inv H2. econstructor; split.
left. eapply plus_left. constructor. apply push_seq. traceEq.
econstructor; eauto. econstructor; eauto.
- (* optimized *)
- inv H10. econstructor; split.
- right; split. apply push_seq. simpl. omega.
- econstructor; eauto. constructor; auto.
+
(* ifthenelse *)
inv H8.
- (* not optimized *)
exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. eapply plus_two. constructor.
apply step_ifthenelse with (v1 := v) (b := b); auto. traceEq.
destruct b; econstructor; eauto.
- (* optimized *)
- exploit tr_top_val_for_test_inv; eauto. intros [c [A [B C]]]. subst.
- econstructor; split.
- left. simpl. eapply plus_left. constructor.
- apply step_makeif with (v1 := v) (b := b); auto. traceEq.
- econstructor; eauto.
- destruct b; auto.
(* while *)
inv H6. inv H1. econstructor; split.
- left. eapply plus_left. eapply step_while_true. constructor.
- auto. eapply star_left. constructor.
+ left. eapply plus_left. constructor.
+ eapply star_left. constructor.
apply push_seq.
- reflexivity. traceEq.
- econstructor; eauto. econstructor; eauto. econstructor; eauto.
+ reflexivity. traceEq. rewrite Kseqlist_app.
+ econstructor; eauto. simpl. econstructor; eauto. econstructor; eauto.
(* while false *)
inv H8.
- exploit tr_top_val_for_test_inv; eauto. intros [b [A [B C]]]. subst.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (v1 := v) (b := false); auto.
- eapply star_two. constructor. apply step_break_while.
+ eapply star_two. constructor. apply step_break_loop1.
reflexivity. reflexivity. traceEq.
constructor; auto. constructor.
(* while true *)
inv H8.
- exploit tr_top_val_for_test_inv; eauto. intros [b [A [B C]]]. subst.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
@@ -1805,31 +1975,32 @@ Proof.
assert (ts = Sskip \/ ts = Scontinue). destruct H; subst s0; inv H7; auto.
inv H8.
econstructor; split.
- left. apply plus_one. apply step_skip_or_continue_while; auto.
+ left. eapply plus_two. apply step_skip_or_continue_loop1; auto.
+ apply step_skip_loop2. traceEq.
constructor; auto. constructor; auto.
(* break while *)
inv H6. inv H7.
econstructor; split.
- left. apply plus_one. apply step_break_while.
+ left. apply plus_one. apply step_break_loop1.
constructor; auto. constructor.
(* dowhile *)
inv H6.
econstructor; split.
- left. apply plus_one.
- apply step_for_true with (Vint Int.one). constructor. auto.
+ left. apply plus_one. apply step_loop.
constructor; auto. constructor; auto.
(* skip_or_continue dowhile *)
assert (ts = Sskip \/ ts = Scontinue). destruct H; subst s0; inv H7; auto.
inv H8. inv H4.
econstructor; split.
- left. eapply plus_left. apply step_skip_or_continue_for2. auto.
+ left. eapply plus_left. apply step_skip_or_continue_loop1. auto.
apply push_seq.
- reflexivity. traceEq.
- econstructor; eauto. econstructor; auto. econstructor; eauto.
+ traceEq.
+ rewrite Kseqlist_app.
+ econstructor; eauto. simpl. econstructor; auto. econstructor; eauto.
(* dowhile false *)
inv H8.
- exploit tr_top_val_for_test_inv; eauto. intros [b [A [B C]]]. subst.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := false); auto.
@@ -1838,7 +2009,7 @@ Proof.
constructor; auto. constructor.
(* dowhile true *)
inv H8.
- exploit tr_top_val_for_test_inv; eauto. intros [b [A [B C]]]. subst.
+ exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
@@ -1848,7 +2019,7 @@ Proof.
(* break dowhile *)
inv H6. inv H7.
econstructor; split.
- left. apply plus_one. apply step_break_for2.
+ left. apply plus_one. apply step_break_loop1.
constructor; auto. constructor.
(* for start *)
@@ -1859,21 +2030,20 @@ Proof.
(* for *)
inv H6; try congruence. inv H2.
econstructor; split.
- left. eapply plus_left. apply step_for_true with (Vint Int.one).
- constructor. auto.
+ left. eapply plus_left. apply step_loop.
eapply star_left. constructor. apply push_seq.
reflexivity. traceEq.
- econstructor; eauto. constructor; auto. econstructor; eauto.
+ rewrite Kseqlist_app. econstructor; eauto. simpl. constructor; auto. econstructor; eauto.
(* for false *)
- inv H8. exploit tr_top_val_for_test_inv; eauto. intros [b [A [B C]]]. subst.
+ inv H8. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (v1 := v) (b := false); auto.
- eapply star_two. constructor. apply step_break_for2.
+ eapply star_two. constructor. apply step_break_loop1.
reflexivity. reflexivity. traceEq.
constructor; auto. constructor.
(* for true *)
- inv H8. exploit tr_top_val_for_test_inv; eauto. intros [b [A [B C]]]. subst.
+ inv H8. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
@@ -1884,12 +2054,12 @@ Proof.
assert (ts = Sskip \/ ts = Scontinue). destruct H; subst x; inv H7; auto.
inv H8.
econstructor; split.
- left. apply plus_one. apply step_skip_or_continue_for2. auto.
+ left. apply plus_one. apply step_skip_or_continue_loop1. auto.
econstructor; eauto. econstructor; auto.
(* break for3 *)
inv H6. inv H7.
econstructor; split.
- left. apply plus_one. apply step_break_for2.
+ left. apply plus_one. apply step_break_loop1.
econstructor; eauto. constructor.
(* skip for4 *)
inv H6. inv H7.
@@ -1982,11 +2152,6 @@ Proof.
(* return *)
inv H3.
- (* none *)
- econstructor; split.
- left; apply plus_one. constructor.
- econstructor; eauto.
- (* some *)
econstructor; split.
left; apply plus_one. constructor.
econstructor; eauto.
diff --git a/cfrontend/SimplExprspec.v b/cfrontend/SimplExprspec.v
index bd7f336..3f9d7e9 100644
--- a/cfrontend/SimplExprspec.v
+++ b/cfrontend/SimplExprspec.v
@@ -46,7 +46,7 @@ Definition final (dst: destination) (a: expr) : list statement :=
match dst with
| For_val => nil
| For_effects => nil
- | For_test tyl s1 s2 => makeif (fold_left Ecast tyl a) s1 s2 :: nil
+ | For_set tyl t => do_set t tyl a
end.
Inductive tr_rvalof: type -> expr -> list statement -> expr -> list ident -> Prop :=
@@ -55,7 +55,7 @@ Inductive tr_rvalof: type -> expr -> list statement -> expr -> list ident -> Pro
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.
+ tr_rvalof ty a (make_set 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,
@@ -78,13 +78,13 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
eval_expr tge e le' m a v) ->
tr_expr le For_val (C.Eval v ty)
nil a tmp
- | tr_val_test: forall le tyl s1 s2 v ty a any tmp,
+ | tr_val_set: forall le tyl t v ty a any tmp,
typeof a = ty ->
(forall tge e le' m,
(forall id, In id tmp -> le'!id = le!id) ->
eval_expr tge e le' m a v) ->
- tr_expr le (For_test tyl s1 s2) (C.Eval v ty)
- (makeif (fold_left Ecast tyl a) s1 s2 :: nil) any tmp
+ tr_expr le (For_set tyl t) (C.Eval v ty)
+ (do_set t tyl a) any tmp
| tr_sizeof: forall le dst ty' ty tmp,
tr_expr le dst (C.Esizeof ty' ty)
(final dst (Esizeof ty' ty))
@@ -122,41 +122,86 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
tr_expr le dst (C.Ecast e1 ty)
(sl1 ++ final dst (Ecast a1 ty))
(Ecast a1 ty) tmp
- | tr_condition_simple: forall le dst e1 e2 e3 ty sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 tmp,
- Cstrategy.simple e2 = true -> Cstrategy.simple e3 = true ->
+ | tr_seqand_val: forall le e1 e2 ty sl1 a1 tmp1 t sl2 a2 tmp2 tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
- tr_expr le For_val e2 sl2 a2 tmp2 ->
- tr_expr le For_val e3 sl3 a3 tmp3 ->
+ tr_expr le (For_set (type_bool :: ty :: nil) t) e2 sl2 a2 tmp2 ->
list_disjoint tmp1 tmp2 ->
- list_disjoint tmp1 tmp3 ->
- incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp ->
- tr_expr le dst (C.Econdition e1 e2 e3 ty)
- (sl1 ++ final dst (Econdition a1 a2 a3 ty))
- (Econdition a1 a2 a3 ty) tmp
+ incl tmp1 tmp -> incl tmp2 tmp -> In t tmp ->
+ tr_expr le For_val (C.Eseqand e1 e2 ty)
+ (sl1 ++ makeif a1 (makeseq sl2)
+ (Sset t (Econst_int Int.zero ty)) :: nil)
+ (Etempvar t ty) tmp
+ | tr_seqand_effects: forall le e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
+ tr_expr le For_val e1 sl1 a1 tmp1 ->
+ tr_expr le For_effects e2 sl2 a2 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
+ incl tmp1 tmp -> incl tmp2 tmp ->
+ tr_expr le For_effects (C.Eseqand e1 e2 ty)
+ (sl1 ++ makeif a1 (makeseq sl2) Sskip :: nil)
+ any tmp
+ | tr_seqand_set: forall le tyl t e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
+ tr_expr le For_val e1 sl1 a1 tmp1 ->
+ tr_expr le (For_set (type_bool :: ty :: tyl) t) e2 sl2 a2 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
+ incl tmp1 tmp -> incl tmp2 tmp -> In t tmp ->
+ tr_expr le (For_set tyl t) (C.Eseqand e1 e2 ty)
+ (sl1 ++ makeif a1 (makeseq sl2)
+ (makeseq (do_set t tyl (Econst_int Int.zero ty))) :: nil)
+ any tmp
+ | tr_seqor_val: forall le e1 e2 ty sl1 a1 tmp1 t sl2 a2 tmp2 tmp,
+ tr_expr le For_val e1 sl1 a1 tmp1 ->
+ tr_expr le (For_set (type_bool :: ty :: nil) t) e2 sl2 a2 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
+ incl tmp1 tmp -> incl tmp2 tmp -> In t tmp ->
+ tr_expr le For_val (C.Eseqor e1 e2 ty)
+ (sl1 ++ makeif a1 (Sset t (Econst_int Int.one ty))
+ (makeseq sl2) :: nil)
+ (Etempvar t ty) tmp
+ | tr_seqor_effects: forall le e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
+ tr_expr le For_val e1 sl1 a1 tmp1 ->
+ tr_expr le For_effects e2 sl2 a2 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
+ incl tmp1 tmp -> incl tmp2 tmp ->
+ tr_expr le For_effects (C.Eseqor e1 e2 ty)
+ (sl1 ++ makeif a1 Sskip (makeseq sl2) :: nil)
+ any tmp
+ | tr_seqor_set: forall le tyl t e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
+ tr_expr le For_val e1 sl1 a1 tmp1 ->
+ tr_expr le (For_set (type_bool :: ty :: tyl) t) e2 sl2 a2 tmp2 ->
+ list_disjoint tmp1 tmp2 ->
+ incl tmp1 tmp -> incl tmp2 tmp -> In t tmp ->
+ tr_expr le (For_set tyl t) (C.Eseqor e1 e2 ty)
+ (sl1 ++ makeif a1 (makeseq (do_set t tyl (Econst_int Int.one ty)))
+ (makeseq sl2) :: nil)
+ any tmp
| tr_condition_val: forall le e1 e2 e3 ty sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 t tmp,
- Cstrategy.simple e2 = false \/ Cstrategy.simple e3 = false ->
tr_expr le For_val e1 sl1 a1 tmp1 ->
- tr_expr le For_val e2 sl2 a2 tmp2 ->
- tr_expr le For_val e3 sl3 a3 tmp3 ->
+ tr_expr le (For_set (ty::nil) t) e2 sl2 a2 tmp2 ->
+ tr_expr le (For_set (ty::nil) t) e3 sl3 a3 tmp3 ->
list_disjoint tmp1 tmp2 ->
list_disjoint tmp1 tmp3 ->
- incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp ->
- In t tmp -> ~In t tmp1 ->
+ incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp -> In t tmp ->
tr_expr le For_val (C.Econdition e1 e2 e3 ty)
- (sl1 ++ makeif a1
- (Ssequence (makeseq sl2) (Sset t (Ecast a2 ty)))
- (Ssequence (makeseq sl3) (Sset t (Ecast a3 ty))) :: nil)
+ (sl1 ++ makeif a1 (makeseq sl2) (makeseq sl3) :: nil)
(Etempvar t ty) tmp
- | tr_condition_effects: forall le dst e1 e2 e3 ty sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 any tmp,
- Cstrategy.simple e2 = false \/ Cstrategy.simple e3 = false ->
- dst <> For_val ->
+ | tr_condition_effects: forall le e1 e2 e3 ty sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 any tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
- tr_expr le (cast_destination ty dst) e2 sl2 a2 tmp2 ->
- tr_expr le (cast_destination ty dst) e3 sl3 a3 tmp3 ->
+ tr_expr le For_effects e2 sl2 a2 tmp2 ->
+ tr_expr le For_effects e3 sl3 a3 tmp3 ->
list_disjoint tmp1 tmp2 ->
list_disjoint tmp1 tmp3 ->
incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp ->
- tr_expr le dst (C.Econdition e1 e2 e3 ty)
+ tr_expr le For_effects (C.Econdition e1 e2 e3 ty)
+ (sl1 ++ makeif a1 (makeseq sl2) (makeseq sl3) :: nil)
+ any tmp
+ | tr_condition_set: forall le tyl t e1 e2 e3 ty sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 any tmp,
+ tr_expr le For_val e1 sl1 a1 tmp1 ->
+ tr_expr le (For_set (ty::tyl) t) e2 sl2 a2 tmp2 ->
+ tr_expr le (For_set (ty::tyl) t) e3 sl3 a3 tmp3 ->
+ list_disjoint tmp1 tmp2 ->
+ list_disjoint tmp1 tmp3 ->
+ incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp -> In t tmp ->
+ tr_expr le (For_set tyl t) (C.Econdition e1 e2 e3 ty)
(sl1 ++ makeif a1 (makeseq sl2) (makeseq sl3) :: nil)
any tmp
| tr_assign_effects: forall le e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
@@ -165,7 +210,7 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
list_disjoint tmp1 tmp2 ->
incl tmp1 tmp -> incl tmp2 tmp ->
tr_expr le For_effects (C.Eassign e1 e2 ty)
- (sl1 ++ sl2 ++ Sassign a1 a2 :: nil)
+ (sl1 ++ sl2 ++ make_assign a1 a2 :: nil)
any tmp
| tr_assign_val: forall le dst e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 t tmp ty1 ty2,
tr_expr le For_val e1 sl1 a1 tmp1 ->
@@ -178,7 +223,7 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
tr_expr le dst (C.Eassign e1 e2 ty)
(sl1 ++ sl2 ++
Sset t a2 ::
- Sassign a1 (Etempvar t ty2) ::
+ make_assign 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 ty1 sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 any tmp,
@@ -189,7 +234,7 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
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 ++ sl3 ++ Sassign a1 (Ebinop op a3 a2 tyres) :: nil)
+ (sl1 ++ sl2 ++ sl3 ++ make_assign 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 sl3 a3 tmp3 t tmp ty1,
tr_expr le For_val e1 sl1 a1 tmp1 ->
@@ -202,7 +247,7 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
tr_expr le dst (C.Eassignop op e1 e2 tyres ty)
(sl1 ++ sl2 ++ sl3 ++
Sset t (Ebinop op a3 a2 tyres) ::
- Sassign a1 (Etempvar t tyres) ::
+ make_assign 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 ty1 sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
@@ -212,7 +257,7 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
incl tmp1 tmp -> incl tmp2 tmp ->
list_disjoint tmp1 tmp2 ->
tr_expr le For_effects (C.Epostincr id e1 ty)
- (sl1 ++ sl2 ++ Sassign a1 (transl_incrdecr id a2 ty1) :: nil)
+ (sl1 ++ sl2 ++ make_assign 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 ->
@@ -220,7 +265,7 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
ty1 = C.typeof e1 ->
tr_expr le dst (C.Epostincr id e1 ty)
(sl1 ++ make_set t a1 ::
- Sassign a1 (transl_incrdecr id (Etempvar t ty1) ty1) ::
+ make_assign a1 (transl_incrdecr id (Etempvar t ty1) ty1) ::
final dst (Etempvar t ty1))
(Etempvar t ty1) tmp
| tr_comma: forall le dst e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 tmp,
@@ -246,16 +291,32 @@ Inductive tr_expr: temp_env -> destination -> C.expr -> list statement -> expr -
tr_expr le dst (C.Ecall e1 el2 ty)
(sl1 ++ sl2 ++ Scall (Some t) a1 al2 :: final dst (Etempvar t ty))
(Etempvar t ty) tmp
- | tr_paren_val: forall le e1 ty sl1 a1 tmp1 t tmp,
- tr_expr le For_val e1 sl1 a1 tmp1 ->
- incl tmp1 tmp -> In t tmp ->
+ | tr_builtin_effects: forall le ef tyargs el ty sl al tmp1 any tmp,
+ tr_exprlist le el sl al tmp1 ->
+ incl tmp1 tmp ->
+ tr_expr le For_effects (C.Ebuiltin ef tyargs el ty)
+ (sl ++ Sbuiltin None ef tyargs al :: nil)
+ any tmp
+ | tr_builtin_val: forall le dst ef tyargs el ty sl al tmp1 t tmp,
+ dst <> For_effects ->
+ tr_exprlist le el sl al tmp1 ->
+ In t tmp -> incl tmp1 tmp ->
+ tr_expr le dst (C.Ebuiltin ef tyargs el ty)
+ (sl ++ Sbuiltin (Some t) ef tyargs al :: final dst (Etempvar t ty))
+ (Etempvar t ty) tmp
+ | tr_paren_val: forall le e1 ty sl1 a1 t tmp,
+ tr_expr le (For_set (ty :: nil) t) e1 sl1 a1 tmp ->
+ In t tmp ->
tr_expr le For_val (C.Eparen e1 ty)
- (sl1 ++ Sset t (Ecast a1 ty) :: nil)
+ sl1
(Etempvar t ty) tmp
- | tr_paren_effects: forall le dst e1 ty sl1 a1 tmp any,
- dst <> For_val ->
- tr_expr le (cast_destination ty dst) e1 sl1 a1 tmp ->
- tr_expr le dst (C.Eparen e1 ty) sl1 any tmp
+ | tr_paren_effects: forall le e1 ty sl1 a1 tmp any,
+ tr_expr le For_effects e1 sl1 a1 tmp ->
+ tr_expr le For_effects (C.Eparen e1 ty) sl1 any tmp
+ | tr_paren_set: forall le tyl t e1 ty sl1 a1 tmp any,
+ tr_expr le (For_set (ty::tyl) t) e1 sl1 a1 tmp ->
+ In t tmp ->
+ tr_expr le (For_set tyl t) (C.Eparen e1 ty) sl1 any tmp
with tr_exprlist: temp_env -> C.exprlist -> list statement -> list expr -> list ident -> Prop :=
| tr_nil: forall le tmp,
@@ -328,15 +389,19 @@ Inductive tr_top: destination -> C.expr -> list statement -> expr -> list ident
| tr_top_val_val: forall v ty a tmp,
typeof a = ty -> eval_expr ge e le m a v ->
tr_top For_val (C.Eval v ty) nil a tmp
- | tr_top_val_test: forall tyl s1 s2 v ty a any tmp,
+(*
+ | tr_top_val_set: forall t tyl v ty a any tmp,
typeof a = ty -> eval_expr ge e le m a v ->
- tr_top (For_test tyl s1 s2) (C.Eval v ty) (makeif (fold_left Ecast tyl a) s1 s2 :: nil) any tmp
+ tr_top (For_set tyl t) (C.Eval v ty) (Sset t (fold_left Ecast tyl a) :: nil) any tmp
+*)
| tr_top_base: forall dst r sl a tmp,
tr_expr le dst r sl a tmp ->
- tr_top dst r sl a tmp
- | tr_top_paren_test: forall tyl s1 s2 r ty sl a tmp,
- tr_top (For_test (ty :: tyl) s1 s2) r sl a tmp ->
- tr_top (For_test tyl s1 s2) (C.Eparen r ty) sl a tmp.
+ tr_top dst r sl a tmp.
+(*
+ | tr_top_paren_test: forall tyl t r ty sl a tmp,
+ tr_top (For_set (ty :: tyl) t) r sl a tmp ->
+ tr_top (For_set tyl t) (C.Eparen r ty) sl a tmp.
+*)
End TR_TOP.
@@ -354,8 +419,8 @@ Inductive tr_expr_stmt: C.expr -> statement -> Prop :=
Inductive tr_if: C.expr -> statement -> statement -> statement -> Prop :=
| tr_if_intro: forall r s1 s2 sl a tmps,
- (forall ge e le m, tr_top ge e le m (For_test nil s1 s2) r sl a tmps) ->
- tr_if r s1 s2 (makeseq sl).
+ (forall ge e le m, tr_top ge e le m For_val r sl a tmps) ->
+ tr_if r s1 s2 (makeseq (sl ++ makeif a s1 s2 :: nil)).
Inductive tr_stmt: C.statement -> statement -> Prop :=
| tr_skip:
@@ -366,31 +431,26 @@ Inductive tr_stmt: C.statement -> statement -> Prop :=
| tr_seq: forall s1 s2 ts1 ts2,
tr_stmt s1 ts1 -> tr_stmt s2 ts2 ->
tr_stmt (C.Ssequence s1 s2) (Ssequence ts1 ts2)
- | tr_ifthenelse_big: forall r s1 s2 s' a ts1 ts2,
+ | tr_ifthenelse: forall r s1 s2 s' a ts1 ts2,
tr_expression r s' a ->
tr_stmt s1 ts1 -> tr_stmt s2 ts2 ->
tr_stmt (C.Sifthenelse r s1 s2) (Ssequence s' (Sifthenelse a ts1 ts2))
- | tr_ifthenelse_small: forall r s1 s2 ts1 ts2 ts,
- tr_stmt s1 ts1 -> tr_stmt s2 ts2 ->
- small_stmt ts1 = true -> small_stmt ts2 = true ->
- tr_if r ts1 ts2 ts ->
- tr_stmt (C.Sifthenelse r s1 s2) ts
| tr_while: forall r s1 s' ts1,
tr_if r Sskip Sbreak s' ->
tr_stmt s1 ts1 ->
tr_stmt (C.Swhile r s1)
- (Swhile expr_true (Ssequence s' ts1))
+ (Sloop (Ssequence s' ts1) Sskip)
| tr_dowhile: forall r s1 s' ts1,
tr_if r Sskip Sbreak s' ->
tr_stmt s1 ts1 ->
tr_stmt (C.Sdowhile r s1)
- (Sfor' expr_true s' ts1)
+ (Sloop ts1 s')
| tr_for_1: forall r s3 s4 s' ts3 ts4,
tr_if r Sskip Sbreak s' ->
tr_stmt s3 ts3 ->
tr_stmt s4 ts4 ->
tr_stmt (C.Sfor C.Sskip r s3 s4)
- (Sfor' expr_true ts3 (Ssequence s' ts4))
+ (Sloop (Ssequence s' ts4) ts3)
| tr_for_2: forall s1 r s3 s4 s' ts1 ts3 ts4,
tr_if r Sskip Sbreak s' ->
s1 <> C.Sskip ->
@@ -398,7 +458,7 @@ Inductive tr_stmt: C.statement -> statement -> Prop :=
tr_stmt s3 ts3 ->
tr_stmt s4 ts4 ->
tr_stmt (C.Sfor s1 r s3 s4)
- (Ssequence ts1 (Sfor' expr_true ts3 (Ssequence s' ts4)))
+ (Ssequence ts1 (Sloop (Ssequence s' ts4) ts3))
| tr_break:
tr_stmt C.Sbreak Sbreak
| tr_continue:
@@ -584,13 +644,49 @@ Proof.
elim (Plt_strict id). apply Plt_Ple_trans with (gen_next g2); auto.
Qed.
+Definition dest_below (dst: destination) (g: generator) : Prop :=
+ match dst with
+ | For_set tyl tmp => Plt tmp g.(gen_next)
+ | _ => True
+ end.
+
+Remark dest_for_val_below: forall g, dest_below For_val g.
+Proof. intros; simpl; auto. Qed.
+
+Remark dest_for_effect_below: forall g, dest_below For_effects g.
+Proof. intros; simpl; auto. Qed.
+
+Lemma dest_below_notin:
+ forall tyl tmp g1 g2 l,
+ dest_below (For_set tyl tmp) g1 -> contained l g1 g2 -> ~In tmp l.
+Proof.
+ simpl; intros. red in H0. red; intros. destruct (H0 _ H1).
+ elim (Plt_strict tmp). apply Plt_Ple_trans with (gen_next g1); auto.
+Qed.
+
+Lemma within_dest_below:
+ forall tyl tmp g1 g2,
+ within tmp g1 g2 -> dest_below (For_set tyl tmp) g2.
+Proof.
+ intros; simpl. destruct H. tauto.
+Qed.
+
+Lemma dest_below_le:
+ forall dst g1 g2,
+ dest_below dst g1 -> Ple g1.(gen_next) g2.(gen_next) -> dest_below dst g2.
+Proof.
+ intros. destruct dst; simpl in *; auto. eapply Plt_Ple_trans; eauto.
+Qed.
+
Hint Resolve gensym_within within_widen contained_widen
contained_cons contained_app contained_disjoint
- contained_notin contained_nil
- incl_refl incl_tl incl_app incl_appl incl_appr
- in_eq in_cons
+ contained_notin contained_nil dest_below_notin
+ incl_refl incl_tl incl_app incl_appl incl_appr incl_same_head
+ in_eq in_cons within_dest_below dest_below_le
Ple_trans Ple_refl: gensym.
+Hint Resolve dest_for_val_below dest_for_effect_below.
+
(** ** Correctness of the translation functions *)
Lemma finish_meets_spec_1:
@@ -615,6 +711,26 @@ Ltac UseFinish :=
repeat rewrite app_ass
end.
+Definition add_dest (dst: destination) (tmps: list ident) :=
+ match dst with
+ | For_set tyl t => t :: tmps
+ | _ => tmps
+ end.
+
+Lemma add_dest_incl:
+ forall dst tmps, incl tmps (add_dest dst tmps).
+Proof.
+ intros. destruct dst; simpl; auto with coqlib.
+Qed.
+
+Lemma tr_expr_add_dest:
+ forall le dst r sl a tmps,
+ tr_expr le dst r sl a tmps ->
+ tr_expr le dst r sl a (add_dest dst tmps).
+Proof.
+ intros. apply tr_expr_monotone with tmps; auto. apply add_dest_incl.
+Qed.
+
Lemma transl_valof_meets_spec:
forall ty a g sl b g' I,
transl_valof ty a g = Res (sl, b) g' I ->
@@ -624,14 +740,6 @@ Proof.
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
@@ -641,84 +749,131 @@ Combined Scheme expr_exprlist_ind from expr_ind2, exprlist_ind2.
Lemma transl_meets_spec:
(forall r dst g sl a g' I,
transl_expr dst r g = Res (sl, a) g' I ->
- exists tmps, (forall le, tr_expr le dst r sl a tmps) /\ contained tmps g g')
+ dest_below dst g ->
+ exists tmps, (forall le, tr_expr le dst r sl a (add_dest dst tmps)) /\ contained tmps g g')
/\
(forall rl g sl al g' I,
transl_exprlist rl g = Res (sl, al) g' I ->
exists tmps, (forall le, tr_exprlist le rl sl al tmps) /\ contained tmps g g').
Proof.
- apply expr_exprlist_ind; intros.
+ apply expr_exprlist_ind; simpl add_dest; intros.
(* val *)
simpl in H. destruct v; monadInv H; exists (@nil ident); split; auto with gensym.
Opaque makeif.
- intros. destruct dst; simpl in H1; inv H1.
+ intros. destruct dst; simpl in *; inv H2.
constructor. auto. intros; constructor.
constructor.
constructor. auto. intros; constructor.
- intros. destruct dst; simpl in H1; inv H1.
+ intros. destruct dst; simpl in *; inv H2.
constructor. auto. intros; constructor.
constructor.
constructor. auto. intros; constructor.
(* var *)
monadInv H; econstructor; split; auto with gensym. UseFinish. constructor.
(* field *)
- monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
- econstructor; split; eauto. constructor; auto.
+ monadInv H0. exploit H; eauto. auto. intros [tmp [A B]]. UseFinish.
+ econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
(* valof *)
monadInv H0. exploit H; eauto. intros [tmp1 [A B]].
exploit transl_valof_meets_spec; eauto. intros [tmp2 [C D]]. UseFinish.
exists (tmp1 ++ tmp2); split.
- econstructor; eauto with gensym.
+ intros; apply tr_expr_add_dest. econstructor; eauto with gensym.
eauto with gensym.
(* deref *)
monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
- econstructor; split; eauto. constructor; auto.
+ econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
(* addrof *)
monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
- econstructor; split; eauto. econstructor; eauto.
+ econstructor; split; eauto. intros; apply tr_expr_add_dest. econstructor; eauto.
(* unop *)
monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
- econstructor; split; eauto. constructor; auto.
+ econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
(* binop *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [C D]]. UseFinish.
exists (tmp1 ++ tmp2); split.
- econstructor; eauto with gensym.
+ intros; apply tr_expr_add_dest. econstructor; eauto with gensym.
eauto with gensym.
(* cast *)
monadInv H0. exploit H; eauto. intros [tmp [A B]]. UseFinish.
- econstructor; split; eauto. constructor; auto.
-(* condition *)
- simpl in H2.
- destruct (Cstrategy.simple r2 && Cstrategy.simple r3) as []_eqn.
- (* simple *)
- monadInv H2. exploit H; eauto. intros [tmp1 [A B]].
- exploit H0; eauto. intros [tmp2 [C D]].
- exploit H1; eauto. intros [tmp3 [E F]].
- UseFinish. destruct (andb_prop _ _ Heqb).
- exists (tmp1 ++ tmp2 ++ tmp3); split.
- intros; eapply tr_condition_simple; eauto with gensym.
- apply contained_app. eauto with gensym.
+ econstructor; split; eauto. intros; apply tr_expr_add_dest. constructor; auto.
+(* seqand *)
+ monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
+ destruct dst; monadInv EQ0.
+ (* for value *)
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
+ simpl add_dest in *.
+ exists (x0 :: tmp1 ++ tmp2); split.
+ intros; eapply tr_seqand_val; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
+ apply contained_cons. eauto with gensym.
+ apply contained_app; eauto with gensym.
+ (* for effects *)
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
+ simpl add_dest in *.
+ exists (tmp1 ++ tmp2); split.
+ intros; eapply tr_seqand_effects; eauto with gensym.
+ apply contained_app; eauto with gensym.
+ (* for set *)
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
+ simpl add_dest in *.
+ exists (tmp1 ++ tmp2); split.
+ intros; eapply tr_seqand_set; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
+ apply contained_app; eauto with gensym.
+(* seqor *)
+ monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
+ destruct dst; monadInv EQ0.
+ (* for value *)
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
+ simpl add_dest in *.
+ exists (x0 :: tmp1 ++ tmp2); split.
+ intros; eapply tr_seqor_val; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
+ apply contained_cons. eauto with gensym.
+ apply contained_app; eauto with gensym.
+ (* for effects *)
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
+ simpl add_dest in *.
+ exists (tmp1 ++ tmp2); split.
+ intros; eapply tr_seqor_effects; eauto with gensym.
+ apply contained_app; eauto with gensym.
+ (* for set *)
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
+ simpl add_dest in *.
+ exists (tmp1 ++ tmp2); split.
+ intros; eapply tr_seqor_set; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
apply contained_app; eauto with gensym.
- (* not simple *)
+(* condition *)
monadInv H2. exploit H; eauto. intros [tmp1 [A B]].
- exploit H0; eauto. intros [tmp2 [C D]].
- exploit H1; eauto. intros [tmp3 [E F]].
- rewrite andb_false_iff in Heqb.
- destruct dst; monadInv EQ3.
+ destruct dst; monadInv EQ0.
(* for value *)
- exists (x2 :: tmp1 ++ tmp2 ++ tmp3); split.
- intros; eapply tr_condition_val; eauto with gensym.
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
+ exploit H1; eauto with gensym. intros [tmp3 [E F]].
+ simpl add_dest in *.
+ exists (x0 :: tmp1 ++ tmp2 ++ tmp3); split.
+ simpl; intros; eapply tr_condition_val; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
apply contained_cons. eauto with gensym.
apply contained_app. eauto with gensym.
apply contained_app; eauto with gensym.
(* for effects *)
+ exploit H0; eauto. intros [tmp2 [C D]].
+ exploit H1; eauto. intros [tmp3 [E F]].
+ simpl add_dest in *.
exists (tmp1 ++ tmp2 ++ tmp3); split.
- intros; eapply tr_condition_effects; eauto with gensym. congruence.
+ intros; eapply tr_condition_effects; eauto with gensym.
apply contained_app; eauto with gensym.
(* for test *)
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
+ exploit H1; eauto with gensym. intros [tmp3 [E F]].
+ simpl add_dest in *.
exists (tmp1 ++ tmp2 ++ tmp3); split.
- intros; eapply tr_condition_effects; eauto with gensym. congruence.
+ intros; eapply tr_condition_set; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
apply contained_app; eauto with gensym.
(* sizeof *)
monadInv H. UseFinish.
@@ -729,7 +884,7 @@ Opaque makeif.
(* assign *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [C D]].
- destruct dst; monadInv EQ2.
+ destruct dst; monadInv EQ2; simpl add_dest in *.
(* for value *)
exists (x1 :: tmp1 ++ tmp2); split.
intros. eapply tr_assign_val with (dst := For_val); eauto with gensym.
@@ -739,17 +894,17 @@ Opaque makeif.
exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
- (* for test *)
+ (* for set *)
exists (x1 :: tmp1 ++ tmp2); split.
repeat rewrite app_ass. simpl.
- intros. eapply tr_assign_val with (dst := For_test tyl s1 s2); eauto with gensym.
+ intros. eapply tr_assign_val with (dst := For_set tyl tmp); eauto with gensym.
apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* assignop *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [C D]].
exploit transl_valof_meets_spec; eauto. intros [tmp3 [E F]].
- destruct dst; monadInv EQ3.
+ destruct dst; monadInv EQ3; simpl add_dest in *.
(* for value *)
exists (x2 :: tmp1 ++ tmp2 ++ tmp3); split.
intros. eapply tr_assignop_val with (dst := For_val); eauto with gensym.
@@ -759,15 +914,15 @@ Opaque makeif.
exists (tmp1 ++ tmp2 ++ tmp3); split.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
- (* for test *)
+ (* for set *)
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.
+ intros. eapply tr_assignop_val with (dst := For_set tyl tmp); eauto with gensym.
apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
(* postincr *)
monadInv H0. exploit H; eauto. intros [tmp1 [A B]].
- destruct dst; monadInv EQ0.
+ destruct dst; monadInv EQ0; simpl add_dest in *.
(* for value *)
exists (x0 :: tmp1); split.
econstructor; eauto with gensym.
@@ -777,21 +932,25 @@ Opaque makeif.
exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
eauto with gensym.
- (* for test *)
+ (* for set *)
repeat rewrite app_ass; simpl.
exists (x0 :: tmp1); split.
econstructor; eauto with gensym.
apply contained_cons; eauto with gensym.
(* comma *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
- exploit H0; eauto. intros [tmp2 [C D]].
+ exploit H0; eauto with gensym. intros [tmp2 [C D]].
exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
+ destruct dst; simpl; eauto with gensym.
+ apply list_disjoint_cons_r; eauto with gensym.
+ simpl. eapply incl_tran. 2: apply add_dest_incl. auto with gensym.
+ destruct dst; simpl; auto with gensym.
apply contained_app; eauto with gensym.
(* call *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [C D]].
- destruct dst; monadInv EQ2.
+ destruct dst; monadInv EQ2; simpl add_dest in *.
(* for value *)
exists (x1 :: tmp1 ++ tmp2); split.
econstructor; eauto with gensym. congruence.
@@ -801,14 +960,28 @@ Opaque makeif.
exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
- (* for test *)
+ (* for set *)
exists (x1 :: tmp1 ++ tmp2); split.
repeat rewrite app_ass. econstructor; eauto with gensym. congruence.
apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
+(* builtin *)
+ monadInv H0. exploit H; eauto. intros [tmp1 [A B]].
+ destruct dst; monadInv EQ0; simpl add_dest in *.
+ (* for value *)
+ exists (x0 :: tmp1); split.
+ econstructor; eauto with gensym. congruence.
+ apply contained_cons; eauto with gensym.
+ (* for effects *)
+ exists tmp1; split.
+ econstructor; eauto with gensym.
+ auto.
+ (* for set *)
+ exists (x0 :: tmp1); split.
+ repeat rewrite app_ass. econstructor; eauto with gensym. congruence.
+ apply contained_cons; eauto with gensym.
(* loc *)
monadInv H.
-
(* paren *)
monadInv H0.
(* nil *)
@@ -824,10 +997,11 @@ Qed.
Lemma transl_expr_meets_spec:
forall r dst g sl a g' I,
transl_expr dst r g = Res (sl, a) g' I ->
+ dest_below dst g ->
exists tmps, forall ge e le m, tr_top ge e le m dst r sl a tmps.
Proof.
intros. exploit (proj1 transl_meets_spec); eauto. intros [tmps [A B]].
- exists tmps; intros. apply tr_top_base. auto.
+ exists (add_dest dst tmps); intros. apply tr_top_base. auto.
Qed.
Lemma transl_expression_meets_spec:
@@ -867,10 +1041,6 @@ Opaque transl_expression transl_expr_stmt.
clear transl_stmt_meets_spec.
induction s; simpl; intros until I; intros TR;
try (monadInv TR); try (constructor; eauto).
- remember (small_stmt x && small_stmt x0). destruct b.
- exploit andb_prop; eauto. intros [A B].
- eapply tr_ifthenelse_small; eauto.
- monadInv EQ2. eapply tr_ifthenelse_big; eauto.
destruct (is_Sskip s1); monadInv EQ4.
apply tr_for_1; eauto.
apply tr_for_2; eauto.