aboutsummaryrefslogtreecommitdiffstats
path: root/cfrontend
diff options
context:
space:
mode:
authorXavier Leroy <xavier.leroy@college-de-france.fr>2021-08-22 13:29:36 +0200
committerXavier Leroy <xavier.leroy@college-de-france.fr>2021-08-22 13:29:36 +0200
commitd2595e3afb8c38a3391a66c3fc3f7a92fff9eff4 (patch)
treef3c8fba9ffffee5924dadd803fcebdc3520c9361 /cfrontend
parentd97caa16d15b0faca8386a060ec2bfaedad3cdab (diff)
parent47fae389c800034e002c9f8a398e9adc79a14b81 (diff)
downloadcompcert-kvx-d2595e3afb8c38a3391a66c3fc3f7a92fff9eff4.tar.gz
compcert-kvx-d2595e3afb8c38a3391a66c3fc3f7a92fff9eff4.zip
Merge branch 'bitfields' (#400)
Diffstat (limited to 'cfrontend')
-rw-r--r--cfrontend/C2C.ml39
-rw-r--r--cfrontend/Cexec.v322
-rw-r--r--cfrontend/Clight.v69
-rw-r--r--cfrontend/ClightBigstep.v6
-rw-r--r--cfrontend/Cop.v54
-rw-r--r--cfrontend/Csem.v106
-rw-r--r--cfrontend/Cshmgen.v139
-rw-r--r--cfrontend/Cshmgenproof.v221
-rw-r--r--cfrontend/Cstrategy.v320
-rw-r--r--cfrontend/Csyntax.v4
-rw-r--r--cfrontend/Ctypes.v637
-rw-r--r--cfrontend/Ctyping.v54
-rw-r--r--cfrontend/Initializers.v328
-rw-r--r--cfrontend/Initializersproof.v1197
-rw-r--r--cfrontend/PrintCsyntax.ml15
-rw-r--r--cfrontend/SimplExpr.v150
-rw-r--r--cfrontend/SimplExprproof.v1017
-rw-r--r--cfrontend/SimplExprspec.v161
-rw-r--r--cfrontend/SimplLocalsproof.v60
19 files changed, 3295 insertions, 1604 deletions
diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml
index 25d3654f..3c71718c 100644
--- a/cfrontend/C2C.ml
+++ b/cfrontend/C2C.ml
@@ -535,7 +535,7 @@ let checkFunctionType env tres targs =
end
end
-let rec convertTyp env t =
+let rec convertTyp env ?bitwidth t =
match t with
| C.TVoid a -> Tvoid
| C.TInt(ik, a) ->
@@ -566,7 +566,21 @@ let rec convertTyp env t =
| C.TUnion(id, a) ->
Tunion(intern_string id.name, convertAttr a)
| C.TEnum(id, a) ->
- convertIkind Cutil.enum_ikind (convertAttr a)
+ let ik =
+ match bitwidth with
+ | None -> Cutil.enum_ikind
+ | Some w ->
+ let info = Env.find_enum env id in
+ let representable sg =
+ List.for_all (fun (_, v, _) -> Cutil.int_representable v w sg)
+ info.Env.ei_members in
+ if representable false then
+ Cutil.unsigned_ikind_of Cutil.enum_ikind
+ else if representable true then
+ Cutil.signed_ikind_of Cutil.enum_ikind
+ else
+ Cutil.enum_ikind in
+ convertIkind ik (convertAttr a)
and convertParams env = function
| [] -> Tnil
@@ -602,9 +616,16 @@ let rec convertTypAnnotArgs env = function
convertTypAnnotArgs env el)
let convertField env f =
- if f.fld_bitfield <> None then
- unsupported "bit field in struct or union (consider adding option [-fbitfields])";
- (intern_string f.fld_name, convertTyp env f.fld_typ)
+ let id = intern_string f.fld_name
+ and ty = convertTyp env ?bitwidth: f.fld_bitfield f.fld_typ in
+ match f.fld_bitfield with
+ | None -> Member_plain(id, ty)
+ | Some w ->
+ match ty with
+ | Tint(sz, sg, attr) ->
+ Member_bitfield(id, sz, sg, attr, Z.of_uint w, f.fld_name = "")
+ | _ ->
+ fatal_error "bitfield must have type int"
let convertCompositedef env su id attr members =
if Cutil.find_custom_attributes ["packed";"__packed__"] attr <> [] then
@@ -707,6 +728,11 @@ let convertFloat f kind =
(** Expressions *)
+let check_volatile_bitfield env e =
+ if Cutil.is_bitfield env e
+ && List.mem AVolatile (Cutil.attributes_of_type env e.etyp) then
+ warning Diagnostics.Unnamed "access to a volatile bit field, the 'volatile' qualifier is ignored"
+
let ezero = Eval(Vint(coqint_of_camlint 0l), type_int32s)
let ewrap = function
@@ -721,6 +747,7 @@ let rec convertExpr env e =
| C.EUnop((C.Oderef|C.Odot _|C.Oarrow _), _)
| C.EBinop(C.Oindex, _, _, _) ->
let l = convertLvalue env e in
+ check_volatile_bitfield env e;
ewrap (Ctyping.evalof l)
| C.EConst(C.CInt(i, k, _)) ->
@@ -790,6 +817,7 @@ let rec convertExpr env e =
if Cutil.is_composite_type env e2.etyp
&& List.mem AVolatile (Cutil.attributes_of_type env e2.etyp) then
warning Diagnostics.Unnamed "assignment of a value of volatile composite type, the 'volatile' qualifier is ignored";
+ check_volatile_bitfield env e1;
ewrap (Ctyping.eassign e1' e2')
| C.EBinop((C.Oadd_assign|C.Osub_assign|C.Omul_assign|C.Odiv_assign|
C.Omod_assign|C.Oand_assign|C.Oor_assign|C.Oxor_assign|
@@ -810,6 +838,7 @@ let rec convertExpr env e =
| _ -> assert false in
let e1' = convertLvalue env e1 in
let e2' = convertExpr env e2 in
+ check_volatile_bitfield env e1;
ewrap (Ctyping.eassignop op' e1' e2')
| C.EBinop(C.Ocomma, e1, e2, _) ->
ewrap (Ctyping.ecomma (convertExpr env e1) (convertExpr env e2))
diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v
index d763c98c..52037ac0 100644
--- a/cfrontend/Cexec.v
+++ b/cfrontend/Cexec.v
@@ -37,6 +37,10 @@ Notation "'do' X , Y , Z <- A ; B" := (match A with Some (X, Y, Z) => B | None =
(at level 200, X ident, Y ident, Z ident, A at level 100, B at level 200)
: option_monad_scope.
+Notation "'do' X , Y , Z , W <- A ; B" := (match A with Some (X, Y, Z, W) => B | None => None end)
+ (at level 200, X ident, Y ident, Z ident, W ident, A at level 100, B at level 200)
+ : option_monad_scope.
+
Notation " 'check' A ; B" := (if A then B else None)
(at level 200, A at level 100, B at level 200)
: option_monad_scope.
@@ -61,14 +65,14 @@ Proof.
intros until ty. destruct a; simpl; congruence.
Qed.
-Definition is_loc (a: expr) : option (block * ptrofs * type) :=
+Definition is_loc (a: expr) : option (block * ptrofs * bitfield * type) :=
match a with
- | Eloc b ofs ty => Some(b, ofs, ty)
+ | Eloc b ofs bf ty => Some(b, ofs, bf, ty)
| _ => None
end.
Lemma is_loc_inv:
- forall a b ofs ty, is_loc a = Some(b, ofs, ty) -> a = Eloc b ofs ty.
+ forall a b ofs bf ty, is_loc a = Some(b, ofs, bf, ty) -> a = Eloc b ofs bf ty.
Proof.
intros until ty. destruct a; simpl; congruence.
Qed.
@@ -209,15 +213,15 @@ Definition do_volatile_load (w: world) (chunk: memory_chunk) (m: mem) (b: block)
Some(w, E0, v).
Definition do_volatile_store (w: world) (chunk: memory_chunk) (m: mem) (b: block) (ofs: ptrofs) (v: val)
- : option (world * trace * mem) :=
+ : option (world * trace * mem * val) :=
if Genv.block_is_volatile ge b then
do id <- Genv.invert_symbol ge b;
do ev <- eventval_of_val (Val.load_result chunk v) (type_of_chunk chunk);
do w' <- nextworld_vstore w chunk id ofs ev;
- Some(w', Event_vstore chunk id ofs ev :: nil, m)
+ Some(w', Event_vstore chunk id ofs ev :: nil, m, v)
else
do m' <- Mem.store chunk m b (Ptrofs.unsigned ofs) v;
- Some(w, E0, m').
+ Some(w, E0, m', v).
Lemma do_volatile_load_sound:
forall w chunk m b ofs w' t v,
@@ -244,21 +248,21 @@ Proof.
Qed.
Lemma do_volatile_store_sound:
- forall w chunk m b ofs v w' t m',
- do_volatile_store w chunk m b ofs v = Some(w', t, m') ->
- volatile_store ge chunk m b ofs v t m' /\ possible_trace w t w'.
+ forall w chunk m b ofs v w' t m' v',
+ do_volatile_store w chunk m b ofs v = Some(w', t, m', v') ->
+ volatile_store ge chunk m b ofs v t m' /\ possible_trace w t w' /\ v' = v.
Proof.
- intros until m'. unfold do_volatile_store. mydestr.
+ intros until v'. unfold do_volatile_store. mydestr.
split. constructor; auto. apply Genv.invert_find_symbol; auto.
apply eventval_of_val_sound; auto.
- econstructor. constructor; eauto. constructor.
- split. constructor; auto. constructor.
+ split. econstructor. constructor; eauto. constructor. auto.
+ split. constructor; auto. split. constructor. auto.
Qed.
Lemma do_volatile_store_complete:
forall w chunk m b ofs v w' t m',
volatile_store ge chunk m b ofs v t m' -> possible_trace w t w' ->
- do_volatile_store w chunk m b ofs v = Some(w', t, m').
+ do_volatile_store w chunk m b ofs v = Some(w', t, m', v).
Proof.
unfold do_volatile_store; intros. inv H; simpl in *.
rewrite H1. rewrite (Genv.find_invert_symbol _ _ H2).
@@ -269,16 +273,31 @@ Qed.
(** Accessing locations *)
-Definition do_deref_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: ptrofs) : option (world * trace * val) :=
- match access_mode ty with
- | By_value chunk =>
+Definition do_deref_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: ptrofs) (bf: bitfield) : option (world * trace * val) :=
+ match bf with
+ | Full =>
+ match access_mode ty with
+ | By_value chunk =>
match type_is_volatile ty with
| false => do v <- Mem.loadv chunk m (Vptr b ofs); Some(w, E0, v)
| true => do_volatile_load w chunk m b ofs
end
- | By_reference => Some(w, E0, Vptr b ofs)
- | By_copy => Some(w, E0, Vptr b ofs)
- | _ => None
+ | By_reference => Some(w, E0, Vptr b ofs)
+ | By_copy => Some(w, E0, Vptr b ofs)
+ | _ => None
+ end
+ | Bits sz sg pos width =>
+ match ty with
+ | Tint sz1 sg1 _ =>
+ check (intsize_eq sz1 sz &&
+ signedness_eq sg1 (if zlt width (bitsize_intsize sz) then Signed else sg) &&
+ zle 0 pos && zlt 0 width && zle width (bitsize_intsize sz) && zle (pos + width) (bitsize_carrier sz));
+ match Mem.loadv (chunk_for_carrier sz) m (Vptr b ofs) with
+ | Some (Vint c) => Some (w, E0, Vint (bitfield_extract sz sg pos width c))
+ | _ => None
+ end
+ | _ => None
+ end
end.
Definition assign_copy_ok (ty: type) (b: block) (ofs: ptrofs) (b': block) (ofs': ptrofs) : Prop :=
@@ -310,75 +329,103 @@ Proof with try (right; intuition lia).
destruct Y... left; intuition lia.
Defined.
-Definition do_assign_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: ptrofs) (v: val): option (world * trace * mem) :=
- match access_mode ty with
- | By_value chunk =>
- match type_is_volatile ty with
- | false => do m' <- Mem.storev chunk m (Vptr b ofs) v; Some(w, E0, m')
- | true => do_volatile_store w chunk m b ofs v
- end
- | By_copy =>
- match v with
- | Vptr b' ofs' =>
- if check_assign_copy ty b ofs b' ofs' then
- do bytes <- Mem.loadbytes m b' (Ptrofs.unsigned ofs') (sizeof ge ty);
- do m' <- Mem.storebytes m b (Ptrofs.unsigned ofs) bytes;
- Some(w, E0, m')
- else None
- | _ => None
- end
- | _ => None
+Definition do_assign_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: ptrofs) (bf: bitfield) (v: val): option (world * trace * mem * val) :=
+ match bf with
+ | Full =>
+ match access_mode ty with
+ | By_value chunk =>
+ match type_is_volatile ty with
+ | false => do m' <- Mem.storev chunk m (Vptr b ofs) v; Some(w, E0, m', v)
+ | true => do_volatile_store w chunk m b ofs v
+ end
+ | By_copy =>
+ match v with
+ | Vptr b' ofs' =>
+ if check_assign_copy ty b ofs b' ofs' then
+ do bytes <- Mem.loadbytes m b' (Ptrofs.unsigned ofs') (sizeof ge ty);
+ do m' <- Mem.storebytes m b (Ptrofs.unsigned ofs) bytes;
+ Some(w, E0, m', v)
+ else None
+ | _ => None
+ end
+ | _ => None
+ end
+ | Bits sz sg pos width =>
+ check (zle 0 pos && zlt 0 width && zle width (bitsize_intsize sz) && zle (pos + width) (bitsize_carrier sz));
+ match ty, v, Mem.loadv (chunk_for_carrier sz) m (Vptr b ofs) with
+ | Tint sz1 sg1 _, Vint n, Some (Vint c) =>
+ check (intsize_eq sz1 sz &&
+ signedness_eq sg1 (if zlt width (bitsize_intsize sz) then Signed else sg));
+ do m' <- Mem.storev (chunk_for_carrier sz) m (Vptr b ofs)
+ (Vint ((Int.bitfield_insert (first_bit sz pos width) width c n)));
+ Some (w, E0, m', Vint (bitfield_normalize sz sg width n))
+ | _, _, _ => None
+ end
end.
Lemma do_deref_loc_sound:
- forall w ty m b ofs w' t v,
- do_deref_loc w ty m b ofs = Some(w', t, v) ->
- deref_loc ge ty m b ofs t v /\ possible_trace w t w'.
+ forall w ty m b ofs bf w' t v,
+ do_deref_loc w ty m b ofs bf = Some(w', t, v) ->
+ deref_loc ge ty m b ofs bf t v /\ possible_trace w t w'.
Proof.
unfold do_deref_loc; intros until v.
- destruct (access_mode ty) eqn:?; mydestr.
+ destruct bf.
+- destruct (access_mode ty) eqn:?; mydestr.
intros. exploit do_volatile_load_sound; eauto. intuition. eapply deref_loc_volatile; eauto.
split. eapply deref_loc_value; eauto. constructor.
split. eapply deref_loc_reference; eauto. constructor.
split. eapply deref_loc_copy; eauto. constructor.
+- mydestr. destruct ty; mydestr. InvBooleans. subst i. destruct v0; mydestr.
+ split. eapply deref_loc_bitfield; eauto. econstructor; eauto. constructor.
Qed.
Lemma do_deref_loc_complete:
- forall w ty m b ofs w' t v,
- deref_loc ge ty m b ofs t v -> possible_trace w t w' ->
- do_deref_loc w ty m b ofs = Some(w', t, v).
+ forall w ty m b ofs bf w' t v,
+ deref_loc ge ty m b ofs bf t v -> possible_trace w t w' ->
+ do_deref_loc w ty m b ofs bf = Some(w', t, v).
Proof.
unfold do_deref_loc; intros. inv H.
- inv H0. rewrite H1; rewrite H2; rewrite H3; auto.
- rewrite H1; rewrite H2. apply do_volatile_load_complete; auto.
- inv H0. rewrite H1. auto.
- inv H0. rewrite H1. auto.
+- inv H0. rewrite H1; rewrite H2; rewrite H3; auto.
+- rewrite H1; rewrite H2. apply do_volatile_load_complete; auto.
+- inv H0. rewrite H1. auto.
+- inv H0. rewrite H1. auto.
+- inv H0. inv H1.
+ unfold proj_sumbool; rewrite ! dec_eq_true, ! zle_true, ! zlt_true by lia. cbn.
+ cbn in H4; rewrite H4. auto.
Qed.
Lemma do_assign_loc_sound:
- forall w ty m b ofs v w' t m',
- do_assign_loc w ty m b ofs v = Some(w', t, m') ->
- assign_loc ge ty m b ofs v t m' /\ possible_trace w t w'.
+ forall w ty m b ofs bf v w' t m' v',
+ do_assign_loc w ty m b ofs bf v = Some(w', t, m', v') ->
+ assign_loc ge ty m b ofs bf v t m' v' /\ possible_trace w t w'.
Proof.
- unfold do_assign_loc; intros until m'.
- destruct (access_mode ty) eqn:?; mydestr.
- intros. exploit do_volatile_store_sound; eauto. intuition. eapply assign_loc_volatile; eauto.
+ unfold do_assign_loc; intros until v'.
+ destruct bf.
+- destruct (access_mode ty) eqn:?; mydestr.
+ intros. exploit do_volatile_store_sound; eauto. intros (P & Q & R). subst v'. intuition. eapply assign_loc_volatile; eauto.
split. eapply assign_loc_value; eauto. constructor.
destruct v; mydestr. destruct a as [P [Q R]].
split. eapply assign_loc_copy; eauto. constructor.
+- mydestr. InvBooleans.
+ destruct ty; mydestr. destruct v; mydestr. destruct v; mydestr. InvBooleans. subst s i.
+ split. eapply assign_loc_bitfield; eauto. econstructor; eauto. constructor.
Qed.
Lemma do_assign_loc_complete:
- forall w ty m b ofs v w' t m',
- assign_loc ge ty m b ofs v t m' -> possible_trace w t w' ->
- do_assign_loc w ty m b ofs v = Some(w', t, m').
+ forall w ty m b ofs bf v w' t m' v',
+ assign_loc ge ty m b ofs bf v t m' v' -> possible_trace w t w' ->
+ do_assign_loc w ty m b ofs bf v = Some(w', t, m', v').
Proof.
unfold do_assign_loc; intros. inv H.
- inv H0. rewrite H1; rewrite H2; rewrite H3; auto.
- rewrite H1; rewrite H2. apply do_volatile_store_complete; auto.
- rewrite H1. destruct (check_assign_copy ty b ofs b' ofs').
+- inv H0. rewrite H1; rewrite H2; rewrite H3; auto.
+- rewrite H1; rewrite H2. apply do_volatile_store_complete; auto.
+- rewrite H1. destruct (check_assign_copy ty b ofs b' ofs').
inv H0. rewrite H5; rewrite H6; auto.
elim n. red; tauto.
+- inv H0. inv H1.
+ unfold proj_sumbool; rewrite ! zle_true, ! zlt_true by lia. cbn.
+ rewrite ! dec_eq_true.
+ cbn in H4; rewrite H4. cbn in H5; rewrite H5. auto.
Qed.
(** External calls *)
@@ -421,7 +468,7 @@ Definition do_ef_volatile_load (chunk: memory_chunk)
Definition do_ef_volatile_store (chunk: memory_chunk)
(w: world) (vargs: list val) (m: mem) : option (world * trace * val * mem) :=
match vargs with
- | Vptr b ofs :: v :: nil => do w',t,m' <- do_volatile_store w chunk m b ofs v; Some(w',t,Vundef,m')
+ | Vptr b ofs :: v :: nil => do w',t,m',v' <- do_volatile_store w chunk m b ofs v; Some(w',t,Vundef,m')
| _ => None
end.
@@ -564,7 +611,7 @@ Proof with try congruence.
exploit do_volatile_load_sound; eauto. intuition. econstructor; eauto.
- (* EF_vstore *)
unfold do_ef_volatile_store. destruct vargs... destruct v... destruct vargs... destruct vargs...
- mydestr. destruct p as [[w'' t''] m'']. mydestr.
+ mydestr. destruct p as [[[w'' t''] m''] v'']. mydestr.
exploit do_volatile_store_sound; eauto. intuition. econstructor; eauto.
- (* EF_malloc *)
unfold do_ef_malloc. destruct vargs... destruct vargs... mydestr.
@@ -710,6 +757,10 @@ Notation "'do' X , Y , Z <- A ; B" := (match A with Some (X, Y, Z) => B | None =
(at level 200, X ident, Y ident, Z ident, A at level 100, B at level 200)
: reducts_monad_scope.
+Notation "'do' X , Y , Z , W <- A ; B" := (match A with Some (X, Y, Z, W) => B | None => stuck end)
+ (at level 200, X ident, Y ident, Z ident, W ident, A at level 100, B at level 200)
+ : reducts_monad_scope.
+
Notation " 'check' A ; B" := (if A then B else stuck)
(at level 200, A at level 100, B at level 200)
: reducts_monad_scope.
@@ -718,21 +769,21 @@ Local Open Scope reducts_monad_scope.
Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
match k, a with
- | LV, Eloc b ofs ty =>
+ | LV, Eloc b ofs bf ty =>
nil
| LV, Evar x ty =>
match e!x with
| Some(b, ty') =>
check type_eq ty ty';
- topred (Lred "red_var_local" (Eloc b Ptrofs.zero ty) m)
+ topred (Lred "red_var_local" (Eloc b Ptrofs.zero Full ty) m)
| None =>
do b <- Genv.find_symbol ge x;
- topred (Lred "red_var_global" (Eloc b Ptrofs.zero ty) m)
+ topred (Lred "red_var_global" (Eloc b Ptrofs.zero Full ty) m)
end
| LV, Ederef r ty =>
match is_val r with
| Some(Vptr b ofs, ty') =>
- topred (Lred "red_deref" (Eloc b ofs ty) m)
+ topred (Lred "red_deref" (Eloc b ofs Full ty) m)
| Some _ =>
stuck
| None =>
@@ -746,11 +797,14 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
do co <- ge.(genv_cenv)!id;
match field_offset ge f (co_members co) with
| Error _ => stuck
- | OK delta => topred (Lred "red_field_struct" (Eloc b (Ptrofs.add ofs (Ptrofs.repr delta)) ty) m)
+ | OK (delta, bf) => topred (Lred "red_field_struct" (Eloc b (Ptrofs.add ofs (Ptrofs.repr delta)) bf ty) m)
end
| Tunion id _ =>
do co <- ge.(genv_cenv)!id;
- topred (Lred "red_field_union" (Eloc b ofs ty) m)
+ match union_field_offset ge f (co_members co) with
+ | Error _ => stuck
+ | OK (delta, bf) => topred (Lred "red_field_union" (Eloc b (Ptrofs.add ofs (Ptrofs.repr delta)) bf ty) m)
+ end
| _ => stuck
end
| Some _ =>
@@ -762,16 +816,19 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
nil
| RV, Evalof l ty =>
match is_loc l with
- | Some(b, ofs, ty') =>
+ | Some(b, ofs, bf, ty') =>
check type_eq ty ty';
- do w',t,v <- do_deref_loc w ty m b ofs;
+ do w',t,v <- do_deref_loc w ty m b ofs bf;
topred (Rred "red_rvalof" (Eval v ty) m t)
| None =>
incontext (fun x => Evalof x ty) (step_expr LV l m)
end
| RV, Eaddrof l ty =>
match is_loc l with
- | Some(b, ofs, ty') => topred (Rred "red_addrof" (Eval (Vptr b ofs) ty) m E0)
+ | Some(b, ofs, bf, ty') =>
+ match bf with Full => topred (Rred "red_addrof" (Eval (Vptr b ofs) ty) m E0)
+ | Bits _ _ _ _ => stuck
+ end
| None => incontext (fun x => Eaddrof x ty) (step_expr LV l m)
end
| RV, Eunop op r1 ty =>
@@ -831,21 +888,21 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
topred (Rred "red_alignof" (Eval (Vptrofs (Ptrofs.repr (alignof ge ty'))) ty) m E0)
| RV, Eassign l1 r2 ty =>
match is_loc l1, is_val r2 with
- | Some(b, ofs, ty1), Some(v2, ty2) =>
+ | Some(b, ofs, bf, ty1), Some(v2, ty2) =>
check type_eq ty1 ty;
do v <- sem_cast v2 ty2 ty1 m;
- do w',t,m' <- do_assign_loc w ty1 m b ofs v;
- topred (Rred "red_assign" (Eval v ty) m' t)
+ do w',t,m',v' <- do_assign_loc w ty1 m b ofs bf v;
+ topred (Rred "red_assign" (Eval v' ty) m' t)
| _, _ =>
incontext2 (fun x => Eassign x r2 ty) (step_expr LV l1 m)
(fun x => Eassign l1 x ty) (step_expr RV r2 m)
end
| RV, Eassignop op l1 r2 tyres ty =>
match is_loc l1, is_val r2 with
- | Some(b, ofs, ty1), Some(v2, ty2) =>
+ | Some(b, ofs, bf, ty1), Some(v2, ty2) =>
check type_eq ty1 ty;
- do w',t,v1 <- do_deref_loc w ty1 m b ofs;
- let r' := Eassign (Eloc b ofs ty1)
+ do w',t,v1 <- do_deref_loc w ty1 m b ofs bf;
+ let r' := Eassign (Eloc b ofs bf ty1)
(Ebinop op (Eval v1 ty1) (Eval v2 ty2) tyres) ty1 in
topred (Rred "red_assignop" r' m t)
| _, _ =>
@@ -854,12 +911,12 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr :=
end
| RV, Epostincr id l ty =>
match is_loc l with
- | Some(b, ofs, ty1) =>
+ | Some(b, ofs, bf, ty1) =>
check type_eq ty1 ty;
- do w',t, v1 <- do_deref_loc w ty m b ofs;
+ do w',t, v1 <- do_deref_loc w ty m b ofs bf;
let op := match id with Incr => Oadd | Decr => Osub end in
let r' :=
- Ecomma (Eassign (Eloc b ofs ty)
+ Ecomma (Eassign (Eloc b ofs bf ty)
(Ebinop op (Eval v1 ty) (Eval (Vint Int.one) type_int32s) (incrdecr_type ty))
ty)
(Eval v1 ty) ty in
@@ -926,8 +983,8 @@ with step_exprlist (rl: exprlist) (m: mem): reducts exprlist :=
Inductive imm_safe_t: kind -> expr -> mem -> Prop :=
| imm_safe_t_val: forall v ty m,
imm_safe_t RV (Eval v ty) m
- | imm_safe_t_loc: forall b ofs ty m,
- imm_safe_t LV (Eloc b ofs ty) m
+ | imm_safe_t_loc: forall b ofs ty bf m,
+ imm_safe_t LV (Eloc b ofs bf ty) m
| imm_safe_t_lred: forall to C l m l' m',
lred ge e l m l' m' ->
context LV to C ->
@@ -961,23 +1018,25 @@ Fixpoint exprlist_all_values (rl: exprlist) : Prop :=
Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
match a with
- | Eloc b ofs ty => False
+ | Eloc b ofs bf ty => False
| Evar x ty =>
exists b,
e!x = Some(b, ty)
\/ (e!x = None /\ Genv.find_symbol ge x = Some b)
| Ederef (Eval v ty1) ty =>
exists b, exists ofs, v = Vptr b ofs
+ | Eaddrof (Eloc b ofs bf ty1) ty =>
+ bf = Full
| Efield (Eval v ty1) f ty =>
exists b, exists ofs, v = Vptr b ofs /\
match ty1 with
- | Tstruct id _ => exists co delta, ge.(genv_cenv)!id = Some co /\ field_offset ge f (co_members co) = OK delta
- | Tunion id _ => exists co, ge.(genv_cenv)!id = Some co
+ | Tstruct id _ => exists co delta bf, ge.(genv_cenv)!id = Some co /\ field_offset ge f (co_members co) = OK (delta, bf)
+ | Tunion id _ => exists co delta bf, ge.(genv_cenv)!id = Some co /\ union_field_offset ge f (co_members co) = OK (delta, bf)
| _ => False
end
| Eval v ty => False
- | Evalof (Eloc b ofs ty') ty =>
- ty' = ty /\ exists t, exists v, exists w', deref_loc ge ty m b ofs t v /\ possible_trace w t w'
+ | Evalof (Eloc b ofs bf ty') ty =>
+ ty' = ty /\ exists t, exists v, exists w', deref_loc ge ty m b ofs bf t v /\ possible_trace w t w'
| Eunop op (Eval v1 ty1) ty =>
exists v, sem_unary_operation op v1 ty1 m = Some v
| Ebinop op (Eval v1 ty1) (Eval v2 ty2) ty =>
@@ -990,15 +1049,15 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
exists b, bool_val v1 ty1 m = Some b
| Econdition (Eval v1 ty1) r1 r2 ty =>
exists b, bool_val v1 ty1 m = Some b
- | Eassign (Eloc b ofs ty1) (Eval v2 ty2) ty =>
- exists v, exists m', exists t, exists w',
- ty = ty1 /\ sem_cast v2 ty2 ty1 m = Some v /\ assign_loc ge ty1 m b ofs v t m' /\ possible_trace w t w'
- | Eassignop op (Eloc b ofs ty1) (Eval v2 ty2) tyres ty =>
+ | Eassign (Eloc b ofs bf ty1) (Eval v2 ty2) ty =>
+ exists v, exists m', exists v', exists t, exists w',
+ ty = ty1 /\ sem_cast v2 ty2 ty1 m = Some v /\ assign_loc ge ty1 m b ofs bf v t m' v' /\ possible_trace w t w'
+ | Eassignop op (Eloc b ofs bf ty1) (Eval v2 ty2) tyres ty =>
exists t, exists v1, exists w',
- ty = ty1 /\ deref_loc ge ty1 m b ofs t v1 /\ possible_trace w t w'
- | Epostincr id (Eloc b ofs ty1) ty =>
+ ty = ty1 /\ deref_loc ge ty1 m b ofs bf t v1 /\ possible_trace w t w'
+ | Epostincr id (Eloc b ofs bf ty1) ty =>
exists t, exists v1, exists w',
- ty = ty1 /\ deref_loc ge ty m b ofs t v1 /\ possible_trace w t w'
+ ty = ty1 /\ deref_loc ge ty m b ofs bf t v1 /\ possible_trace w t w'
| Ecomma (Eval v ty1) r2 ty =>
typeof r2 = ty
| Eparen (Eval v1 ty1) tycast ty =>
@@ -1026,8 +1085,8 @@ Proof.
exists b; auto.
exists b; auto.
exists b; exists ofs; auto.
- exists b; exists ofs; split; auto. exists co, delta; auto.
- exists b; exists ofs; split; auto. exists co; auto.
+ exists b; exists ofs; split; auto. exists co, delta, bf; auto.
+ exists b; exists ofs; split; auto. exists co, delta, bf; auto.
Qed.
Lemma rred_invert:
@@ -1041,7 +1100,7 @@ Proof.
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 v; exists m'; exists v'; exists t; exists w'; auto.
exists t; exists v1; exists w'; auto.
exists t; exists v1; exists w'; auto.
exists v; auto.
@@ -1076,7 +1135,7 @@ Proof.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
- auto.
+ destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
destruct e1; auto; destruct (C a); auto; contradiction.
@@ -1102,7 +1161,7 @@ Lemma imm_safe_t_inv:
forall k a m,
imm_safe_t k a m ->
match a with
- | Eloc _ _ _ => True
+ | Eloc _ _ _ _ => True
| Eval _ _ => True
| _ => invert_expr_prop a m
end.
@@ -1228,7 +1287,7 @@ Qed.
Lemma not_invert_ok:
forall k a m,
match a with
- | Eloc _ _ _ => False
+ | Eloc _ _ _ _ => False
| Eval _ _ => False
| _ => invert_expr_prop a m -> False
end ->
@@ -1369,18 +1428,19 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
destruct ty'...
(* top struct *)
destruct (ge.(genv_cenv)!i0) as [co|] eqn:?...
- destruct (field_offset ge f (co_members co)) as [delta|] eqn:?...
+ destruct (field_offset ge f (co_members co)) as [[delta bf]|] eqn:?...
apply topred_ok; auto. eapply red_field_struct; eauto.
(* top union *)
destruct (ge.(genv_cenv)!i0) as [co|] eqn:?...
+ destruct (union_field_offset ge f (co_members co)) as [[delta bf]|] eqn:?...
apply topred_ok; auto. eapply red_field_union; eauto.
(* in depth *)
eapply incontext_ok; eauto.
(* Evalof *)
- destruct (is_loc a) as [[[b ofs] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ Heqo).
+ destruct (is_loc a) as [[[[b ofs] bf] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ _ Heqo).
(* top *)
destruct (type_eq ty ty')... subst ty'.
- destruct (do_deref_loc w ty m b ofs) as [[[w' t] v] | ] eqn:?.
+ destruct (do_deref_loc w ty m b ofs bf) as [[[w' t] v] | ] eqn:?.
exploit do_deref_loc_sound; eauto. intros [A B].
apply topred_ok; auto. red. split. apply red_rvalof; auto. exists w'; auto.
apply not_invert_ok; simpl; intros; myinv. exploit do_deref_loc_complete; eauto. congruence.
@@ -1393,8 +1453,9 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
(* depth *)
eapply incontext_ok; eauto.
(* Eaddrof *)
- destruct (is_loc a) as [[[b ofs] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ Heqo).
+ destruct (is_loc a) as [[[[b ofs] bf ] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ _ Heqo).
(* top *)
+ destruct bf...
apply topred_ok; auto. split. apply red_addrof; auto. exists w; constructor.
(* depth *)
eapply incontext_ok; eauto.
@@ -1450,26 +1511,26 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
(* alignof *)
apply topred_ok; auto. split. apply red_alignof. exists w; constructor.
(* assign *)
- destruct (is_loc a1) as [[[b ofs] ty1] | ] eqn:?.
+ destruct (is_loc a1) as [[[[b ofs] bf] ty1] | ] eqn:?.
destruct (is_val a2) as [[v2 ty2] | ] eqn:?.
- rewrite (is_loc_inv _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
+ rewrite (is_loc_inv _ _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
(* top *)
destruct (type_eq ty1 ty)... subst ty1.
destruct (sem_cast v2 ty2 ty m) as [v|] eqn:?...
- destruct (do_assign_loc w ty m b ofs v) as [[[w' t] m']|] eqn:?.
+ destruct (do_assign_loc w ty m b ofs bf v) as [[[[w' t] m'] v']|] eqn:?.
exploit do_assign_loc_sound; eauto. intros [P Q].
- apply topred_ok; auto. split. apply red_assign; auto. exists w'; auto.
+ apply topred_ok; auto. split. eapply red_assign; eauto. exists w'; auto.
apply not_invert_ok; simpl; intros; myinv. exploit do_assign_loc_complete; eauto. congruence.
(* depth *)
eapply incontext2_ok; eauto.
eapply incontext2_ok; eauto.
(* assignop *)
- destruct (is_loc a1) as [[[b ofs] ty1] | ] eqn:?.
+ destruct (is_loc a1) as [[[[b ofs] bf] ty1] | ] eqn:?.
destruct (is_val a2) as [[v2 ty2] | ] eqn:?.
- rewrite (is_loc_inv _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
+ rewrite (is_loc_inv _ _ _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0).
(* top *)
destruct (type_eq ty1 ty)... subst ty1.
- destruct (do_deref_loc w ty m b ofs) as [[[w' t] v] | ] eqn:?.
+ destruct (do_deref_loc w ty m b ofs bf) as [[[w' t] v] | ] eqn:?.
exploit do_deref_loc_sound; eauto. intros [A B].
apply topred_ok; auto. red. split. apply red_assignop; auto. exists w'; auto.
apply not_invert_ok; simpl; intros; myinv. exploit do_deref_loc_complete; eauto. congruence.
@@ -1477,10 +1538,10 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence;
eapply incontext2_ok; eauto.
eapply incontext2_ok; eauto.
(* postincr *)
- destruct (is_loc a) as [[[b ofs] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ Heqo).
+ destruct (is_loc a) as [[[[b ofs] bf] ty'] | ] eqn:?. rewrite (is_loc_inv _ _ _ _ _ Heqo).
(* top *)
destruct (type_eq ty' ty)... subst ty'.
- destruct (do_deref_loc w ty m b ofs) as [[[w' t] v] | ] eqn:?.
+ destruct (do_deref_loc w ty m b ofs bf) as [[[w' t] v] | ] eqn:?.
exploit do_deref_loc_sound; eauto. intros [A B].
apply topred_ok; auto. red. split. apply red_postincr; auto. exists w'; auto.
apply not_invert_ok; simpl; intros; myinv. exploit do_deref_loc_complete; eauto. congruence.
@@ -1576,7 +1637,7 @@ Proof.
(* field struct *)
rewrite H, H0; econstructor; eauto.
(* field union *)
- rewrite H; econstructor; eauto.
+ rewrite H, H0; econstructor; eauto.
Qed.
Lemma rred_topred:
@@ -1587,7 +1648,7 @@ Proof.
induction 1; simpl; intros.
(* valof *)
rewrite dec_eq_true.
- rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0). econstructor; eauto.
+ rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ _ H H0). econstructor; eauto.
(* addrof *)
inv H. econstructor; eauto.
(* unop *)
@@ -1609,13 +1670,13 @@ Proof.
(* alignof *)
inv H. econstructor; eauto.
(* assign *)
- rewrite dec_eq_true. rewrite H. rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ H0 H1).
+ rewrite dec_eq_true. rewrite H. rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ _ _ H0 H1).
econstructor; eauto.
(* assignop *)
- rewrite dec_eq_true. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0).
+ rewrite dec_eq_true. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ _ H H0).
econstructor; eauto.
(* postincr *)
- rewrite dec_eq_true. subst. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H1).
+ rewrite dec_eq_true. subst. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ _ H H1).
econstructor; eauto.
(* comma *)
inv H0. rewrite dec_eq_true. econstructor; eauto.
@@ -1664,10 +1725,10 @@ Proof.
Qed.
Lemma reducts_incl_loc:
- forall (A: Type) a m b ofs ty (C: expr -> A) res,
- is_loc a = Some(b, ofs, ty) -> reducts_incl C (step_expr LV a m) res.
+ forall (A: Type) a m b ofs ty bf (C: expr -> A) res,
+ is_loc a = Some(b, ofs, bf, ty) -> reducts_incl C (step_expr LV a m) res.
Proof.
- intros. rewrite (is_loc_inv _ _ _ _ H). apply reducts_incl_nil.
+ intros. rewrite (is_loc_inv _ _ _ _ _ H). apply reducts_incl_nil.
Qed.
Lemma reducts_incl_listval:
@@ -1732,10 +1793,10 @@ Proof.
destruct (is_val (C a)) as [[v ty']|] eqn:?; eauto.
(* valof *)
eapply reducts_incl_trans with (C' := fun x => Evalof x ty); eauto.
- destruct (is_loc (C a)) as [[[b ofs] ty']|] eqn:?; eauto.
+ destruct (is_loc (C a)) as [[[[b ofs] bf] ty']|] eqn:?; eauto.
(* addrof *)
eapply reducts_incl_trans with (C' := fun x => Eaddrof x ty); eauto.
- destruct (is_loc (C a)) as [[[b ofs] ty']|] eqn:?; eauto.
+ destruct (is_loc (C a)) as [[[[b ofs] bf] ty']|] eqn:?; eauto.
(* unop *)
eapply reducts_incl_trans with (C' := fun x => Eunop op x ty); eauto.
destruct (is_val (C a)) as [[v ty']|] eqn:?; eauto.
@@ -1760,21 +1821,21 @@ Proof.
destruct (is_val (C a)) as [[v ty']|] eqn:?; eauto.
(* assign left *)
eapply reducts_incl_trans with (C' := fun x => Eassign x e2 ty); eauto.
- destruct (is_loc (C a)) as [[[b ofs] ty']|] eqn:?; eauto.
+ destruct (is_loc (C a)) as [[[[b ofs] bf] ty']|] eqn:?; eauto.
(* assign right *)
eapply reducts_incl_trans with (C' := fun x => Eassign e1 x ty); eauto.
- destruct (is_loc e1) as [[[b ofs] ty1]|] eqn:?; eauto.
+ destruct (is_loc e1) as [[[[b ofs] bf] ty1]|] eqn:?; eauto.
destruct (is_val (C a)) as [[v2 ty2]|] eqn:?; eauto.
(* assignop left *)
eapply reducts_incl_trans with (C' := fun x => Eassignop op x e2 tyres ty); eauto.
- destruct (is_loc (C a)) as [[[b ofs] ty']|] eqn:?; eauto.
+ destruct (is_loc (C a)) as [[[[b ofs] bf] ty']|] eqn:?; eauto.
(* assignop right *)
eapply reducts_incl_trans with (C' := fun x => Eassignop op e1 x tyres ty); eauto.
- destruct (is_loc e1) as [[[b ofs] ty1]|] eqn:?; eauto.
+ destruct (is_loc e1) as [[[[b ofs] bf] ty1]|] eqn:?; eauto.
destruct (is_val (C a)) as [[v2 ty2]|] eqn:?; eauto.
(* postincr *)
eapply reducts_incl_trans with (C' := fun x => Epostincr id x ty); eauto.
- destruct (is_loc (C a)) as [[[b ofs] ty']|] eqn:?; eauto.
+ destruct (is_loc (C a)) as [[[[b ofs] bf] ty']|] eqn:?; eauto.
(* call left *)
eapply reducts_incl_trans with (C' := fun x => Ecall x el ty); eauto.
destruct (is_val (C a)) as [[v ty']|] eqn:?; eauto.
@@ -1915,7 +1976,7 @@ Function sem_bind_parameters (w: world) (e: env) (m: mem) (l: list (ident * type
match PTree.get id e with
| Some (b, ty') =>
check (type_eq ty ty');
- do w', t, m1 <- do_assign_loc w ty m b Ptrofs.zero v1;
+ do w', t, m1, v' <- do_assign_loc w ty m b Ptrofs.zero Full v1;
match t with nil => sem_bind_parameters w e m1 params lv | _ => None end
| None => None
end
@@ -1935,10 +1996,11 @@ Lemma sem_bind_parameters_complete : forall w e m l lv m',
bind_parameters ge e m l lv m' ->
sem_bind_parameters w e m l lv = Some m'.
Proof.
+Local Opaque do_assign_loc.
induction 1; simpl; auto.
rewrite H. rewrite dec_eq_true.
assert (possible_trace w E0 w) by constructor.
- rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ H0 H2).
+ rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ _ _ H0 H2).
simpl. auto.
Qed.
diff --git a/cfrontend/Clight.v b/cfrontend/Clight.v
index 3b21be28..d15bc90a 100644
--- a/cfrontend/Clight.v
+++ b/cfrontend/Clight.v
@@ -197,36 +197,42 @@ 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].
+(** [deref_loc ty m b ofs bf v] computes the value of a datum
+ of type [ty] residing in memory [m] at block [b], offset [ofs],
+ and bitfield designation [bf].
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: ptrofs) : val -> Prop :=
+Inductive deref_loc (ty: type) (m: mem) (b: block) (ofs: ptrofs) :
+ bitfield -> 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 ty m b ofs Full v
| deref_loc_reference:
access_mode ty = By_reference ->
- deref_loc ty m b ofs (Vptr b ofs)
+ deref_loc ty m b ofs Full (Vptr b ofs)
| deref_loc_copy:
access_mode ty = By_copy ->
- deref_loc ty m b ofs (Vptr b ofs).
+ deref_loc ty m b ofs Full (Vptr b ofs)
+ | deref_loc_bitfield: forall sz sg pos width v,
+ load_bitfield ty sz sg pos width m (Vptr b ofs) v ->
+ deref_loc ty m b ofs (Bits sz sg pos width) v.
-(** Symmetrically, [assign_loc ty m b ofs v m'] returns the
+(** Symmetrically, [assign_loc ty m b ofs bf 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].
+ of type [ty] residing in memory [m] at block [b], offset [ofs],
+ bitfield designation [bf].
This is allowed only if [ty] indicates an access by value or by copy.
[m'] is the updated memory state. *)
Inductive assign_loc (ce: composite_env) (ty: type) (m: mem) (b: block) (ofs: ptrofs):
- val -> mem -> Prop :=
+ bitfield -> 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 ce ty m b ofs v m'
+ assign_loc ce ty m b ofs Full v m'
| assign_loc_copy: forall b' ofs' bytes m',
access_mode ty = By_copy ->
(sizeof ce ty > 0 -> (alignof_blockcopy ce ty | Ptrofs.unsigned ofs')) ->
@@ -236,7 +242,10 @@ Inductive assign_loc (ce: composite_env) (ty: type) (m: mem) (b: block) (ofs: pt
\/ Ptrofs.unsigned ofs + sizeof ce ty <= Ptrofs.unsigned ofs' ->
Mem.loadbytes m b' (Ptrofs.unsigned ofs') (sizeof ce ty) = Some bytes ->
Mem.storebytes m b (Ptrofs.unsigned ofs) bytes = Some m' ->
- assign_loc ce ty m b ofs (Vptr b' ofs') m'.
+ assign_loc ce ty m b ofs Full (Vptr b' ofs') m'
+ | assign_loc_bitfield: forall sz sg pos width v m' v',
+ store_bitfield ty sz sg pos width m (Vptr b ofs) v m' v' ->
+ assign_loc ce ty m b ofs (Bits sz sg pos width) v m'.
Section SEMANTICS.
@@ -275,7 +284,7 @@ Inductive bind_parameters (e: env):
| bind_parameters_cons:
forall m id ty params v1 vl b m1 m2,
PTree.get id e = Some(b, ty) ->
- assign_loc ge ty m b Ptrofs.zero v1 m1 ->
+ assign_loc ge ty m b Ptrofs.zero Full v1 m1 ->
bind_parameters e m1 params vl m2 ->
bind_parameters e m ((id, ty) :: params) (v1 :: vl) m2.
@@ -369,7 +378,7 @@ Inductive eval_expr: expr -> val -> Prop :=
le!id = Some v ->
eval_expr (Etempvar id ty) v
| eval_Eaddrof: forall a ty loc ofs,
- eval_lvalue a loc ofs ->
+ eval_lvalue a loc ofs Full ->
eval_expr (Eaddrof a ty) (Vptr loc ofs)
| eval_Eunop: forall op a ty v1 v,
eval_expr a v1 ->
@@ -388,37 +397,39 @@ Inductive eval_expr: expr -> val -> Prop :=
eval_expr (Esizeof ty1 ty) (Vptrofs (Ptrofs.repr (sizeof ge ty1)))
| eval_Ealignof: forall ty1 ty,
eval_expr (Ealignof ty1 ty) (Vptrofs (Ptrofs.repr (alignof ge ty1)))
- | eval_Elvalue: forall a loc ofs v,
- eval_lvalue a loc ofs ->
- deref_loc (typeof a) m loc ofs v ->
+ | eval_Elvalue: forall a loc ofs bf v,
+ eval_lvalue a loc ofs bf ->
+ deref_loc (typeof a) m loc ofs bf v ->
eval_expr a v
(** [eval_lvalue ge e m a b ofs] defines the evaluation of expression [a]
in l-value position. The result is the memory location [b, ofs]
- that contains the value of the expression [a]. *)
+ that contains the value of the expression [a], and the bitfield [bf]
+ designation. *)
-with eval_lvalue: expr -> block -> ptrofs -> Prop :=
+with eval_lvalue: expr -> block -> ptrofs -> bitfield -> Prop :=
| eval_Evar_local: forall id l ty,
e!id = Some(l, ty) ->
- eval_lvalue (Evar id ty) l Ptrofs.zero
+ eval_lvalue (Evar id ty) l Ptrofs.zero Full
| eval_Evar_global: forall id l ty,
e!id = None ->
Genv.find_symbol ge id = Some l ->
- eval_lvalue (Evar id ty) l Ptrofs.zero
+ eval_lvalue (Evar id ty) l Ptrofs.zero Full
| eval_Ederef: forall a ty l ofs,
eval_expr a (Vptr l ofs) ->
- eval_lvalue (Ederef a ty) l ofs
- | eval_Efield_struct: forall a i ty l ofs id co att delta,
+ eval_lvalue (Ederef a ty) l ofs Full
+ | eval_Efield_struct: forall a i ty l ofs id co att delta bf,
eval_expr a (Vptr l ofs) ->
typeof a = Tstruct id att ->
ge.(genv_cenv)!id = Some co ->
- field_offset ge i (co_members co) = OK delta ->
- eval_lvalue (Efield a i ty) l (Ptrofs.add ofs (Ptrofs.repr delta))
- | eval_Efield_union: forall a i ty l ofs id co att,
+ field_offset ge i (co_members co) = OK (delta, bf) ->
+ eval_lvalue (Efield a i ty) l (Ptrofs.add ofs (Ptrofs.repr delta)) bf
+ | eval_Efield_union: forall a i ty l ofs id co att delta bf,
eval_expr a (Vptr l ofs) ->
typeof a = Tunion id att ->
ge.(genv_cenv)!id = Some co ->
- eval_lvalue (Efield a i ty) l ofs.
+ union_field_offset ge i (co_members co) = OK (delta, bf) ->
+ eval_lvalue (Efield a i ty) l (Ptrofs.add ofs (Ptrofs.repr delta)) bf.
Scheme eval_expr_ind2 := Minimality for eval_expr Sort Prop
with eval_lvalue_ind2 := Minimality for eval_lvalue Sort Prop.
@@ -547,11 +558,11 @@ Variable function_entry: function -> list val -> mem -> env -> temp_env -> mem -
Inductive step: state -> trace -> state -> Prop :=
- | step_assign: forall f a1 a2 k e le m loc ofs v2 v m',
- eval_lvalue e le m a1 loc ofs ->
+ | step_assign: forall f a1 a2 k e le m loc ofs bf v2 v m',
+ eval_lvalue e le m a1 loc ofs bf ->
eval_expr e le m a2 v2 ->
sem_cast v2 (typeof a2) (typeof a1) m = Some v ->
- assign_loc ge (typeof a1) m loc ofs v m' ->
+ assign_loc ge (typeof a1) m loc ofs bf v m' ->
step (State f (Sassign a1 a2) k e le m)
E0 (State f Sskip k e le m')
diff --git a/cfrontend/ClightBigstep.v b/cfrontend/ClightBigstep.v
index 644c4c6c..51487fa2 100644
--- a/cfrontend/ClightBigstep.v
+++ b/cfrontend/ClightBigstep.v
@@ -79,11 +79,11 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env ->
| exec_Sskip: forall e le m,
exec_stmt e le m Sskip
E0 le m Out_normal
- | exec_Sassign: forall e le m a1 a2 loc ofs v2 v m',
- eval_lvalue ge e le m a1 loc ofs ->
+ | exec_Sassign: forall e le m a1 a2 loc ofs bf v2 v m',
+ eval_lvalue ge e le m a1 loc ofs bf ->
eval_expr ge e le m a2 v2 ->
sem_cast v2 (typeof a2) (typeof a1) m = Some v ->
- assign_loc ge (typeof a1) m loc ofs v m' ->
+ assign_loc ge (typeof a1) m loc ofs bf v m' ->
exec_stmt e le m (Sassign a1 a2)
E0 le m' Out_normal
| exec_Sset: forall e le m id a v,
diff --git a/cfrontend/Cop.v b/cfrontend/Cop.v
index f6524124..0d7bcc3a 100644
--- a/cfrontend/Cop.v
+++ b/cfrontend/Cop.v
@@ -1084,6 +1084,60 @@ Definition incrdecr_type (ty: type) :=
| _ => Tvoid
end.
+(** ** Accessing bit fields *)
+
+Definition chunk_for_carrier (sz: intsize) : memory_chunk :=
+ match sz with
+ | I8 | IBool => Mint8unsigned
+ | I16 => Mint16unsigned
+ | I32 => Mint32
+ end.
+
+(** For bitfield accesses, bits are numbered differently on
+ little-endian and on big-endian machines: bit 0 is the least
+ significant bit in little-endian, and the most significant bit in
+ big-endian. *)
+
+Definition bitsize_carrier (sz: intsize) : Z :=
+ match sz with
+ | I8 | IBool => 8
+ | I16 => 16
+ | I32 => 32
+ end.
+
+Definition first_bit (sz: intsize) (pos width: Z) : Z :=
+ if Archi.big_endian
+ then bitsize_carrier sz - pos - width
+ else pos.
+
+Definition bitfield_extract (sz: intsize) (sg: signedness) (pos width: Z) (c: int) : int :=
+ if intsize_eq sz IBool || signedness_eq sg Unsigned
+ then Int.unsigned_bitfield_extract (first_bit sz pos width) width c
+ else Int.signed_bitfield_extract (first_bit sz pos width) width c.
+
+Definition bitfield_normalize (sz: intsize) (sg: signedness) (width: Z) (n: int) : int :=
+ if intsize_eq sz IBool || signedness_eq sg Unsigned
+ then Int.zero_ext width n
+ else Int.sign_ext width n.
+
+Inductive load_bitfield: type -> intsize -> signedness -> Z -> Z -> mem -> val -> val -> Prop :=
+ | load_bitfield_intro: forall sz sg1 attr sg pos width m addr c,
+ 0 <= pos -> 0 < width <= bitsize_intsize sz -> pos + width <= bitsize_carrier sz ->
+ sg1 = (if zlt width (bitsize_intsize sz) then Signed else sg) ->
+ Mem.loadv (chunk_for_carrier sz) m addr = Some (Vint c) ->
+ load_bitfield (Tint sz sg1 attr) sz sg pos width m addr
+ (Vint (bitfield_extract sz sg pos width c)).
+
+Inductive store_bitfield: type -> intsize -> signedness -> Z -> Z -> mem -> val -> val -> mem -> val -> Prop :=
+ | store_bitfield_intro: forall sz sg1 attr sg pos width m addr c n m',
+ 0 <= pos -> 0 < width <= bitsize_intsize sz -> pos + width <= bitsize_carrier sz ->
+ sg1 = (if zlt width (bitsize_intsize sz) then Signed else sg) ->
+ Mem.loadv (chunk_for_carrier sz) m addr = Some (Vint c) ->
+ Mem.storev (chunk_for_carrier sz) m addr
+ (Vint (Int.bitfield_insert (first_bit sz pos width) width c n)) = Some m' ->
+ store_bitfield (Tint sz sg1 attr) sz sg pos width m addr (Vint n)
+ m' (Vint (bitfield_normalize sz sg width n)).
+
(** * Compatibility with extensions and injections *)
Section GENERIC_INJECTION.
diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v
index 724c1c9e..6698c56f 100644
--- a/cfrontend/Csem.v
+++ b/cfrontend/Csem.v
@@ -46,49 +46,59 @@ Section SEMANTICS.
Variable ge: genv.
-(** [deref_loc ty m b ofs t v] computes the value of a datum
- of type [ty] residing in memory [m] at block [b], offset [ofs].
+(** [deref_loc ty m b ofs bf t v] computes the value of a datum
+ of type [ty] residing in memory [m] at block [b], offset [ofs],
+ and bitfield designation [bf].
If the type [ty] indicates an access by value, the corresponding
memory load is performed. If the type [ty] indicates an access by
reference, the pointer [Vptr b ofs] is returned. [v] is the value
returned, and [t] the trace of observables (nonempty if this is
a volatile access). *)
-Inductive deref_loc (ty: type) (m: mem) (b: block) (ofs: ptrofs) : trace -> val -> Prop :=
+Inductive deref_loc (ty: type) (m: mem) (b: block) (ofs: ptrofs) :
+ bitfield -> trace -> val -> Prop :=
| deref_loc_value: forall chunk v,
access_mode ty = By_value chunk ->
type_is_volatile ty = false ->
Mem.loadv chunk m (Vptr b ofs) = Some v ->
- deref_loc ty m b ofs E0 v
+ deref_loc ty m b ofs Full E0 v
| deref_loc_volatile: forall chunk t v,
access_mode ty = By_value chunk -> type_is_volatile ty = true ->
volatile_load ge chunk m b ofs t v ->
- deref_loc ty m b ofs t v
+ deref_loc ty m b ofs Full t v
| deref_loc_reference:
access_mode ty = By_reference ->
- deref_loc ty m b ofs E0 (Vptr b ofs)
+ deref_loc ty m b ofs Full E0 (Vptr b ofs)
| deref_loc_copy:
access_mode ty = By_copy ->
- deref_loc ty m b ofs E0 (Vptr b ofs).
+ deref_loc ty m b ofs Full E0 (Vptr b ofs)
+ | deref_loc_bitfield: forall sz sg pos width v,
+ load_bitfield ty sz sg pos width m (Vptr b ofs) v ->
+ deref_loc ty m b ofs (Bits sz sg pos width) E0 v.
-(** Symmetrically, [assign_loc ty m b ofs v t m'] returns the
+(** Symmetrically, [assign_loc ty m b ofs bf v t m' v'] returns the
memory state after storing the value [v] in the datum
- of type [ty] residing in memory [m] at block [b], offset [ofs].
+ of type [ty] residing in memory [m] at block [b], offset [ofs],
+ and bitfield designation [bf].
This is allowed only if [ty] indicates an access by value or by copy.
[m'] is the updated memory state and [t] the trace of observables
- (nonempty if this is a volatile store). *)
+ (nonempty if this is a volatile store).
+ [v'] is the result value of the assignment. It is equal to [v]
+ if [bf] is [Full], and to [v] normalized to the width and signedness
+ of the bitfield [bf] otherwise.
+*)
Inductive assign_loc (ty: type) (m: mem) (b: block) (ofs: ptrofs):
- val -> trace -> mem -> Prop :=
+ bitfield -> val -> trace -> mem -> val -> Prop :=
| assign_loc_value: forall v chunk m',
access_mode ty = By_value chunk ->
type_is_volatile ty = false ->
Mem.storev chunk m (Vptr b ofs) v = Some m' ->
- assign_loc ty m b ofs v E0 m'
+ assign_loc ty m b ofs Full v E0 m' v
| assign_loc_volatile: forall v chunk t m',
access_mode ty = By_value chunk -> type_is_volatile ty = true ->
volatile_store ge chunk m b ofs v t m' ->
- assign_loc ty m b ofs v t m'
+ assign_loc ty m b ofs Full v t m' v
| assign_loc_copy: forall b' ofs' bytes m',
access_mode ty = By_copy ->
(alignof_blockcopy ge ty | Ptrofs.unsigned ofs') ->
@@ -98,7 +108,10 @@ Inductive assign_loc (ty: type) (m: mem) (b: block) (ofs: ptrofs):
\/ Ptrofs.unsigned ofs + sizeof ge ty <= Ptrofs.unsigned ofs' ->
Mem.loadbytes m b' (Ptrofs.unsigned ofs') (sizeof ge ty) = Some bytes ->
Mem.storebytes m b (Ptrofs.unsigned ofs) bytes = Some m' ->
- assign_loc ty m b ofs (Vptr b' ofs') E0 m'.
+ assign_loc ty m b ofs Full (Vptr b' ofs') E0 m' (Vptr b' ofs')
+ | assign_loc_bitfield: forall sz sg pos width v m' v',
+ store_bitfield ty sz sg pos width m (Vptr b ofs) v m' v' ->
+ assign_loc ty m b ofs (Bits sz sg pos width) v E0 m' v'.
(** Allocation of function-local variables.
[alloc_variables e1 m1 vars e2 m2] allocates one memory block
@@ -131,9 +144,9 @@ Inductive bind_parameters (e: env):
forall m,
bind_parameters e m nil nil m
| bind_parameters_cons:
- forall m id ty params v1 vl b m1 m2,
+ forall m id ty params v1 vl v1' b m1 m2,
PTree.get id e = Some(b, ty) ->
- assign_loc ty m b Ptrofs.zero v1 E0 m1 ->
+ assign_loc ty m b Ptrofs.zero Full v1 E0 m1 v1' ->
bind_parameters e m1 params vl m2 ->
bind_parameters e m ((id, ty) :: params) (v1 :: vl) m2.
@@ -202,34 +215,35 @@ Inductive lred: expr -> mem -> expr -> mem -> Prop :=
| red_var_local: forall x ty m b,
e!x = Some(b, ty) ->
lred (Evar x ty) m
- (Eloc b Ptrofs.zero ty) m
+ (Eloc b Ptrofs.zero Full ty) m
| red_var_global: forall x ty m b,
e!x = None ->
Genv.find_symbol ge x = Some b ->
lred (Evar x ty) m
- (Eloc b Ptrofs.zero ty) m
+ (Eloc b Ptrofs.zero Full ty) m
| red_deref: forall b ofs ty1 ty m,
lred (Ederef (Eval (Vptr b ofs) ty1) ty) m
- (Eloc b ofs ty) m
- | red_field_struct: forall b ofs id co a f ty m delta,
+ (Eloc b ofs Full ty) m
+ | red_field_struct: forall b ofs id co a f ty m delta bf,
ge.(genv_cenv)!id = Some co ->
- field_offset ge f (co_members co) = OK delta ->
+ field_offset ge f (co_members co) = OK (delta, bf) ->
lred (Efield (Eval (Vptr b ofs) (Tstruct id a)) f ty) m
- (Eloc b (Ptrofs.add ofs (Ptrofs.repr delta)) ty) m
- | red_field_union: forall b ofs id co a f ty m,
+ (Eloc b (Ptrofs.add ofs (Ptrofs.repr delta)) bf ty) m
+ | red_field_union: forall b ofs id co a f ty m delta bf,
ge.(genv_cenv)!id = Some co ->
+ union_field_offset ge f (co_members co) = OK (delta, bf) ->
lred (Efield (Eval (Vptr b ofs) (Tunion id a)) f ty) m
- (Eloc b ofs ty) m.
+ (Eloc b (Ptrofs.add ofs (Ptrofs.repr delta)) bf ty) m.
(** Head reductions for r-values *)
Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop :=
- | red_rvalof: forall b ofs ty m t v,
- deref_loc ty m b ofs t v ->
- rred (Evalof (Eloc b ofs ty) ty) m
+ | red_rvalof: forall b ofs bf ty m t v,
+ deref_loc ty m b ofs bf t v ->
+ rred (Evalof (Eloc b ofs bf ty) ty) m
t (Eval v ty) m
| red_addrof: forall b ofs ty1 ty m,
- rred (Eaddrof (Eloc b ofs ty1) ty) m
+ rred (Eaddrof (Eloc b ofs Full ty1) ty) m
E0 (Eval (Vptr b ofs) ty) m
| red_unop: forall op v1 ty1 ty m v,
sem_unary_operation op v1 ty1 m = Some v ->
@@ -269,21 +283,21 @@ Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop :=
| red_alignof: forall ty1 ty m,
rred (Ealignof ty1 ty) m
E0 (Eval (Vptrofs (Ptrofs.repr (alignof ge ty1))) ty) m
- | red_assign: forall b ofs ty1 v2 ty2 m v t m',
+ | red_assign: forall b ofs ty1 bf v2 ty2 m v t m' v',
sem_cast v2 ty2 ty1 m = Some v ->
- assign_loc ty1 m b ofs v t m' ->
- rred (Eassign (Eloc b ofs ty1) (Eval v2 ty2) ty1) m
- t (Eval v ty1) m'
- | red_assignop: forall op b ofs ty1 v2 ty2 tyres m t v1,
- deref_loc ty1 m b ofs t v1 ->
- rred (Eassignop op (Eloc b ofs ty1) (Eval v2 ty2) tyres ty1) m
- t (Eassign (Eloc b ofs ty1)
+ assign_loc ty1 m b ofs bf v t m' v' ->
+ rred (Eassign (Eloc b ofs bf ty1) (Eval v2 ty2) ty1) m
+ t (Eval v' ty1) m'
+ | red_assignop: forall op b ofs ty1 bf v2 ty2 tyres m t v1,
+ deref_loc ty1 m b ofs bf t v1 ->
+ rred (Eassignop op (Eloc b ofs bf ty1) (Eval v2 ty2) tyres ty1) m
+ t (Eassign (Eloc b ofs bf ty1)
(Ebinop op (Eval v1 ty1) (Eval v2 ty2) tyres) ty1) m
- | red_postincr: forall id b ofs ty m t v1 op,
- deref_loc ty m b ofs t v1 ->
+ | red_postincr: forall id b ofs ty bf m t v1 op,
+ deref_loc ty m b ofs bf t v1 ->
op = match id with Incr => Oadd | Decr => Osub end ->
- rred (Epostincr id (Eloc b ofs ty) ty) m
- t (Ecomma (Eassign (Eloc b ofs ty)
+ rred (Epostincr id (Eloc b ofs bf ty) ty) m
+ t (Ecomma (Eassign (Eloc b ofs bf ty)
(Ebinop op (Eval v1 ty)
(Eval (Vint Int.one) type_int32s)
(incrdecr_type ty))
@@ -409,8 +423,8 @@ with contextlist: kind -> (expr -> exprlist) -> Prop :=
Inductive imm_safe: kind -> expr -> mem -> Prop :=
| imm_safe_val: forall v ty m,
imm_safe RV (Eval v ty) m
- | imm_safe_loc: forall b ofs ty m,
- imm_safe LV (Eloc b ofs ty) m
+ | imm_safe_loc: forall b ofs bf ty m,
+ imm_safe LV (Eloc b ofs bf ty) m
| imm_safe_lred: forall to C e m e' m',
lred e m e' m' ->
context LV to C ->
@@ -839,10 +853,10 @@ Lemma semantics_single_events:
Proof.
unfold semantics; intros; red; simpl; intros.
set (ge := globalenv p) in *.
- assert (DEREF: forall chunk m b ofs t v, deref_loc ge chunk m b ofs t v -> (length t <= 1)%nat).
- intros. inv H0; simpl; try lia. inv H3; simpl; try lia.
- 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 lia. inv H3; simpl; try lia.
+ assert (DEREF: forall chunk m b ofs bf t v, deref_loc ge chunk m b ofs bf t v -> (length t <= 1)%nat).
+ { intros. inv H0; simpl; try lia. inv H3; simpl; try lia. }
+ assert (ASSIGN: forall chunk m b ofs bf t v m' v', assign_loc ge chunk m b ofs bf v t m' v' -> (length t <= 1)%nat).
+ { intros. inv H0; simpl; try lia. inv H3; simpl; try lia. }
destruct H.
inv H; simpl; try lia. inv H0; eauto; simpl; try lia.
eapply external_call_trace_length; eauto.
diff --git a/cfrontend/Cshmgen.v b/cfrontend/Cshmgen.v
index 5bd12d00..9e804176 100644
--- a/cfrontend/Cshmgen.v
+++ b/cfrontend/Cshmgen.v
@@ -372,19 +372,56 @@ Definition make_cmp (c: comparison) (e1: expr) (ty1: type) (e2: expr) (ty2: type
e1 ty1 e2 ty2
end.
+(** Auxiliary for translating bitfield accesses *)
+
+Definition make_extract_bitfield (sz: intsize) (sg: signedness) (pos width: Z)
+ (addr: expr) : res expr :=
+ if zle 0 pos && zlt 0 width && zle (pos + width) (bitsize_carrier sz) then
+ let amount1 := Int.repr (Int.zwordsize - first_bit sz pos width - width) in
+ let amount2 := Int.repr (Int.zwordsize - width) in
+ let e1 := Eload (chunk_for_carrier sz) addr in
+ let e2 := Ebinop Oshl e1 (make_intconst amount1) in
+ let e3 := Ebinop (if intsize_eq sz IBool
+ || signedness_eq sg Unsigned then Oshru else Oshr)
+ e2 (make_intconst amount2) in
+ OK e3
+ else
+ Error(msg "Cshmgen.extract_bitfield").
+
(** [make_load addr ty_res] loads a value of type [ty_res] from
- the memory location denoted by the Csharpminor expression [addr].
+ the memory location denoted by the Csharpminor expression [addr]
+ and the bitfield designator [bf].
If [ty_res] is an array or function type, returns [addr] instead,
as consistent with C semantics.
*)
-Definition make_load (addr: expr) (ty_res: type) :=
- match (access_mode ty_res) with
- | By_value chunk => OK (Eload chunk addr)
- | By_reference => OK addr
- | By_copy => OK addr
- | By_nothing => Error (msg "Cshmgen.make_load")
- end.
+Definition make_load (addr: expr) (ty_res: type) (bf: bitfield) :=
+ match bf with
+ | Full =>
+ match access_mode ty_res with
+ | By_value chunk => OK (Eload chunk addr)
+ | By_reference => OK addr
+ | By_copy => OK addr
+ | By_nothing => Error (msg "Cshmgen.make_load")
+ end
+ | Bits sz sg pos width =>
+ make_extract_bitfield sz sg pos width addr
+ end.
+
+(** Auxiliary for translating bitfield updates *)
+
+Definition make_store_bitfield (sz: intsize) (sg: signedness) (pos width: Z)
+ (addr val: expr) : res stmt :=
+ if zle 0 pos && zlt 0 width && zle (pos + width) (bitsize_carrier sz) then
+ let amount := first_bit sz pos width in
+ let mask := Int.shl (Int.repr (two_p width - 1)) (Int.repr amount) in
+ let e1 := Eload (chunk_for_carrier sz) addr in
+ let e2 := Ebinop Oshl val (make_intconst (Int.repr amount)) in
+ let e3 := Ebinop Oor (Ebinop Oand e2 (make_intconst mask))
+ (Ebinop Oand e1 (make_intconst (Int.not mask))) in
+ OK (Sstore (chunk_for_carrier sz) addr e3)
+ else
+ Error(msg "Cshmgen.make_store_bitfield").
(** [make_memcpy dst src ty] returns a [memcpy] builtin appropriate for
by-copy assignment of a value of Clight type [ty]. *)
@@ -394,16 +431,21 @@ Definition make_memcpy (ce: composite_env) (dst src: expr) (ty: type) :=
OK (Sbuiltin None (EF_memcpy sz (Ctypes.alignof_blockcopy ce ty))
(dst :: src :: nil)).
-(** [make_store addr ty rhs] stores the value of the
+(** [make_store addr ty bf rhs] stores the value of the
Csharpminor expression [rhs] into the memory location denoted by the
Csharpminor expression [addr].
- [ty] is the type of the memory location. *)
-
-Definition make_store (ce: composite_env) (addr: expr) (ty: type) (rhs: expr) :=
- match access_mode ty with
- | By_value chunk => OK (Sstore chunk addr rhs)
- | By_copy => make_memcpy ce addr rhs ty
- | _ => Error (msg "Cshmgen.make_store")
+ [ty] is the type of the memory location and [bf] a bitfield designator. *)
+
+Definition make_store (ce: composite_env) (addr: expr) (ty: type) (bf: bitfield) (rhs: expr) :=
+ match bf with
+ | Full =>
+ match access_mode ty with
+ | By_value chunk => OK (Sstore chunk addr rhs)
+ | By_copy => make_memcpy ce addr rhs ty
+ | _ => Error (msg "Cshmgen.make_store")
+ end
+ | Bits sz sg pos width =>
+ make_store_bitfield sz sg pos width addr rhs
end.
(** ** Translation of operators *)
@@ -441,23 +483,27 @@ Definition transl_binop (ce: composite_env)
(** ** Translation of field accesses *)
-Definition make_field_access (ce: composite_env) (ty: type) (f: ident) (a: expr) : res expr :=
- match ty with
- | Tstruct id _ =>
- match ce!id with
- | None =>
- Error (MSG "Undefined struct " :: CTX id :: nil)
- | Some co =>
- do ofs <- field_offset ce f (co_members co);
- OK (if Archi.ptr64
- then Ebinop Oaddl a (make_longconst (Int64.repr ofs))
- else Ebinop Oadd a (make_intconst (Int.repr ofs)))
- end
- | Tunion id _ =>
- OK a
- | _ =>
- Error(msg "Cshmgen.make_field_access")
- end.
+Definition make_field_access (ce: composite_env) (ty: type) (f: ident) (a: expr) : res (expr * bitfield) :=
+ do (ofs, bf) <-
+ match ty with
+ | Tstruct id _ =>
+ match ce!id with
+ | None => Error (MSG "Undefined struct " :: CTX id :: nil)
+ | Some co => field_offset ce f (co_members co)
+ end
+ | Tunion id _ =>
+ match ce!id with
+ | None => Error (MSG "Undefined union " :: CTX id :: nil)
+ | Some co => union_field_offset ce f (co_members co)
+ end
+ | _ =>
+ Error(msg "Cshmgen.make_field_access")
+ end;
+ let a' :=
+ if Archi.ptr64
+ then Ebinop Oaddl a (make_longconst (Int64.repr ofs))
+ else Ebinop Oadd a (make_intconst (Int.repr ofs)) in
+ OK (a', bf).
(** * Translation of expressions *)
@@ -476,14 +522,18 @@ Fixpoint transl_expr (ce: composite_env) (a: Clight.expr) {struct a} : res expr
| Clight.Econst_long n _ =>
OK(make_longconst n)
| Clight.Evar id ty =>
- make_load (Eaddrof id) ty
+ make_load (Eaddrof id) ty Full
| Clight.Etempvar id ty =>
OK(Evar id)
| Clight.Ederef b ty =>
do tb <- transl_expr ce b;
- make_load tb ty
+ make_load tb ty Full
| Clight.Eaddrof b _ =>
- transl_lvalue ce b
+ do (tb, bf) <- transl_lvalue ce b;
+ match bf with
+ | Full => OK tb
+ | Bits _ _ _ _ => Error (msg "Cshmgen.transl_expr: addrof bitfield")
+ end
| Clight.Eunop op b _ =>
do tb <- transl_expr ce b;
transl_unop op tb (typeof b)
@@ -496,8 +546,8 @@ Fixpoint transl_expr (ce: composite_env) (a: Clight.expr) {struct a} : res expr
make_cast (typeof b) ty tb
| Clight.Efield b i ty =>
do tb <- transl_expr ce b;
- do addr <- make_field_access ce (typeof b) i tb;
- make_load addr ty
+ do (addr, bf) <- make_field_access ce (typeof b) i tb;
+ make_load addr ty bf
| Clight.Esizeof ty' ty =>
do sz <- sizeof ce ty'; OK(make_ptrofsconst sz)
| Clight.Ealignof ty' ty =>
@@ -506,15 +556,16 @@ Fixpoint transl_expr (ce: composite_env) (a: Clight.expr) {struct a} : res expr
(** [transl_lvalue a] returns the Csharpminor code that evaluates
[a] as a lvalue, that is, code that returns the memory address
- where the value of [a] is stored.
+ where the value of [a] is stored. It also returns the bitfield to be
+ accessed at this address, if appropriate.
*)
-with transl_lvalue (ce: composite_env) (a: Clight.expr) {struct a} : res expr :=
+with transl_lvalue (ce: composite_env) (a: Clight.expr) {struct a} : res (expr * bitfield) :=
match a with
| Clight.Evar id _ =>
- OK (Eaddrof id)
+ OK (Eaddrof id, Full)
| Clight.Ederef b _ =>
- transl_expr ce b
+ do tb <- transl_expr ce b; OK (tb, Full)
| Clight.Efield b i ty =>
do tb <- transl_expr ce b;
make_field_access ce (typeof b) i tb
@@ -618,10 +669,10 @@ Fixpoint transl_statement (ce: composite_env) (tyret: type) (nbrk ncnt: nat)
| Clight.Sskip =>
OK Sskip
| Clight.Sassign b c =>
- do tb <- transl_lvalue ce b;
+ do (tb, bf) <- transl_lvalue ce b;
do tc <- transl_expr ce c;
do tc' <- make_cast (typeof c) (typeof b) tc;
- make_store ce tb (typeof b) tc'
+ make_store ce tb (typeof b) bf tc'
| Clight.Sset x b =>
do tb <- transl_expr ce b;
OK(Sset x tb)
diff --git a/cfrontend/Cshmgenproof.v b/cfrontend/Cshmgenproof.v
index 6e653e01..8e396e2a 100644
--- a/cfrontend/Cshmgenproof.v
+++ b/cfrontend/Cshmgenproof.v
@@ -115,6 +115,21 @@ Proof.
destruct (prog_comp_env cunit)!i as [co|] eqn:X; try discriminate; erewrite H1 by eauto; auto.
Qed.
+Lemma union_field_offset_stable:
+ forall (cunit prog: Clight.program) id co f,
+ linkorder cunit prog ->
+ cunit.(prog_comp_env)!id = Some co ->
+ prog.(prog_comp_env)!id = Some co /\
+ union_field_offset prog.(prog_comp_env) f (co_members co) = union_field_offset cunit.(prog_comp_env) f (co_members co).
+Proof.
+ intros.
+ assert (C: composite_consistent cunit.(prog_comp_env) co).
+ { apply build_composite_env_consistent with cunit.(prog_types) id; auto.
+ apply prog_comp_env_eq. }
+ destruct H as [_ A].
+ split. auto. apply Ctypes.union_field_offset_stable; eauto using co_consistent_complete.
+Qed.
+
Lemma field_offset_stable:
forall (cunit prog: Clight.program) id co f,
linkorder cunit prog ->
@@ -127,38 +142,11 @@ Proof.
{ apply build_composite_env_consistent with cunit.(prog_types) id; auto.
apply prog_comp_env_eq. }
destruct H as [_ A].
- split. auto. generalize (co_consistent_complete _ _ C).
- unfold field_offset. generalize 0. induction (co_members co) as [ | [f1 t1] m]; simpl; intros.
-- auto.
-- InvBooleans.
- rewrite ! (alignof_stable _ _ A) by auto.
- rewrite ! (sizeof_stable _ _ A) by auto.
- destruct (ident_eq f f1); eauto.
+ split. auto. apply Ctypes.field_offset_stable; eauto using co_consistent_complete.
Qed.
(** * Properties of the translation functions *)
-(** Transformation of expressions and statements. *)
-
-Lemma transl_expr_lvalue:
- forall ge e le m a loc ofs ce ta,
- Clight.eval_lvalue ge e le m a loc ofs ->
- transl_expr ce a = OK ta ->
- (exists tb, transl_lvalue ce a = OK tb /\ make_load tb (typeof a) = OK ta).
-Proof.
- intros until ta; intros EVAL TR. inv EVAL; simpl in TR.
- (* var local *)
- exists (Eaddrof id); auto.
- (* var global *)
- exists (Eaddrof id); auto.
- (* deref *)
- monadInv TR. exists x; auto.
- (* field struct *)
- monadInv TR. exists x0; split; auto. simpl; rewrite EQ; auto.
- (* field union *)
- monadInv TR. exists x0; split; auto. simpl; rewrite EQ; auto.
-Qed.
-
(** Properties of labeled statements *)
Lemma transl_lbl_stmt_1:
@@ -940,28 +928,83 @@ Proof.
eapply make_cmp_correct; eauto.
Qed.
+Remark int_ltu_true:
+ forall x, 0 <= x < Int.zwordsize -> Int.ltu (Int.repr x) Int.iwordsize = true.
+Proof.
+ intros. unfold Int.ltu. rewrite Int.unsigned_repr_wordsize, Int.unsigned_repr, zlt_true by (generalize Int.wordsize_max_unsigned; lia).
+ auto.
+Qed.
+
+Remark first_bit_range: forall sz pos width,
+ 0 <= pos -> 0 < width -> pos + width <= bitsize_carrier sz ->
+ 0 <= first_bit sz pos width < Int.zwordsize
+ /\ 0 <= Int.zwordsize - first_bit sz pos width - width < Int.zwordsize.
+Proof.
+ intros.
+ assert (bitsize_carrier sz <= Int.zwordsize) by (destruct sz; compute; congruence).
+ unfold first_bit; destruct Archi.big_endian; lia.
+Qed.
+
Lemma make_load_correct:
- forall addr ty code b ofs v e le m,
- make_load addr ty = OK code ->
+ forall addr ty bf code b ofs v e le m,
+ make_load addr ty bf = OK code ->
eval_expr ge e le m addr (Vptr b ofs) ->
- deref_loc ty m b ofs v ->
+ deref_loc ty m b ofs bf v ->
eval_expr ge e le m code v.
Proof.
unfold make_load; intros until m; intros MKLOAD EVEXP DEREF.
inv DEREF.
- (* scalar *)
+- (* scalar *)
rewrite H in MKLOAD. inv MKLOAD. apply eval_Eload with (Vptr b ofs); auto.
- (* by reference *)
+- (* by reference *)
rewrite H in MKLOAD. inv MKLOAD. auto.
- (* by copy *)
+- (* by copy *)
rewrite H in MKLOAD. inv MKLOAD. auto.
+- (* by bitfield *)
+ inv H.
+ unfold make_extract_bitfield in MKLOAD. unfold bitfield_extract.
+ exploit (first_bit_range sz pos width); eauto. lia. intros [A1 A2].
+ set (amount1 := Int.repr (Int.zwordsize - first_bit sz pos width - width)) in MKLOAD.
+ set (amount2 := Int.repr (Int.zwordsize - width)) in MKLOAD.
+ destruct (zle 0 pos && zlt 0 width && zle (pos + width) (bitsize_carrier sz)); inv MKLOAD.
+ set (e1 := Eload (chunk_for_carrier sz) addr).
+ assert (E1: eval_expr ge e le m e1 (Vint c)) by (econstructor; eauto).
+ set (e2 := Ebinop Oshl e1 (make_intconst amount1)).
+ assert (E2: eval_expr ge e le m e2 (Vint (Int.shl c amount1))).
+ { econstructor; eauto using make_intconst_correct. cbn.
+ unfold amount1 at 1; rewrite int_ltu_true by lia. auto. }
+ econstructor; eauto using make_intconst_correct.
+ destruct (Ctypes.intsize_eq sz IBool || Ctypes.signedness_eq sg Unsigned); cbn.
+ + unfold amount2 at 1; rewrite int_ltu_true by lia.
+ rewrite Int.unsigned_bitfield_extract_by_shifts by lia. auto.
+ + unfold amount2 at 1; rewrite int_ltu_true by lia.
+ rewrite Int.signed_bitfield_extract_by_shifts by lia. auto.
+Qed.
+
+Lemma make_store_bitfield_correct:
+ forall f sz sg pos width dst src ty k e le m b ofs v m' s,
+ eval_expr ge e le m dst (Vptr b ofs) ->
+ eval_expr ge e le m src v ->
+ assign_loc prog.(prog_comp_env) ty m b ofs (Bits sz sg pos width) v m' ->
+ make_store_bitfield sz sg pos width dst src = OK s ->
+ step ge (State f s k e le m) E0 (State f Sskip k e le m').
+Proof.
+ intros until s; intros DST SRC ASG MK.
+ inv ASG. inv H5. unfold make_store_bitfield in MK.
+ destruct (zle 0 pos && zlt 0 width && zle (pos + width) (bitsize_carrier sz)); inv MK.
+ econstructor; eauto.
+ exploit (first_bit_range sz pos width); eauto. lia. intros [A1 A2].
+ rewrite Int.bitfield_insert_alternative by lia.
+ set (amount := first_bit sz pos width).
+ set (mask := Int.shl (Int.repr (two_p width - 1)) (Int.repr amount)).
+ repeat econstructor; eauto. cbn. rewrite int_ltu_true by lia. auto.
Qed.
Lemma make_memcpy_correct:
forall f dst src ty k e le m b ofs v m' s,
eval_expr ge e le m dst (Vptr b ofs) ->
eval_expr ge e le m src v ->
- assign_loc prog.(prog_comp_env) ty m b ofs v m' ->
+ assign_loc prog.(prog_comp_env) ty m b ofs Full v m' ->
access_mode ty = By_copy ->
make_memcpy cunit.(prog_comp_env) dst src ty = OK s ->
step ge (State f s k e le m) E0 (State f Sskip k e le m').
@@ -979,21 +1022,23 @@ Proof.
Qed.
Lemma make_store_correct:
- forall addr ty rhs code e le m b ofs v m' f k,
- make_store cunit.(prog_comp_env) addr ty rhs = OK code ->
+ forall addr ty bf rhs code e le m b ofs v m' f k,
+ make_store cunit.(prog_comp_env) addr ty bf rhs = OK code ->
eval_expr ge e le m addr (Vptr b ofs) ->
eval_expr ge e le m rhs v ->
- assign_loc prog.(prog_comp_env) ty m b ofs v m' ->
+ assign_loc prog.(prog_comp_env) ty m b ofs bf 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.
inversion ASSIGN; subst.
- (* nonvolatile scalar *)
+- (* nonvolatile scalar *)
rewrite H in MKSTORE; inv MKSTORE.
econstructor; eauto.
- (* by copy *)
+- (* by copy *)
rewrite H in MKSTORE.
eapply make_memcpy_correct with (b := b) (v := Vptr b' ofs'); eauto.
+- (* bitfield *)
+ eapply make_store_bitfield_correct; eauto.
Qed.
Lemma make_normalization_correct:
@@ -1212,15 +1257,51 @@ Variable m: mem.
Variable te: Csharpminor.env.
Hypothesis MENV: match_env e te.
+Lemma transl_expr_lvalue:
+ forall a loc ofs bf ta,
+ Clight.eval_lvalue ge e le m a loc ofs bf ->
+ transl_expr cunit.(prog_comp_env) a = OK ta ->
+ exists tb, transl_lvalue cunit.(prog_comp_env) a = OK (tb, bf)
+ /\ make_load tb (typeof a) bf = OK ta.
+Proof.
+ intros until ta; intros EVAL TR. inv EVAL; simpl in TR.
+- (* var local *)
+ exists (Eaddrof id); auto.
+- (* var global *)
+ exists (Eaddrof id); auto.
+- (* deref *)
+ monadInv TR. cbn; rewrite EQ. exists x; auto.
+- (* field struct *)
+ monadInv TR.
+ assert (x1 = bf).
+ { rewrite H0 in EQ1. unfold make_field_access in EQ1.
+ destruct ((prog_comp_env cunit)!id) as [co'|] eqn:E; try discriminate.
+ monadInv EQ1.
+ exploit field_offset_stable. eexact LINK. eauto. instantiate (1 := i). intros [A B].
+ simpl in H1, H2. congruence. }
+ subst x1.
+ exists x0; split; auto. simpl; rewrite EQ; auto.
+- (* field union *)
+ monadInv TR.
+ assert (x1 = bf).
+ { rewrite H0 in EQ1. unfold make_field_access in EQ1.
+ destruct ((prog_comp_env cunit)!id) as [co'|] eqn:E; try discriminate.
+ monadInv EQ1.
+ exploit union_field_offset_stable. eexact LINK. eauto. instantiate (1 := i). intros [A B].
+ simpl in H1, H2. congruence. }
+ subst x1.
+ exists x0; split; auto. simpl; rewrite EQ; auto.
+Qed.
+
Lemma transl_expr_lvalue_correct:
(forall a v,
Clight.eval_expr ge e le m a v ->
forall ta (TR: transl_expr cunit.(prog_comp_env) a = OK ta) ,
Csharpminor.eval_expr tge te le m ta v)
-/\(forall a b ofs,
- Clight.eval_lvalue ge e le m a b ofs ->
- forall ta (TR: transl_lvalue cunit.(prog_comp_env) a = OK ta),
- Csharpminor.eval_expr tge te le m ta (Vptr b ofs)).
+/\(forall a b ofs bf,
+ Clight.eval_lvalue ge e le m a b ofs bf ->
+ forall ta bf' (TR: transl_lvalue cunit.(prog_comp_env) a = OK (ta, bf')),
+ bf = bf' /\ Csharpminor.eval_expr tge te le m ta (Vptr b ofs)).
Proof.
apply eval_expr_lvalue_ind; intros; try (monadInv TR).
- (* const int *)
@@ -1234,7 +1315,7 @@ Proof.
- (* temp var *)
constructor; auto.
- (* addrof *)
- simpl in TR. auto.
+ destruct x0; inv EQ0. apply H0 in EQ. destruct EQ. auto.
- (* unop *)
eapply transl_unop_correct; eauto.
- (* binop *)
@@ -1247,31 +1328,43 @@ Proof.
rewrite (transl_alignof _ _ _ _ LINK EQ). apply make_ptrofsconst_correct.
- (* rvalue out of lvalue *)
exploit transl_expr_lvalue; eauto. intros [tb [TRLVAL MKLOAD]].
+ apply H0 in TRLVAL; destruct TRLVAL.
eapply make_load_correct; eauto.
- (* var local *)
exploit (me_local _ _ MENV); eauto. intros EQ.
- econstructor. eapply eval_var_addr_local. eauto.
+ split; auto. econstructor. eapply eval_var_addr_local. eauto.
- (* var global *)
- econstructor. eapply eval_var_addr_global.
+ split; auto. econstructor. eapply eval_var_addr_global.
eapply match_env_globals; eauto.
rewrite symbols_preserved. auto.
- (* deref *)
- simpl in TR. eauto.
+ eauto.
- (* field struct *)
unfold make_field_access in EQ0. rewrite H1 in EQ0.
- destruct (prog_comp_env cunit)!id as [co'|] eqn:CO; monadInv EQ0.
+ destruct (prog_comp_env cunit)!id as [co'|] eqn:CO; try discriminate; monadInv EQ0.
exploit field_offset_stable. eexact LINK. eauto. instantiate (1 := i). intros [A B].
- rewrite <- B in EQ1.
+ rewrite <- B in EQ1.
assert (x0 = delta) by (unfold ge in *; simpl in *; congruence).
- subst x0.
+ assert (bf' = bf) by (unfold ge in *; simpl in *; congruence).
+ subst x0 bf'. split; auto.
destruct Archi.ptr64 eqn:SF.
+ eapply eval_Ebinop; eauto using make_longconst_correct.
simpl. rewrite SF. apply f_equal. apply f_equal. apply f_equal. auto with ptrofs.
+ eapply eval_Ebinop; eauto using make_intconst_correct.
simpl. rewrite SF. apply f_equal. apply f_equal. apply f_equal. auto with ptrofs.
- (* field union *)
- unfold make_field_access in EQ0; rewrite H1 in EQ0; monadInv EQ0.
- auto.
+ unfold make_field_access in EQ0. rewrite H1 in EQ0.
+ destruct (prog_comp_env cunit)!id as [co'|] eqn:CO; try discriminate; monadInv EQ0.
+ exploit union_field_offset_stable. eexact LINK. eauto. instantiate (1 := i). intros [A B].
+ rewrite <- B in EQ1.
+ assert (x0 = delta) by (unfold ge in *; simpl in *; congruence).
+ assert (bf' = bf) by (unfold ge in *; simpl in *; congruence).
+ subst x0 bf'. split; auto.
+ destruct Archi.ptr64 eqn:SF.
++ eapply eval_Ebinop; eauto using make_longconst_correct.
+ simpl. rewrite SF. apply f_equal. apply f_equal. apply f_equal. auto with ptrofs.
++ eapply eval_Ebinop; eauto using make_intconst_correct.
+ simpl. rewrite SF. apply f_equal. apply f_equal. apply f_equal. auto with ptrofs.
Qed.
Lemma transl_expr_correct:
@@ -1282,10 +1375,10 @@ Lemma transl_expr_correct:
Proof (proj1 transl_expr_lvalue_correct).
Lemma transl_lvalue_correct:
- forall a b ofs,
- Clight.eval_lvalue ge e le m a b ofs ->
- forall ta, transl_lvalue cunit.(prog_comp_env) a = OK ta ->
- Csharpminor.eval_expr tge te le m ta (Vptr b ofs).
+ forall a b ofs bf,
+ Clight.eval_lvalue ge e le m a b ofs bf ->
+ forall ta bf', transl_lvalue cunit.(prog_comp_env) a = OK (ta, bf') ->
+ bf = bf' /\ Csharpminor.eval_expr tge te le m ta (Vptr b ofs).
Proof (proj2 transl_expr_lvalue_correct).
Lemma transl_arglist_correct:
@@ -1468,7 +1561,11 @@ Proof.
auto.
- (* assign *)
unfold make_store, make_memcpy in EQ3.
+ destruct x0.
destruct (access_mode (typeof e)); monadInv EQ3; auto.
+ unfold make_store_bitfield in EQ3.
+ destruct (zle 0 pos && zlt 0 width && zle (pos + width) (bitsize_carrier sz));
+ monadInv EQ3; auto.
- (* set *)
auto.
- (* call *)
@@ -1568,11 +1665,17 @@ Proof.
assert (SAME: ts' = ts /\ tk' = tk).
{ inversion MTR. auto.
subst ts. unfold make_store, make_memcpy in EQ3.
- destruct (access_mode (typeof a1)); monadInv EQ3; auto. }
+ destruct x0.
+ destruct (access_mode (typeof a1)); monadInv EQ3; auto.
+ unfold make_store_bitfield in EQ3.
+ destruct (zle 0 pos && zlt 0 width && zle (pos + width) (bitsize_carrier sz));
+ monadInv EQ3; auto.
+ }
destruct SAME; subst ts' tk'.
+ exploit transl_lvalue_correct; eauto. intros [A B]; subst x0.
econstructor; split.
apply plus_one. eapply make_store_correct; eauto.
- eapply transl_lvalue_correct; eauto. eapply make_cast_correct; eauto.
+ eapply make_cast_correct; eauto.
eapply transl_expr_correct; eauto.
eapply match_states_skip; eauto.
diff --git a/cfrontend/Cstrategy.v b/cfrontend/Cstrategy.v
index 6365f85c..ce965672 100644
--- a/cfrontend/Cstrategy.v
+++ b/cfrontend/Cstrategy.v
@@ -49,7 +49,7 @@ Variable ge: genv.
Fixpoint simple (a: expr) : bool :=
match a with
- | Eloc _ _ _ => true
+ | Eloc _ _ _ _ => true
| Evar _ _ => true
| Ederef r _ => simple r
| Efield r _ _ => simple r
@@ -86,41 +86,42 @@ Section SIMPLE_EXPRS.
Variable e: env.
Variable m: mem.
-Inductive eval_simple_lvalue: expr -> block -> ptrofs -> Prop :=
- | esl_loc: forall b ofs ty,
- eval_simple_lvalue (Eloc b ofs ty) b ofs
+Inductive eval_simple_lvalue: expr -> block -> ptrofs -> bitfield -> Prop :=
+ | esl_loc: forall b ofs ty bf,
+ eval_simple_lvalue (Eloc b ofs bf ty) b ofs bf
| esl_var_local: forall x ty b,
e!x = Some(b, ty) ->
- eval_simple_lvalue (Evar x ty) b Ptrofs.zero
+ eval_simple_lvalue (Evar x ty) b Ptrofs.zero Full
| esl_var_global: forall x ty b,
e!x = None ->
Genv.find_symbol ge x = Some b ->
- eval_simple_lvalue (Evar x ty) b Ptrofs.zero
+ eval_simple_lvalue (Evar x ty) b Ptrofs.zero Full
| esl_deref: forall r ty b ofs,
eval_simple_rvalue r (Vptr b ofs) ->
- eval_simple_lvalue (Ederef r ty) b ofs
- | esl_field_struct: forall r f ty b ofs id co a delta,
+ eval_simple_lvalue (Ederef r ty) b ofs Full
+ | esl_field_struct: forall r f ty b ofs id co a delta bf,
eval_simple_rvalue r (Vptr b ofs) ->
typeof r = Tstruct id a ->
ge.(genv_cenv)!id = Some co ->
- field_offset ge f (co_members co) = OK delta ->
- eval_simple_lvalue (Efield r f ty) b (Ptrofs.add ofs (Ptrofs.repr delta))
- | esl_field_union: forall r f ty b ofs id co a,
+ field_offset ge f (co_members co) = OK (delta, bf) ->
+ eval_simple_lvalue (Efield r f ty) b (Ptrofs.add ofs (Ptrofs.repr delta)) bf
+ | esl_field_union: forall r f ty b ofs id co a delta bf,
eval_simple_rvalue r (Vptr b ofs) ->
typeof r = Tunion id a ->
+ union_field_offset ge f (co_members co) = OK (delta, bf) ->
ge.(genv_cenv)!id = Some co ->
- eval_simple_lvalue (Efield r f ty) b ofs
+ eval_simple_lvalue (Efield r f ty) b (Ptrofs.add ofs (Ptrofs.repr delta)) bf
with eval_simple_rvalue: expr -> val -> Prop :=
| esr_val: forall v ty,
eval_simple_rvalue (Eval v ty) v
- | esr_rvalof: forall b ofs l ty v,
- eval_simple_lvalue l b ofs ->
+ | esr_rvalof: forall b ofs bf l ty v,
+ eval_simple_lvalue l b ofs bf ->
ty = typeof l -> type_is_volatile ty = false ->
- deref_loc ge ty m b ofs E0 v ->
+ deref_loc ge ty m b ofs bf E0 v ->
eval_simple_rvalue (Evalof l ty) v
| esr_addrof: forall b ofs l ty,
- eval_simple_lvalue l b ofs ->
+ eval_simple_lvalue l b ofs Full ->
eval_simple_rvalue (Eaddrof l ty) (Vptr b ofs)
| esr_unop: forall op r1 ty v1 v,
eval_simple_rvalue r1 v1 ->
@@ -240,10 +241,10 @@ Inductive estep: state -> trace -> state -> Prop :=
estep (ExprState f r k e m)
E0 (ExprState f (Eval v ty) k e m)
- | step_rvalof_volatile: forall f C l ty k e m b ofs t v,
+ | step_rvalof_volatile: forall f C l ty k e m b ofs bf t v,
leftcontext RV RV C ->
- eval_simple_lvalue e m l b ofs ->
- deref_loc ge ty m b ofs t v ->
+ eval_simple_lvalue e m l b ofs bf ->
+ deref_loc ge ty m b ofs bf t v ->
ty = typeof l -> type_is_volatile ty = true ->
estep (ExprState f (C (Evalof l ty)) k e m)
t (ExprState f (C (Eval v ty)) k e m)
@@ -281,68 +282,68 @@ Inductive estep: state -> trace -> state -> Prop :=
estep (ExprState f (C (Econdition r1 r2 r3 ty)) k e m)
E0 (ExprState f (C (Eparen (if b then r2 else r3) ty ty)) k e m)
- | step_assign: forall f C l r ty k e m b ofs v v' t m',
+ | step_assign: forall f C l r ty k e m b ofs bf v v1 t m' v',
leftcontext RV RV C ->
- eval_simple_lvalue e m l b ofs ->
+ eval_simple_lvalue e m l b ofs bf ->
eval_simple_rvalue e m r v ->
- sem_cast v (typeof r) (typeof l) m = Some v' ->
- assign_loc ge (typeof l) m b ofs v' t m' ->
+ sem_cast v (typeof r) (typeof l) m = Some v1 ->
+ assign_loc ge (typeof l) m b ofs bf v1 t m' v' ->
ty = typeof l ->
estep (ExprState f (C (Eassign l r ty)) k e m)
t (ExprState f (C (Eval v' ty)) k e m')
- | step_assignop: forall f C op l r tyres ty k e m b ofs v1 v2 v3 v4 t1 t2 m' t,
+ | step_assignop: forall f C op l r tyres ty k e m b ofs bf v1 v2 v3 v4 t1 t2 m' v' t,
leftcontext RV RV C ->
- eval_simple_lvalue e m l b ofs ->
- deref_loc ge (typeof l) m b ofs t1 v1 ->
+ eval_simple_lvalue e m l b ofs bf ->
+ deref_loc ge (typeof l) m b ofs bf t1 v1 ->
eval_simple_rvalue e m r v2 ->
sem_binary_operation ge op v1 (typeof l) v2 (typeof r) m = Some v3 ->
sem_cast v3 tyres (typeof l) m = Some v4 ->
- assign_loc ge (typeof l) m b ofs v4 t2 m' ->
+ assign_loc ge (typeof l) m b ofs bf v4 t2 m' v' ->
ty = typeof l ->
t = t1 ** t2 ->
estep (ExprState f (C (Eassignop op l r tyres ty)) k e m)
- t (ExprState f (C (Eval v4 ty)) k e m')
+ t (ExprState f (C (Eval v' ty)) k e m')
- | step_assignop_stuck: forall f C op l r tyres ty k e m b ofs v1 v2 t,
+ | step_assignop_stuck: forall f C op l r tyres ty k e m b ofs bf v1 v2 t,
leftcontext RV RV C ->
- eval_simple_lvalue e m l b ofs ->
- deref_loc ge (typeof l) m b ofs t v1 ->
+ eval_simple_lvalue e m l b ofs bf ->
+ deref_loc ge (typeof l) m b ofs bf t v1 ->
eval_simple_rvalue e m r v2 ->
match sem_binary_operation ge op v1 (typeof l) v2 (typeof r) m with
| None => True
| Some v3 =>
match sem_cast v3 tyres (typeof l) m with
| None => True
- | Some v4 => forall t2 m', ~(assign_loc ge (typeof l) m b ofs v4 t2 m')
+ | Some v4 => forall t2 m' v', ~(assign_loc ge (typeof l) m b ofs bf v4 t2 m' v')
end
end ->
ty = typeof l ->
estep (ExprState f (C (Eassignop op l r tyres ty)) k e m)
t Stuckstate
- | step_postincr: forall f C id l ty k e m b ofs v1 v2 v3 t1 t2 m' t,
+ | step_postincr: forall f C id l ty k e m b ofs bf v1 v2 v3 t1 t2 m' v' t,
leftcontext RV RV C ->
- eval_simple_lvalue e m l b ofs ->
- deref_loc ge ty m b ofs t1 v1 ->
+ eval_simple_lvalue e m l b ofs bf ->
+ deref_loc ge ty m b ofs bf t1 v1 ->
sem_incrdecr ge id v1 ty m = Some v2 ->
sem_cast v2 (incrdecr_type ty) ty m = Some v3 ->
- assign_loc ge ty m b ofs v3 t2 m' ->
+ assign_loc ge ty m b ofs bf v3 t2 m' v' ->
ty = typeof l ->
t = t1 ** t2 ->
estep (ExprState f (C (Epostincr id l ty)) k e m)
t (ExprState f (C (Eval v1 ty)) k e m')
- | step_postincr_stuck: forall f C id l ty k e m b ofs v1 t,
+ | step_postincr_stuck: forall f C id l ty k e m b ofs bf v1 t,
leftcontext RV RV C ->
- eval_simple_lvalue e m l b ofs ->
- deref_loc ge ty m b ofs t v1 ->
+ eval_simple_lvalue e m l b ofs bf ->
+ deref_loc ge ty m b ofs bf t v1 ->
match sem_incrdecr ge id v1 ty m with
| None => True
| Some v2 =>
match sem_cast v2 (incrdecr_type ty) ty m with
| None => True
- | Some v3 => forall t2 m', ~(assign_loc ge (typeof l) m b ofs v3 t2 m')
+ | Some v3 => forall t2 m' v', ~(assign_loc ge (typeof l) m b ofs bf v3 t2 m' v')
end
end ->
ty = typeof l ->
@@ -452,7 +453,7 @@ Qed.
Definition expr_kind (a: expr) : kind :=
match a with
- | Eloc _ _ _ => LV
+ | Eloc _ _ _ _ => LV
| Evar _ _ => LV
| Ederef _ _ => LV
| Efield _ _ _ => LV
@@ -518,23 +519,25 @@ Fixpoint exprlist_all_values (rl: exprlist) : Prop :=
Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
match a with
- | Eloc b ofs ty => False
+ | Eloc b ofs bf ty => False
| Evar x ty =>
exists b,
e!x = Some(b, ty)
\/ (e!x = None /\ Genv.find_symbol ge x = Some b)
| Ederef (Eval v ty1) ty =>
exists b, exists ofs, v = Vptr b ofs
+ | Eaddrof (Eloc b ofs bf ty) ty' =>
+ bf = Full
| Efield (Eval v ty1) f ty =>
exists b, exists ofs, v = Vptr b ofs /\
match ty1 with
- | Tstruct id _ => exists co delta, ge.(genv_cenv)!id = Some co /\ field_offset ge f (co_members co) = Errors.OK delta
- | Tunion id _ => exists co, ge.(genv_cenv)!id = Some co
+ | Tstruct id _ => exists co delta bf, ge.(genv_cenv)!id = Some co /\ field_offset ge f (co_members co) = Errors.OK (delta, bf)
+ | Tunion id _ => exists co delta bf, ge.(genv_cenv)!id = Some co /\ union_field_offset ge f (co_members co) = Errors.OK (delta, bf)
| _ => False
end
| Eval v ty => False
- | Evalof (Eloc b ofs ty') ty =>
- ty' = ty /\ exists t, exists v, deref_loc ge ty m b ofs t v
+ | Evalof (Eloc b ofs bf ty') ty =>
+ ty' = ty /\ exists t, exists v, deref_loc ge ty m b ofs bf t v
| Eunop op (Eval v1 ty1) ty =>
exists v, sem_unary_operation op v1 ty1 m = Some v
| Ebinop op (Eval v1 ty1) (Eval v2 ty2) ty =>
@@ -547,17 +550,17 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop :=
exists b, bool_val v1 ty1 m = Some b
| Econdition (Eval v1 ty1) r1 r2 ty =>
exists b, bool_val v1 ty1 m = Some b
- | Eassign (Eloc b ofs ty1) (Eval v2 ty2) ty =>
- exists v, exists m', exists t,
- ty = ty1 /\ sem_cast v2 ty2 ty1 m = Some v /\ assign_loc ge ty1 m b ofs v t m'
- | Eassignop op (Eloc b ofs ty1) (Eval v2 ty2) tyres ty =>
- exists t, exists v1,
+ | Eassign (Eloc b ofs bf ty1) (Eval v2 ty2) ty =>
+ exists v m' v' t,
+ ty = ty1 /\ sem_cast v2 ty2 ty1 m = Some v /\ assign_loc ge ty1 m b ofs bf v t m' v'
+ | Eassignop op (Eloc b ofs bf ty1) (Eval v2 ty2) tyres ty =>
+ exists t v1,
ty = ty1
- /\ deref_loc ge ty1 m b ofs t v1
- | Epostincr id (Eloc b ofs ty1) ty =>
- exists t, exists v1,
+ /\ deref_loc ge ty1 m b ofs bf t v1
+ | Epostincr id (Eloc b ofs bf ty1) ty =>
+ exists t v1,
ty = ty1
- /\ deref_loc ge ty m b ofs t v1
+ /\ deref_loc ge ty m b ofs bf t v1
| Ecomma (Eval v ty1) r2 ty =>
typeof r2 = ty
| Eparen (Eval v1 ty1) ty2 ty =>
@@ -584,8 +587,8 @@ Proof.
exists b; auto.
exists b; auto.
exists b; exists ofs; auto.
- exists b; exists ofs; split; auto. exists co, delta; auto.
- exists b; exists ofs; split; auto. exists co; auto.
+ exists b; exists ofs; split; auto. exists co, delta, bf; auto.
+ exists b; exists ofs; split; auto. exists co, delta, bf; auto.
Qed.
Lemma rred_invert:
@@ -599,7 +602,7 @@ Proof.
exists true; auto. exists false; auto.
exists true; auto. exists false; auto.
exists b; auto.
- exists v; exists m'; exists t; auto.
+ exists v; exists m'; exists v'; exists t; auto.
exists t; exists v1; auto.
exists t; exists v1; auto.
exists v; auto.
@@ -634,7 +637,7 @@ Proof.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
- auto.
+ destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
destruct (C a); auto; contradiction.
destruct e1; auto; destruct (C a); auto; contradiction.
@@ -660,7 +663,7 @@ Lemma imm_safe_inv:
forall k a m,
imm_safe ge e k a m ->
match a with
- | Eloc _ _ _ => True
+ | Eloc _ _ _ _ => True
| Eval _ _ => True
| _ => invert_expr_prop a m
end.
@@ -685,7 +688,7 @@ Lemma safe_inv:
safe (ExprState f (C a) K e m) ->
context k RV C ->
match a with
- | Eloc _ _ _ => True
+ | Eloc _ _ _ _ => True
| Eval _ _ => True
| _ => invert_expr_prop a m
end.
@@ -709,10 +712,10 @@ Lemma eval_simple_steps:
forall C, context RV RV C ->
star Csem.step ge (ExprState f (C a) k e m)
E0 (ExprState f (C (Eval v (typeof a))) k e m))
-/\ (forall a b ofs, eval_simple_lvalue e m a b ofs ->
+/\ (forall a b ofs bf, eval_simple_lvalue e m a b ofs bf ->
forall C, context LV RV C ->
star Csem.step ge (ExprState f (C a) k e m)
- E0 (ExprState f (C (Eloc b ofs (typeof a))) k e m)).
+ E0 (ExprState f (C (Eloc b ofs bf (typeof a))) k e m)).
Proof.
Ltac Steps REC C' := eapply star_trans; [apply (REC C'); eauto | idtac | simpl; reflexivity].
@@ -760,10 +763,10 @@ Lemma eval_simple_rvalue_steps:
Proof (proj1 eval_simple_steps).
Lemma eval_simple_lvalue_steps:
- forall a b ofs, eval_simple_lvalue e m a b ofs ->
+ forall a b ofs bf, eval_simple_lvalue e m a b ofs bf ->
forall C, context LV RV C ->
star Csem.step ge (ExprState f (C a) k e m)
- E0 (ExprState f (C (Eloc b ofs (typeof a))) k e m).
+ E0 (ExprState f (C (Eloc b ofs bf (typeof a))) k e m).
Proof (proj2 eval_simple_steps).
Corollary eval_simple_rvalue_safe:
@@ -776,10 +779,10 @@ Proof.
Qed.
Corollary eval_simple_lvalue_safe:
- forall C a b ofs,
- eval_simple_lvalue e m a b ofs ->
+ forall C a b ofs bf,
+ eval_simple_lvalue e m a b ofs bf ->
context LV RV C -> safe (ExprState f (C a) k e m) ->
- safe (ExprState f (C (Eloc b ofs (typeof a))) k e m).
+ safe (ExprState f (C (Eloc b ofs bf (typeof a))) k e m).
Proof.
intros. eapply safe_steps; eauto. eapply eval_simple_lvalue_steps; eauto.
Qed.
@@ -788,15 +791,15 @@ Lemma simple_can_eval:
forall a from C,
simple a = true -> context from RV C -> safe (ExprState f (C a) k e m) ->
match from with
- | LV => exists b, exists ofs, eval_simple_lvalue e m a b ofs
+ | LV => exists b ofs bf, eval_simple_lvalue e m a b ofs bf
| RV => exists v, eval_simple_rvalue e m a v
end.
Proof.
Ltac StepL REC C' a :=
- let b := fresh "b" in let ofs := fresh "ofs" in
+ let b := fresh "b" in let ofs := fresh "ofs" in let bf := fresh "bf" in
let E := fresh "E" in let S := fresh "SAFE" in
- exploit (REC LV C'); eauto; intros [b [ofs E]];
- assert (S: safe (ExprState f (C' (Eloc b ofs (typeof a))) k e m)) by
+ exploit (REC LV C'); eauto; intros (b & ofs & bf & E);
+ assert (S: safe (ExprState f (C' (Eloc b ofs bf (typeof a))) k e m)) by
(eapply (eval_simple_lvalue_safe C'); eauto);
simpl in S.
Ltac StepR REC C' a :=
@@ -809,51 +812,52 @@ Ltac StepR REC C' a :=
induction a; intros from C S CTX SAFE;
generalize (safe_expr_kind _ _ _ _ _ _ _ CTX SAFE); intro K; subst;
simpl in S; try discriminate; simpl.
-(* val *)
+- (* val *)
exists v; constructor.
-(* var *)
+- (* var *)
exploit safe_inv; eauto; simpl. intros [b A].
- exists b; exists Ptrofs.zero.
+ exists b, Ptrofs.zero, Full.
intuition. apply esl_var_local; auto. apply esl_var_global; auto.
-(* field *)
+- (* field *)
StepR IHa (fun x => C(Efield x f0 ty)) a.
exploit safe_inv. eexact SAFE0. eauto. simpl.
intros [b [ofs [EQ TY]]]. subst v. destruct (typeof a) eqn:?; try contradiction.
- destruct TY as (co & delta & CE & OFS). exists b; exists (Ptrofs.add ofs (Ptrofs.repr delta)); econstructor; eauto.
- destruct TY as (co & CE). exists b; exists ofs; econstructor; eauto.
-(* valof *)
+ destruct TY as (co & delta & bf & CE & OFS). exists b, (Ptrofs.add ofs (Ptrofs.repr delta)), bf; eapply esl_field_struct; eauto.
+ destruct TY as (co & delta & bf & CE & OFS). exists b, (Ptrofs.add ofs (Ptrofs.repr delta)), bf; eapply esl_field_union; eauto.
+- (* valof *)
destruct (andb_prop _ _ S) as [S1 S2]. clear S. rewrite negb_true_iff in S2.
StepL IHa (fun x => C(Evalof x ty)) a.
exploit safe_inv. eexact SAFE0. eauto. simpl. intros [TY [t [v LOAD]]].
assert (t = E0). inv LOAD; auto. congruence. subst t.
exists v; econstructor; eauto. congruence.
-(* deref *)
+- (* deref *)
StepR IHa (fun x => C(Ederef x ty)) a.
exploit safe_inv. eexact SAFE0. eauto. simpl. intros [b [ofs EQ]].
- subst v. exists b; exists ofs; econstructor; eauto.
-(* addrof *)
+ subst v. exists b, ofs, Full; econstructor; eauto.
+- (* addrof *)
StepL IHa (fun x => C(Eaddrof x ty)) a.
+ exploit safe_inv. eexact SAFE0. eauto. simpl. intros EQ; subst bf.
exists (Vptr b ofs); econstructor; eauto.
-(* unop *)
+- (* unop *)
StepR IHa (fun x => C(Eunop op x ty)) a.
exploit safe_inv. eexact SAFE0. eauto. simpl. intros [v' EQ].
exists v'; econstructor; eauto.
-(* binop *)
+- (* binop *)
destruct (andb_prop _ _ S) as [S1 S2]; clear S.
StepR IHa1 (fun x => C(Ebinop op x a2 ty)) a1.
StepR IHa2 (fun x => C(Ebinop op (Eval v (typeof a1)) x ty)) a2.
exploit safe_inv. eexact SAFE1. eauto. simpl. intros [v' EQ].
exists v'; econstructor; eauto.
-(* cast *)
+- (* cast *)
StepR IHa (fun x => C(Ecast x ty)) a.
exploit safe_inv. eexact SAFE0. eauto. simpl. intros [v' CAST].
exists v'; econstructor; eauto.
-(* sizeof *)
+- (* sizeof *)
econstructor; econstructor.
-(* alignof *)
+- (* alignof *)
econstructor; econstructor.
-(* loc *)
- exists b; exists ofs; constructor.
+- (* loc *)
+ exists b, ofs, bf; constructor.
Qed.
Lemma simple_can_eval_rval:
@@ -869,11 +873,11 @@ Qed.
Lemma simple_can_eval_lval:
forall l C,
simple l = true -> context LV RV C -> safe (ExprState f (C l) k e m) ->
- exists b, exists ofs, eval_simple_lvalue e m l b ofs
- /\ safe (ExprState f (C (Eloc b ofs (typeof l))) k e m).
+ exists b ofs bf, eval_simple_lvalue e m l b ofs bf
+ /\ safe (ExprState f (C (Eloc b ofs bf (typeof l))) k e m).
Proof.
- intros. exploit (simple_can_eval l LV); eauto. intros [b [ofs A]].
- exists b; exists ofs; split; auto. eapply eval_simple_lvalue_safe; eauto.
+ intros. exploit (simple_can_eval l LV); eauto. intros (b & ofs & bf & A).
+ exists b, ofs, bf; split; auto. eapply eval_simple_lvalue_safe; eauto.
Qed.
Fixpoint rval_list (vl: list val) (rl: exprlist) : exprlist :=
@@ -1178,18 +1182,18 @@ Proof.
eapply star_plus_trans.
eapply eval_simple_lvalue_steps with (C := fun x => C(Eassign x r (typeof l))); eauto.
eapply plus_right.
- eapply eval_simple_rvalue_steps with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eassign (Eloc b ofs bf (typeof l)) x (typeof l))); eauto.
left; apply step_rred; eauto. econstructor; eauto.
reflexivity. auto.
(* assignop *)
eapply star_plus_trans.
eapply eval_simple_lvalue_steps with (C := fun x => C(Eassignop op x r tyres (typeof l))); eauto.
eapply star_plus_trans.
- eapply eval_simple_rvalue_steps with (C := fun x => C(Eassignop op (Eloc b ofs (typeof l)) x tyres (typeof l))); eauto.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eassignop op (Eloc b ofs bf (typeof l)) x tyres (typeof l))); eauto.
eapply plus_left.
left; apply step_rred; auto. econstructor; eauto.
eapply star_left.
- left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto. econstructor; eauto.
+ left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs bf (typeof l)) x (typeof l))); eauto. econstructor; eauto.
apply star_one.
left; apply step_rred; auto. econstructor; eauto.
reflexivity. reflexivity. reflexivity. traceEq.
@@ -1197,19 +1201,19 @@ Proof.
eapply star_plus_trans.
eapply eval_simple_lvalue_steps with (C := fun x => C(Eassignop op x r tyres (typeof l))); eauto.
eapply star_plus_trans.
- eapply eval_simple_rvalue_steps with (C := fun x => C(Eassignop op (Eloc b ofs (typeof l)) x tyres (typeof l))); eauto.
+ eapply eval_simple_rvalue_steps with (C := fun x => C(Eassignop op (Eloc b ofs bf (typeof l)) x tyres (typeof l))); eauto.
eapply plus_left.
left; apply step_rred; auto. econstructor; eauto.
destruct (sem_binary_operation ge op v1 (typeof l) v2 (typeof r) m) as [v3|] eqn:?.
eapply star_left.
- left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto. econstructor; eauto.
+ left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs bf (typeof l)) x (typeof l))); eauto. econstructor; eauto.
apply star_one.
left; eapply step_stuck; eauto.
- red; intros. exploit imm_safe_inv; eauto. simpl. intros [v4' [m' [t' [A [B D]]]]].
+ red; intros. exploit imm_safe_inv; eauto. simpl. intros [v4' [m' [v' [t' [A [B D]]]]]].
rewrite B in H4. eelim H4; eauto.
reflexivity.
apply star_one.
- left; eapply step_stuck with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto.
+ left; eapply step_stuck with (C := fun x => C(Eassign (Eloc b ofs bf (typeof l)) x (typeof l))); eauto.
red; intros. exploit imm_safe_inv; eauto. simpl. intros [v3 A]. congruence.
reflexivity.
reflexivity. traceEq.
@@ -1219,7 +1223,7 @@ Proof.
eapply plus_left.
left; apply step_rred; auto. econstructor; eauto.
eapply star_left.
- left; apply step_rred with (C := fun x => C (Ecomma (Eassign (Eloc b ofs (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto.
+ left; apply step_rred with (C := fun x => C (Ecomma (Eassign (Eloc b ofs bf (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto.
econstructor. instantiate (1 := v2). destruct id; assumption.
eapply star_left.
left; apply step_rred with (C := fun x => C (Ecomma x (Eval v1 (typeof l)) (typeof l))); eauto.
@@ -1238,15 +1242,15 @@ Proof.
destruct id; auto.
destruct (sem_incrdecr ge id v1 (typeof l) m) as [v2|].
eapply star_left.
- left; apply step_rred with (C := fun x => C (Ecomma (Eassign (Eloc b ofs (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto.
+ left; apply step_rred with (C := fun x => C (Ecomma (Eassign (Eloc b ofs bf (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto.
econstructor; eauto.
apply star_one.
left; eapply step_stuck with (C := fun x => C (Ecomma x (Eval v1 (typeof l)) (typeof l))); eauto.
- red; intros. exploit imm_safe_inv; eauto. simpl. intros [v3 [m' [t' [A [B D]]]]].
+ red; intros. exploit imm_safe_inv; eauto. simpl. intros [v3 [m' [v' [t' [A [B D]]]]]].
rewrite B in H3. eelim H3; eauto.
reflexivity.
apply star_one.
- left; eapply step_stuck with (C := fun x => C (Ecomma (Eassign (Eloc b ofs (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto.
+ left; eapply step_stuck with (C := fun x => C (Ecomma (Eassign (Eloc b ofs bf (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto.
red; intros. exploit imm_safe_inv; eauto. simpl. intros [v2 A]. congruence.
reflexivity.
traceEq.
@@ -1291,7 +1295,7 @@ Proof.
(* valof volatile *)
destruct Q.
exploit (simple_can_eval_lval f k e m b (fun x => C(Evalof x ty))); eauto.
- intros [b1 [ofs [E1 S1]]].
+ intros (b1 & ofs & bf & E1 & S1).
exploit safe_inv. eexact S1. eauto. simpl. intros [A [t [v B]]].
econstructor; econstructor; eapply step_rvalof_volatile; eauto. congruence.
(* seqand *)
@@ -1316,40 +1320,40 @@ Proof.
(* assign *)
destruct Q.
exploit (simple_can_eval_lval f k e m b1 (fun x => C(Eassign x b2 ty))); eauto.
- intros [b [ofs [E1 S1]]].
- exploit (simple_can_eval_rval f k e m b2 (fun x => C(Eassign (Eloc b ofs (typeof b1)) x ty))); eauto.
+ intros (b & ofs & bf & E1 & S1).
+ exploit (simple_can_eval_rval f k e m b2 (fun x => C(Eassign (Eloc b ofs bf (typeof b1)) x ty))); eauto.
intros [v [E2 S2]].
- exploit safe_inv. eexact S2. eauto. simpl. intros [v' [m' [t [A [B D]]]]].
+ exploit safe_inv. eexact S2. eauto. simpl. intros [v1 [m' [v' [t [A [B D]]]]]].
econstructor; econstructor; eapply step_assign; eauto.
(* assignop *)
destruct Q.
exploit (simple_can_eval_lval f k e m b1 (fun x => C(Eassignop op x b2 tyres ty))); eauto.
- intros [b [ofs [E1 S1]]].
- exploit (simple_can_eval_rval f k e m b2 (fun x => C(Eassignop op (Eloc b ofs (typeof b1)) x tyres ty))); eauto.
+ intros (b & ofs & bf & E1 & S1).
+ exploit (simple_can_eval_rval f k e m b2 (fun x => C(Eassignop op (Eloc b ofs bf (typeof b1)) x tyres ty))); eauto.
intros [v [E2 S2]].
exploit safe_inv. eexact S2. eauto. simpl. intros [t1 [v1 [A B]]].
destruct (sem_binary_operation ge op v1 (typeof b1) v (typeof b2) m) as [v3|] eqn:?.
destruct (sem_cast v3 tyres (typeof b1) m) as [v4|] eqn:?.
- destruct (classic (exists t2, exists m', assign_loc ge (typeof b1) m b ofs v4 t2 m')).
- destruct H2 as [t2 [m' D]].
+ destruct (classic (exists t2 m' v', assign_loc ge (typeof b1) m b ofs bf v4 t2 m' v')).
+ destruct H2 as [t2 [m' [v' D]]].
econstructor; econstructor; eapply step_assignop; eauto.
econstructor; econstructor; eapply step_assignop_stuck; eauto.
- rewrite Heqo. rewrite Heqo0. intros; red; intros. elim H2. exists t2; exists m'; auto.
+ rewrite Heqo. rewrite Heqo0. intros; red; intros. elim H2. exists t2, m', v'; auto.
econstructor; econstructor; eapply step_assignop_stuck; eauto.
rewrite Heqo. rewrite Heqo0. auto.
econstructor; econstructor; eapply step_assignop_stuck; eauto.
rewrite Heqo. auto.
(* postincr *)
exploit (simple_can_eval_lval f k e m b (fun x => C(Epostincr id x ty))); eauto.
- intros [b1 [ofs [E1 S1]]].
+ intros (b1 & ofs & bf & E1 & S1).
exploit safe_inv. eexact S1. eauto. simpl. intros [t [v1 [A B]]].
destruct (sem_incrdecr ge id v1 ty m) as [v2|] eqn:?.
destruct (sem_cast v2 (incrdecr_type ty) ty m) as [v3|] eqn:?.
- destruct (classic (exists t2, exists m', assign_loc ge ty m b1 ofs v3 t2 m')).
- destruct H0 as [t2 [m' D]].
+ destruct (classic (exists t2 m' v', assign_loc ge ty m b1 ofs bf v3 t2 m' v')).
+ destruct H0 as [t2 [m' [v' D]]].
econstructor; econstructor; eapply step_postincr; eauto.
econstructor; econstructor; eapply step_postincr_stuck; eauto.
- rewrite Heqo. rewrite Heqo0. intros; red; intros. elim H0. exists t2; exists m'; congruence.
+ rewrite Heqo. rewrite Heqo0. intros; red; intros. elim H0. exists t2, m', v'; congruence.
econstructor; econstructor; eapply step_postincr_stuck; eauto.
rewrite Heqo. rewrite Heqo0. auto.
econstructor; econstructor; eapply step_postincr_stuck; eauto.
@@ -1440,18 +1444,18 @@ Definition semantics (p: program) :=
(** This semantics is receptive to changes in events. *)
Remark deref_loc_trace:
- forall ge ty m b ofs t v,
- deref_loc ge ty m b ofs t v ->
+ forall ge ty m b ofs bf t v,
+ deref_loc ge ty m b ofs bf t v ->
match t with nil => True | ev :: nil => True | _ => False end.
Proof.
intros. inv H; simpl; auto. inv H2; simpl; auto.
Qed.
Remark deref_loc_receptive:
- forall ge ty m b ofs ev1 t1 v ev2,
- deref_loc ge ty m b ofs (ev1 :: t1) v ->
+ forall ge ty m b ofs bf ev1 t1 v ev2,
+ deref_loc ge ty m b ofs bf (ev1 :: t1) v ->
match_traces ge (ev1 :: nil) (ev2 :: nil) ->
- t1 = nil /\ exists v', deref_loc ge ty m b ofs (ev2 :: nil) v'.
+ t1 = nil /\ exists v', deref_loc ge ty m b ofs bf (ev2 :: nil) v'.
Proof.
intros.
assert (t1 = nil). exploit deref_loc_trace; eauto. destruct t1; simpl; tauto.
@@ -1460,16 +1464,16 @@ Proof.
Qed.
Remark assign_loc_trace:
- forall ge ty m b ofs t v m',
- assign_loc ge ty m b ofs v t m' ->
+ forall ge ty m b ofs bf t v m' v',
+ assign_loc ge ty m b ofs bf v t m' v' ->
match t with nil => True | ev :: nil => output_event ev | _ => False end.
Proof.
intros. inv H; simpl; auto. inv H2; simpl; auto.
Qed.
Remark assign_loc_receptive:
- forall ge ty m b ofs ev1 t1 v m' ev2,
- assign_loc ge ty m b ofs v (ev1 :: t1) m' ->
+ forall ge ty m b ofs bf ev1 t1 v m' v' ev2,
+ assign_loc ge ty m b ofs bf v (ev1 :: t1) m' v' ->
match_traces ge (ev1 :: nil) (ev2 :: nil) ->
ev1 :: t1 = ev2 :: nil.
Proof.
@@ -1499,11 +1503,11 @@ Proof.
inv H10. exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t0.
destruct (sem_binary_operation ge op v1' (typeof l) v2 (typeof r) m) as [v3'|] eqn:?.
destruct (sem_cast v3' tyres (typeof l) m) as [v4'|] eqn:?.
- destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v4' t2' m'')).
- destruct H1 as [t2' [m'' P]].
+ destruct (classic (exists t2' m'' v'', assign_loc ge (typeof l) m b ofs bf v4' t2' m'' v'')).
+ destruct H1 as [t2' [m'' [v'' P]]].
econstructor; econstructor. left; eapply step_assignop with (v1 := v1'); eauto. simpl; reflexivity.
econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
- rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t0; exists m'0; auto.
+ rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t0, m'0, v'0; auto.
econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0; auto.
econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
@@ -1512,11 +1516,11 @@ Proof.
exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t1.
destruct (sem_binary_operation ge op v1' (typeof l) v2 (typeof r) m) as [v3'|] eqn:?.
destruct (sem_cast v3' tyres (typeof l) m) as [v4'|] eqn:?.
- destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v4' t2' m'')).
- destruct H1 as [t2' [m'' P]].
+ destruct (classic (exists t2' m'' v'', assign_loc ge (typeof l) m b ofs bf v4' t2' m'' v'')).
+ destruct H1 as [t2' [m'' [v'' P]]].
econstructor; econstructor. left; eapply step_assignop with (v1 := v1'); eauto. simpl; reflexivity.
econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
- rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t2; exists m'; auto.
+ rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t2, m', v'; auto.
econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0; auto.
econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto.
@@ -1528,11 +1532,11 @@ Proof.
inv H9. exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t0.
destruct (sem_incrdecr ge id v1' (typeof l) m) as [v2'|] eqn:?.
destruct (sem_cast v2' (incrdecr_type (typeof l)) (typeof l) m) as [v3'|] eqn:?.
- destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v3' t2' m'')).
- destruct H1 as [t2' [m'' P]].
+ destruct (classic (exists t2' m'' v'', assign_loc ge (typeof l) m b ofs bf v3' t2' m'' v'')).
+ destruct H1 as [t2' [m'' [v'' P]]].
econstructor; econstructor. left; eapply step_postincr with (v1 := v1'); eauto. simpl; reflexivity.
econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
- rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t0; exists m'0; auto.
+ rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t0, m'0, v'0; auto.
econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0; auto.
econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
@@ -1541,11 +1545,11 @@ Proof.
exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t1.
destruct (sem_incrdecr ge id v1' (typeof l) m) as [v2'|] eqn:?.
destruct (sem_cast v2' (incrdecr_type (typeof l)) (typeof l) m) as [v3'|] eqn:?.
- destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v3' t2' m'')).
- destruct H1 as [t2' [m'' P]].
+ destruct (classic (exists t2' m'' v'', assign_loc ge (typeof l) m b ofs bf v3' t2' m'' v'')).
+ destruct H1 as [t2' [m'' [v'' P]]].
econstructor; econstructor. left; eapply step_postincr with (v1 := v1'); eauto. simpl; reflexivity.
econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
- rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t2; exists m'; auto.
+ rewrite Heqo; rewrite Heqo0. intros; red; intros; elim H1. exists t2, m', v'; auto.
econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
rewrite Heqo; rewrite Heqo0; auto.
econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto.
@@ -1671,11 +1675,11 @@ with eval_expr: env -> mem -> kind -> expr -> trace -> mem -> expr -> Prop :=
type_is_volatile (typeof a) = false ->
eval_expr e m LV a t m' a' ->
eval_expr e m RV (Evalof a ty) t m' (Evalof a' ty)
- | eval_valof_volatile: forall e m a t1 m' a' ty b ofs t2 v,
+ | eval_valof_volatile: forall e m a t1 m' a' ty b ofs bf t2 v,
type_is_volatile (typeof a) = true ->
eval_expr e m LV a t1 m' a' ->
- eval_simple_lvalue ge e m' a' b ofs ->
- deref_loc ge (typeof a) m' b ofs t2 v ->
+ eval_simple_lvalue ge e m' a' b ofs bf ->
+ deref_loc ge (typeof a) m' b ofs bf t2 v ->
ty = typeof a ->
eval_expr e m RV (Evalof a ty) (t1 ** t2) m' (Eval v ty)
| eval_deref: forall e m a t m' a' ty,
@@ -1723,32 +1727,32 @@ with eval_expr: env -> mem -> kind -> expr -> trace -> mem -> expr -> Prop :=
eval_expr e m RV (Esizeof ty' ty) E0 m (Esizeof ty' ty)
| eval_alignof: forall e m ty' ty,
eval_expr e m RV (Ealignof ty' ty) E0 m (Ealignof ty' ty)
- | eval_assign: forall e m l r ty t1 m1 l' t2 m2 r' b ofs v v' t3 m3,
+ | eval_assign: forall e m l r ty t1 m1 l' t2 m2 r' b ofs bf v v1 v' t3 m3,
eval_expr e m LV l t1 m1 l' -> eval_expr e m1 RV r t2 m2 r' ->
- eval_simple_lvalue ge e m2 l' b ofs ->
+ eval_simple_lvalue ge e m2 l' b ofs bf ->
eval_simple_rvalue ge e m2 r' v ->
- sem_cast v (typeof r) (typeof l) m2 = Some v' ->
- assign_loc ge (typeof l) m2 b ofs v' t3 m3 ->
+ sem_cast v (typeof r) (typeof l) m2 = Some v1 ->
+ assign_loc ge (typeof l) m2 b ofs bf v1 t3 m3 v' ->
ty = typeof l ->
eval_expr e m RV (Eassign l r ty) (t1**t2**t3) m3 (Eval v' ty)
- | eval_assignop: forall e m op l r tyres ty t1 m1 l' t2 m2 r' b ofs
- v1 v2 v3 v4 t3 t4 m3,
+ | eval_assignop: forall e m op l r tyres ty t1 m1 l' t2 m2 r' b ofs bf
+ v1 v2 v3 v4 v' t3 t4 m3,
eval_expr e m LV l t1 m1 l' -> eval_expr e m1 RV r t2 m2 r' ->
- eval_simple_lvalue ge e m2 l' b ofs ->
- deref_loc ge (typeof l) m2 b ofs t3 v1 ->
+ eval_simple_lvalue ge e m2 l' b ofs bf ->
+ deref_loc ge (typeof l) m2 b ofs bf t3 v1 ->
eval_simple_rvalue ge e m2 r' v2 ->
sem_binary_operation ge op v1 (typeof l) v2 (typeof r) m2 = Some v3 ->
sem_cast v3 tyres (typeof l) m2 = Some v4 ->
- assign_loc ge (typeof l) m2 b ofs v4 t4 m3 ->
+ assign_loc ge (typeof l) m2 b ofs bf v4 t4 m3 v' ->
ty = typeof l ->
- eval_expr e m RV (Eassignop op l r tyres ty) (t1**t2**t3**t4) m3 (Eval v4 ty)
- | eval_postincr: forall e m id l ty t1 m1 l' b ofs v1 v2 v3 m2 t2 t3,
+ eval_expr e m RV (Eassignop op l r tyres ty) (t1**t2**t3**t4) m3 (Eval v' ty)
+ | eval_postincr: forall e m id l ty t1 m1 l' b ofs bf v1 v2 v3 v' m2 t2 t3,
eval_expr e m LV l t1 m1 l' ->
- eval_simple_lvalue ge e m1 l' b ofs ->
- deref_loc ge ty m1 b ofs t2 v1 ->
+ eval_simple_lvalue ge e m1 l' b ofs bf ->
+ deref_loc ge ty m1 b ofs bf t2 v1 ->
sem_incrdecr ge id v1 ty m1 = Some v2 ->
sem_cast v2 (incrdecr_type ty) ty m1 = Some v3 ->
- assign_loc ge ty m1 b ofs v3 t3 m2 ->
+ assign_loc ge ty m1 b ofs bf v3 t3 m2 v' ->
ty = typeof l ->
eval_expr e m RV (Epostincr id l ty) (t1**t2**t3) m2 (Eval v1 ty)
| eval_comma: forall e m r1 r2 ty t1 m1 r1' v1 t2 m2 r2',
@@ -2311,7 +2315,7 @@ Proof.
simpl; intuition.
eapply star_trans. eexact D.
eapply star_right. eexact G.
- left. eapply step_assign; eauto. congruence. rewrite B; eauto. congruence.
+ left. eapply step_assign with (v1 := v1); eauto. congruence. rewrite B; eauto. congruence.
reflexivity. traceEq.
(* assignop *)
exploit (H0 (fun x => C(Eassignop op x r tyres ty))).
@@ -2322,7 +2326,7 @@ Proof.
eapply star_trans. eexact D.
eapply star_right. eexact G.
left. eapply step_assignop; eauto.
- rewrite B; eauto. rewrite B; rewrite F; eauto. congruence. rewrite B; eauto. congruence.
+ rewrite B; eauto. rewrite B; rewrite F; eauto. rewrite B; eauto. rewrite B; eauto. congruence.
reflexivity. traceEq.
(* postincr *)
exploit (H0 (fun x => C(Epostincr id x ty))).
@@ -2656,7 +2660,7 @@ Proof (proj2 (proj2 (proj2 (proj2 bigstep_to_steps)))).
Fixpoint esize (a: expr) : nat :=
match a with
- | Eloc _ _ _ => 1%nat
+ | Eloc _ _ _ _ => 1%nat
| Evar _ _ => 1%nat
| Ederef r1 _ => S(esize r1)
| Efield l1 _ _ => S(esize l1)
diff --git a/cfrontend/Csyntax.v b/cfrontend/Csyntax.v
index 19bc2ec3..5b8a62be 100644
--- a/cfrontend/Csyntax.v
+++ b/cfrontend/Csyntax.v
@@ -56,7 +56,7 @@ Inductive expr : Type :=
(**r function call [r1(rargs)] *)
| Ebuiltin (ef: external_function) (tyargs: typelist) (rargs: exprlist) (ty: type)
(**r builtin function call *)
- | Eloc (b: block) (ofs: ptrofs) (ty: type)
+ | Eloc (b: block) (ofs: ptrofs) (bf: bitfield) (ty: type)
(**r memory location, result of evaluating a l-value *)
| Eparen (r: expr) (tycast: type) (ty: type) (**r marked subexpression *)
@@ -117,7 +117,7 @@ Definition Eselection (r1 r2 r3: expr) (ty: type) :=
Definition typeof (a: expr) : type :=
match a with
- | Eloc _ _ ty => ty
+ | Eloc _ _ _ ty => ty
| Evar _ ty => ty
| Ederef _ ty => ty
| Efield _ _ ty => ty
diff --git a/cfrontend/Ctypes.v b/cfrontend/Ctypes.v
index bcd8d350..7b2c7354 100644
--- a/cfrontend/Ctypes.v
+++ b/cfrontend/Ctypes.v
@@ -20,6 +20,8 @@ Require Import Axioms Coqlib Maps Errors.
Require Import AST Linking.
Require Archi.
+Local Open Scope error_monad_scope.
+
(** * Syntax of types *)
(** Compcert C types are similar to those of C. They include numeric types,
@@ -84,21 +86,28 @@ Proof.
decide equality.
Defined.
+Lemma signedness_eq: forall (s1 s2: signedness), {s1=s2} + {s1<>s2}.
+Proof.
+ decide equality.
+Defined.
+
+Lemma attr_eq: forall (a1 a2: attr), {a1=a2} + {a1<>a2}.
+Proof.
+ decide equality. decide equality. apply N.eq_dec. apply bool_dec.
+Defined.
+
Lemma type_eq: forall (ty1 ty2: type), {ty1=ty2} + {ty1<>ty2}
with typelist_eq: forall (tyl1 tyl2: typelist), {tyl1=tyl2} + {tyl1<>tyl2}.
Proof.
- assert (forall (x y: signedness), {x=y} + {x<>y}) by decide equality.
assert (forall (x y: floatsize), {x=y} + {x<>y}) by decide equality.
- assert (forall (x y: attr), {x=y} + {x<>y}).
- { decide equality. decide equality. apply N.eq_dec. apply bool_dec. }
- generalize ident_eq zeq bool_dec ident_eq intsize_eq; intros.
+ generalize ident_eq zeq bool_dec ident_eq intsize_eq signedness_eq attr_eq; intros.
decide equality.
decide equality.
decide equality.
decide equality.
Defined.
-Opaque type_eq typelist_eq.
+Global Opaque intsize_eq signedness_eq attr_eq type_eq typelist_eq.
(** Extract the attributes of a type. *)
@@ -150,17 +159,53 @@ Definition attr_union (a1 a2: attr) : attr :=
Definition merge_attributes (ty: type) (a: attr) : type :=
change_attributes (attr_union a) ty.
+(** Maximal size in bits of a bitfield of type [sz]. *)
+
+Definition bitsize_intsize (sz: intsize) : Z :=
+ match sz with
+ | I8 => 8
+ | I16 => 16
+ | I32 => 32
+ | IBool => 1
+ end.
+
(** Syntax for [struct] and [union] definitions. [struct] and [union]
are collectively called "composites". Each compilation unit
comes with a list of top-level definitions of composites. *)
Inductive struct_or_union : Type := Struct | Union.
-Definition members : Type := list (ident * type).
+Inductive member : Type :=
+ | Member_plain (id: ident) (t: type)
+ | Member_bitfield (id: ident) (sz: intsize) (sg: signedness) (a: attr)
+ (width: Z) (padding: bool).
+
+Definition members : Type := list member.
Inductive composite_definition : Type :=
Composite (id: ident) (su: struct_or_union) (m: members) (a: attr).
+Definition name_member (m: member) : ident :=
+ match m with
+ | Member_plain id _ => id
+ | Member_bitfield id _ _ _ _ _ => id
+ end.
+
+Definition type_member (m: member) : type :=
+ match m with
+ | Member_plain _ t => t
+ | Member_bitfield _ sz sg a w _ =>
+ (* An unsigned bitfield of width < size of type reads with a signed type *)
+ let sg' := if zlt w (bitsize_intsize sz) then Signed else sg in
+ Tint sz sg' a
+ end.
+
+Definition member_is_padding (m: member) : bool :=
+ match m with
+ | Member_plain _ _ => false
+ | Member_bitfield _ _ _ _ _ p => p
+ end.
+
Definition name_composite_def (c: composite_definition) : ident :=
match c with Composite id su m a => id end.
@@ -168,7 +213,9 @@ Definition composite_def_eq (x y: composite_definition): {x=y} + {x<>y}.
Proof.
decide equality.
- decide equality. decide equality. apply N.eq_dec. apply bool_dec.
-- apply list_eq_dec. decide equality. apply type_eq. apply ident_eq.
+- apply list_eq_dec. decide equality.
+ apply type_eq. apply ident_eq.
+ apply bool_dec. apply zeq. apply attr_eq. apply signedness_eq. apply intsize_eq. apply ident_eq.
- decide equality.
- apply ident_eq.
Defined.
@@ -194,6 +241,13 @@ Record composite : Type := {
Definition composite_env : Type := PTree.t composite.
+(** Access modes for members of structs or unions: either a plain field
+ or a bitfield *)
+
+Inductive bitfield : Type :=
+ | Full
+ | Bits (sz: intsize) (sg: signedness) (pos: Z) (width: Z).
+
(** * Operations over types *)
(** ** Conversions *)
@@ -389,41 +443,205 @@ Proof.
- destruct (env!i). apply co_sizeof_alignof. apply Z.divide_0_r.
Qed.
+(** ** Layout of struct fields *)
+
+Section LAYOUT.
+
+Variable env: composite_env.
+
+Definition bitalignof (t: type) := alignof env t * 8.
+
+Definition bitsizeof (t: type) := sizeof env t * 8.
+
+Definition bitalignof_intsize (sz: intsize) : Z :=
+ match sz with
+ | I8 | IBool => 8
+ | I16 => 16
+ | I32 => 32
+ end.
+
+Definition next_field (pos: Z) (m: member) : Z :=
+ match m with
+ | Member_plain _ t =>
+ align pos (bitalignof t) + bitsizeof t
+ | Member_bitfield _ sz _ _ w _ =>
+ let s := bitalignof_intsize sz in
+ if zle w 0 then
+ align pos s
+ else
+ let curr := floor pos s in
+ let next := curr + s in
+ if zle (pos + w) next then pos + w else next + w
+ end.
+
+Definition layout_field (pos: Z) (m: member) : res (Z * bitfield) :=
+ match m with
+ | Member_plain _ t =>
+ OK (align pos (bitalignof t) / 8, Full)
+ | Member_bitfield _ sz sg _ w _ =>
+ if zle w 0 then Error (msg "accessing zero-width bitfield")
+ else if zlt (bitsize_intsize sz) w then Error (msg "bitfield too wide")
+ else
+ let s := bitalignof_intsize sz in
+ let start := floor pos s in
+ let next := start + s in
+ if zle (pos + w) next then
+ OK (start / 8, Bits sz sg (pos - start) w)
+ else
+ OK (next / 8, Bits sz sg 0 w)
+ end.
+
+(** Some properties *)
+
+Lemma bitalignof_intsize_pos:
+ forall sz, bitalignof_intsize sz > 0.
+Proof.
+ destruct sz; simpl; lia.
+Qed.
+
+Lemma next_field_incr:
+ forall pos m, pos <= next_field pos m.
+Proof.
+ intros. unfold next_field. destruct m.
+- set (al := bitalignof t).
+ assert (A: al > 0).
+ { unfold al, bitalignof. generalize (alignof_pos env t). lia. }
+ assert (pos <= align pos al) by (apply align_le; auto).
+ assert (bitsizeof t >= 0).
+ { unfold bitsizeof. generalize (sizeof_pos env t). lia. }
+ lia.
+- set (s := bitalignof_intsize sz).
+ assert (A: s > 0) by (apply bitalignof_intsize_pos).
+ destruct (zle width 0).
++ apply align_le; auto.
++ generalize (floor_interval pos s A).
+ set (start := floor pos s). intros B.
+ destruct (zle (pos + width) (start + s)); lia.
+Qed.
+
+Definition layout_start (p: Z) (bf: bitfield) :=
+ p * 8 + match bf with Full => 0 | Bits sz sg pos w => pos end.
+
+Definition layout_width (t: type) (bf: bitfield) :=
+ match bf with Full => bitsizeof t | Bits sz sg pos w => w end.
+
+Lemma layout_field_range: forall pos m ofs bf,
+ layout_field pos m = OK (ofs, bf) ->
+ pos <= layout_start ofs bf
+ /\ layout_start ofs bf + layout_width (type_member m) bf <= next_field pos m.
+Proof.
+ intros until bf; intros L. unfold layout_start, layout_width. destruct m; simpl in L.
+- inv L. simpl.
+ set (al := bitalignof t).
+ set (q := align pos al).
+ assert (A: al > 0).
+ { unfold al, bitalignof. generalize (alignof_pos env t). lia. }
+ assert (B: pos <= q) by (apply align_le; auto).
+ assert (C: (al | q)) by (apply align_divides; auto).
+ assert (D: (8 | q)).
+ { apply Z.divide_transitive with al; auto. apply Z.divide_factor_r. }
+ assert (E: q / 8 * 8 = q).
+ { destruct D as (n & E). rewrite E. rewrite Z.div_mul by lia. auto. }
+ rewrite E. lia.
+- unfold next_field.
+ destruct (zle width 0); try discriminate.
+ destruct (zlt (bitsize_intsize sz) width); try discriminate.
+ set (s := bitalignof_intsize sz) in *.
+ assert (A: s > 0) by (apply bitalignof_intsize_pos).
+ generalize (floor_interval pos s A). set (p := floor pos s) in *. intros B.
+ assert (C: (s | p)) by (apply floor_divides; auto).
+ assert (D: (8 | s)).
+ { exists (s / 8). unfold s. destruct sz; reflexivity. }
+ assert (E: (8 | p)) by (apply Z.divide_transitive with s; auto).
+ assert (F: (8 | p + s)) by (apply Z.divide_add_r; auto).
+ assert (G: p / 8 * 8 = p).
+ { destruct E as (n & EQ). rewrite EQ. rewrite Z.div_mul by lia. auto. }
+ assert (H: (p + s) / 8 * 8 = p + s).
+ { destruct F as (n & EQ). rewrite EQ. rewrite Z.div_mul by lia. auto. }
+ destruct (zle (pos + width) (p + s)); inv L; lia.
+Qed.
+
+Definition layout_alignment (t: type) (bf: bitfield) :=
+ match bf with
+ | Full => alignof env t
+ | Bits sz _ _ _ => bitalignof_intsize sz / 8
+ end.
+
+Lemma layout_field_alignment: forall pos m ofs bf,
+ layout_field pos m = OK (ofs, bf) ->
+ (layout_alignment (type_member m) bf | ofs).
+Proof.
+ intros until bf; intros L. destruct m; simpl in L.
+- inv L; simpl.
+ set (q := align pos (bitalignof t)).
+ assert (A: (bitalignof t | q)).
+ { apply align_divides. unfold bitalignof. generalize (alignof_pos env t). lia. }
+ destruct A as [n E]. exists n. rewrite E. unfold bitalignof. rewrite Z.mul_assoc, Z.div_mul by lia. auto.
+- destruct (zle width 0); try discriminate.
+ destruct (zlt (bitsize_intsize sz) width); try discriminate.
+ set (s := bitalignof_intsize sz) in *.
+ assert (A: s > 0) by (apply bitalignof_intsize_pos).
+ set (p := floor pos s) in *.
+ assert (C: (s | p)) by (apply floor_divides; auto).
+ assert (D: (8 | s)).
+ { exists (s / 8). unfold s. destruct sz; reflexivity. }
+ assert (E: forall n, (s | n) -> (s / 8 | n / 8)).
+ { intros. destruct H as [n1 E1], D as [n2 E2]. rewrite E1, E2.
+ rewrite Z.mul_assoc, ! Z.div_mul by lia. exists n1; auto. }
+ destruct (zle (pos + width) (p + s)); inv L; simpl; fold s.
+ + apply E. auto.
+ + apply E. apply Z.divide_add_r; auto using Z.divide_refl.
+Qed.
+
+End LAYOUT.
+
(** ** Size and alignment for composite definitions *)
(** The alignment for a structure or union is the max of the alignment
- of its members. *)
+ of its members. Padding bitfields are ignored. *)
-Fixpoint alignof_composite (env: composite_env) (m: members) : Z :=
- match m with
+Fixpoint alignof_composite (env: composite_env) (ms: members) : Z :=
+ match ms with
| nil => 1
- | (id, t) :: m' => Z.max (alignof env t) (alignof_composite env m')
+ | m :: ms =>
+ if member_is_padding m
+ then alignof_composite env ms
+ else Z.max (alignof env (type_member m)) (alignof_composite env ms)
end.
(** The size of a structure corresponds to its layout: fields are
laid out consecutively, and padding is inserted to align
- each field to the alignment for its type. *)
+ each field to the alignment for its type. Bitfields are packed
+ as described above. *)
-Fixpoint sizeof_struct (env: composite_env) (cur: Z) (m: members) : Z :=
- match m with
+Fixpoint bitsizeof_struct (env: composite_env) (cur: Z) (ms: members) : Z :=
+ match ms with
| nil => cur
- | (id, t) :: m' => sizeof_struct env (align cur (alignof env t) + sizeof env t) m'
+ | m :: ms => bitsizeof_struct env (next_field env cur m) ms
end.
+Definition bytes_of_bits (n: Z) := (n + 7) / 8.
+
+Definition sizeof_struct (env: composite_env) (m: members) : Z :=
+ bytes_of_bits (bitsizeof_struct env 0 m).
+
(** The size of an union is the max of the sizes of its members. *)
-Fixpoint sizeof_union (env: composite_env) (m: members) : Z :=
- match m with
+Fixpoint sizeof_union (env: composite_env) (ms: members) : Z :=
+ match ms with
| nil => 0
- | (id, t) :: m' => Z.max (sizeof env t) (sizeof_union env m')
+ | m :: ms => Z.max (sizeof env (type_member m)) (sizeof_union env ms)
end.
+(** Some properties *)
+
Lemma alignof_composite_two_p:
forall env m, exists n, alignof_composite env m = two_power_nat n.
Proof.
- induction m as [|[id t]]; simpl.
+ induction m; simpl.
- exists 0%nat; auto.
-- apply Z.max_case; auto. apply alignof_two_p.
+- destruct (member_is_padding a); auto.
+ apply Z.max_case; auto. apply alignof_two_p.
Qed.
Lemma alignof_composite_pos:
@@ -435,94 +653,113 @@ Proof.
rewrite EQ; apply two_power_nat_pos.
Qed.
-Lemma sizeof_struct_incr:
- forall env m cur, cur <= sizeof_struct env cur m.
+Lemma bitsizeof_struct_incr:
+ forall env m cur, cur <= bitsizeof_struct env cur m.
Proof.
- induction m as [|[id t]]; simpl; intros.
+ induction m; simpl; intros.
- lia.
-- apply Z.le_trans with (align cur (alignof env t)).
- apply align_le. apply alignof_pos.
- apply Z.le_trans with (align cur (alignof env t) + sizeof env t).
- generalize (sizeof_pos env t); lia.
- apply IHm.
+- apply Z.le_trans with (next_field env cur a).
+ apply next_field_incr. apply IHm.
Qed.
Lemma sizeof_union_pos:
forall env m, 0 <= sizeof_union env m.
Proof.
- induction m as [|[id t]]; simpl; extlia.
+ induction m; simpl; extlia.
Qed.
-(** ** Byte offset for a field of a structure *)
+(** ** Byte offset and bitfield designator for a field of a structure *)
-(** [field_offset env id fld] returns the byte offset for field [id]
- in a structure whose members are [fld]. Fields are laid out
- consecutively, and padding is inserted to align each field to the
- alignment for its type. *)
-
-Fixpoint field_offset_rec (env: composite_env) (id: ident) (fld: members) (pos: Z)
- {struct fld} : res Z :=
- match fld with
+Fixpoint field_type (id: ident) (ms: members) {struct ms} : res type :=
+ match ms with
| nil => Error (MSG "Unknown field " :: CTX id :: nil)
- | (id', t) :: fld' =>
- if ident_eq id id'
- then OK (align pos (alignof env t))
- else field_offset_rec env id fld' (align pos (alignof env t) + sizeof env t)
+ | m :: ms => if ident_eq id (name_member m) then OK (type_member m) else field_type id ms
end.
-Definition field_offset (env: composite_env) (id: ident) (fld: members) : res Z :=
- field_offset_rec env id fld 0.
+(** [field_offset env id fld] returns the byte offset for field [id]
+ in a structure whose members are [fld]. It also returns a
+ bitfield designator, giving the location of the bits to access
+ within the storage unit for the bitfield. *)
-Fixpoint field_type (id: ident) (fld: members) {struct fld} : res type :=
- match fld with
+Fixpoint field_offset_rec (env: composite_env) (id: ident) (ms: members) (pos: Z)
+ {struct ms} : res (Z * bitfield) :=
+ match ms with
| nil => Error (MSG "Unknown field " :: CTX id :: nil)
- | (id', t) :: fld' => if ident_eq id id' then OK t else field_type id fld'
+ | m :: ms =>
+ if ident_eq id (name_member m)
+ then layout_field env pos m
+ else field_offset_rec env id ms (next_field env pos m)
end.
+Definition field_offset (env: composite_env) (id: ident) (ms: members) : res (Z * bitfield) :=
+ field_offset_rec env id ms 0.
+
(** Some sanity checks about field offsets. First, field offsets are
within the range of acceptable offsets. *)
Remark field_offset_rec_in_range:
- forall env id ofs ty fld pos,
- field_offset_rec env id fld pos = OK ofs -> field_type id fld = OK ty ->
- pos <= ofs /\ ofs + sizeof env ty <= sizeof_struct env pos fld.
+ forall env id ofs bf ty ms pos,
+ field_offset_rec env id ms pos = OK (ofs, bf) -> field_type id ms = OK ty ->
+ pos <= layout_start ofs bf
+ /\ layout_start ofs bf + layout_width env ty bf <= bitsizeof_struct env pos ms.
Proof.
- intros until ty. induction fld as [|[i t]]; simpl; intros.
+ induction ms as [ | m ms]; simpl; intros.
- discriminate.
-- destruct (ident_eq id i); intros.
- inv H. inv H0. split.
- apply align_le. apply alignof_pos. apply sizeof_struct_incr.
- exploit IHfld; eauto. intros [A B]. split; auto.
- eapply Z.le_trans; eauto. apply Z.le_trans with (align pos (alignof env t)).
- apply align_le. apply alignof_pos. generalize (sizeof_pos env t). lia.
+- destruct (ident_eq id (name_member m)).
+ + inv H0.
+ exploit layout_field_range; eauto.
+ generalize (bitsizeof_struct_incr env ms (next_field env pos m)).
+ lia.
+ + exploit IHms; eauto.
+ generalize (next_field_incr env pos m).
+ lia.
Qed.
-Lemma field_offset_in_range:
- forall env fld id ofs ty,
- field_offset env id fld = OK ofs -> field_type id fld = OK ty ->
- 0 <= ofs /\ ofs + sizeof env ty <= sizeof_struct env 0 fld.
+Lemma field_offset_in_range_gen:
+ forall env ms id ofs bf ty,
+ field_offset env id ms = OK (ofs, bf) -> field_type id ms = OK ty ->
+ 0 <= layout_start ofs bf
+ /\ layout_start ofs bf + layout_width env ty bf <= bitsizeof_struct env 0 ms.
Proof.
intros. eapply field_offset_rec_in_range; eauto.
Qed.
+Corollary field_offset_in_range:
+ forall env ms id ofs ty,
+ field_offset env id ms = OK (ofs, Full) -> field_type id ms = OK ty ->
+ 0 <= ofs /\ ofs + sizeof env ty <= sizeof_struct env ms.
+Proof.
+ intros. exploit field_offset_in_range_gen; eauto.
+ unfold layout_start, layout_width, bitsizeof, sizeof_struct. intros [A B].
+ assert (C: forall x y, x * 8 <= y -> x <= bytes_of_bits y).
+ { unfold bytes_of_bits; intros.
+ assert (P: 8 > 0) by lia.
+ generalize (Z_div_mod_eq (y + 7) 8 P) (Z_mod_lt (y + 7) 8 P).
+ lia. }
+ split. lia. apply C. lia.
+Qed.
+
(** Second, two distinct fields do not overlap *)
Lemma field_offset_no_overlap:
- forall env id1 ofs1 ty1 id2 ofs2 ty2 fld,
- field_offset env id1 fld = OK ofs1 -> field_type id1 fld = OK ty1 ->
- field_offset env id2 fld = OK ofs2 -> field_type id2 fld = OK ty2 ->
+ forall env id1 ofs1 bf1 ty1 id2 ofs2 bf2 ty2 fld,
+ field_offset env id1 fld = OK (ofs1, bf1) -> field_type id1 fld = OK ty1 ->
+ field_offset env id2 fld = OK (ofs2, bf2) -> field_type id2 fld = OK ty2 ->
id1 <> id2 ->
- ofs1 + sizeof env ty1 <= ofs2 \/ ofs2 + sizeof env ty2 <= ofs1.
+ layout_start ofs1 bf1 + layout_width env ty1 bf1 <= layout_start ofs2 bf2
+ \/ layout_start ofs2 bf2 + layout_width env ty2 bf2 <= layout_start ofs1 bf1.
Proof.
intros until fld. unfold field_offset. generalize 0 as pos.
- induction fld as [|[i t]]; simpl; intros.
+ induction fld as [|m fld]; simpl; intros.
- discriminate.
-- destruct (ident_eq id1 i); destruct (ident_eq id2 i).
+- destruct (ident_eq id1 (name_member m)); destruct (ident_eq id2 (name_member m)).
+ congruence.
-+ subst i. inv H; inv H0.
- exploit field_offset_rec_in_range. eexact H1. eauto. tauto.
-+ subst i. inv H1; inv H2.
- exploit field_offset_rec_in_range. eexact H. eauto. tauto.
++ inv H0.
+ exploit field_offset_rec_in_range; eauto.
+ exploit layout_field_range; eauto. lia.
++ inv H2.
+ exploit field_offset_rec_in_range; eauto.
+ exploit layout_field_range; eauto. lia.
+ eapply IHfld; eauto.
Qed.
@@ -530,31 +767,90 @@ Qed.
are the same. *)
Lemma field_offset_prefix:
- forall env id ofs fld2 fld1,
- field_offset env id fld1 = OK ofs ->
- field_offset env id (fld1 ++ fld2) = OK ofs.
+ forall env id ofs bf fld2 fld1,
+ field_offset env id fld1 = OK (ofs, bf) ->
+ field_offset env id (fld1 ++ fld2) = OK (ofs, bf).
Proof.
intros until fld1. unfold field_offset. generalize 0 as pos.
- induction fld1 as [|[i t]]; simpl; intros.
+ induction fld1 as [|m fld1]; simpl; intros.
- discriminate.
-- destruct (ident_eq id i); auto.
+- destruct (ident_eq id (name_member m)); auto.
Qed.
(** Fourth, the position of each field respects its alignment. *)
-Lemma field_offset_aligned:
- forall env id fld ofs ty,
- field_offset env id fld = OK ofs -> field_type id fld = OK ty ->
- (alignof env ty | ofs).
+Lemma field_offset_aligned_gen:
+ forall env id fld ofs bf ty,
+ field_offset env id fld = OK (ofs, bf) -> field_type id fld = OK ty ->
+ (layout_alignment env ty bf | ofs).
Proof.
intros until ty. unfold field_offset. generalize 0 as pos. revert fld.
- induction fld as [|[i t]]; simpl; intros.
+ induction fld as [|m fld]; simpl; intros.
- discriminate.
-- destruct (ident_eq id i).
-+ inv H; inv H0. apply align_divides. apply alignof_pos.
+- destruct (ident_eq id (name_member m)).
++ inv H0. eapply layout_field_alignment; eauto.
+ eauto.
Qed.
+Corollary field_offset_aligned:
+ forall env id fld ofs ty,
+ field_offset env id fld = OK (ofs, Full) -> field_type id fld = OK ty ->
+ (alignof env ty | ofs).
+Proof.
+ intros. exploit field_offset_aligned_gen; eauto.
+Qed.
+
+(** [union_field_offset env id ms] returns the byte offset and
+ bitfield designator for accessing a member named [id] of a union
+ whose members are [ms]. The byte offset is always 0. *)
+
+Fixpoint union_field_offset (env: composite_env) (id: ident) (ms: members)
+ {struct ms} : res (Z * bitfield) :=
+ match ms with
+ | nil => Error (MSG "Unknown field " :: CTX id :: nil)
+ | m :: ms =>
+ if ident_eq id (name_member m)
+ then layout_field env 0 m
+ else union_field_offset env id ms
+ end.
+
+(** Some sanity checks about union field offsets. First, field offsets
+ fit within the size of the union. *)
+
+Lemma union_field_offset_in_range_gen:
+ forall env id ofs bf ty ms,
+ union_field_offset env id ms = OK (ofs, bf) -> field_type id ms = OK ty ->
+ ofs = 0 /\ 0 <= layout_start ofs bf /\ layout_start ofs bf + layout_width env ty bf <= sizeof_union env ms * 8.
+Proof.
+ induction ms as [ | m ms]; simpl; intros.
+- discriminate.
+- destruct (ident_eq id (name_member m)).
+ + inv H0. set (ty := type_member m) in *.
+ destruct m; simpl in H.
+ * inv H. unfold layout_start, layout_width.
+ rewrite align_same. change (0 / 8) with 0. unfold bitsizeof. lia.
+ unfold bitalignof. generalize (alignof_pos env t). lia.
+ apply Z.divide_0_r.
+ * destruct (zle width 0); try discriminate.
+ destruct (zlt (bitsize_intsize sz) width); try discriminate.
+ assert (A: bitsize_intsize sz <= bitalignof_intsize sz <= sizeof env ty * 8).
+ { unfold ty, type_member; destruct sz; simpl; lia. }
+ rewrite zle_true in H by lia. inv H.
+ unfold layout_start, layout_width.
+ unfold floor; rewrite Z.div_0_l by lia.
+ lia.
+ + exploit IHms; eauto. lia.
+Qed.
+
+Corollary union_field_offset_in_range:
+ forall env ms id ofs ty,
+ union_field_offset env id ms = OK (ofs, Full) -> field_type id ms = OK ty ->
+ ofs = 0 /\ sizeof env ty <= sizeof_union env ms.
+Proof.
+ intros. exploit union_field_offset_in_range_gen; eauto.
+ unfold layout_start, layout_width, bitsizeof. lia.
+Qed.
+
(** ** Access modes *)
(** The [access_mode] function describes how a l-value of the given
@@ -712,7 +1008,8 @@ Fixpoint rank_type (ce: composite_env) (t: type) : nat :=
Fixpoint rank_members (ce: composite_env) (m: members) : nat :=
match m with
| nil => 0%nat
- | (id, t) :: m => Init.Nat.max (rank_type ce t) (rank_members ce m)
+ | Member_plain _ t :: m => Init.Nat.max (rank_type ce t) (rank_members ce m)
+ | Member_bitfield _ _ _ _ _ _ :: m => rank_members ce m
end.
(** ** C types and back-end types *)
@@ -766,7 +1063,7 @@ Definition signature_of_type (args: typelist) (res: type) (cc: calling_conventio
Definition sizeof_composite (env: composite_env) (su: struct_or_union) (m: members) : Z :=
match su with
- | Struct => sizeof_struct env 0 m
+ | Struct => sizeof_struct env m
| Union => sizeof_union env m
end.
@@ -774,21 +1071,23 @@ Lemma sizeof_composite_pos:
forall env su m, 0 <= sizeof_composite env su m.
Proof.
intros. destruct su; simpl.
- apply sizeof_struct_incr.
- apply sizeof_union_pos.
+- unfold sizeof_struct, bytes_of_bits.
+ assert (0 <= bitsizeof_struct env 0 m) by apply bitsizeof_struct_incr.
+ change 0 with (0 / 8) at 1. apply Z.div_le_mono; lia.
+- apply sizeof_union_pos.
Qed.
-Fixpoint complete_members (env: composite_env) (m: members) : bool :=
- match m with
+Fixpoint complete_members (env: composite_env) (ms: members) : bool :=
+ match ms with
| nil => true
- | (id, t) :: m' => complete_type env t && complete_members env m'
+ | m :: ms => complete_type env (type_member m) && complete_members env ms
end.
Lemma complete_member:
- forall env id t m,
- In (id, t) m -> complete_members env m = true -> complete_type env t = true.
+ forall env m ms,
+ In m ms -> complete_members env ms = true -> complete_type env (type_member m) = true.
Proof.
- induction m as [|[id1 t1] m]; simpl; intuition auto.
+ induction ms as [|m1 ms]; simpl; intuition auto.
InvBooleans; inv H1; auto.
InvBooleans; eauto.
Qed.
@@ -852,8 +1151,6 @@ Defined.
must precede all uses of this composite, unless the use is under
a pointer or function type. *)
-Local Open Scope error_monad_scope.
-
Fixpoint add_composite_definitions (env: composite_env) (defs: list composite_definition) : res composite_env :=
match defs with
| nil => OK env
@@ -916,52 +1213,88 @@ Proof.
Qed.
Lemma alignof_composite_stable:
- forall m, complete_members env m = true -> alignof_composite env' m = alignof_composite env m.
+ forall ms, complete_members env ms = true -> alignof_composite env' ms = alignof_composite env ms.
Proof.
- induction m as [|[id t]]; simpl; intros.
+ induction ms as [|m ms]; simpl; intros.
auto.
- InvBooleans. rewrite alignof_stable by auto. rewrite IHm by auto. auto.
+ InvBooleans. rewrite alignof_stable by auto. rewrite IHms by auto. auto.
Qed.
-Lemma sizeof_struct_stable:
- forall m pos, complete_members env m = true -> sizeof_struct env' pos m = sizeof_struct env pos m.
+Remark next_field_stable: forall pos m,
+ complete_type env (type_member m) = true -> next_field env' pos m = next_field env pos m.
Proof.
- induction m as [|[id t]]; simpl; intros.
+ destruct m; simpl; intros.
+- unfold bitalignof, bitsizeof. rewrite alignof_stable, sizeof_stable by auto. auto.
+- auto.
+Qed.
+
+Lemma bitsizeof_struct_stable:
+ forall ms pos, complete_members env ms = true -> bitsizeof_struct env' pos ms = bitsizeof_struct env pos ms.
+Proof.
+ induction ms as [|m ms]; simpl; intros.
auto.
- InvBooleans. rewrite alignof_stable by auto. rewrite sizeof_stable by auto.
- rewrite IHm by auto. auto.
+ InvBooleans. rewrite next_field_stable by auto. apply IHms; auto.
Qed.
Lemma sizeof_union_stable:
- forall m, complete_members env m = true -> sizeof_union env' m = sizeof_union env m.
+ forall ms, complete_members env ms = true -> sizeof_union env' ms = sizeof_union env ms.
Proof.
- induction m as [|[id t]]; simpl; intros.
+ induction ms as [|m ms]; simpl; intros.
auto.
- InvBooleans. rewrite sizeof_stable by auto. rewrite IHm by auto. auto.
+ InvBooleans. rewrite sizeof_stable by auto. rewrite IHms by auto. auto.
Qed.
Lemma sizeof_composite_stable:
- forall su m, complete_members env m = true -> sizeof_composite env' su m = sizeof_composite env su m.
+ forall su ms, complete_members env ms = true -> sizeof_composite env' su ms = sizeof_composite env su ms.
Proof.
intros. destruct su; simpl.
- apply sizeof_struct_stable; auto.
+ unfold sizeof_struct. f_equal. apply bitsizeof_struct_stable; auto.
apply sizeof_union_stable; auto.
Qed.
Lemma complete_members_stable:
- forall m, complete_members env m = true -> complete_members env' m = true.
+ forall ms, complete_members env ms = true -> complete_members env' ms = true.
Proof.
- induction m as [|[id t]]; simpl; intros.
+ induction ms as [|m ms]; simpl; intros.
auto.
- InvBooleans. rewrite complete_type_stable by auto. rewrite IHm by auto. auto.
+ InvBooleans. rewrite complete_type_stable by auto. rewrite IHms by auto. auto.
Qed.
Lemma rank_members_stable:
- forall m, complete_members env m = true -> rank_members env' m = rank_members env m.
+ forall ms, complete_members env ms = true -> rank_members env' ms = rank_members env ms.
Proof.
- induction m as [|[id t]]; simpl; intros.
+ induction ms as [|m ms]; simpl; intros.
auto.
- InvBooleans. f_equal; auto. apply rank_type_stable; auto.
+ InvBooleans. destruct m; auto. f_equal; auto. apply rank_type_stable; auto.
+Qed.
+
+Remark layout_field_stable: forall pos m,
+ complete_type env (type_member m) = true -> layout_field env' pos m = layout_field env pos m.
+Proof.
+ destruct m; simpl; intros.
+- unfold bitalignof. rewrite alignof_stable by auto. auto.
+- auto.
+Qed.
+
+Lemma field_offset_stable:
+ forall f ms, complete_members env ms = true -> field_offset env' f ms = field_offset env f ms.
+Proof.
+ intros until ms. unfold field_offset. generalize 0.
+ induction ms as [|m ms]; simpl; intros.
+- auto.
+- InvBooleans. destruct (ident_eq f (name_member m)).
+ apply layout_field_stable; auto.
+ rewrite next_field_stable by auto. apply IHms; auto.
+Qed.
+
+Lemma union_field_offset_stable:
+ forall f ms, complete_members env ms = true -> union_field_offset env' f ms = union_field_offset env f ms.
+Proof.
+ induction ms as [|m ms]; simpl; intros.
+- auto.
+- InvBooleans. destruct (ident_eq f (name_member m)).
+ apply layout_field_stable; auto.
+ apply IHms; auto.
Qed.
End STABILITY.
@@ -1091,19 +1424,23 @@ Qed.
is strictly greater than the ranks of its member types. *)
Remark rank_type_members:
- forall ce id t m, In (id, t) m -> (rank_type ce t <= rank_members ce m)%nat.
+ forall ce m ms, In m ms -> (rank_type ce (type_member m) <= rank_members ce ms)%nat.
Proof.
- induction m; simpl; intros; intuition auto.
- subst a. extlia.
- extlia.
+ induction ms; simpl; intros.
+- tauto.
+- destruct a; destruct H; subst; simpl.
+ + lia.
+ + apply IHms in H. lia.
+ + lia.
+ + apply IHms; auto.
Qed.
Lemma rank_struct_member:
- forall ce id a co id1 t1,
+ forall ce id a co m,
composite_env_consistent ce ->
ce!id = Some co ->
- In (id1, t1) (co_members co) ->
- (rank_type ce t1 < rank_type ce (Tstruct id a))%nat.
+ In m (co_members co) ->
+ (rank_type ce (type_member m) < rank_type ce (Tstruct id a))%nat.
Proof.
intros; simpl. rewrite H0.
erewrite co_consistent_rank by eauto.
@@ -1112,11 +1449,11 @@ Proof.
Qed.
Lemma rank_union_member:
- forall ce id a co id1 t1,
+ forall ce id a co m,
composite_env_consistent ce ->
ce!id = Some co ->
- In (id1, t1) (co_members co) ->
- (rank_type ce t1 < rank_type ce (Tunion id a))%nat.
+ In m (co_members co) ->
+ (rank_type ce (type_member m) < rank_type ce (Tunion id a))%nat.
Proof.
intros; simpl. rewrite H0.
erewrite co_consistent_rank by eauto.
@@ -1516,6 +1853,57 @@ Global Opaque Linker_program.
(** ** Commutation between linking and program transformations *)
+Section LINK_MATCH_PROGRAM_GEN.
+
+Context {F G: Type}.
+Variable match_fundef: program F -> fundef F -> fundef G -> Prop.
+
+Hypothesis link_match_fundef:
+ forall ctx1 ctx2 f1 tf1 f2 tf2 f,
+ link f1 f2 = Some f ->
+ match_fundef ctx1 f1 tf1 -> match_fundef ctx2 f2 tf2 ->
+ exists tf, link tf1 tf2 = Some tf /\ (match_fundef ctx1 f tf \/ match_fundef ctx2 f tf).
+
+Let match_program (p: program F) (tp: program G) : Prop :=
+ Linking.match_program_gen match_fundef eq p p tp
+ /\ prog_types tp = prog_types p.
+
+Theorem link_match_program_gen:
+ forall p1 p2 tp1 tp2 p,
+ link p1 p2 = Some p -> match_program p1 tp1 -> match_program p2 tp2 ->
+ exists tp, link tp1 tp2 = Some tp /\ match_program p tp.
+Proof.
+ intros until p; intros L [M1 T1] [M2 T2].
+ exploit link_linkorder; eauto. intros [LO1 LO2].
+Local Transparent Linker_program.
+ simpl in L; unfold link_program in L.
+ destruct (link (program_of_program p1) (program_of_program p2)) as [pp|] eqn:LP; try discriminate.
+ assert (A: exists tpp,
+ link (program_of_program tp1) (program_of_program tp2) = Some tpp
+ /\ Linking.match_program_gen match_fundef eq p pp tpp).
+ { eapply Linking.link_match_program; eauto.
+ - intros.
+ Local Transparent Linker_types.
+ simpl in *. destruct (type_eq v1 v2); inv H. exists v; rewrite dec_eq_true; auto.
+ }
+ destruct A as (tpp & TLP & MP).
+ simpl; unfold link_program. rewrite TLP.
+ destruct (lift_option (link (prog_types p1) (prog_types p2))) as [[typs EQ]|EQ]; try discriminate.
+ destruct (link_build_composite_env (prog_types p1) (prog_types p2) typs
+ (prog_comp_env p1) (prog_comp_env p2) (prog_comp_env_eq p1)
+ (prog_comp_env_eq p2) EQ) as (env & P & Q).
+ rewrite <- T1, <- T2 in EQ.
+ destruct (lift_option (link (prog_types tp1) (prog_types tp2))) as [[ttyps EQ']|EQ']; try congruence.
+ assert (ttyps = typs) by congruence. subst ttyps.
+ destruct (link_build_composite_env (prog_types tp1) (prog_types tp2) typs
+ (prog_comp_env tp1) (prog_comp_env tp2) (prog_comp_env_eq tp1)
+ (prog_comp_env_eq tp2) EQ') as (tenv & R & S).
+ assert (tenv = env) by congruence. subst tenv.
+ econstructor; split; eauto. inv L. split; auto.
+Qed.
+
+End LINK_MATCH_PROGRAM_GEN.
+
Section LINK_MATCH_PROGRAM.
Context {F G: Type}.
@@ -1571,3 +1959,4 @@ Local Transparent Linker_program.
Qed.
End LINK_MATCH_PROGRAM.
+
diff --git a/cfrontend/Ctyping.v b/cfrontend/Ctyping.v
index 5f0a3e5b..c930a407 100644
--- a/cfrontend/Ctyping.v
+++ b/cfrontend/Ctyping.v
@@ -410,8 +410,8 @@ Inductive wt_rvalue : expr -> Prop :=
wt_rvalue (Eparen r tycast ty)
with wt_lvalue : expr -> Prop :=
- | wt_Eloc: forall b ofs ty,
- wt_lvalue (Eloc b ofs ty)
+ | wt_Eloc: forall b ofs bf ty,
+ wt_lvalue (Eloc b ofs bf ty)
| wt_Evar: forall x ty,
e!x = Some ty ->
wt_lvalue (Evar x ty)
@@ -440,7 +440,7 @@ Definition wt_expr_kind (k: kind) (a: expr) :=
Definition expr_kind (a: expr) : kind :=
match a with
- | Eloc _ _ _ => LV
+ | Eloc _ _ _ _ => LV
| Evar _ _ => LV
| Ederef _ _ => LV
| Efield _ _ _ => LV
@@ -596,7 +596,7 @@ Fixpoint check_arguments (el: exprlist) (tyl: typelist) : res unit :=
Definition check_rval (e: expr) : res unit :=
match e with
- | Eloc _ _ _ | Evar _ _ | Ederef _ _ | Efield _ _ _ =>
+ | Eloc _ _ _ _ | Evar _ _ | Ederef _ _ | Efield _ _ _ =>
Error (msg "not a r-value")
| _ =>
OK tt
@@ -604,7 +604,7 @@ Definition check_rval (e: expr) : res unit :=
Definition check_lval (e: expr) : res unit :=
match e with
- | Eloc _ _ _ | Evar _ _ | Ederef _ _ | Efield _ _ _ =>
+ | Eloc _ _ _ _ | Evar _ _ | Ederef _ _ | Efield _ _ _ =>
OK tt
| _ =>
Error (msg "not a l-value")
@@ -846,7 +846,7 @@ Fixpoint retype_expr (ce: composite_env) (e: typenv) (a: expr) : res expr :=
do r1' <- retype_expr ce e r1; do rl' <- retype_exprlist ce e rl; ecall r1' rl'
| Ebuiltin ef tyargs rl tyres =>
do rl' <- retype_exprlist ce e rl; ebuiltin ef tyargs rl' tyres
- | Eloc _ _ _ =>
+ | Eloc _ _ _ _ =>
Error (msg "Eloc in source")
| Eparen _ _ _ =>
Error (msg "Eparen in source")
@@ -984,6 +984,7 @@ Lemma binarith_type_cast:
forall t1 t2 m t,
binarith_type t1 t2 m = OK t -> wt_cast t1 t /\ wt_cast t2 t.
Proof.
+Local Transparent Ctypes.intsize_eq.
unfold wt_cast, binarith_type, classify_binarith; intros; DestructCases;
simpl; split; try congruence;
try (destruct Archi.ptr64; congruence).
@@ -1656,9 +1657,31 @@ Proof.
unfold Mptr in *. destruct Archi.ptr64 eqn:SF; auto with ty.
Qed.
+Remark wt_bitfield_normalize: forall sz sg width sg1 n,
+ 0 < width <= bitsize_intsize sz ->
+ sg1 = (if zlt width (bitsize_intsize sz) then Signed else sg) ->
+ wt_int (bitfield_normalize sz sg width n) sz sg1.
+Proof.
+ intros. destruct sz; cbn in *.
+ + destruct sg.
+ * replace sg1 with Signed by (destruct zlt; auto).
+ apply Int.sign_ext_widen; lia.
+ * subst sg1; destruct zlt.
+ ** apply Int.sign_zero_ext_widen; lia.
+ ** apply Int.zero_ext_widen; lia.
+ + destruct sg.
+ * replace sg1 with Signed by (destruct zlt; auto).
+ apply Int.sign_ext_widen; lia.
+ * subst sg1; destruct zlt.
+ ** apply Int.sign_zero_ext_widen; lia.
+ ** apply Int.zero_ext_widen; lia.
+ + auto.
+ + apply Int.zero_ext_widen; lia.
+Qed.
+
Lemma wt_deref_loc:
- forall ge ty m b ofs t v,
- deref_loc ge ty m b ofs t v ->
+ forall ge ty m b ofs bf t v,
+ deref_loc ge ty m b ofs bf t v ->
wt_val v ty.
Proof.
induction 1.
@@ -1680,6 +1703,19 @@ Proof.
destruct ty; simpl in H; try discriminate; auto with ty.
destruct i; destruct s; discriminate.
destruct f; discriminate.
+- (* bitfield *)
+ inv H. constructor.
+ apply wt_bitfield_normalize. lia. auto.
+Qed.
+
+Lemma wt_assign_loc:
+ forall ge ty m b ofs bf v t m' v',
+ assign_loc ge ty m b ofs bf v t m' v' ->
+ wt_val v ty -> wt_val v' ty.
+Proof.
+ induction 1; intros; auto.
+- inv H. constructor.
+ apply wt_bitfield_normalize. lia. auto.
Qed.
Lemma wt_cast_self:
@@ -1770,7 +1806,7 @@ Proof.
- (* condition *) constructor. destruct b; auto. destruct b; auto. red; auto.
- (* sizeof *) unfold size_t, Vptrofs; destruct Archi.ptr64; constructor; auto with ty.
- (* alignof *) unfold size_t, Vptrofs; destruct Archi.ptr64; constructor; auto with ty.
-- (* assign *) inversion H5. constructor. eapply pres_sem_cast; eauto.
+- (* assign *) inversion H5. constructor. eapply wt_assign_loc; eauto. eapply pres_sem_cast; eauto.
- (* assignop *) subst tyres l r. constructor. auto.
constructor. constructor. eapply wt_deref_loc; eauto.
auto. auto. auto.
diff --git a/cfrontend/Initializers.v b/cfrontend/Initializers.v
index 77d6cfea..32fbf46b 100644
--- a/cfrontend/Initializers.v
+++ b/cfrontend/Initializers.v
@@ -114,18 +114,23 @@ Fixpoint constval (ce: composite_env) (a: expr) : res val :=
| Ederef r ty =>
constval ce r
| Efield l f ty =>
- match typeof l with
- | Tstruct id _ =>
- do co <- lookup_composite ce id;
- do delta <- field_offset ce f (co_members co);
- do v <- constval ce l;
+ do (delta, bf) <-
+ match typeof l with
+ | Tstruct id _ =>
+ do co <- lookup_composite ce id; field_offset ce f (co_members co)
+ | Tunion id _ =>
+ do co <- lookup_composite ce id; union_field_offset ce f (co_members co)
+ | _ =>
+ Error (msg "ill-typed field access")
+ end;
+ do v <- constval ce l;
+ match bf with
+ | Full =>
OK (if Archi.ptr64
then Val.addl v (Vlong (Int64.repr delta))
else Val.add v (Vint (Int.repr delta)))
- | Tunion id _ =>
- constval ce l
- | _ =>
- Error(msg "ill-typed field access")
+ | Bits _ _ _ _ =>
+ Error(msg "taking the address of a bitfield")
end
| Eparen r tycast ty =>
do v <- constval ce r; do_cast v (typeof r) tycast
@@ -138,6 +143,183 @@ Fixpoint constval (ce: composite_env) (a: expr) : res val :=
Definition constval_cast (ce: composite_env) (a: expr) (ty: type): res val :=
do v <- constval ce a; do_cast v (typeof a) ty.
+(** * Building and recording initialization data *)
+
+(** The following [state] type is the output of the translation of
+ initializers. It contains the list of initialization data
+ generated so far, the corresponding position in bytes, and the
+ total size expected for the final initialization data, in bytes. *)
+
+Record state : Type := {
+ init: list init_data; (**r reversed *)
+ curr: Z; (**r current position for head of [init] *)
+ total_size: Z (**r total expected size *)
+}.
+
+(** A state [s] can also be viewed as a memory block. The size of
+ the block is [s.(total_size)], it is initialized with zero bytes,
+ then filled with the initialization data [rev s.(init)] like
+ [Genv.store_init_data_list] does. *)
+
+Definition initial_state (sz: Z) : state :=
+ {| init := nil; curr := 0; total_size := sz |}.
+
+(** We now define abstract "store" operations that operate
+ directly on the state, but whose behavior mimic those of
+ storing in the corresponding memory block. To initialize
+ bitfields, we also need an abstract "load" operation.
+ The operations are optimized for stores that occur at increasing
+ positions, like those that take place during initialization. *)
+
+(** Initialization from bytes *)
+
+Definition int_of_byte (b: byte) := Int.repr (Byte.unsigned b).
+
+Definition Init_byte (b: byte) := Init_int8 (int_of_byte b).
+
+(** Add a list of bytes to a reversed initialization data list. *)
+
+Fixpoint add_rev_bytes (l: list byte) (il: list init_data) :=
+ match l with
+ | nil => il
+ | b :: l => add_rev_bytes l (Init_byte b :: il)
+ end.
+
+(** Add [n] zero bytes to an initialization data list. *)
+
+Definition add_zeros (n: Z) (il: list init_data) :=
+ Z.iter n (fun l => Init_int8 Int.zero :: l) il.
+
+(** Make sure the [depth] positions at the top of [il] are bytes,
+ that is, [Init_int8] items. Other numerical items are split
+ into bytes. [Init_addrof] items cannot be split and result in
+ an error. *)
+
+Fixpoint normalize (il: list init_data) (depth: Z) : res (list init_data) :=
+ if zle depth 0 then OK il else
+ match il with
+ | nil =>
+ Error (msg "normalize: empty list")
+ | Init_int8 n :: il =>
+ do il' <- normalize il (depth - 1);
+ OK (Init_int8 n :: il')
+ | Init_int16 n :: il =>
+ do il' <- normalize il (depth - 2);
+ OK (add_rev_bytes (encode_int 2%nat (Int.unsigned n)) il')
+ | Init_int32 n :: il =>
+ do il' <- normalize il (depth - 4);
+ OK (add_rev_bytes (encode_int 4%nat (Int.unsigned n)) il')
+ | Init_int64 n :: il =>
+ do il' <- normalize il (depth - 8);
+ OK (add_rev_bytes (encode_int 8%nat (Int64.unsigned n)) il')
+ | Init_float32 f :: il =>
+ do il' <- normalize il (depth - 4);
+ OK (add_rev_bytes (encode_int 4%nat (Int.unsigned (Float32.to_bits f))) il')
+ | Init_float64 f :: il =>
+ do il' <- normalize il (depth - 8);
+ OK (add_rev_bytes (encode_int 8%nat (Int64.unsigned (Float.to_bits f))) il')
+ | Init_addrof _ _ :: il =>
+ Error (msg "normalize: Init_addrof")
+ | Init_space n :: il =>
+ let n := Z.max 0 n in
+ if zle n depth then
+ do il' <- normalize il (depth - n);
+ OK (add_zeros n il')
+ else
+ OK (add_zeros depth (Init_space (n - depth) :: il))
+ end.
+
+(** Split [il] into [depth] bytes and the initialization list that follows.
+ The bytes are returned reversed. *)
+
+Fixpoint decompose_rec (accu: list byte) (il: list init_data) (depth: Z) : res (list byte * list init_data) :=
+ if zle depth 0 then OK (accu, il) else
+ match il with
+ | Init_int8 n :: il => decompose_rec (Byte.repr (Int.unsigned n) :: accu) il (depth - 1)
+ | _ => Error (msg "decompose: wrong shape")
+ end.
+
+Definition decompose (il: list init_data) (depth: Z) : res (list byte * list init_data) :=
+ decompose_rec nil il depth.
+
+(** Decompose an initialization list in three parts:
+ [depth] bytes (reversed), [sz] bytes (reversed),
+ and the remainder of the initialization list. *)
+
+Definition trisection (il: list init_data) (depth sz: Z) : res (list byte * list byte * list init_data) :=
+ do il0 <- normalize il (depth + sz);
+ do (bytes1, il1) <- decompose il0 depth;
+ do (bytes2, il2) <- decompose il1 sz;
+ OK (bytes1, bytes2, il2).
+
+(** Graphically: [rev il] is equal to
+<<
+ <---sz---><--depth-->
++----------------+---------+---------+
+| | | |
++----------------+---------+---------+
+ rev il2 bytes2 bytes1
+>>
+*)
+
+(** Add padding if necessary so that position [pos] is within the state. *)
+
+Definition pad_to (s: state) (pos: Z) : state :=
+ if zle pos s.(curr)
+ then s
+ else {| init := Init_space (pos - s.(curr)) :: s.(init);
+ curr := pos;
+ total_size := s.(total_size) |}.
+
+(** Store the initialization data [i] at position [pos] in state [s]. *)
+
+Definition store_data (s: state) (pos: Z) (i: init_data) : res state :=
+ let sz := init_data_size i in
+ assertion (zle 0 pos && zle (pos + sz) s.(total_size));
+ if zle s.(curr) pos then
+ OK {| init := i :: (if zlt s.(curr) pos
+ then Init_space (pos - s.(curr)) :: s.(init)
+ else s.(init));
+ curr := pos + sz;
+ total_size := s.(total_size) |}
+ else
+ let s' := pad_to s (pos + sz) in
+ do x3 <- trisection s'.(init) (s'.(curr) - (pos + sz)) sz;
+ let '(bytes1, _, il2) := x3 in
+ OK {| init := add_rev_bytes bytes1 (i :: il2);
+ curr := s'.(curr);
+ total_size := s'.(total_size) |}.
+
+(** Store the integer [n] of size [isz] at position [pos] in state [s]. *)
+
+Definition init_data_for_carrier (isz: intsize) (n: int) :=
+ match isz with
+ | I8 | IBool => Init_int8 n
+ | I16 => Init_int16 n
+ | I32 => Init_int32 n
+ end.
+
+Definition store_int (s: state) (pos: Z) (isz: intsize) (n: int) : res state :=
+ store_data s pos (init_data_for_carrier isz n).
+
+(** Load the integer of size [isz] at position [pos] in state [s]. *)
+
+Definition load_int (s: state) (pos: Z) (isz: intsize) : res int :=
+ let chunk := chunk_for_carrier isz in
+ let sz := size_chunk chunk in
+ assertion (zle 0 pos && zle (pos + sz) s.(total_size));
+ let s' := pad_to s (pos + sz) in
+ do x3 <- trisection s'.(init) (s'.(curr) - (pos + sz)) sz;
+ let '(_, bytes2, _) := x3 in
+ OK (Int.repr (decode_int bytes2)).
+
+(** Extract the final initialization data from a state. *)
+
+Definition init_data_list_of_state (s: state) : res (list init_data) :=
+ assertion (zle s.(curr) s.(total_size));
+ let s' := pad_to s s.(total_size) in
+ OK (List.rev' s'.(init)).
+
(** * Translation of initializers *)
Inductive initializer :=
@@ -149,6 +331,11 @@ with initializer_list :=
| Init_nil
| Init_cons (i: initializer) (il: initializer_list).
+Definition length_initializer_list (il: initializer_list) :=
+ let fix length (accu: Z) (il: initializer_list) : Z :=
+ match il with Init_nil => accu | Init_cons _ il => length (Z.succ accu) il end
+ in length 0 il.
+
(** Translate an initializing expression [a] for a scalar variable
of type [ty]. Return the corresponding initialization datum. *)
@@ -170,69 +357,116 @@ Definition transl_init_single (ce: composite_env) (ty: type) (a: expr) : res ini
| _, _ => Error (msg "type mismatch in initializer")
end.
-(** Translate an initializer [i] for a variable of type [ty].
- [transl_init ce ty i] returns the appropriate list of initialization data.
- The intermediate functions [transl_init_rec], [transl_init_array]
- and [transl_init_struct] append initialization data to the given
- list [k], and build the list of initialization data in reverse order,
- so as to remain tail-recursive. *)
+(** Initialize a bitfield [Bits sz sg p w] with expression [a]. *)
-Definition padding (frm to: Z) (k: list init_data) : list init_data :=
- if zlt frm to then Init_space (to - frm) :: k else k.
+Definition transl_init_bitfield (ce: composite_env) (s: state)
+ (ty: type) (sz: intsize) (p w: Z)
+ (i: initializer) (pos: Z) : res state :=
+ match i with
+ | Init_single a =>
+ do v <- constval_cast ce a ty;
+ match v with
+ | Vint n =>
+ do c <- load_int s pos sz;
+ let c' := Int.bitfield_insert (first_bit sz p w) w c n in
+ store_int s pos sz c'
+ | Vundef =>
+ Error (msg "undefined operation in bitfield initializer")
+ | _ =>
+ Error (msg "type mismatch in bitfield initializer")
+ end
+ | _ =>
+ Error (msg "bitfield initialized by composite initializer")
+ end.
+
+(** Padding bitfields and bitfields with zero width are not initialized. *)
+
+Definition member_not_initialized (m: member) : bool :=
+ match m with
+ | Member_plain _ _ => false
+ | Member_bitfield _ _ _ _ w p => p || zle w 0
+ end.
-Fixpoint transl_init_rec (ce: composite_env) (ty: type) (i: initializer)
- (k: list init_data) {struct i} : res (list init_data) :=
+(** Translate an initializer [i] for a variable of type [ty]
+ and store the corresponding list of initialization data in state [s]
+ at position [pos]. Return the updated state. *)
+
+Fixpoint transl_init_rec (ce: composite_env) (s: state)
+ (ty: type) (i: initializer) (pos: Z)
+ {struct i} : res state :=
match i, ty with
| Init_single a, _ =>
- do d <- transl_init_single ce ty a; OK (d :: k)
+ do d <- transl_init_single ce ty a; store_data s pos d
| Init_array il, Tarray tyelt nelt _ =>
- transl_init_array ce tyelt il (Z.max 0 nelt) k
+ assertion (zle (length_initializer_list il) (Z.max 0 nelt));
+ transl_init_array ce s tyelt il pos
| Init_struct il, Tstruct id _ =>
do co <- lookup_composite ce id;
match co_su co with
- | Struct => transl_init_struct ce ty (co_members co) il 0 k
+ | Struct => transl_init_struct ce s (co_members co) il pos 0
| Union => Error (MSG "struct/union mismatch on " :: CTX id :: nil)
end
| Init_union f i1, Tunion id _ =>
do co <- lookup_composite ce id;
match co_su co with
| Struct => Error (MSG "union/struct mismatch on " :: CTX id :: nil)
- | Union =>
- do ty1 <- field_type f (co_members co);
- do k1 <- transl_init_rec ce ty1 i1 k;
- OK (padding (sizeof ce ty1) (sizeof ce ty) k1)
+ | Union => do ty1 <- field_type f (co_members co);
+ do (delta, layout) <- union_field_offset ce f (co_members co);
+ match layout with
+ | Full =>
+ transl_init_rec ce s ty1 i1 (pos + delta)
+ | Bits sz sg p w =>
+ transl_init_bitfield ce s ty1 sz p w i1 (pos + delta)
+ end
end
| _, _ =>
Error (msg "wrong type for compound initializer")
end
-with transl_init_array (ce: composite_env) (ty: type) (il: initializer_list) (sz: Z)
- (k: list init_data) {struct il} : res (list init_data) :=
+with transl_init_array (ce: composite_env) (s: state)
+ (tyelt: type) (il: initializer_list) (pos: Z)
+ {struct il} : res state :=
match il with
| Init_nil =>
- if zeq sz 0 then OK k
- else if zle 0 sz then OK (Init_space (sz * sizeof ce ty) :: k)
- else Error (msg "wrong number of elements in array initializer")
+ OK s
| Init_cons i1 il' =>
- do k1 <- transl_init_rec ce ty i1 k;
- transl_init_array ce ty il' (sz - 1) k1
+ do s1 <- transl_init_rec ce s tyelt i1 pos;
+ transl_init_array ce s1 tyelt il' (pos + sizeof ce tyelt)
end
-with transl_init_struct (ce: composite_env) (ty: type)
- (fl: members) (il: initializer_list) (pos: Z)
- (k: list init_data)
- {struct il} : res (list init_data) :=
- match il, fl with
- | Init_nil, nil =>
- OK (padding pos (sizeof ce ty) k)
- | Init_cons i1 il', (_, ty1) :: fl' =>
- let pos1 := align pos (alignof ce ty1) in
- do k1 <- transl_init_rec ce ty1 i1 (padding pos pos1 k);
- transl_init_struct ce ty fl' il' (pos1 + sizeof ce ty1) k1
- | _, _ =>
- Error (msg "wrong number of elements in struct initializer")
+with transl_init_struct (ce: composite_env) (s: state)
+ (ms: members) (il: initializer_list)
+ (base: Z) (pos: Z)
+ {struct il} : res state :=
+ match il with
+ | Init_nil =>
+ OK s
+ | Init_cons i1 il' =>
+ let fix init (ms: members) (pos: Z) {struct ms} : res state :=
+ match ms with
+ | nil =>
+ Error (msg "too many elements in struct initializer")
+ | m :: ms' =>
+ if member_not_initialized m then
+ init ms' (next_field ce pos m)
+ else
+ do (delta, layout) <- layout_field ce pos m;
+ do s1 <-
+ match layout with
+ | Full =>
+ transl_init_rec ce s (type_member m) i1 (base + delta)
+ | Bits sz sg p w =>
+ transl_init_bitfield ce s (type_member m) sz p w i1 (base + delta)
+ end;
+ transl_init_struct ce s1 ms' il' base (next_field ce pos m)
+ end in
+ init ms pos
end.
+(** The entry point. *)
+
Definition transl_init (ce: composite_env) (ty: type) (i: initializer)
: res (list init_data) :=
- do k <- transl_init_rec ce ty i nil; OK (List.rev' k).
+ let s0 := initial_state (sizeof ce ty) in
+ do s1 <- transl_init_rec ce s0 ty i 0;
+ init_data_list_of_state s1.
diff --git a/cfrontend/Initializersproof.v b/cfrontend/Initializersproof.v
index 10ccbeff..00f7e331 100644
--- a/cfrontend/Initializersproof.v
+++ b/cfrontend/Initializersproof.v
@@ -12,7 +12,7 @@
(** Compile-time evaluation of initializers for global C variables. *)
-Require Import Coqlib Maps.
+Require Import Zwf Coqlib Maps.
Require Import Errors Integers Floats Values AST Memory Globalenvs Events Smallstep.
Require Import Ctypes Cop Csyntax Csem.
Require Import Initializers.
@@ -30,7 +30,7 @@ Variable ge: genv.
Fixpoint simple (a: expr) : Prop :=
match a with
- | Eloc _ _ _ => True
+ | Eloc _ _ _ _ => True
| Evar _ _ => True
| Ederef r _ => simple r
| Efield l1 _ _ => simple l1
@@ -65,38 +65,38 @@ Section SIMPLE_EXPRS.
Variable e: env.
Variable m: mem.
-Inductive eval_simple_lvalue: expr -> block -> ptrofs -> Prop :=
- | esl_loc: forall b ofs ty,
- eval_simple_lvalue (Eloc b ofs ty) b ofs
+Inductive eval_simple_lvalue: expr -> block -> ptrofs -> bitfield -> Prop :=
+ | esl_loc: forall b ofs bf ty,
+ eval_simple_lvalue (Eloc b ofs bf ty) b ofs bf
| esl_var_local: forall x ty b,
e!x = Some(b, ty) ->
- eval_simple_lvalue (Evar x ty) b Ptrofs.zero
+ eval_simple_lvalue (Evar x ty) b Ptrofs.zero Full
| esl_var_global: forall x ty b,
e!x = None ->
Genv.find_symbol ge x = Some b ->
- eval_simple_lvalue (Evar x ty) b Ptrofs.zero
+ eval_simple_lvalue (Evar x ty) b Ptrofs.zero Full
| esl_deref: forall r ty b ofs,
eval_simple_rvalue r (Vptr b ofs) ->
- eval_simple_lvalue (Ederef r ty) b ofs
- | esl_field_struct: forall r f ty b ofs id co a delta,
+ eval_simple_lvalue (Ederef r ty) b ofs Full
+ | esl_field_struct: forall r f ty b ofs id co a delta bf,
eval_simple_rvalue r (Vptr b ofs) ->
- typeof r = Tstruct id a -> ge.(genv_cenv)!id = Some co -> field_offset ge f (co_members co) = OK delta ->
- eval_simple_lvalue (Efield r f ty) b (Ptrofs.add ofs (Ptrofs.repr delta))
- | esl_field_union: forall r f ty b ofs id a,
+ typeof r = Tstruct id a -> ge.(genv_cenv)!id = Some co -> field_offset ge f (co_members co) = OK (delta, bf) ->
+ eval_simple_lvalue (Efield r f ty) b (Ptrofs.add ofs (Ptrofs.repr delta)) bf
+ | esl_field_union: forall r f ty b ofs id co a delta bf,
eval_simple_rvalue r (Vptr b ofs) ->
- typeof r = Tunion id a ->
- eval_simple_lvalue (Efield r f ty) b ofs
+ typeof r = Tunion id a -> ge.(genv_cenv)!id = Some co -> union_field_offset ge f (co_members co) = OK (delta, bf) ->
+ eval_simple_lvalue (Efield r f ty) b (Ptrofs.add ofs (Ptrofs.repr delta)) bf
with eval_simple_rvalue: expr -> val -> Prop :=
| esr_val: forall v ty,
eval_simple_rvalue (Eval v ty) v
- | esr_rvalof: forall b ofs l ty v,
- eval_simple_lvalue l b ofs ->
+ | esr_rvalof: forall b ofs bf l ty v,
+ eval_simple_lvalue l b ofs bf ->
ty = typeof l ->
- deref_loc ge ty m b ofs E0 v ->
+ deref_loc ge ty m b ofs bf E0 v ->
eval_simple_rvalue (Evalof l ty) v
| esr_addrof: forall b ofs l ty,
- eval_simple_lvalue l b ofs ->
+ eval_simple_lvalue l b ofs Full ->
eval_simple_rvalue (Eaddrof l ty) (Vptr b ofs)
| esr_unop: forall op r1 ty v1 v,
eval_simple_rvalue r1 v1 ->
@@ -153,7 +153,7 @@ End SIMPLE_EXPRS.
Definition compat_eval (k: kind) (e: env) (a a': expr) (m: mem) : Prop :=
typeof a = typeof a' /\
match k with
- | LV => forall b ofs, eval_simple_lvalue e m a' b ofs -> eval_simple_lvalue e m a b ofs
+ | LV => forall b ofs bf, eval_simple_lvalue e m a' b ofs bf -> eval_simple_lvalue e m a b ofs bf
| RV => forall v, eval_simple_rvalue e m a' v -> eval_simple_rvalue e m a v
end.
@@ -167,7 +167,7 @@ Lemma lred_compat:
forall e l m l' m', lred ge e l m l' m' ->
m = m' /\ compat_eval LV e l l' m.
Proof.
- induction 1; simpl; split; auto; split; auto; intros bx ofsx EV; inv EV.
+ induction 1; simpl; split; auto; split; auto; intros bx ofsx bf' EV; inv EV.
apply esl_var_local; auto.
apply esl_var_global; auto.
constructor. constructor.
@@ -365,6 +365,22 @@ Proof.
intros. eapply bool_val_inj; eauto. intros. rewrite mem_empty_not_weak_valid_pointer in H2; discriminate.
Qed.
+Lemma add_offset_match:
+ forall v b ofs delta,
+ Val.inject inj v (Vptr b ofs) ->
+ Val.inject inj
+ (if Archi.ptr64
+ then Val.addl v (Vlong (Int64.repr delta))
+ else Val.add v (Vint (Int.repr delta)))
+ (Vptr b (Ptrofs.add ofs (Ptrofs.repr delta))).
+Proof.
+ intros. inv H.
+- rewrite Ptrofs.add_assoc. rewrite (Ptrofs.add_commut (Ptrofs.repr delta0)).
+ unfold Val.addl, Val.add; destruct Archi.ptr64 eqn:SF;
+ econstructor; eauto; rewrite ! Ptrofs.add_assoc; f_equal; f_equal; symmetry; auto with ptrofs.
+- unfold Val.addl, Val.add; destruct Archi.ptr64; auto.
+Qed.
+
(** Soundness of [constval] with respect to the big-step semantics *)
Lemma constval_rvalue:
@@ -374,20 +390,22 @@ Lemma constval_rvalue:
constval ge a = OK v' ->
Val.inject inj v' v
with constval_lvalue:
- forall m a b ofs,
- eval_simple_lvalue empty_env m a b ofs ->
+ forall m a b ofs bf,
+ eval_simple_lvalue empty_env m a b ofs bf ->
forall v',
constval ge a = OK v' ->
- Val.inject inj v' (Vptr b ofs).
+ bf = Full /\ Val.inject inj v' (Vptr b ofs).
Proof.
(* rvalue *)
induction 1; intros vres CV; simpl in CV; try (monadInv CV).
(* val *)
destruct v; monadInv CV; constructor.
(* rval *)
- inv H1; rewrite H2 in CV; try congruence. eauto. eauto.
+ assert (constval ge l = OK vres) by (destruct (access_mode ty); congruence).
+ exploit constval_lvalue; eauto. intros [A B]. subst bf.
+ inv H1; rewrite H3 in CV; congruence.
(* addrof *)
- eauto.
+ eapply constval_lvalue; eauto.
(* unop *)
destruct (sem_unary_operation op x (typeof r1) Mem.empty) as [v1'|] eqn:E; inv EQ0.
exploit (sem_unary_operation_inj inj Mem.empty m).
@@ -438,28 +456,31 @@ Proof.
(* lvalue *)
induction 1; intros v' CV; simpl in CV; try (monadInv CV).
(* var local *)
- unfold empty_env in H. rewrite PTree.gempty in H. congruence.
+ split; auto. unfold empty_env in H. rewrite PTree.gempty in H. congruence.
(* var_global *)
- econstructor. unfold inj. rewrite H0. eauto. auto.
+ split; auto. econstructor. unfold inj. rewrite H0. eauto. auto.
(* deref *)
- eauto.
+ split; eauto.
(* field struct *)
- rewrite H0 in CV. monadInv CV. unfold lookup_composite in EQ; rewrite H1 in EQ; monadInv EQ.
- exploit constval_rvalue; eauto. intro MV. inv MV.
- replace x0 with delta by congruence. rewrite Ptrofs.add_assoc. rewrite (Ptrofs.add_commut (Ptrofs.repr delta0)).
- simpl; destruct Archi.ptr64 eqn:SF;
- econstructor; eauto; rewrite ! Ptrofs.add_assoc; f_equal; f_equal; symmetry; auto with ptrofs.
- destruct Archi.ptr64; auto.
+ rewrite H0 in EQ. monadInv EQ. destruct x0; monadInv EQ2.
+ unfold lookup_composite in EQ0; rewrite H1 in EQ0; monadInv EQ0.
+ exploit constval_rvalue; eauto. intro MV.
+ split. congruence.
+ replace x with delta by congruence.
+ apply (add_offset_match _ _ _ _ MV).
(* field union *)
- rewrite H0 in CV. eauto.
+ rewrite H0 in EQ. monadInv EQ. destruct x0; monadInv EQ2.
+ unfold lookup_composite in EQ0; rewrite H1 in EQ0; monadInv EQ0.
+ exploit constval_rvalue; eauto. intro MV.
+ split. congruence.
+ replace x with delta by congruence.
+ apply (add_offset_match _ _ _ _ MV).
Qed.
Lemma constval_simple:
forall a v, constval ge a = OK v -> simple a.
Proof.
induction a; simpl; intros vx CV; try (monadInv CV); eauto.
- destruct (typeof a); discriminate || eauto.
- monadInv CV. eauto.
destruct (access_mode ty); discriminate || eauto.
intuition eauto.
Qed.
@@ -476,331 +497,757 @@ Proof.
intros [A [B C]]. intuition. eapply constval_rvalue; eauto.
Qed.
-(** * Relational specification of the translation of initializers *)
-
-Definition tr_padding (frm to: Z) : list init_data :=
- if zlt frm to then Init_space (to - frm) :: nil else nil.
-
-Inductive tr_init: type -> initializer -> list init_data -> Prop :=
- | tr_init_sgl: forall ty a d,
- transl_init_single ge ty a = OK d ->
- tr_init ty (Init_single a) (d :: nil)
- | tr_init_arr: forall tyelt nelt attr il d,
- tr_init_array tyelt il (Z.max 0 nelt) d ->
- tr_init (Tarray tyelt nelt attr) (Init_array il) d
- | tr_init_str: forall id attr il co d,
- lookup_composite ge id = OK co -> co_su co = Struct ->
- tr_init_struct (Tstruct id attr) (co_members co) il 0 d ->
- tr_init (Tstruct id attr) (Init_struct il) d
- | tr_init_uni: forall id attr f i1 co ty1 d,
- lookup_composite ge id = OK co -> co_su co = Union -> field_type f (co_members co) = OK ty1 ->
- tr_init ty1 i1 d ->
- tr_init (Tunion id attr) (Init_union f i1)
- (d ++ tr_padding (sizeof ge ty1) (sizeof ge (Tunion id attr)))
-
-with tr_init_array: type -> initializer_list -> Z -> list init_data -> Prop :=
- | tr_init_array_nil_0: forall ty,
- tr_init_array ty Init_nil 0 nil
- | tr_init_array_nil_pos: forall ty sz,
- 0 < sz ->
- tr_init_array ty Init_nil sz (Init_space (sz * sizeof ge ty) :: nil)
- | tr_init_array_cons: forall ty i il sz d1 d2,
- tr_init ty i d1 -> tr_init_array ty il (sz - 1) d2 ->
- tr_init_array ty (Init_cons i il) sz (d1 ++ d2)
-
-with tr_init_struct: type -> members -> initializer_list -> Z -> list init_data -> Prop :=
- | tr_init_struct_nil: forall ty pos,
- tr_init_struct ty nil Init_nil pos (tr_padding pos (sizeof ge ty))
- | tr_init_struct_cons: forall ty f1 ty1 fl i1 il pos d1 d2,
- let pos1 := align pos (alignof ge ty1) in
- tr_init ty1 i1 d1 ->
- tr_init_struct ty fl il (pos1 + sizeof ge ty1) d2 ->
- tr_init_struct ty ((f1, ty1) :: fl) (Init_cons i1 il)
- pos (tr_padding pos pos1 ++ d1 ++ d2).
-
-Lemma transl_padding_spec:
- forall frm to k, padding frm to k = rev (tr_padding frm to) ++ k.
-Proof.
- unfold padding, tr_padding; intros.
- destruct (zlt frm to); auto.
-Qed.
-
-Lemma transl_init_rec_spec:
- forall i ty k res,
- transl_init_rec ge ty i k = OK res ->
- exists d, tr_init ty i d /\ res = rev d ++ k
-
-with transl_init_array_spec:
- forall il ty sz k res,
- transl_init_array ge ty il sz k = OK res ->
- exists d, tr_init_array ty il sz d /\ res = rev d ++ k
-
-with transl_init_struct_spec:
- forall il ty fl pos k res,
- transl_init_struct ge ty fl il pos k = OK res ->
- exists d, tr_init_struct ty fl il pos d /\ res = rev d ++ k.
-
-Proof.
-Local Opaque sizeof.
-- destruct i; intros until res; intros TR; simpl in TR.
-+ monadInv TR. exists (x :: nil); split; auto. constructor; auto.
-+ destruct ty; try discriminate.
- destruct (transl_init_array_spec _ _ _ _ _ TR) as (d & A & B).
- exists d; split; auto. constructor; auto.
-+ destruct ty; try discriminate. monadInv TR. destruct (co_su x) eqn:SU; try discriminate.
- destruct (transl_init_struct_spec _ _ _ _ _ _ EQ0) as (d & A & B).
- exists d; split; auto. econstructor; eauto.
-+ destruct ty; try discriminate.
- monadInv TR. destruct (co_su x) eqn:SU; monadInv EQ0.
- destruct (transl_init_rec_spec _ _ _ _ EQ0) as (d & A & B).
- exists (d ++ tr_padding (sizeof ge x0) (sizeof ge (Tunion i0 a))); split.
- econstructor; eauto.
- rewrite rev_app_distr, app_ass, B. apply transl_padding_spec.
-
-- destruct il; intros until res; intros TR; simpl in TR.
-+ destruct (zeq sz 0).
- inv TR. exists (@nil init_data); split; auto. constructor.
- destruct (zle 0 sz).
- inv TR. econstructor; split. constructor. lia. auto.
- discriminate.
-+ monadInv TR.
- destruct (transl_init_rec_spec _ _ _ _ EQ) as (d1 & A1 & B1).
- destruct (transl_init_array_spec _ _ _ _ _ EQ0) as (d2 & A2 & B2).
- exists (d1 ++ d2); split. econstructor; eauto.
- subst res x. rewrite rev_app_distr, app_ass. auto.
+(** * Correctness of operations over the initialization state *)
+
+(** ** Properties of the in-memory bytes denoted by initialization data *)
+
+Local Notation boid := (Genv.bytes_of_init_data (genv_genv ge)).
+Local Notation boidl := (Genv.bytes_of_init_data_list (genv_genv ge)).
+
+Lemma boidl_app: forall il2 il1,
+ boidl (il1 ++ il2) = boidl il1 ++ boidl il2.
+Proof.
+ induction il1 as [ | il il1]; simpl. auto. rewrite app_ass. f_equal; auto.
+Qed.
+
+Corollary boidl_rev_cons: forall i il,
+ boidl (rev il ++ i :: nil) = boidl (rev il) ++ boid i.
+Proof.
+ intros. rewrite boidl_app. simpl. rewrite <- app_nil_end. auto.
+Qed.
+
+Definition byte_of_int (n: int) := Byte.repr (Int.unsigned n).
+
+Lemma byte_of_int_of_byte: forall b, byte_of_int (int_of_byte b) = b.
+Proof.
+ intros. unfold int_of_byte, byte_of_int.
+ rewrite Int.unsigned_repr, Byte.repr_unsigned. auto.
+ assert(Byte.max_unsigned < Int.max_unsigned) by reflexivity.
+ generalize (Byte.unsigned_range_2 b). lia.
+Qed.
+
+Lemma inj_bytes_1: forall n,
+ inj_bytes (encode_int 1 n) = Byte (Byte.repr n) :: nil.
+Proof.
+ intros. unfold encode_int, bytes_of_int, rev_if_be. destruct Archi.big_endian; auto.
+Qed.
+
+Lemma inj_bytes_byte: forall b,
+ inj_bytes (encode_int 1 (Int.unsigned (int_of_byte b))) = Byte b :: nil.
+Proof.
+ intros. rewrite inj_bytes_1. do 2 f_equal. apply byte_of_int_of_byte.
+Qed.
+
+Lemma boidl_init_ints8: forall l,
+ boidl (map Init_int8 l) = inj_bytes (map byte_of_int l).
+Proof.
+ induction l as [ | i l]; simpl. auto. rewrite inj_bytes_1; simpl. f_equal; auto.
+Qed.
+
+Lemma boidl_init_bytes: forall l,
+ boidl (map Init_byte l) = inj_bytes l.
+Proof.
+ induction l as [ | b l]; simpl. auto. rewrite inj_bytes_byte, IHl. auto.
+Qed.
+
+Lemma boidl_ints8: forall i n,
+ boidl (repeat (Init_int8 i) n) = repeat (Byte (byte_of_int i)) n.
+Proof.
+ induction n; simpl. auto. rewrite inj_bytes_1. simpl; f_equal; auto.
+Qed.
+
+(** ** Properties of operations over list of initialization data *)
+
+Lemma add_rev_bytes_spec: forall l il,
+ add_rev_bytes l il = List.map Init_byte (List.rev l) ++ il.
+Proof.
+ induction l as [ | b l]; intros; simpl.
+- auto.
+- rewrite IHl. rewrite map_app. simpl. rewrite app_ass. auto.
+Qed.
+
+Lemma add_rev_bytes_spec': forall l il,
+ List.rev (add_rev_bytes l il) = List.rev il ++ List.map Init_byte l.
+Proof.
+ intros. rewrite add_rev_bytes_spec. rewrite rev_app_distr, map_rev, rev_involutive. auto.
+Qed.
+
+Lemma add_zeros_spec: forall n il,
+ 0 <= n ->
+ add_zeros n il = List.repeat (Init_int8 Int.zero) (Z.to_nat n) ++ il.
+Proof.
+ intros.
+ unfold add_zeros; rewrite iter_nat_of_Z by auto; rewrite Zabs2Nat.abs_nat_nonneg by auto.
+ induction (Z.to_nat n); simpl. auto. f_equal; auto.
+Qed.
+
+Lemma decompose_spec: forall il depth bl il',
+ decompose il depth = OK (bl, il') ->
+ exists nl, il = List.map Init_int8 nl ++ il'
+ /\ bl = List.map byte_of_int (rev nl)
+ /\ List.length nl = Z.to_nat depth.
+Proof.
+ assert (REC: forall il accu depth bl il',
+ decompose_rec accu il depth = OK (bl, il') ->
+ exists nl, il = List.map Init_int8 nl ++ il'
+ /\ bl = List.map byte_of_int (rev nl) ++ accu
+ /\ List.length nl = Z.to_nat depth).
+ { induction il as [ | i il ]; intros until il'; intros D; simpl in D.
+ - destruct (zle depth 0); inv D.
+ exists (@nil int); simpl. rewrite Z_to_nat_neg by auto. auto.
+ - destruct (zle depth 0).
+ + inv D. exists (@nil int); simpl. rewrite Z_to_nat_neg by auto. auto.
+ + destruct i; try discriminate.
+ apply IHil in D; destruct D as (nl & P & Q & R).
+ exists (i :: nl); simpl; split. congruence. split.
+ rewrite map_app. simpl. rewrite app_ass. exact Q.
+ rewrite R, <- Z2Nat.inj_succ by lia. f_equal; lia.
+ }
+ intros. apply REC in H. destruct H as (nl & P & Q & R). rewrite app_nil_r in Q.
+ exists nl; auto.
+Qed.
+
+Lemma list_repeat_app: forall (A: Type) (a: A) n2 n1,
+ List.repeat a n1 ++ List.repeat a n2 = List.repeat a (n1 + n2)%nat.
+Proof.
+ induction n1; simpl; congruence.
+Qed.
+
+Lemma list_rev_repeat: forall (A: Type) (a: A) n,
+ rev (List.repeat a n) = List.repeat a n.
+Proof.
+ induction n; simpl. auto. rewrite IHn. change (a :: nil) with (repeat a 1%nat).
+ rewrite list_repeat_app. rewrite Nat.add_comm. auto.
+Qed.
+
+Lemma normalize_boidl: forall il depth il',
+ normalize il depth = OK il' ->
+ boidl (rev il') = boidl (rev il).
+Proof.
+ induction il as [ | i il]; simpl; intros depth il' AT.
+- destruct (zle depth 0); inv AT. auto.
+- destruct (zle depth 0). inv AT. auto.
+ destruct i;
+ try (monadInv AT; simpl;
+ rewrite ? add_rev_bytes_spec', ? boidl_rev_cons, ? boidl_app, ? boidl_init_bytes;
+ erewrite IHil by eauto; reflexivity).
+ set (n := Z.max 0 z) in *. destruct (zle n depth); monadInv AT.
+ + rewrite add_zeros_spec, rev_app_distr, ! boidl_app by lia.
+ erewrite IHil by eauto. f_equal.
+ rewrite list_rev_repeat. simpl. rewrite app_nil_r, boidl_ints8.
+ f_equal. unfold n. apply Z.max_case_strong; intros; auto. rewrite ! Z_to_nat_neg by lia. auto.
+ + rewrite add_zeros_spec, rev_app_distr, !boidl_app by lia.
+ simpl. rewrite boidl_rev_cons, list_rev_repeat. simpl.
+ rewrite app_ass, app_nil_r, !boidl_ints8. f_equal.
+ rewrite list_repeat_app. f_equal. rewrite <- Z2Nat.inj_add by lia.
+ unfold n. apply Z.max_case_strong; intros; f_equal; lia.
+Qed.
+
+Lemma trisection_boidl: forall il depth sz bytes1 bytes2 il',
+ trisection il depth sz = OK (bytes1, bytes2, il') ->
+ boidl (rev il) = boidl (rev il') ++ inj_bytes bytes2 ++ inj_bytes bytes1
+ /\ length bytes1 = Z.to_nat depth
+ /\ length bytes2 = Z.to_nat sz.
+Proof.
+ unfold trisection; intros. monadInv H.
+ apply normalize_boidl in EQ. rewrite <- EQ.
+ apply decompose_spec in EQ1. destruct EQ1 as (nl1 & A1 & B1 & C1).
+ apply decompose_spec in EQ0. destruct EQ0 as (nl2 & A2 & B2 & C2).
+ split.
+- rewrite A1, A2, !rev_app_distr, !boidl_app, app_ass.
+ rewrite <- !map_rev, !boidl_init_ints8. rewrite <- B1, <- B2. auto.
+- rewrite B1, B2, !map_length, !rev_length. auto.
+Qed.
+
+Lemma store_init_data_loadbytes:
+ forall m b p i m',
+ Genv.store_init_data ge m b p i = Some m' ->
+ match i with Init_space _ => False | _ => True end ->
+ Mem.loadbytes m' b p (init_data_size i) = Some (boid i).
+Proof.
+ intros; destruct i; simpl in H; try apply (Mem.loadbytes_store_same _ _ _ _ _ _ H).
+- contradiction.
+- rewrite Genv.init_data_size_addrof. simpl.
+ destruct (Genv.find_symbol ge i) as [b'|]; try discriminate.
+ rewrite (Mem.loadbytes_store_same _ _ _ _ _ _ H).
+ unfold encode_val, Mptr; destruct Archi.ptr64; reflexivity.
+Qed.
+
+(** ** Validity and size of initialization data *)
+
+Definition idvalid (i: init_data) : Prop :=
+ match i with
+ | Init_addrof symb ofs => exists b, Genv.find_symbol ge symb = Some b
+ | _ => True
+ end.
+
+Fixpoint idlvalid (p: Z) (il: list init_data) {struct il} : Prop :=
+ match il with
+ | nil => True
+ | i1 :: il =>
+ (Genv.init_data_alignment i1 | p)
+ /\ idvalid i1
+ /\ idlvalid (p + init_data_size i1) il
+ end.
+
+Lemma idlvalid_app: forall l2 l1 pos,
+ idlvalid pos (l1 ++ l2) <-> idlvalid pos l1 /\ idlvalid (pos + init_data_list_size l1) l2.
+Proof.
+ induction l1 as [ | d l1]; intros; simpl.
+- rewrite Z.add_0_r; tauto.
+- rewrite IHl1. rewrite Z.add_assoc. tauto.
+Qed.
+
+Lemma add_rev_bytes_valid: forall il bl,
+ idlvalid 0 (rev il) -> idlvalid 0 (rev (add_rev_bytes bl il)).
+Proof.
+ intros. rewrite add_rev_bytes_spec, rev_app_distr, idlvalid_app. split; auto.
+ generalize (rev bl) (0 + init_data_list_size (rev il)). induction l; simpl; intros.
+ auto.
+ rewrite idlvalid_app; split; auto. simpl. auto using Z.divide_1_l.
+Qed.
+
+Lemma add_zeros_valid: forall il n,
+ 0 <= n -> idlvalid 0 (rev il) -> idlvalid 0 (rev (add_zeros n il)).
+Proof.
+ intros. rewrite add_zeros_spec, rev_app_distr, idlvalid_app by auto.
+ split; auto.
+ generalize (Z.to_nat n) (0 + init_data_list_size (rev il)). induction n0; simpl; intros.
+ auto.
+ rewrite idlvalid_app; split; auto. simpl. auto using Z.divide_1_l.
+Qed.
-- destruct il; intros until res; intros TR; simpl in TR.
-+ destruct fl; inv TR. econstructor; split. constructor. apply transl_padding_spec.
-+ destruct fl as [ | [f1 ty1] fl ]; monadInv TR.
- destruct (transl_init_rec_spec _ _ _ _ EQ) as (d1 & A1 & B1).
- destruct (transl_init_struct_spec _ _ _ _ _ _ EQ0) as (d2 & A2 & B2).
- exists (tr_padding pos (align pos (alignof ge ty1)) ++ d1 ++ d2); split.
- econstructor; eauto.
- rewrite ! rev_app_distr. subst res x. rewrite ! app_ass. rewrite transl_padding_spec. auto.
+Lemma normalize_valid: forall il depth il',
+ normalize il depth = OK il' -> idlvalid 0 (rev il) -> idlvalid 0 (rev il').
+Proof.
+ induction il as [ | i il]; simpl; intros.
+- destruct (zle depth 0); inv H. simpl. tauto.
+- destruct (zle depth 0). inv H. auto.
+ rewrite idlvalid_app in H0; destruct H0.
+ destruct i; try (monadInv H; apply add_rev_bytes_valid; eapply IHil; eauto).
+ + monadInv H. simpl. rewrite idlvalid_app; split. eauto. simpl; auto using Z.divide_1_l.
+ + destruct (zle (Z.max 0 z)); monadInv H.
+ * apply add_zeros_valid. lia. eauto.
+ * apply add_zeros_valid. lia. simpl. rewrite idlvalid_app; split. auto. simpl; auto using Z.divide_1_l.
Qed.
-Theorem transl_init_spec:
- forall ty i d, transl_init ge ty i = OK d -> tr_init ty i d.
+Lemma trisection_valid: forall il depth sz bytes1 bytes2 il',
+ trisection il depth sz = OK (bytes1, bytes2, il') ->
+ idlvalid 0 (rev il) ->
+ idlvalid 0 (rev il').
Proof.
- unfold transl_init; intros. monadInv H.
- exploit transl_init_rec_spec; eauto. intros (d & A & B).
- subst x. unfold rev'; rewrite <- rev_alt.
- rewrite rev_app_distr; simpl. rewrite rev_involutive. auto.
+ unfold trisection; intros. monadInv H.
+ apply decompose_spec in EQ1. destruct EQ1 as (nl1 & A1 & B1 & C1).
+ apply decompose_spec in EQ0. destruct EQ0 as (nl2 & A2 & B2 & C2).
+ exploit normalize_valid; eauto. rewrite A1, A2, !rev_app_distr, !idlvalid_app.
+ tauto.
+Qed.
+
+Lemma init_data_size_boid: forall i,
+ init_data_size i = Z.of_nat (length (boid i)).
+Proof.
+ intros. destruct i; simpl; rewrite ?length_inj_bytes, ?encode_int_length; auto.
+- rewrite repeat_length. rewrite Z_to_nat_max; auto.
+- destruct (Genv.find_symbol ge i), Archi.ptr64; reflexivity.
+Qed.
+
+Lemma init_data_list_size_boidl: forall il,
+ init_data_list_size il = Z.of_nat (length (boidl il)).
+Proof.
+ induction il as [ | i il]; simpl. auto.
+ rewrite app_length, init_data_size_boid. lia.
+Qed.
+
+Lemma init_data_list_size_app: forall l1 l2,
+ init_data_list_size (l1 ++ l2) = init_data_list_size l1 + init_data_list_size l2.
+Proof.
+ induction l1 as [ | i l1]; intros; simpl. auto. rewrite IHl1; lia.
+Qed.
+
+(** ** Memory areas that are initialized to zeros *)
+
+Definition reads_as_zeros (m: mem) (b: block) (from to: Z) : Prop :=
+ forall i, from <= i < to -> Mem.loadbytes m b i 1 = Some (Byte Byte.zero :: nil).
+
+Lemma reads_as_zeros_mono: forall m b from1 from2 to1 to2,
+ reads_as_zeros m b from1 to1 -> from1 <= from2 -> to2 <= to1 ->
+ reads_as_zeros m b from2 to2.
+Proof.
+ intros; red; intros. apply H; lia.
+Qed.
+
+Remark reads_as_zeros_unchanged:
+ forall (P: block -> Z -> Prop) m b from to m',
+ reads_as_zeros m b from to ->
+ Mem.unchanged_on P m m' ->
+ (forall i, from <= i < to -> P b i) ->
+ reads_as_zeros m' b from to.
+Proof.
+ intros; red; intros. eapply Mem.loadbytes_unchanged_on; eauto.
+ intros; apply H1. lia.
+Qed.
+
+Lemma reads_as_zeros_loadbytes: forall m b from to,
+ reads_as_zeros m b from to ->
+ forall len pos, from <= pos -> pos + len <= to -> 0 <= len ->
+ Mem.loadbytes m b pos len = Some (repeat (Byte Byte.zero) (Z.to_nat len)).
+Proof.
+ intros until to; intros RZ.
+ induction len using (well_founded_induction (Zwf_well_founded 0)).
+ intros. destruct (zeq len 0).
+- subst len. rewrite Mem.loadbytes_empty by lia. auto.
+- replace (Z.to_nat len) with (S (Z.to_nat (len - 1))).
+ change (repeat (Byte Byte.zero) (S (Z.to_nat (len - 1))))
+ with ((Byte Byte.zero :: nil) ++ repeat (Byte Byte.zero) (Z.to_nat (len - 1))).
+ replace len with (1 + (len - 1)) at 1 by lia.
+ apply Mem.loadbytes_concat; try lia.
+ + apply RZ. lia.
+ + apply H; unfold Zwf; lia.
+ + rewrite <- Z2Nat.inj_succ by lia. f_equal; lia.
+Qed.
+
+Lemma reads_as_zeros_equiv: forall m b from to,
+ reads_as_zeros m b from to <-> Genv.readbytes_as_zero m b from (to - from).
+Proof.
+ intros; split; intros.
+- red; intros. set (len := Z.of_nat n).
+ replace n with (Z.to_nat len) by apply Nat2Z.id.
+ eapply reads_as_zeros_loadbytes; eauto. lia. lia.
+- red; intros. red in H. apply (H i 1%nat). lia. lia.
+Qed.
+
+(** ** Semantic correctness of state operations *)
+
+(** Semantic interpretation of states. *)
+
+Record match_state (s: state) (m: mem) (b: block) : Prop := {
+ match_range:
+ 0 <= s.(curr) <= s.(total_size);
+ match_contents:
+ Mem.loadbytes m b 0 s.(curr) = Some (boidl (rev s.(init)));
+ match_valid:
+ idlvalid 0 (rev s.(init));
+ match_uninitialized:
+ reads_as_zeros m b s.(curr) s.(total_size)
+}.
+
+Lemma match_size: forall s m b,
+ match_state s m b ->
+ init_data_list_size (rev s.(init)) = s.(curr).
+Proof.
+ intros. rewrite init_data_list_size_boidl.
+ erewrite Mem.loadbytes_length by (eapply match_contents; eauto).
+ apply Z2Nat.id. eapply match_range; eauto.
+Qed.
+
+Lemma curr_pad_to: forall s pos,
+ curr s <= curr (pad_to s pos) /\ pos <= curr (pad_to s pos).
+Proof.
+ unfold pad_to; intros. destruct (zle pos (curr s)); simpl; lia.
+Qed.
+
+Lemma total_size_pad_to: forall s pos,
+ total_size (pad_to s pos) = total_size s.
+Proof.
+ unfold pad_to; intros. destruct (zle pos (curr s)); auto.
+Qed.
+
+Lemma pad_to_correct: forall pos s m b,
+ match_state s m b -> pos <= s.(total_size) ->
+ match_state (pad_to s pos) m b.
+Proof.
+ intros. unfold pad_to. destruct (zle pos (curr s)); auto.
+ destruct H; constructor; simpl; intros.
+- lia.
+- rewrite boidl_rev_cons. simpl.
+ replace pos with (s.(curr) + (pos - s.(curr))) at 1 by lia.
+ apply Mem.loadbytes_concat; try lia.
+ * auto.
+ * eapply reads_as_zeros_loadbytes; eauto. lia. lia. lia.
+- rewrite idlvalid_app. split; auto. simpl. intuition auto using Z.divide_1_l.
+- eapply reads_as_zeros_mono; eauto; lia.
+Qed.
+
+Lemma trisection_correct: forall s m b pos sz bytes1 bytes2 il,
+ match_state s m b ->
+ trisection s.(init) (s.(curr) - (pos + sz)) sz = OK (bytes1, bytes2, il) ->
+ 0 <= pos -> pos + sz <= s.(curr) -> 0 <= sz ->
+ Mem.loadbytes m b 0 pos = Some (boidl (rev il))
+ /\ Mem.loadbytes m b pos sz = Some (inj_bytes bytes2)
+ /\ Mem.loadbytes m b (pos + sz) (s.(curr) - (pos + sz)) = Some (inj_bytes bytes1).
+Proof.
+ intros. apply trisection_boidl in H0. destruct H0 as (A & B & C).
+ set (depth := curr s - (pos + sz)) in *.
+ pose proof (match_contents _ _ _ H) as D.
+ replace (curr s) with ((pos + sz) + depth) in D by lia.
+ exploit Mem.loadbytes_split. eexact D. lia. lia.
+ rewrite Z.add_0_l. intros (bytes0 & bytes1' & LB0 & LB1 & E1).
+ exploit Mem.loadbytes_split. eexact LB0. lia. lia.
+ rewrite Z.add_0_l. intros (bytes3 & bytes2' & LB3 & LB2 & E2).
+ rewrite A in E1. rewrite <- app_ass in E1.
+ exploit list_append_injective_r. eexact E1.
+ { unfold inj_bytes; rewrite map_length. erewrite Mem.loadbytes_length; eauto. }
+ intros (E3 & E4).
+ rewrite E2 in E3.
+ exploit list_append_injective_r. eexact E3.
+ { unfold inj_bytes; rewrite map_length. erewrite Mem.loadbytes_length; eauto. }
+ intros (E5 & E6).
+ intuition congruence.
+Qed.
+
+Remark decode_int_zero_ext: forall n bytes,
+ 0 <= n <= 4 -> n = Z.of_nat (length bytes) ->
+ Int.zero_ext (n * 8) (Int.repr (decode_int bytes)) = Int.repr (decode_int bytes).
+Proof.
+ intros.
+ assert (0 <= decode_int bytes < two_p (n * 8)).
+ { rewrite H0. replace (length bytes) with (length (rev_if_be bytes)).
+ apply int_of_bytes_range.
+ apply rev_if_be_length. }
+ assert (two_p (n * 8) <= Int.modulus).
+ { apply (two_p_monotone (n * 8) 32); lia. }
+ unfold Int.zero_ext.
+ rewrite Int.unsigned_repr by (unfold Int.max_unsigned; lia).
+ rewrite Zbits.Zzero_ext_mod by lia.
+ rewrite Zmod_small by auto. auto.
+Qed.
+
+Theorem load_int_correct: forall s m b pos isz i v,
+ match_state s m b ->
+ load_int s pos isz = OK i ->
+ Mem.load (chunk_for_carrier isz) m b pos = Some v ->
+ v = Vint i.
+Proof.
+ intros until v; intros MS RI LD.
+ exploit Mem.load_valid_access. eauto. intros [PERM ALIGN].
+ unfold load_int in RI.
+ set (chunk := chunk_for_carrier isz) in *.
+ set (sz := size_chunk chunk) in *.
+ assert (sz > 0) by (apply size_chunk_pos).
+ set (s1 := pad_to s (pos + sz)) in *.
+ assert (pos + sz <= curr s1) by (apply curr_pad_to).
+ monadInv RI. InvBooleans. destruct x as [[bytes1 bytes2] il].
+ assert (MS': match_state s1 m b) by (apply pad_to_correct; auto).
+ exploit trisection_correct; eauto. lia.
+ intros (L1 & L2 & L3).
+ assert (LEN: Z.of_nat (length bytes2) = sz).
+ { apply Mem.loadbytes_length in L2. unfold inj_bytes in L2.
+ rewrite map_length in L2. rewrite L2. apply Z2Nat.id; lia. }
+ exploit Mem.loadbytes_load. eexact L2. exact ALIGN. rewrite LD.
+ unfold decode_val. rewrite proj_inj_bytes. intros E; inv E; inv EQ0.
+ unfold chunk, chunk_for_carrier; destruct isz; f_equal.
+ - apply (decode_int_zero_ext 1). lia. auto.
+ - apply (decode_int_zero_ext 2). lia. auto.
+ - apply (decode_int_zero_ext 1). lia. auto.
+Qed.
+
+Remark loadbytes_concat_3: forall m b ofs1 len1 l1 ofs2 len2 l2 ofs3 len3 l3 len,
+ Mem.loadbytes m b ofs1 len1 = Some l1 ->
+ Mem.loadbytes m b ofs2 len2 = Some l2 ->
+ Mem.loadbytes m b ofs3 len3 = Some l3 ->
+ ofs2 = ofs1 + len1 -> ofs3 = ofs2 + len2 -> 0 <= len1 -> 0 <= len2 -> 0 <= len3 ->
+ len = len1 + len2 + len3 ->
+ Mem.loadbytes m b ofs1 len = Some (l1 ++ l2 ++ l3).
+Proof.
+ intros. rewrite H7, <- Z.add_assoc. apply Mem.loadbytes_concat. auto.
+ apply Mem.loadbytes_concat. rewrite <- H2; auto. rewrite <- H2, <- H3; auto.
+ lia. lia. lia. lia.
+Qed.
+
+Theorem store_data_correct: forall s m b pos i s' m',
+ match_state s m b ->
+ store_data s pos i = OK s' ->
+ Genv.store_init_data ge m b pos i = Some m' ->
+ match i with Init_space _ => False | _ => True end ->
+ match_state s' m' b.
+Proof.
+ intros until m'; intros MS ST SI NOSPACE.
+ exploit Genv.store_init_data_aligned; eauto. intros ALIGN.
+ assert (VALID: idvalid i).
+ { destruct i; simpl; auto. simpl in SI. destruct (Genv.find_symbol ge i); try discriminate. exists b0; auto. }
+ unfold store_data in ST.
+ set (sz := init_data_size i) in *.
+ assert (sz >= 0) by (apply init_data_size_pos).
+ set (s1 := pad_to s (pos + sz)) in *.
+ monadInv ST. InvBooleans.
+ assert (U: Mem.unchanged_on (fun b i => ~(pos <= i < pos + sz)) m m').
+ { eapply Genv.store_init_data_unchanged. eauto. tauto. }
+ exploit store_init_data_loadbytes; eauto. fold sz. intros D.
+ destruct (zle (curr s) pos).
+- inv ST.
+ set (il := if zlt (curr s) pos then Init_space (pos - curr s) :: init s else init s).
+ assert (IL: boidl (rev il) = boidl (rev (init s)) ++ repeat (Byte Byte.zero) (Z.to_nat (pos - curr s))).
+ { unfold il; destruct (zlt (curr s) pos).
+ - simpl rev. rewrite boidl_rev_cons. simpl. auto.
+ - rewrite Z_to_nat_neg by lia. simpl. rewrite app_nil_r; auto.
+ }
+ constructor; simpl; intros.
+ + lia.
+ + rewrite boidl_rev_cons, IL, app_ass.
+ apply loadbytes_concat_3 with (len1 := curr s) (ofs2 := curr s) (len2 := pos - curr s) (ofs3 := pos) (len3 := sz); try lia.
+ * eapply Mem.loadbytes_unchanged_on; eauto.
+ intros. simpl. lia.
+ eapply match_contents; eauto.
+ * eapply Mem.loadbytes_unchanged_on; eauto.
+ intros. simpl. lia.
+ eapply reads_as_zeros_loadbytes. eapply match_uninitialized; eauto. lia. lia. lia.
+ * exact D.
+ * eapply match_range; eauto.
+ + rewrite idlvalid_app; split.
+ * unfold il; destruct (zlt (curr s) pos).
+ ** simpl; rewrite idlvalid_app; split. eapply match_valid; eauto. simpl. auto using Z.divide_1_l.
+ ** eapply match_valid; eauto.
+ * simpl.
+ replace (init_data_list_size (rev il)) with pos. tauto.
+ unfold il; destruct (zlt (curr s) pos).
+ ** simpl; rewrite init_data_list_size_app; simpl.
+ erewrite match_size by eauto. lia.
+ ** erewrite match_size by eauto. lia.
+ + eapply reads_as_zeros_unchanged; eauto.
+ eapply reads_as_zeros_mono. eapply match_uninitialized; eauto. lia. lia.
+ intros. simpl. lia.
+- monadInv ST. destruct x as [[bytes1 bytes2] il]. inv EQ0.
+ assert (pos + sz <= curr s1) by (apply curr_pad_to).
+ assert (MS': match_state s1 m b) by (apply pad_to_correct; auto).
+ exploit trisection_correct; eauto. lia.
+ intros (L1 & L2 & L3).
+ constructor; simpl; intros.
+ + eapply match_range; eauto.
+ + rewrite add_rev_bytes_spec, rev_app_distr; simpl; rewrite app_ass; simpl.
+ rewrite <- map_rev, rev_involutive.
+ rewrite boidl_app. simpl. rewrite boidl_init_bytes.
+ apply loadbytes_concat_3 with (len1 := pos) (ofs2 := pos) (len2 := sz) (ofs3 := pos + sz)
+ (len3 := curr s1 - (pos + sz)); try lia.
+ * eapply Mem.loadbytes_unchanged_on; eauto.
+ intros. simpl. lia.
+ * exact D.
+ * eapply Mem.loadbytes_unchanged_on; eauto.
+ intros. simpl. lia.
+ + apply add_rev_bytes_valid. simpl; rewrite idlvalid_app; split.
+ * eapply trisection_valid; eauto. eapply match_valid; eauto.
+ * rewrite init_data_list_size_boidl. erewrite Mem.loadbytes_length by eauto.
+ rewrite Z2Nat.id by lia. simpl. tauto.
+ + eapply reads_as_zeros_unchanged; eauto. eapply match_uninitialized; eauto.
+ intros. simpl. lia.
+Qed.
+
+Corollary store_int_correct: forall s m b pos isz n s' m',
+ match_state s m b ->
+ store_int s pos isz n = OK s' ->
+ Mem.store (chunk_for_carrier isz) m b pos (Vint n) = Some m' ->
+ match_state s' m' b.
+Proof.
+ intros. eapply store_data_correct; eauto.
+- destruct isz; exact H1.
+- destruct isz; exact I.
+Qed.
+
+Theorem init_data_list_of_state_correct: forall s m b il b' m1,
+ match_state s m b ->
+ init_data_list_of_state s = OK il ->
+ Mem.range_perm m1 b' 0 s.(total_size) Cur Writable ->
+ reads_as_zeros m1 b' 0 s.(total_size) ->
+ exists m2,
+ Genv.store_init_data_list ge m1 b' 0 il = Some m2
+ /\ Mem.loadbytes m2 b' 0 (init_data_list_size il) = Mem.loadbytes m b 0 s.(total_size).
+Proof.
+ intros. unfold init_data_list_of_state in H0; monadInv H0. rename l into LE.
+ set (s1 := pad_to s s.(total_size)) in *.
+ assert (MS1: match_state s1 m b) by (apply pad_to_correct; auto; lia).
+ apply reads_as_zeros_equiv in H2. rewrite Z.sub_0_r in H2.
+ assert (R: rev' (init s1) = rev (init s1)).
+ { unfold rev'. rewrite <- rev_alt. auto. }
+ assert (C: curr s1 = total_size s).
+ { unfold s1, pad_to. destruct zle; simpl; lia. }
+ assert (A: Genv.init_data_list_aligned 0 (rev (init s1))).
+ { exploit match_valid; eauto. generalize (rev (init s1)) 0.
+ induction l as [ | i l]; simpl; intuition. }
+ assert (B: forall id ofs, In (Init_addrof id ofs) (rev (init s1)) ->
+ exists b, Genv.find_symbol ge id = Some b).
+ { intros id ofs. exploit match_valid; eauto. generalize (rev (init s1)) 0.
+ induction l as [ | i l]; simpl; intuition eauto. subst i; assumption. }
+ exploit Genv.store_init_data_list_exists.
+ 2: eexact A. 2: eexact B.
+ erewrite match_size by eauto. rewrite C. eauto.
+ intros (m2 & ST). exists m2; split.
+- rewrite R. auto.
+- rewrite R. transitivity (Some (boidl (rev (init s1)))).
+ + eapply Genv.store_init_data_list_loadbytes; eauto.
+ erewrite match_size, C by eauto. auto.
+ + symmetry. rewrite <- C. eapply match_contents; eauto.
+Qed.
+
+(** ** Total size properties *)
+
+Lemma total_size_store_data: forall s pos i s',
+ store_data s pos i = OK s' -> total_size s' = total_size s.
+Proof.
+ unfold store_data; intros. monadInv H. destruct (zle (curr s) pos); monadInv H.
+- auto.
+- destruct x as [[bytes1 bytes2] il2]. inv EQ0. simpl. apply total_size_pad_to.
+Qed.
+
+Lemma total_size_transl_init_bitfield: forall ce s ty sz p w i pos s',
+ transl_init_bitfield ce s ty sz p w i pos = OK s' -> total_size s' = total_size s.
+Proof.
+ unfold transl_init_bitfield; intros. destruct i; monadInv H. destruct x; monadInv EQ0.
+ eapply total_size_store_data. eexact EQ2.
+Qed.
+
+Lemma total_size_transl_init_rec: forall ce s ty i pos s',
+ transl_init_rec ce s ty i pos = OK s' -> total_size s' = total_size s
+with total_size_transl_init_array: forall ce s tyelt il pos s',
+ transl_init_array ce s tyelt il pos = OK s' -> total_size s' = total_size s
+with total_size_transl_init_struct: forall ce s ms il base pos s',
+ transl_init_struct ce s ms il base pos = OK s' -> total_size s' = total_size s.
+Proof.
+- destruct i; simpl; intros.
+ + monadInv H; eauto using total_size_store_data.
+ + destruct ty; monadInv H. eauto.
+ + destruct ty; monadInv H. destruct (co_su x); try discriminate. eauto.
+ + destruct ty; monadInv H. destruct (co_su x); monadInv EQ0. destruct x2.
+ * eauto.
+ * eauto using total_size_transl_init_bitfield.
+- destruct il; simpl; intros.
+ + inv H; auto.
+ + monadInv H. transitivity (total_size x); eauto.
+- destruct il; simpl; intros.
+ + inv H; auto.
+ + revert ms pos H. induction ms; intros.
+ * inv H.
+ * destruct (member_not_initialized a). eapply IHms; eauto.
+ monadInv H. transitivity (total_size x1). eauto.
+ destruct x0; eauto using total_size_transl_init_bitfield.
Qed.
(** * Soundness of the translation of initializers *)
(** Soundness for single initializers. *)
-Theorem transl_init_single_steps:
- forall ty a data f m v1 ty1 m' v chunk b ofs m'',
+Inductive exec_assign: mem -> block -> Z -> bitfield -> type -> val -> mem -> Prop :=
+ | exec_assign_full: forall m b ofs ty v m' chunk,
+ access_mode ty = By_value chunk ->
+ Mem.store chunk m b ofs v = Some m' ->
+ exec_assign m b ofs Full ty v m'
+ | exec_assign_bits: forall m b ofs sz sg sg1 attr pos width ty n m' c,
+ type_is_volatile ty = false ->
+ 0 <= pos -> 0 < width -> pos + width <= bitsize_intsize sz ->
+ sg1 = (if zlt width (bitsize_intsize sz) then Signed else sg) ->
+ Mem.load (chunk_for_carrier sz) m b ofs = Some (Vint c) ->
+ Mem.store (chunk_for_carrier sz) m b ofs
+ (Vint (Int.bitfield_insert (first_bit sz pos width) width c n)) = Some m' ->
+ exec_assign m b ofs (Bits sz sg pos width) (Tint sz sg1 attr) (Vint n) m'.
+
+Lemma transl_init_single_sound:
+ forall ty a data f m v1 ty1 m' v b ofs m'',
transl_init_single ge ty a = OK data ->
star step ge (ExprState f a Kstop empty_env m) E0 (ExprState f (Eval v1 ty1) Kstop empty_env m') ->
sem_cast v1 ty1 ty m' = Some v ->
- access_mode ty = By_value chunk ->
- Mem.store chunk m' b ofs v = Some m'' ->
- Genv.store_init_data ge m b ofs data = Some m''.
+ exec_assign m' b ofs Full ty v m'' ->
+ Genv.store_init_data ge m b ofs data = Some m''
+ /\ match data with Init_space _ => False | _ => True end.
Proof.
- intros. monadInv H. monadInv EQ.
+ intros until m''; intros TR STEPS CAST ASG.
+ monadInv TR. monadInv EQ.
exploit constval_steps; eauto. intros [A [B C]]. subst m' ty1.
exploit sem_cast_match; eauto. intros D.
- unfold Genv.store_init_data.
- inv D.
+ inv ASG. rename H into A. unfold Genv.store_init_data. inv D.
- (* int *)
remember Archi.ptr64 as ptr64. destruct ty; try discriminate EQ0.
+ destruct i0; inv EQ0.
- destruct s; simpl in H2; inv H2. rewrite <- Mem.store_signed_unsigned_8; auto. auto.
- destruct s; simpl in H2; inv H2. rewrite <- Mem.store_signed_unsigned_16; auto. auto.
- simpl in H2; inv H2. assumption.
- simpl in H2; inv H2. assumption.
-+ destruct ptr64; inv EQ0. simpl in H2; unfold Mptr in H2; rewrite <- Heqptr64 in H2; inv H2. assumption.
+ destruct s; simpl in A; inv A. rewrite <- Mem.store_signed_unsigned_8; auto. auto.
+ destruct s; simpl in A; inv A. rewrite <- Mem.store_signed_unsigned_16; auto. auto.
+ simpl in A; inv A. auto.
+ simpl in A; inv A. auto.
++ destruct ptr64; inv EQ0. simpl in A; unfold Mptr in A; rewrite <- Heqptr64 in A; inv A. auto.
- (* Long *)
- remember Archi.ptr64 as ptr64. destruct ty; inv EQ0.
-+ simpl in H2; inv H2. assumption.
-+ simpl in H2; unfold Mptr in H2; destruct Archi.ptr64; inv H4.
- inv H2; assumption.
+ remember Archi.ptr64 as ptr64. destruct ty; monadInv EQ0.
++ simpl in A; inv A. auto.
++ simpl in A; unfold Mptr in A; rewrite <- Heqptr64 in A; inv A. auto.
- (* float *)
destruct ty; try discriminate.
- destruct f1; inv EQ0; simpl in H2; inv H2; assumption.
+ destruct f1; inv EQ0; simpl in A; inv A; auto.
- (* single *)
destruct ty; try discriminate.
- destruct f1; inv EQ0; simpl in H2; inv H2; assumption.
+ destruct f1; inv EQ0; simpl in A; inv A; auto.
- (* pointer *)
unfold inj in H.
- assert (data = Init_addrof b1 ofs1 /\ chunk = Mptr).
+ assert (X: data = Init_addrof b1 ofs1 /\ chunk = Mptr).
{ remember Archi.ptr64 as ptr64.
destruct ty; inversion EQ0.
- destruct i; inv H5. unfold Mptr. destruct Archi.ptr64; inv H6; inv H2; auto.
- subst ptr64. unfold Mptr. destruct Archi.ptr64; inv H5; inv H2; auto.
- inv H2. auto. }
- destruct H4; subst. destruct (Genv.find_symbol ge b1); inv H.
- rewrite Ptrofs.add_zero in H3. auto.
+ - destruct i; monadInv H2. unfold Mptr. rewrite <- Heqptr64. inv A; auto.
+ - monadInv H2. unfold Mptr. rewrite <- Heqptr64. inv A; auto.
+ - inv A; auto.
+ }
+ destruct X; subst. destruct (Genv.find_symbol ge b1); inv H.
+ rewrite Ptrofs.add_zero in H0. auto.
- (* undef *)
discriminate.
Qed.
-(** Size properties for initializers. *)
-
-Lemma transl_init_single_size:
- forall ty a data,
- transl_init_single ge ty a = OK data ->
- init_data_size data = sizeof ge ty.
-Proof.
- intros. monadInv H. monadInv EQ. remember Archi.ptr64 as ptr64. destruct x.
-- monadInv EQ0.
-- destruct ty; try discriminate.
- destruct i0; inv EQ0; auto.
- destruct ptr64; inv EQ0.
-Local Transparent sizeof.
- unfold sizeof. rewrite <- Heqptr64; auto.
-- destruct ty; inv EQ0; auto.
- unfold sizeof. destruct Archi.ptr64; inv H0; auto.
-- destruct ty; try discriminate.
- destruct f0; inv EQ0; auto.
-- destruct ty; try discriminate.
- destruct f0; inv EQ0; auto.
-- destruct ty; try discriminate.
- destruct i0; inv EQ0; auto.
- destruct Archi.ptr64 eqn:SF; inv H0. simpl. rewrite SF; auto.
- destruct ptr64; inv EQ0. simpl. rewrite <- Heqptr64; auto.
- inv EQ0. unfold init_data_size, sizeof. auto.
-Qed.
-
-Notation idlsize := init_data_list_size.
-
-Remark padding_size:
- forall frm to, frm <= to -> idlsize (tr_padding frm to) = to - frm.
-Proof.
- unfold tr_padding; intros. destruct (zlt frm to).
- simpl. extlia.
- simpl. lia.
-Qed.
-
-Remark idlsize_app:
- forall d1 d2, idlsize (d1 ++ d2) = idlsize d1 + idlsize d2.
-Proof.
- induction d1; simpl; intros.
- auto.
- rewrite IHd1. lia.
-Qed.
-
-Remark union_field_size:
- forall f ty fl, field_type f fl = OK ty -> sizeof ge ty <= sizeof_union ge fl.
-Proof.
- induction fl as [|[i t]]; simpl; intros.
-- inv H.
-- destruct (ident_eq f i).
- + inv H. extlia.
- + specialize (IHfl H). extlia.
-Qed.
-
-Hypothesis ce_consistent: composite_env_consistent ge.
-
-Lemma tr_init_size:
- forall i ty data,
- tr_init ty i data ->
- idlsize data = sizeof ge ty
-with tr_init_array_size:
- forall ty il sz data,
- tr_init_array ty il sz data ->
- idlsize data = sizeof ge ty * sz
-with tr_init_struct_size:
- forall ty fl il pos data,
- tr_init_struct ty fl il pos data ->
- sizeof_struct ge pos fl <= sizeof ge ty ->
- idlsize data + pos = sizeof ge ty.
-Proof.
-Local Opaque sizeof.
-- destruct 1; simpl.
-+ erewrite transl_init_single_size by eauto. lia.
-+ Local Transparent sizeof. simpl. eapply tr_init_array_size; eauto.
-+ replace (idlsize d) with (idlsize d + 0) by lia.
- eapply tr_init_struct_size; eauto. simpl.
- unfold lookup_composite in H. destruct (ge.(genv_cenv)!id) as [co'|] eqn:?; inv H.
- erewrite co_consistent_sizeof by (eapply ce_consistent; eauto).
- unfold sizeof_composite. rewrite H0. apply align_le.
- destruct (co_alignof_two_p co) as [n EQ]. rewrite EQ. apply two_power_nat_pos.
-+ rewrite idlsize_app, padding_size.
- exploit tr_init_size; eauto. intros EQ; rewrite EQ. lia.
- simpl. unfold lookup_composite in H. destruct (ge.(genv_cenv)!id) as [co'|] eqn:?; inv H.
- apply Z.le_trans with (sizeof_union ge (co_members co)).
- eapply union_field_size; eauto.
- erewrite co_consistent_sizeof by (eapply ce_consistent; eauto).
- unfold sizeof_composite. rewrite H0. apply align_le.
- destruct (co_alignof_two_p co) as [n EQ]. rewrite EQ. apply two_power_nat_pos.
-
-- destruct 1; simpl.
-+ lia.
-+ rewrite Z.mul_comm.
- assert (0 <= sizeof ge ty * sz).
- { apply Zmult_gt_0_le_0_compat. lia. generalize (sizeof_pos ge ty); lia. }
- extlia.
-+ rewrite idlsize_app.
- erewrite tr_init_size by eauto.
- erewrite tr_init_array_size by eauto.
- ring.
-
-- destruct 1; simpl; intros.
-+ rewrite padding_size by auto. lia.
-+ rewrite ! idlsize_app, padding_size.
- erewrite tr_init_size by eauto.
- rewrite <- (tr_init_struct_size _ _ _ _ _ H0 H1). lia.
- unfold pos1. apply align_le. apply alignof_pos.
-Qed.
+(* Hypothesis ce_consistent: composite_env_consistent ge. *)
(** A semantics for general initializers *)
Definition dummy_function := mkfunction Tvoid cc_default nil nil Sskip.
-Fixpoint fields_of_struct (fl: members) (pos: Z) : list (Z * type) :=
- match fl with
- | nil => nil
- | (id1, ty1) :: fl' =>
- (align pos (alignof ge ty1), ty1) :: fields_of_struct fl' (align pos (alignof ge ty1) + sizeof ge ty1)
+Fixpoint initialized_fields_of_struct (ms: members) (pos: Z) : res (list (Z * bitfield * type)) :=
+ match ms with
+ | nil =>
+ OK nil
+ | m :: ms' =>
+ let pos' := next_field ge.(genv_cenv) pos m in
+ if member_not_initialized m
+ then initialized_fields_of_struct ms' pos'
+ else
+ do ofs_bf <- layout_field ge.(genv_cenv) pos m;
+ do l <- initialized_fields_of_struct ms' pos';
+ OK ((ofs_bf, type_member m) :: l)
end.
-Inductive exec_init: mem -> block -> Z -> type -> initializer -> mem -> Prop :=
- | exec_init_single: forall m b ofs ty a v1 ty1 chunk m' v m'',
+Inductive exec_init: mem -> block -> Z -> bitfield -> type -> initializer -> mem -> Prop :=
+ | exec_init_single_: forall m b ofs bf ty a v1 ty1 m' v m'',
star step ge (ExprState dummy_function a Kstop empty_env m)
E0 (ExprState dummy_function (Eval v1 ty1) Kstop empty_env m') ->
sem_cast v1 ty1 ty m' = Some v ->
- access_mode ty = By_value chunk ->
- Mem.store chunk m' b ofs v = Some m'' ->
- exec_init m b ofs ty (Init_single a) m''
+ exec_assign m' b ofs bf ty v m'' ->
+ exec_init m b ofs bf ty (Init_single a) m''
| exec_init_array_: forall m b ofs ty sz a il m',
exec_init_array m b ofs ty sz il m' ->
- exec_init m b ofs (Tarray ty sz a) (Init_array il) m'
- | exec_init_struct: forall m b ofs id a il co m',
+ exec_init m b ofs Full (Tarray ty sz a) (Init_array il) m'
+ | exec_init_struct_: forall m b ofs id a il co flds m',
ge.(genv_cenv)!id = Some co -> co_su co = Struct ->
- exec_init_list m b ofs (fields_of_struct (co_members co) 0) il m' ->
- exec_init m b ofs (Tstruct id a) (Init_struct il) m'
- | exec_init_union: forall m b ofs id a f i ty co m',
+ initialized_fields_of_struct (co_members co) 0 = OK flds ->
+ exec_init_struct m b ofs flds il m' ->
+ exec_init m b ofs Full (Tstruct id a) (Init_struct il) m'
+ | exec_init_union_: forall m b ofs id a f i co ty pos bf m',
ge.(genv_cenv)!id = Some co -> co_su co = Union ->
field_type f (co_members co) = OK ty ->
- exec_init m b ofs ty i m' ->
- exec_init m b ofs (Tunion id a) (Init_union f i) m'
+ union_field_offset ge f (co_members co) = OK (pos, bf) ->
+ exec_init m b (ofs + pos) bf ty i m' ->
+ exec_init m b ofs Full (Tunion id a) (Init_union f i) m'
with exec_init_array: mem -> block -> Z -> type -> Z -> initializer_list -> mem -> Prop :=
| exec_init_array_nil: forall m b ofs ty sz,
sz >= 0 ->
exec_init_array m b ofs ty sz Init_nil m
| exec_init_array_cons: forall m b ofs ty sz i1 il m' m'',
- exec_init m b ofs ty i1 m' ->
+ exec_init m b ofs Full ty i1 m' ->
exec_init_array m' b (ofs + sizeof ge ty) ty (sz - 1) il m'' ->
exec_init_array m b ofs ty sz (Init_cons i1 il) m''
-with exec_init_list: mem -> block -> Z -> list (Z * type) -> initializer_list -> mem -> Prop :=
- | exec_init_list_nil: forall m b ofs,
- exec_init_list m b ofs nil Init_nil m
- | exec_init_list_cons: forall m b ofs pos ty l i1 il m' m'',
- exec_init m b (ofs + pos) ty i1 m' ->
- exec_init_list m' b ofs l il m'' ->
- exec_init_list m b ofs ((pos, ty) :: l) (Init_cons i1 il) m''.
+with exec_init_struct: mem -> block -> Z -> list (Z * bitfield * type) -> initializer_list -> mem -> Prop :=
+ | exec_init_struct_nil: forall m b ofs,
+ exec_init_struct m b ofs nil Init_nil m
+ | exec_init_struct_cons: forall m b ofs pos bf ty l i1 il m' m'',
+ exec_init m b (ofs + pos) bf ty i1 m' ->
+ exec_init_struct m' b ofs l il m'' ->
+ exec_init_struct m b ofs ((pos, bf, ty) :: l) (Init_cons i1 il) m''.
Scheme exec_init_ind3 := Minimality for exec_init Sort Prop
with exec_init_array_ind3 := Minimality for exec_init_array Sort Prop
- with exec_init_list_ind3 := Minimality for exec_init_list Sort Prop.
-Combined Scheme exec_init_scheme from exec_init_ind3, exec_init_array_ind3, exec_init_list_ind3.
+ with exec_init_struct_ind3 := Minimality for exec_init_struct Sort Prop.
+Combined Scheme exec_init_scheme from exec_init_ind3, exec_init_array_ind3, exec_init_struct_ind3.
Remark exec_init_array_length:
forall m b ofs ty sz il m',
@@ -809,87 +1256,95 @@ Proof.
induction 1; lia.
Qed.
-Lemma store_init_data_list_app:
- forall data1 m b ofs m' data2 m'',
- Genv.store_init_data_list ge m b ofs data1 = Some m' ->
- Genv.store_init_data_list ge m' b (ofs + idlsize data1) data2 = Some m'' ->
- Genv.store_init_data_list ge m b ofs (data1 ++ data2) = Some m''.
-Proof.
- induction data1; simpl; intros.
- inv H. rewrite Z.add_0_r in H0. auto.
- destruct (Genv.store_init_data ge m b ofs a); try discriminate.
- rewrite Z.add_assoc in H0. eauto.
-Qed.
-
-Remark store_init_data_list_padding:
- forall frm to b ofs m,
- Genv.store_init_data_list ge m b ofs (tr_padding frm to) = Some m.
+Lemma transl_init_rec_sound:
+ (forall m b ofs bf ty i m',
+ exec_init m b ofs bf ty i m' ->
+ forall s s',
+ match_state s m b ->
+ match bf with
+ | Full => transl_init_rec ge s ty i ofs
+ | Bits sz sg p w => transl_init_bitfield ge s ty sz p w i ofs
+ end = OK s' ->
+ match_state s' m' b)
+/\ (forall m b ofs ty sz il m',
+ exec_init_array m b ofs ty sz il m' ->
+ forall s s',
+ match_state s m b ->
+ transl_init_array ge s ty il ofs = OK s' ->
+ match_state s' m' b)
+/\ (forall m b ofs flds il m',
+ exec_init_struct m b ofs flds il m' ->
+ forall s s' ms pos,
+ match_state s m b ->
+ initialized_fields_of_struct ms pos = OK flds ->
+ transl_init_struct ge s ms il ofs pos = OK s' ->
+ match_state s' m' b).
Proof.
- intros. unfold tr_padding. destruct (zlt frm to); auto.
-Qed.
-
-Lemma tr_init_sound:
- (forall m b ofs ty i m', exec_init m b ofs ty i m' ->
- forall data, tr_init ty i data ->
- Genv.store_init_data_list ge m b ofs data = Some m')
-/\(forall m b ofs ty sz il m', exec_init_array m b ofs ty sz il m' ->
- forall data, tr_init_array ty il sz data ->
- Genv.store_init_data_list ge m b ofs data = Some m')
-/\(forall m b ofs l il m', exec_init_list m b ofs l il m' ->
- forall ty fl data pos,
- l = fields_of_struct fl pos ->
- tr_init_struct ty fl il pos data ->
- Genv.store_init_data_list ge m b (ofs + pos) data = Some m').
-Proof.
-Local Opaque sizeof.
- apply exec_init_scheme; simpl; intros.
+ apply exec_init_scheme.
- (* single *)
- inv H3. simpl. erewrite transl_init_single_steps by eauto. auto.
+ intros until m''; intros STEP CAST ASG s s' MS TR. destruct bf; monadInv TR.
+ + (* full *)
+ exploit transl_init_single_sound; eauto. intros [P Q].
+ eapply store_data_correct; eauto.
+ + (* bitfield *)
+ destruct x; monadInv EQ0. monadInv EQ.
+ exploit constval_steps; eauto. intros [A [B C]]. subst m' ty1.
+ exploit sem_cast_match; eauto. intros D.
+ inv ASG. inv D.
+ set (f := first_bit sz pos width) in *.
+ assert (E: Vint c = Vint x) by (eapply load_int_correct; eauto).
+ inv E.
+ eapply store_int_correct; eauto.
- (* array *)
- inv H1. replace (Z.max 0 sz) with sz in H7. eauto.
- assert (sz >= 0) by (eapply exec_init_array_length; eauto). extlia.
+ intros. monadInv H2. eauto.
- (* struct *)
- inv H3. unfold lookup_composite in H7. rewrite H in H7. inv H7.
- replace ofs with (ofs + 0) by lia. eauto.
+ intros. monadInv H5. unfold lookup_composite in EQ. rewrite H in EQ. inv EQ.
+ rewrite H0 in EQ0. eauto.
- (* union *)
- inv H4. unfold lookup_composite in H9. rewrite H in H9. inv H9. rewrite H1 in H12; inv H12.
- eapply store_init_data_list_app. eauto.
- apply store_init_data_list_padding.
-
-- (* array, empty *)
- inv H0; auto.
-- (* array, nonempty *)
- inv H3.
- eapply store_init_data_list_app.
- eauto.
- rewrite (tr_init_size _ _ _ H7). eauto.
-
-- (* struct, empty *)
- inv H0. apply store_init_data_list_padding.
-- (* struct, nonempty *)
- inv H4. simpl in H3; inv H3.
- eapply store_init_data_list_app. apply store_init_data_list_padding.
- rewrite padding_size.
- replace (ofs + pos0 + (pos2 - pos0)) with (ofs + pos2) by lia.
- eapply store_init_data_list_app.
+ intros. monadInv H6. unfold lookup_composite in EQ. rewrite H in EQ. inv EQ. rewrite H0 in EQ0.
+ rewrite H1, H2 in EQ0. simpl in EQ0. eauto.
+- (* array nil *)
+ intros. monadInv H1. auto.
+- (* array cons *)
+ intros. monadInv H4. eauto.
+- (* struct nil *)
+ intros. monadInv H1. auto.
+- (* struct cons *)
+ intros. simpl in H5. revert H4 H5. generalize pos0. induction ms as [ | m1 ms]. discriminate.
+ simpl. destruct (member_not_initialized m1).
+ intros; eapply IHms; eauto.
+ clear IHms. intros. monadInv H5. rewrite EQ in H4. monadInv H4. inv EQ0.
eauto.
- rewrite (tr_init_size _ _ _ H9).
- rewrite <- Z.add_assoc. eapply H2. eauto. eauto.
- apply align_le. apply alignof_pos.
Qed.
End SOUNDNESS.
Theorem transl_init_sound:
- forall p m b ty i m' data,
- exec_init (globalenv p) m b 0 ty i m' ->
+ forall p m b ty i m1 data,
+ let sz := sizeof (prog_comp_env p) ty in
+ Mem.range_perm m b 0 sz Cur Writable ->
+ reads_as_zeros m b 0 sz ->
+ exec_init (globalenv p) m b 0 Full ty i m1 ->
transl_init (prog_comp_env p) ty i = OK data ->
- Genv.store_init_data_list (globalenv p) m b 0 data = Some m'.
+ exists m2,
+ Genv.store_init_data_list (globalenv p) m b 0 data = Some m2
+ /\ Mem.loadbytes m2 b 0 (init_data_list_size data) = Mem.loadbytes m1 b 0 sz.
Proof.
intros.
set (ge := globalenv p) in *.
- change (prog_comp_env p) with (genv_cenv ge) in H0.
- destruct (tr_init_sound ge) as (A & B & C).
- eapply build_composite_env_consistent. apply prog_comp_env_eq.
- eapply A; eauto. apply transl_init_spec; auto.
+ change (prog_comp_env p) with (genv_cenv ge) in *.
+ unfold transl_init in H2; monadInv H2.
+ fold sz in EQ. set (s0 := initial_state sz) in *.
+ assert (match_state ge s0 m b).
+ { constructor; simpl.
+ - generalize (sizeof_pos ge ty). fold sz. lia.
+ - apply Mem.loadbytes_empty. lia.
+ - auto.
+ - assumption.
+ }
+ assert (match_state ge x m1 b).
+ { eapply (proj1 (transl_init_rec_sound ge)); eauto. }
+ assert (total_size x = sz).
+ { change sz with s0.(total_size). eapply total_size_transl_init_rec; eauto. }
+ rewrite <- H4. eapply init_data_list_of_state_correct; eauto; rewrite H4; auto.
Qed.
diff --git a/cfrontend/PrintCsyntax.ml b/cfrontend/PrintCsyntax.ml
index 75c85d60..16cdfc41 100644
--- a/cfrontend/PrintCsyntax.ml
+++ b/cfrontend/PrintCsyntax.ml
@@ -204,7 +204,7 @@ let rec expr p (prec, e) =
then fprintf p "@[<hov 2>("
else fprintf p "@[<hov 2>";
begin match e with
- | Eloc(b, ofs, _) ->
+ | Eloc(b, ofs, _, _) ->
fprintf p "<loc%a>" !print_pointer_hook (b, ofs)
| Evar(id, _) ->
fprintf p "%s" (extern_atom id)
@@ -534,13 +534,18 @@ let struct_or_union = function Struct -> "struct" | Union -> "union"
let declare_composite p (Composite(id, su, m, a)) =
fprintf p "%s %s;@ " (struct_or_union su) (extern_atom id)
+let print_member p = function
+ | Member_plain(id, ty) ->
+ fprintf p "@ %s;" (name_cdecl (extern_atom id) ty)
+ | Member_bitfield(id, sz, sg, attr, w, _is_padding) ->
+ fprintf p "@ %s : %s;"
+ (name_cdecl (extern_atom id) (Tint(sz, sg, attr)))
+ (Z.to_string w)
+
let define_composite p (Composite(id, su, m, a)) =
fprintf p "@[<v 2>%s %s%s {"
(struct_or_union su) (extern_atom id) (attributes a);
- List.iter
- (fun (fid, fty) ->
- fprintf p "@ %s;" (name_cdecl (extern_atom fid) fty))
- m;
+ List.iter (print_member p) m;
fprintf p "@;<0 -2>};@]@ @ "
let print_program p prog =
diff --git a/cfrontend/SimplExpr.v b/cfrontend/SimplExpr.v
index c7e57a54..bb1dbe38 100644
--- a/cfrontend/SimplExpr.v
+++ b/cfrontend/SimplExpr.v
@@ -13,17 +13,8 @@
(** Translation from Compcert C to Clight.
Side effects are pulled out of Compcert C expressions. *)
-Require Import Coqlib.
-Require Import Errors.
-Require Import Integers.
-Require Import Floats.
-Require Import Values.
-Require Import Memory.
-Require Import AST.
-Require Import Ctypes.
-Require Import Cop.
-Require Import Csyntax.
-Require Import Clight.
+Require Import Coqlib Maps Integers Floats Values AST Memory Errors.
+Require Import Ctypes Cop Csyntax Clight.
Local Open Scope string_scope.
Local Open Scope list_scope.
@@ -71,8 +62,6 @@ Notation "'do' ( X , Y ) <- A ; B" := (bind2 A (fun X Y => B))
(at level 200, X ident, Y ident, A at level 100, B at level 200)
: gensym_monad_scope.
-Local Open Scope gensym_monad_scope.
-
Parameter first_unused_ident: unit -> ident.
Definition initial_generator (x: unit) : generator :=
@@ -96,6 +85,12 @@ Fixpoint makeseq_rec (s: statement) (l: list statement) : statement :=
Definition makeseq (l: list statement) : statement :=
makeseq_rec Sskip l.
+Section SIMPL_EXPR.
+
+Local Open Scope gensym_monad_scope.
+
+Variable ce: composite_env.
+
(** Smart constructor for [if ... then ... else]. *)
Fixpoint eval_simpl_expr (a: expr) : option val :=
@@ -144,16 +139,64 @@ Definition transl_incrdecr (id: incr_or_decr) (a: expr) (ty: type) : expr :=
| Decr => Ebinop Osub a (Econst_int Int.one type_int32s) (incrdecr_type ty)
end.
-(** Generate a [Sset] or [Sbuiltin] operation as appropriate
- to dereference a l-value [l] and store its result in temporary variable [id]. *)
+(** Given a simple l-value expression [l], determine whether it
+ designates a bitfield. *)
+
+Definition is_bitfield_access_aux
+ (fn: composite_env -> ident -> members -> res (Z * bitfield))
+ (id: ident) (fld: ident) : mon bitfield :=
+ match ce!id with
+ | None => error (MSG "unknown composite " :: CTX id :: nil)
+ | Some co =>
+ match fn ce fld (co_members co) with
+ | OK (_, bf) => ret bf
+ | Error _ => error (MSG "unknown field " :: CTX fld :: nil)
+ end
+ end.
-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
+Definition is_bitfield_access (l: expr) : mon bitfield :=
+ match l with
+ | Efield r f _ =>
+ match typeof r with
+ | Tstruct id _ => is_bitfield_access_aux field_offset id f
+ | Tunion id _ => is_bitfield_access_aux union_field_offset id f
+ | _ => error (msg "is_bitfield_access")
+ end
+ | _ => ret Full
+ end.
+
+(** According to the CompCert C semantics, an access to a l-value of
+ volatile-qualified type can either
+ - produce an event in the trace of observable events, or
+ - produce no event and behave as if no volatile qualifier was there.
+
+ The latter case, where the volatile qualifier is ignored, happens if
+ - the l-value is a struct or union
+ - the l-value is an access to a bit field.
+
+ The [chunk_for_volatile_type] function distinguishes between the two
+ cases. It returns [Some chunk] if the semantics is to produce
+ an observable event of the [Event_vload chunk] or [Event_vstore chunk]
+ kind. It returns [None] if the semantics is that of a non-volatile
+ access. *)
+
+Definition chunk_for_volatile_type (ty: type) (bf: bitfield) : option memory_chunk :=
+ if type_is_volatile ty then
+ match access_mode ty with
+ | By_value chunk =>
+ match bf with
+ | Full => Some chunk
+ | Bits _ _ _ _ => None
+ end
+ | _ => None
+ end
else None.
-Definition make_set (id: ident) (l: expr) : statement :=
- match chunk_for_volatile_type (typeof l) with
+(** Generate a [Sset] or [Sbuiltin] operation as appropriate
+ to dereference a l-value [l] and store its result in temporary variable [id]. *)
+
+Definition make_set (bf: bitfield) (id: ident) (l: expr) : statement :=
+ match chunk_for_volatile_type (typeof l) bf with
| None => Sset id l
| Some chunk =>
let typtr := Tpointer (typeof l) noattr in
@@ -165,13 +208,15 @@ Definition make_set (id: ident) (l: expr) : statement :=
Definition transl_valof (ty: type) (l: expr) : mon (list statement * expr) :=
if type_is_volatile ty
- then do t <- gensym ty; ret (make_set t l :: nil, Etempvar t ty)
+ then do t <- gensym ty;
+ do bf <- is_bitfield_access l;
+ ret (make_set bf t l :: nil, Etempvar t ty)
else ret (nil, l).
(** Translation of an assignment. *)
-Definition make_assign (l r: expr) : statement :=
- match chunk_for_volatile_type (typeof l) with
+Definition make_assign (bf: bitfield) (l r: expr) : statement :=
+ match chunk_for_volatile_type (typeof l) bf with
| None =>
Sassign l r
| Some chunk =>
@@ -181,6 +226,30 @@ Definition make_assign (l r: expr) : statement :=
(Eaddrof l typtr :: r :: nil)
end.
+(** Translation of the value of an assignment expression.
+ For non-bitfield assignments, it's the value of the right-hand side
+ converted to the type of the left-hand side.
+ For assignments to bitfields, an additional normalization to
+ the width and signedness of the bitfield is required. *)
+
+Definition make_normalize (sz: intsize) (sg: signedness) (width: Z) (r: expr) :=
+ let intconst (n: Z) := Econst_int (Int.repr n) type_int32s in
+ if intsize_eq sz IBool || signedness_eq sg Unsigned then
+ let mask := two_p width - 1 in
+ Ebinop Oand r (intconst mask) (typeof r)
+ else
+ let amount := Int.zwordsize - width in
+ Ebinop Oshr
+ (Ebinop Oshl r (intconst amount) type_int32s)
+ (intconst amount)
+ (typeof r).
+
+Definition make_assign_value (bf: bitfield) (r: expr): expr :=
+ match bf with
+ | Full => r
+ | Bits sz sg pos width => make_normalize sz sg width r
+ end.
+
(** Translation of expressions. Return a pair [(sl, a)] of
a list of statements [sl] and a pure expression [a].
- If the [dst] argument is [For_val], the statements [sl]
@@ -229,7 +298,7 @@ Definition sd_seqbool_set (ty: type) (sd: set_destination) :=
Fixpoint transl_expr (dst: destination) (a: Csyntax.expr) : mon (list statement * expr) :=
match a with
- | Csyntax.Eloc b ofs ty =>
+ | Csyntax.Eloc b ofs bf ty =>
error (msg "SimplExpr.transl_expr: Eloc")
| Csyntax.Evar x ty =>
ret (finish dst nil (Evar x ty))
@@ -335,16 +404,17 @@ Fixpoint transl_expr (dst: destination) (a: Csyntax.expr) : mon (list statement
| Csyntax.Eassign l1 r2 ty =>
do (sl1, a1) <- transl_expr For_val l1;
do (sl2, a2) <- transl_expr For_val r2;
+ do bf <- is_bitfield_access a1;
let ty1 := Csyntax.typeof l1 in
let ty2 := Csyntax.typeof r2 in
match dst with
| For_val | For_set _ =>
do t <- gensym ty1;
ret (finish dst
- (sl1 ++ sl2 ++ Sset t (Ecast a2 ty1) :: make_assign a1 (Etempvar t ty1) :: nil)
- (Etempvar t ty1))
+ (sl1 ++ sl2 ++ Sset t (Ecast a2 ty1) :: make_assign bf a1 (Etempvar t ty1) :: nil)
+ (make_assign_value bf (Etempvar t ty1)))
| For_effects =>
- ret (sl1 ++ sl2 ++ make_assign a1 a2 :: nil,
+ ret (sl1 ++ sl2 ++ make_assign bf a1 a2 :: nil,
dummy_expr)
end
| Csyntax.Eassignop op l1 r2 tyres ty =>
@@ -352,31 +422,33 @@ Fixpoint transl_expr (dst: destination) (a: Csyntax.expr) : mon (list statement
do (sl1, a1) <- transl_expr For_val l1;
do (sl2, a2) <- transl_expr For_val r2;
do (sl3, a3) <- transl_valof ty1 a1;
+ do bf <- is_bitfield_access a1;
match dst with
| For_val | For_set _ =>
do t <- gensym ty1;
ret (finish dst
(sl1 ++ sl2 ++ sl3 ++
Sset t (Ecast (Ebinop op a3 a2 tyres) ty1) ::
- make_assign a1 (Etempvar t ty1) :: nil)
- (Etempvar t ty1))
+ make_assign bf a1 (Etempvar t ty1) :: nil)
+ (make_assign_value bf (Etempvar t ty1)))
| For_effects =>
- ret (sl1 ++ sl2 ++ sl3 ++ make_assign a1 (Ebinop op a3 a2 tyres) :: nil,
+ ret (sl1 ++ sl2 ++ sl3 ++ make_assign bf a1 (Ebinop op a3 a2 tyres) :: nil,
dummy_expr)
end
| Csyntax.Epostincr id l1 ty =>
let ty1 := Csyntax.typeof l1 in
do (sl1, a1) <- transl_expr For_val l1;
+ do bf <- is_bitfield_access a1;
match dst with
| For_val | For_set _ =>
do t <- gensym ty1;
ret (finish dst
- (sl1 ++ make_set t a1 ::
- make_assign a1 (transl_incrdecr id (Etempvar t ty1) ty1) :: nil)
+ (sl1 ++ make_set bf t a1 ::
+ make_assign bf a1 (transl_incrdecr id (Etempvar t ty1) ty1) :: nil)
(Etempvar t ty1))
| For_effects =>
do (sl2, a2) <- transl_valof ty1 a1;
- ret (sl1 ++ sl2 ++ make_assign a1 (transl_incrdecr id a2 ty1) :: nil,
+ ret (sl1 ++ sl2 ++ make_assign bf a1 (transl_incrdecr id a2 ty1) :: nil,
dummy_expr)
end
| Csyntax.Ecomma r1 r2 ty =>
@@ -424,12 +496,6 @@ Definition transl_expression (r: Csyntax.expr) : mon (statement * expr) :=
Definition transl_expr_stmt (r: Csyntax.expr) : mon statement :=
do (sl, a) <- transl_expr For_effects r; ret (makeseq sl).
-(*
-Definition transl_if (r: Csyntax.expr) (s1 s2: statement) : mon statement :=
- do (sl, a) <- transl_expr For_val r;
- ret (makeseq (sl ++ makeif a s1 s2 :: nil)).
-*)
-
Definition transl_if (r: Csyntax.expr) (s1 s2: statement) : mon statement :=
do (sl, a) <- transl_expr For_val r;
ret (makeseq (sl ++ makeif a s1 s2 :: nil)).
@@ -533,8 +599,12 @@ Definition transl_fundef (fd: Csyntax.fundef) : res fundef :=
OK (External ef targs tres cc)
end.
+End SIMPL_EXPR.
+
+Local Open Scope error_monad_scope.
+
Definition transl_program (p: Csyntax.program) : res program :=
- do p1 <- AST.transform_partial_program transl_fundef p;
+ do p1 <- AST.transform_partial_program (transl_fundef p.(prog_comp_env)) p;
OK {| prog_defs := AST.prog_defs p1;
prog_public := AST.prog_public p1;
prog_main := AST.prog_main p1;
diff --git a/cfrontend/SimplExprproof.v b/cfrontend/SimplExprproof.v
index 2d059ddd..67bf0e51 100644
--- a/cfrontend/SimplExprproof.v
+++ b/cfrontend/SimplExprproof.v
@@ -22,7 +22,7 @@ Require Import SimplExpr SimplExprspec.
(** ** Relational specification of the translation. *)
Definition match_prog (p: Csyntax.program) (tp: Clight.program) :=
- match_program (fun ctx f tf => tr_fundef f tf) eq p tp
+ match_program_gen tr_fundef eq p p tp
/\ prog_types tp = prog_types p.
Lemma transf_program_match:
@@ -30,8 +30,9 @@ Lemma transf_program_match:
Proof.
unfold transl_program; intros. monadInv H. split; auto.
unfold program_of_program; simpl. destruct x; simpl.
- eapply match_transform_partial_program_contextual. eexact EQ.
- intros. apply transl_fundef_spec; auto.
+ eapply match_transform_partial_program2; eauto.
+- intros. apply transl_fundef_spec; auto.
+- intros. inv H. auto.
Qed.
(** ** Semantic preservation *)
@@ -65,25 +66,19 @@ Proof (Genv.senv_match (proj1 TRANSL)).
Lemma function_ptr_translated:
forall b f,
Genv.find_funct_ptr ge b = Some f ->
- exists tf,
- Genv.find_funct_ptr tge b = Some tf /\ tr_fundef f tf.
-Proof.
- intros.
- edestruct (Genv.find_funct_ptr_match (proj1 TRANSL)) as (ctx & tf & A & B & C); eauto.
-Qed.
+ exists cu tf,
+ Genv.find_funct_ptr tge b = Some tf /\ tr_fundef cu f tf /\ linkorder cu prog.
+Proof (Genv.find_funct_ptr_match (proj1 TRANSL)).
Lemma functions_translated:
forall v f,
Genv.find_funct ge v = Some f ->
- exists tf,
- Genv.find_funct tge v = Some tf /\ tr_fundef f tf.
-Proof.
- intros.
- edestruct (Genv.find_funct_match (proj1 TRANSL)) as (ctx & tf & A & B & C); eauto.
-Qed.
+ exists cu tf,
+ Genv.find_funct tge v = Some tf /\ tr_fundef cu f tf /\ linkorder cu prog.
+Proof (Genv.find_funct_match (proj1 TRANSL)).
Lemma type_of_fundef_preserved:
- forall f tf, tr_fundef f tf ->
+ forall cu f tf, tr_fundef cu f tf ->
type_of_fundef tf = Csyntax.type_of_fundef f.
Proof.
intros. inv H.
@@ -92,7 +87,7 @@ Proof.
Qed.
Lemma function_return_preserved:
- forall f tf, tr_function f tf ->
+ forall ce f tf, tr_function ce f tf ->
fn_return tf = Csyntax.fn_return f.
Proof.
intros. inv H; auto.
@@ -100,10 +95,16 @@ Qed.
(** Properties of smart constructors. *)
+Section TRANSLATION.
+
+Variable cunit: Csyntax.program.
+Hypothesis LINKORDER: linkorder cunit prog.
+Let ce := cunit.(prog_comp_env).
+
Lemma eval_Ederef':
forall ge e le m a t l ofs,
eval_expr ge e le m a (Vptr l ofs) ->
- eval_lvalue ge e le m (Ederef' a t) l ofs.
+ eval_lvalue ge e le m (Ederef' a t) l ofs Full.
Proof.
intros. unfold Ederef'; destruct a; auto using eval_Ederef.
destruct (type_eq t (typeof a)); auto using eval_Ederef.
@@ -120,7 +121,7 @@ Qed.
Lemma eval_Eaddrof':
forall ge e le m a t l ofs,
- eval_lvalue ge e le m a l ofs ->
+ eval_lvalue ge e le m a l ofs Full ->
eval_expr ge e le m (Eaddrof' a t) (Vptr l ofs).
Proof.
intros. unfold Eaddrof'; destruct a; auto using eval_Eaddrof.
@@ -134,12 +135,45 @@ Proof.
unfold Eaddrof'; intros; destruct a; auto. destruct (type_eq t (typeof a)); auto.
Qed.
+Lemma eval_make_normalize:
+ forall ge e le m a n sz sg sg1 attr width,
+ 0 < width -> width <= bitsize_intsize sz ->
+ typeof a = Tint sz sg1 attr ->
+ eval_expr ge e le m a (Vint n) ->
+ eval_expr ge e le m (make_normalize sz sg width a) (Vint (bitfield_normalize sz sg width n)).
+Proof.
+ intros. unfold make_normalize, bitfield_normalize.
+ assert (bitsize_intsize sz <= Int.zwordsize) by (destruct sz; compute; congruence).
+ destruct (intsize_eq sz IBool || signedness_eq sg Unsigned).
+- rewrite Int.zero_ext_and by lia. econstructor. eauto. econstructor.
+ rewrite H1; simpl. unfold sem_and, sem_binarith.
+ assert (A: exists sg2, classify_binarith (Tint sz sg1 attr) type_int32s = bin_case_i sg2).
+ { unfold classify_binarith. unfold type_int32s. destruct sz, sg1; econstructor; eauto. }
+ destruct A as (sg2 & A); rewrite A.
+ unfold binarith_type.
+ assert (B: forall i sz0 sg0 attr0,
+ sem_cast (Vint i) (Tint sz0 sg0 attr0) (Tint I32 sg2 noattr) m = Some (Vint i)).
+ { intros. unfold sem_cast, classify_cast. destruct Archi.ptr64; reflexivity. }
+ unfold type_int32s; rewrite ! B. auto.
+- rewrite Int.sign_ext_shr_shl by lia.
+ set (amount := Int.repr (Int.zwordsize - width)).
+ assert (LT: Int.ltu amount Int.iwordsize = true).
+ { unfold Int.ltu. rewrite Int.unsigned_repr_wordsize. apply zlt_true.
+ unfold amount; rewrite Int.unsigned_repr. lia.
+ assert (Int.zwordsize < Int.max_unsigned) by reflexivity. lia. }
+ econstructor.
+ econstructor. eauto. econstructor.
+ rewrite H1. unfold sem_binary_operation, sem_shl, sem_shift. rewrite LT. destruct sz, sg1; reflexivity.
+ econstructor.
+ unfold sem_binary_operation, sem_shr, sem_shift. rewrite LT. reflexivity.
+Qed.
+
(** Translation of simple expressions. *)
Lemma tr_simple_nil:
- (forall le dst r sl a tmps, tr_expr le dst r sl a tmps ->
+ (forall le dst r sl a tmps, tr_expr ce le dst r sl a tmps ->
dst = For_val \/ dst = For_effects -> simple r = true -> sl = nil)
-/\(forall le rl sl al tmps, tr_exprlist le rl sl al tmps ->
+/\(forall le rl sl al tmps, tr_exprlist ce le rl sl al tmps ->
simplelist rl = true -> sl = nil).
Proof.
assert (A: forall dst a, dst = For_val \/ dst = For_effects -> final dst a = nil).
@@ -160,52 +194,104 @@ Proof.
Qed.
Lemma tr_simple_expr_nil:
- forall le dst r sl a tmps, tr_expr le dst r sl a tmps ->
+ forall le dst r sl a tmps, tr_expr ce le dst r sl a tmps ->
dst = For_val \/ dst = For_effects -> simple r = true -> sl = nil.
Proof (proj1 tr_simple_nil).
Lemma tr_simple_exprlist_nil:
- forall le rl sl al tmps, tr_exprlist le rl sl al tmps ->
+ forall le rl sl al tmps, tr_exprlist ce le rl sl al tmps ->
simplelist rl = true -> sl = nil.
Proof (proj2 tr_simple_nil).
(** Translation of [deref_loc] and [assign_loc] operations. *)
Remark deref_loc_translated:
- forall 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
+ forall ty m b ofs bf t v,
+ Csem.deref_loc ge ty m b ofs bf t v ->
+ match chunk_for_volatile_type ty bf with
+ | None => t = E0 /\ Clight.deref_loc ty m b ofs bf v
+ | Some chunk => bf = Full /\ volatile_load tge chunk m b ofs t v
end.
Proof.
intros. unfold chunk_for_volatile_type. inv H.
- (* By_value, not volatile *)
+- (* 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. apply senv_preserved.
- (* By reference *)
+- (* By_value, volatile *)
+ rewrite H0, H1. split; auto. eapply volatile_load_preserved with (ge1 := ge); auto. apply senv_preserved.
+- (* By reference *)
rewrite H0. destruct (type_is_volatile ty); split; auto; eapply deref_loc_reference; eauto.
- (* By copy *)
+- (* By copy *)
rewrite H0. destruct (type_is_volatile ty); split; auto; eapply deref_loc_copy; eauto.
+- (* Bitfield *)
+ destruct (type_is_volatile ty); [destruct (access_mode ty)|]; auto using deref_loc_bitfield.
Qed.
Remark assign_loc_translated:
- forall 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 tge ty m b ofs v m'
- | Some chunk => volatile_store tge chunk m b ofs v t m'
+ forall ty m b ofs bf v t m' v',
+ Csem.assign_loc ge ty m b ofs bf v t m' v' ->
+ match chunk_for_volatile_type ty bf with
+ | None => t = E0 /\ Clight.assign_loc tge ty m b ofs bf v m'
+ | Some chunk => bf = Full /\ volatile_store tge chunk m b ofs v t m'
end.
Proof.
intros. unfold chunk_for_volatile_type. inv H.
- (* By_value, not volatile *)
+- (* 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. apply senv_preserved.
- (* By copy *)
+- (* By_value, volatile *)
+ rewrite H0, H1. split; auto. eapply volatile_store_preserved with (ge1 := ge); auto. apply senv_preserved.
+- (* By copy *)
rewrite H0. rewrite <- comp_env_preserved in *.
destruct (type_is_volatile ty); split; auto; eapply assign_loc_copy; eauto.
+- (* Bitfield *)
+ destruct (type_is_volatile ty); [destruct (access_mode ty)|]; eauto using assign_loc_bitfield.
+Qed.
+
+(** Bitfield accesses *)
+
+Lemma is_bitfield_access_sound: forall e le m a b ofs bf bf',
+ eval_lvalue tge e le m a b ofs bf ->
+ tr_is_bitfield_access ce a bf' ->
+ bf' = bf.
+Proof.
+ assert (A: forall id co co',
+ tge.(genv_cenv)!id = Some co -> ce!id = Some co' ->
+ co' = co /\ complete_members ce (co_members co) = true).
+ { intros. rewrite comp_env_preserved in H.
+ assert (ge.(Csem.genv_cenv) ! id = Some co') by (apply LINKORDER; auto).
+ replace co' with co in * by congruence.
+ split; auto. apply co_consistent_complete.
+ eapply build_composite_env_consistent. eapply prog_comp_env_eq. eauto.
+ }
+ induction 1; simpl; auto.
+- rewrite H0. intros (co' & delta' & E1 & E2). rewrite comp_env_preserved in H2.
+ exploit A; eauto. intros (E3 & E4). subst co'.
+ assert (field_offset ge i (co_members co) = field_offset ce i (co_members co)).
+ { apply field_offset_stable. apply LINKORDER. auto. }
+ congruence.
+- rewrite H0. intros (co' & delta' & E1 & E2). rewrite comp_env_preserved in H2.
+ exploit A; eauto. intros (E3 & E4). subst co'.
+ assert (union_field_offset ge i (co_members co) = union_field_offset ce i (co_members co)).
+ { apply union_field_offset_stable. apply LINKORDER. auto. }
+ congruence.
+Qed.
+
+Lemma make_assign_value_sound:
+ forall ty m b ofs bf v t m' v',
+ Csem.assign_loc ge ty m b ofs bf v t m' v' ->
+ forall tge e le m'' r,
+ typeof r = ty ->
+ eval_expr tge e le m'' r v ->
+ eval_expr tge e le m'' (make_assign_value bf r) v'.
+Proof.
+ unfold make_assign_value; destruct 1; intros; auto.
+ inv H. eapply eval_make_normalize; eauto; lia.
+Qed.
+
+Lemma typeof_make_assign_value: forall bf r,
+ typeof (make_assign_value bf r) = typeof r.
+Proof.
+ intros. destruct bf; simpl; auto. unfold make_normalize.
+ destruct (intsize_eq sz IBool || signedness_eq sg Unsigned); auto.
Qed.
(** Evaluation of simple expressions and of their translation *)
@@ -215,7 +301,7 @@ Lemma tr_simple:
(forall r v,
eval_simple_rvalue ge e m r v ->
forall le dst sl a tmps,
- tr_expr le dst r sl a tmps ->
+ tr_expr ce le dst r sl a tmps ->
match dst with
| For_val => sl = nil /\ Csyntax.typeof r = typeof a /\ eval_expr tge e le m a v
| For_effects => sl = nil
@@ -225,11 +311,11 @@ Lemma tr_simple:
/\ eval_expr tge e le m b v
end)
/\
- (forall l b ofs,
- eval_simple_lvalue ge e m l b ofs ->
+ (forall l b ofs bf,
+ eval_simple_lvalue ge e m l b ofs bf ->
forall le sl a tmps,
- tr_expr le For_val l sl a tmps ->
- sl = nil /\ Csyntax.typeof l = typeof a /\ eval_lvalue tge e le m a b ofs).
+ tr_expr ce le For_val l sl a tmps ->
+ sl = nil /\ Csyntax.typeof l = typeof a /\ eval_lvalue tge e le m a b ofs bf).
Proof.
Opaque makeif.
intros e m.
@@ -306,7 +392,7 @@ Lemma tr_simple_rvalue:
forall e m r v,
eval_simple_rvalue ge e m r v ->
forall le dst sl a tmps,
- tr_expr le dst r sl a tmps ->
+ tr_expr ce le dst r sl a tmps ->
match dst with
| For_val => sl = nil /\ Csyntax.typeof r = typeof a /\ eval_expr tge e le m a v
| For_effects => sl = nil
@@ -320,18 +406,18 @@ Proof.
Qed.
Lemma tr_simple_lvalue:
- forall e m l b ofs,
- eval_simple_lvalue ge e m l b ofs ->
+ forall e m l b ofs bf,
+ eval_simple_lvalue ge e m l b ofs bf ->
forall le sl a tmps,
- tr_expr le For_val l sl a tmps ->
- sl = nil /\ Csyntax.typeof l = typeof a /\ eval_lvalue tge e le m a b ofs.
+ tr_expr ce le For_val l sl a tmps ->
+ sl = nil /\ Csyntax.typeof l = typeof a /\ eval_lvalue tge e le m a b ofs bf.
Proof.
intros e m. exact (proj2 (tr_simple e m)).
Qed.
Lemma tr_simple_exprlist:
forall le rl sl al tmps,
- tr_exprlist le rl sl al tmps ->
+ tr_exprlist ce le rl sl al tmps ->
forall e m tyl vl,
eval_simple_list ge e m rl tyl vl ->
sl = nil /\ eval_exprlist tge e le m al tyl vl.
@@ -362,29 +448,29 @@ Lemma tr_expr_leftcontext_rec:
(
forall from to C, leftcontext from to C ->
forall le e dst sl a tmps,
- tr_expr le dst (C e) sl a tmps ->
+ tr_expr ce le dst (C e) sl a tmps ->
exists dst', exists sl1, exists sl2, exists a', exists tmp',
- tr_expr le dst' e sl1 a' tmp'
+ tr_expr ce le dst' e sl1 a' tmp'
/\ sl = sl1 ++ sl2
/\ incl tmp' tmps
/\ (forall le' e' sl3,
- tr_expr le' dst' e' sl3 a' tmp' ->
+ tr_expr ce le' dst' e' sl3 a' tmp' ->
(forall id, ~In id tmp' -> le'!id = le!id) ->
Csyntax.typeof e' = Csyntax.typeof e ->
- tr_expr le' dst (C e') (sl3 ++ sl2) a tmps)
+ tr_expr ce le' dst (C e') (sl3 ++ sl2) a tmps)
) /\ (
forall from C, leftcontextlist from C ->
forall le e sl a tmps,
- tr_exprlist le (C e) sl a tmps ->
+ tr_exprlist ce le (C e) sl a tmps ->
exists dst', exists sl1, exists sl2, exists a', exists tmp',
- tr_expr le dst' e sl1 a' tmp'
+ tr_expr ce le dst' e sl1 a' tmp'
/\ sl = sl1 ++ sl2
/\ incl tmp' tmps
/\ (forall le' e' sl3,
- tr_expr le' dst' e' sl3 a' tmp' ->
+ tr_expr ce le' dst' e' sl3 a' tmp' ->
(forall id, ~In id tmp' -> le'!id = le!id) ->
Csyntax.typeof e' = Csyntax.typeof e ->
- tr_exprlist le' (C e') (sl3 ++ sl2) a tmps)
+ tr_exprlist ce le' (C e') (sl3 ++ sl2) a tmps)
).
Proof.
@@ -553,7 +639,7 @@ Ltac UNCHANGED :=
red; auto.
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
- auto. auto. auto.
+ auto. auto. auto. auto.
+ (* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1. rewrite app_ass. eauto.
@@ -561,7 +647,7 @@ Ltac UNCHANGED :=
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
auto. auto. auto. auto. auto. auto.
- eapply typeof_context; eauto.
+ eapply typeof_context. eauto. auto. eauto.
auto.
- (* assign right *)
inv H2.
@@ -573,7 +659,7 @@ Ltac UNCHANGED :=
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.
+ apply S; auto. auto. auto. auto. auto.
+ (* for val *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
@@ -583,7 +669,7 @@ Ltac UNCHANGED :=
econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
apply S; auto. auto. auto. auto. auto. auto. auto. auto.
- eapply typeof_context; eauto.
+ eapply typeof_context; eauto. auto.
- (* assignop left *)
inv H1.
+ (* for effects *)
@@ -593,7 +679,7 @@ Ltac UNCHANGED :=
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
symmetry; eapply typeof_context; eauto. eauto.
- auto. auto. auto. auto. auto. auto.
+ auto. auto. auto. auto. auto. auto. auto.
+ (* for val *)
exploit H0; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
TR. subst sl1. rewrite app_ass. eauto.
@@ -601,7 +687,7 @@ Ltac UNCHANGED :=
intros. rewrite <- app_ass. econstructor. apply S; auto.
eapply tr_expr_invariant; eauto. UNCHANGED.
eauto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto.
- eapply typeof_context; eauto.
+ eapply typeof_context; eauto. auto.
- (* assignop right *)
inv H2.
+ (* for effects *)
@@ -611,7 +697,7 @@ Ltac UNCHANGED :=
red; auto.
intros. rewrite <- app_ass. change (sl0 ++ sl2') with (nil ++ sl0 ++ sl2'). rewrite app_ass. econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
- apply S; auto. auto. eauto. auto. auto. auto. auto. auto. auto.
+ apply S; auto. auto. eauto. auto. auto. auto. auto. auto. auto. auto.
+ (* for val *)
assert (sl1 = nil) by (eapply tr_simple_expr_nil; eauto). subst sl1; simpl.
exploit H1; eauto. intros [dst' [sl1' [sl2' [a' [tmp' [P [Q [R S]]]]]]]].
@@ -619,7 +705,7 @@ Ltac UNCHANGED :=
red; auto.
intros. rewrite <- app_ass. change (sl0 ++ sl2') with (nil ++ sl0 ++ sl2'). rewrite app_ass. econstructor.
eapply tr_expr_invariant; eauto. UNCHANGED.
- apply S; auto. eauto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto.
+ apply S; auto. eauto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto. auto.
- (* postincr *)
inv H1.
+ (* for effects *)
@@ -725,35 +811,35 @@ Qed.
Theorem tr_expr_leftcontext:
forall C le r dst sl a tmps,
leftcontext RV RV C ->
- tr_expr le dst (C r) sl a tmps ->
+ tr_expr ce le dst (C r) sl a tmps ->
exists dst', exists sl1, exists sl2, exists a', exists tmp',
- tr_expr le dst' r sl1 a' tmp'
+ tr_expr ce le dst' r sl1 a' tmp'
/\ sl = sl1 ++ sl2
/\ incl tmp' tmps
/\ (forall le' r' sl3,
- tr_expr le' dst' r' sl3 a' tmp' ->
+ tr_expr ce le' dst' r' sl3 a' tmp' ->
(forall id, ~In id tmp' -> le'!id = le!id) ->
Csyntax.typeof r' = Csyntax.typeof r ->
- tr_expr le' dst (C r') (sl3 ++ sl2) a tmps).
+ tr_expr ce le' dst (C r') (sl3 ++ sl2) a tmps).
Proof.
intros. eapply (proj1 tr_expr_leftcontext_rec); eauto.
Qed.
Theorem tr_top_leftcontext:
forall e le m dst rtop sl a tmps,
- tr_top tge e le m dst rtop sl a tmps ->
+ tr_top ce tge e le m dst rtop sl a tmps ->
forall r C,
rtop = C r ->
leftcontext RV RV C ->
exists dst', exists sl1, exists sl2, exists a', exists tmp',
- tr_top tge e le m dst' r sl1 a' tmp'
+ tr_top ce tge e le m dst' r sl1 a' tmp'
/\ sl = sl1 ++ sl2
/\ incl tmp' tmps
/\ (forall le' m' r' sl3,
- tr_expr le' dst' r' sl3 a' tmp' ->
+ tr_expr ce le' dst' r' sl3 a' tmp' ->
(forall id, ~In id tmp' -> le'!id = le!id) ->
Csyntax.typeof r' = Csyntax.typeof r ->
- tr_top tge e le' m' dst (C r') (sl3 ++ sl2) a tmps).
+ tr_top ce tge e le' m' dst (C r') (sl3 ++ sl2) a tmps).
Proof.
induction 1; intros.
(* val for val *)
@@ -835,17 +921,18 @@ Proof.
Qed.
Lemma step_make_set:
- forall id a ty m b ofs t v e le f k,
- Csem.deref_loc ge ty m b ofs t v ->
- eval_lvalue tge e le m a b ofs ->
+ forall id a ty m b ofs bf t v e le f k,
+ Csem.deref_loc ge ty m b ofs bf t v ->
+ eval_lvalue tge e le m a b ofs bf ->
typeof a = ty ->
- step1 tge (State f (make_set id a) k e le m)
+ step1 tge (State f (make_set bf id a) k e le m)
t (State f Sskip k e (PTree.set id v le) m).
Proof.
intros. exploit deref_loc_translated; eauto. rewrite <- H1.
- unfold make_set. destruct (chunk_for_volatile_type (typeof a)) as [chunk|].
+ unfold make_set. destruct (chunk_for_volatile_type (typeof a) bf) as [chunk|].
(* volatile case *)
- intros. change (PTree.set id v le) with (set_opttemp (Some id) v le). econstructor.
+ intros [A B]. subst bf.
+ 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.
@@ -854,19 +941,19 @@ Proof.
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 ->
+ forall a1 a2 ty m b ofs bf t v m' v' v2 e le f k,
+ Csem.assign_loc ge ty m b ofs bf v t m' v' ->
+ eval_lvalue tge e le m a1 b ofs bf ->
eval_expr tge e le m a2 v2 ->
sem_cast v2 (typeof a2) ty m = Some v ->
typeof a1 = ty ->
- step1 tge (State f (make_assign a1 a2) k e le m)
+ step1 tge (State f (make_assign bf 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|].
+ unfold make_assign. destruct (chunk_for_volatile_type (typeof a1) bf) as [chunk|].
(* volatile case *)
- intros. change le with (set_opttemp None Vundef le) at 2. econstructor.
+ intros [A B]. subst bf. 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.
@@ -900,10 +987,10 @@ Proof.
Qed.
Lemma step_tr_rvalof:
- forall ty m b ofs t v e le a sl a' tmp f k,
- 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 ->
+ forall ty m b ofs bf t v e le a sl a' tmp f k,
+ Csem.deref_loc ge ty m b ofs bf t v ->
+ eval_lvalue tge e le m a b ofs bf ->
+ tr_rvalof ce ty a sl a' tmp ->
typeof a = ty ->
exists le',
star step1 tge (State f Sskip (Kseqlist sl k) e le m)
@@ -920,141 +1007,149 @@ Proof.
split. eapply eval_Elvalue; eauto.
auto.
(* volatile *)
- intros. exists (PTree.set t0 v le); split.
+ intros.
+ exploit is_bitfield_access_sound; eauto. intros EQ; subst bf0.
+ 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.
+End TRANSLATION.
+
+
(** Matching between continuations *)
-Inductive match_cont : Csem.cont -> cont -> Prop :=
- | match_Kstop:
- match_cont Csem.Kstop Kstop
- | match_Kseq: forall s k ts tk,
- tr_stmt s ts ->
- match_cont k tk ->
- match_cont (Csem.Kseq s k) (Kseq ts tk)
- | match_Kwhile2: forall r s k s' ts tk,
- tr_if r Sskip Sbreak s' ->
- tr_stmt s ts ->
- match_cont k tk ->
- match_cont (Csem.Kwhile2 r s k)
- (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)
- (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)
- (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)
- (Kloop2 (Ssequence s' ts) ts3 tk)
- | match_Kswitch2: forall k tk,
- match_cont k tk ->
- match_cont (Csem.Kswitch2 k) (Kswitch tk)
- | match_Kcall: forall f e C ty k optid tf le sl tk a dest tmps,
- tr_function f tf ->
- leftcontext RV RV C ->
- (forall v m, tr_top tge e (set_opttemp optid v le) m dest (C (Csyntax.Eval v ty)) sl a tmps) ->
- match_cont_exp dest a k tk ->
- match_cont (Csem.Kcall f e C ty k)
- (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 ->
+Inductive match_cont : composite_env -> Csem.cont -> cont -> Prop :=
+ | match_Kstop: forall ce,
+ match_cont ce Csem.Kstop Kstop
+ | match_Kseq: forall ce s k ts tk,
+ tr_stmt ce s ts ->
+ match_cont ce k tk ->
+ match_cont ce (Csem.Kseq s k) (Kseq ts tk)
+ | match_Kwhile2: forall ce r s k s' ts tk,
+ tr_if ce r Sskip Sbreak s' ->
+ tr_stmt ce s ts ->
+ match_cont ce k tk ->
+ match_cont ce (Csem.Kwhile2 r s k)
+ (Kloop1 (Ssequence s' ts) Sskip tk)
+ | match_Kdowhile1: forall ce r s k s' ts tk,
+ tr_if ce r Sskip Sbreak s' ->
+ tr_stmt ce s ts ->
+ match_cont ce k tk ->
+ match_cont ce (Csem.Kdowhile1 r s k)
+ (Kloop1 ts s' tk)
+ | match_Kfor3: forall ce r s3 s k ts3 s' ts tk,
+ tr_if ce r Sskip Sbreak s' ->
+ tr_stmt ce s3 ts3 ->
+ tr_stmt ce s ts ->
+ match_cont ce k tk ->
+ match_cont ce (Csem.Kfor3 r s3 s k)
+ (Kloop1 (Ssequence s' ts) ts3 tk)
+ | match_Kfor4: forall ce r s3 s k ts3 s' ts tk,
+ tr_if ce r Sskip Sbreak s' ->
+ tr_stmt ce s3 ts3 ->
+ tr_stmt ce s ts ->
+ match_cont ce k tk ->
+ match_cont ce (Csem.Kfor4 r s3 s k)
+ (Kloop2 (Ssequence s' ts) ts3 tk)
+ | match_Kswitch2: forall ce k tk,
+ match_cont ce k tk ->
+ match_cont ce (Csem.Kswitch2 k) (Kswitch tk)
+ | match_Kcall: forall f e C ty k optid tf le sl tk a dest tmps cu ce,
+ linkorder cu prog ->
+ tr_function cu.(prog_comp_env) f tf ->
leftcontext RV RV C ->
- (forall v m, tr_top tge e (PTree.set dst 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 (Some dst) tf e le (Kseqlist sl tk))
-*)
-
-with match_cont_exp : destination -> expr -> Csem.cont -> cont -> Prop :=
- | match_Kdo: forall k a tk,
- match_cont k tk ->
- match_cont_exp For_effects a (Csem.Kdo k) tk
- | match_Kifthenelse_empty: forall a k tk,
- match_cont k tk ->
- match_cont_exp For_val a (Csem.Kifthenelse Csyntax.Sskip Csyntax.Sskip k) (Kseq Sskip tk)
- | match_Kifthenelse_1: forall a s1 s2 k ts1 ts2 tk,
- 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_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_val a
+ (forall v m, tr_top cu.(prog_comp_env) tge e (set_opttemp optid v le) m dest (C (Csyntax.Eval v ty)) sl a tmps) ->
+ match_cont_exp cu.(prog_comp_env) dest a k tk ->
+ match_cont ce (Csem.Kcall f e C ty k)
+ (Kcall optid tf e le (Kseqlist sl tk))
+
+with match_cont_exp : composite_env -> destination -> expr -> Csem.cont -> cont -> Prop :=
+ | match_Kdo: forall ce k a tk,
+ match_cont ce k tk ->
+ match_cont_exp ce For_effects a (Csem.Kdo k) tk
+ | match_Kifthenelse_empty: forall ce a k tk,
+ match_cont ce k tk ->
+ match_cont_exp ce For_val a (Csem.Kifthenelse Csyntax.Sskip Csyntax.Sskip k) (Kseq Sskip tk)
+ | match_Kifthenelse_1: forall ce a s1 s2 k ts1 ts2 tk,
+ tr_stmt ce s1 ts1 -> tr_stmt ce s2 ts2 ->
+ match_cont ce k tk ->
+ match_cont_exp ce For_val a (Csem.Kifthenelse s1 s2 k) (Kseq (Sifthenelse a ts1 ts2) tk)
+ | match_Kwhile1: forall ce r s k s' a ts tk,
+ tr_if ce r Sskip Sbreak s' ->
+ tr_stmt ce s ts ->
+ match_cont ce k tk ->
+ match_cont_exp ce For_val a
(Csem.Kwhile1 r s k)
(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_val a
+ | match_Kdowhile2: forall ce r s k s' a ts tk,
+ tr_if ce r Sskip Sbreak s' ->
+ tr_stmt ce s ts ->
+ match_cont ce k tk ->
+ match_cont_exp ce For_val a
(Csem.Kdowhile2 r s k)
(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_val a
+ | match_Kfor2: forall ce r s3 s k s' a ts3 ts tk,
+ tr_if ce r Sskip Sbreak s' ->
+ tr_stmt ce s3 ts3 ->
+ tr_stmt ce s ts ->
+ match_cont ce k tk ->
+ match_cont_exp ce For_val a
(Csem.Kfor2 r s3 s k)
(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 ->
- match_cont_exp For_val a (Csem.Kswitch1 ls k) (Kseq (Sswitch a tls) tk)
- | match_Kreturn: forall k a tk,
- match_cont k tk ->
- match_cont_exp For_val a (Csem.Kreturn k) (Kseq (Sreturn (Some a)) tk).
-
-Lemma match_cont_call:
- forall k tk,
- match_cont k tk ->
- match_cont (Csem.call_cont k) (call_cont tk).
+ | match_Kswitch1: forall ce ls k a tls tk,
+ tr_lblstmts ce ls tls ->
+ match_cont ce k tk ->
+ match_cont_exp ce For_val a (Csem.Kswitch1 ls k) (Kseq (Sswitch a tls) tk)
+ | match_Kreturn: forall ce k a tk,
+ match_cont ce k tk ->
+ match_cont_exp ce For_val a (Csem.Kreturn k) (Kseq (Sreturn (Some a)) tk).
+
+Lemma match_cont_is_call_cont:
+ forall ce k tk,
+ match_cont ce k tk -> Csem.is_call_cont k ->
+ forall ce', match_cont ce' k tk.
Proof.
- induction 1; simpl; auto. constructor. econstructor; eauto.
+ destruct 1; simpl; intros; try contradiction; econstructor; eauto.
+Qed.
+
+Lemma match_cont_call_cont:
+ forall ce k tk,
+ match_cont ce k tk ->
+ forall ce', match_cont ce' (Csem.call_cont k) (call_cont tk).
+Proof.
+ induction 1; simpl; auto; intros; econstructor; eauto.
Qed.
(** Matching between states *)
Inductive match_states: Csem.state -> state -> Prop :=
- | match_exprstates: forall f r k e m tf sl tk le dest a tmps,
- tr_function f tf ->
- tr_top tge e le m dest r sl a tmps ->
- match_cont_exp dest a k tk ->
+ | match_exprstates: forall f r k e m tf sl tk le dest a tmps cu
+ (LINK: linkorder cu prog)
+ (TRF: tr_function cu.(prog_comp_env) f tf)
+ (TR: tr_top cu.(prog_comp_env) tge e le m dest r sl a tmps)
+ (MK: match_cont_exp cu.(prog_comp_env) dest a k tk),
match_states (Csem.ExprState f r k e m)
(State tf Sskip (Kseqlist sl tk) e le m)
- | match_regularstates: forall f s k e m tf ts tk le,
- tr_function f tf ->
- tr_stmt s ts ->
- match_cont k tk ->
+ | match_regularstates: forall f s k e m tf ts tk le cu
+ (LINK: linkorder cu prog)
+ (TRF: tr_function cu.(prog_comp_env) f tf)
+ (TR: tr_stmt cu.(prog_comp_env) s ts)
+ (MK: match_cont cu.(prog_comp_env) k tk),
match_states (Csem.State f s k e m)
(State tf ts tk e le m)
- | match_callstates: forall fd args k m tfd tk,
- tr_fundef fd tfd ->
- match_cont k tk ->
+ | match_callstates: forall fd args k m tfd tk cu
+ (LINK: linkorder cu prog)
+ (TR: tr_fundef cu fd tfd)
+ (MK: forall ce, match_cont ce k tk),
match_states (Csem.Callstate fd args k m)
(Callstate tfd args tk m)
- | match_returnstates: forall res k m tk,
- match_cont k tk ->
+ | match_returnstates: forall res k m tk
+ (MK: forall ce, match_cont ce k tk),
match_states (Csem.Returnstate res k m)
(Returnstate res tk m)
| match_stuckstate: forall S,
@@ -1063,21 +1158,22 @@ Inductive match_states: Csem.state -> state -> Prop :=
(** Additional results on translation of statements *)
Lemma tr_select_switch:
- forall n ls tls,
- tr_lblstmts ls tls ->
- tr_lblstmts (Csem.select_switch n ls) (select_switch n tls).
+ forall ce n ls tls,
+ tr_lblstmts ce ls tls ->
+ tr_lblstmts ce (Csem.select_switch n ls) (select_switch n tls).
Proof.
+ intros ce.
assert (DFL: forall ls tls,
- tr_lblstmts ls tls ->
- tr_lblstmts (Csem.select_switch_default ls) (select_switch_default tls)).
+ tr_lblstmts ce ls tls ->
+ tr_lblstmts ce (Csem.select_switch_default ls) (select_switch_default tls)).
{ induction 1; simpl. constructor. destruct c; auto. constructor; auto. }
assert (CASE: forall n ls tls,
- tr_lblstmts ls tls ->
+ tr_lblstmts ce ls tls ->
match Csem.select_switch_case n ls with
| None =>
select_switch_case n tls = None
| Some ls' =>
- exists tls', select_switch_case n tls = Some tls' /\ tr_lblstmts ls' tls'
+ exists tls', select_switch_case n tls = Some tls' /\ tr_lblstmts ce ls' tls'
end).
{ induction 1; simpl; intros.
auto.
@@ -1091,9 +1187,9 @@ Proof.
Qed.
Lemma tr_seq_of_labeled_statement:
- forall ls tls,
- tr_lblstmts ls tls ->
- tr_stmt (Csem.seq_of_labeled_statement ls) (seq_of_labeled_statement tls).
+ forall ce ls tls,
+ tr_lblstmts ce ls tls ->
+ tr_stmt ce (Csem.seq_of_labeled_statement ls) (seq_of_labeled_statement tls).
Proof.
induction 1; simpl; constructor; auto.
Qed.
@@ -1102,6 +1198,7 @@ Qed.
Section FIND_LABEL.
+Variable ce: composite_env.
Variable lbl: label.
Definition nolabel (s: statement) : Prop :=
@@ -1137,21 +1234,21 @@ Proof.
Qed.
Lemma make_set_nolabel:
- forall t a, nolabel (make_set t a).
+ forall bf t a, nolabel (make_set bf t a).
Proof.
unfold make_set; intros; red; intros.
- destruct (chunk_for_volatile_type (typeof a)); auto.
+ destruct (chunk_for_volatile_type (typeof a) bf); auto.
Qed.
Lemma make_assign_nolabel:
- forall l r, nolabel (make_assign l r).
+ forall bf l r, nolabel (make_assign bf l r).
Proof.
unfold make_assign; intros; red; intros.
- destruct (chunk_for_volatile_type (typeof l)); auto.
+ destruct (chunk_for_volatile_type (typeof l) bf); auto.
Qed.
Lemma tr_rvalof_nolabel:
- forall ty a sl a' tmp, tr_rvalof ty a sl a' tmp -> nolabel_list sl.
+ forall ce ty a sl a' tmp, tr_rvalof ce ty a sl a' tmp -> nolabel_list sl.
Proof.
destruct 1; simpl; intuition. apply make_set_nolabel.
Qed.
@@ -1177,16 +1274,16 @@ Ltac NoLabelTac :=
| [ 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 (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_list sl)
-/\(forall le rl sl al tmps, tr_exprlist le rl sl al tmps -> nolabel_list sl).
+ (forall le dst r sl a tmps, tr_expr ce le dst r sl a tmps -> nolabel_list sl)
+/\(forall le rl sl al tmps, tr_exprlist ce le rl sl al tmps -> nolabel_list sl).
Proof.
apply tr_expr_exprlist; intros; NoLabelTac.
apply nolabel_do_set.
@@ -1200,14 +1297,14 @@ 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_list sl.
+ tr_top ce tge e le m dst r sl a tmps -> nolabel_list sl.
Proof.
induction 1; intros; NoLabelTac.
eapply (proj1 tr_find_label_expr); eauto.
Qed.
Lemma tr_find_label_expression:
- forall r s a, tr_expression r s a -> forall k, find_label lbl s k = None.
+ forall r s a, tr_expression ce r s a -> forall k, find_label lbl s k = None.
Proof.
intros. inv H.
assert (nolabel (makeseq sl)). apply makeseq_nolabel.
@@ -1216,7 +1313,7 @@ Proof.
Qed.
Lemma tr_find_label_expr_stmt:
- forall r s, tr_expr_stmt r s -> forall k, find_label lbl s k = None.
+ forall r s, tr_expr_stmt ce r s -> forall k, find_label lbl s k = None.
Proof.
intros. inv H.
assert (nolabel (makeseq sl)). apply makeseq_nolabel.
@@ -1226,7 +1323,7 @@ Qed.
Lemma tr_find_label_if:
forall r s,
- tr_if r Sskip Sbreak s ->
+ tr_if ce r Sskip Sbreak s ->
forall k, find_label lbl s k = None.
Proof.
intros. inv H.
@@ -1241,29 +1338,29 @@ Qed.
Lemma tr_find_label:
forall s k ts tk
- (TR: tr_stmt s ts)
- (MC: match_cont k tk),
+ (TR: tr_stmt ce s ts)
+ (MC: match_cont ce k tk),
match Csem.find_label lbl s k with
| None =>
find_label lbl ts tk = None
| Some (s', k') =>
exists ts', exists tk',
find_label lbl ts tk = Some (ts', tk')
- /\ tr_stmt s' ts'
- /\ match_cont k' tk'
+ /\ tr_stmt ce s' ts'
+ /\ match_cont ce k' tk'
end
with tr_find_label_ls:
forall s k ts tk
- (TR: tr_lblstmts s ts)
- (MC: match_cont k tk),
+ (TR: tr_lblstmts ce s ts)
+ (MC: match_cont ce k tk),
match Csem.find_label_ls lbl s k with
| None =>
find_label_ls lbl ts tk = None
| Some (s', k') =>
exists ts', exists tk',
find_label_ls lbl ts tk = Some (ts', tk')
- /\ tr_stmt s' ts'
- /\ match_cont k' tk'
+ /\ tr_stmt ce s' ts'
+ /\ match_cont ce k' tk'
end.
Proof.
induction s; intros; inversion TR; subst; clear TR; simpl.
@@ -1362,7 +1459,7 @@ The following measure decreases for these stuttering steps. *)
Fixpoint esize (a: Csyntax.expr) : nat :=
match a with
- | Csyntax.Eloc _ _ _ => 1%nat
+ | Csyntax.Eloc _ _ _ _ => 1%nat
| Csyntax.Evar _ _ => 1%nat
| Csyntax.Ederef r1 _ => S(esize r1)
| Csyntax.Efield l1 _ _ => S(esize l1)
@@ -1423,12 +1520,12 @@ Qed.
(** Forward simulation for expressions. *)
Lemma tr_val_gen:
- forall le dst v ty a tmp,
+ forall ce le dst v ty a 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 dst (Csyntax.Eval v ty) (final dst a) a tmp.
+ tr_expr ce le dst (Csyntax.Eval v ty) (final dst a) a tmp.
Proof.
intros. destruct dst; simpl; econstructor; auto.
Qed.
@@ -1441,43 +1538,53 @@ Lemma estep_simulation:
(star step1 tge S1' t S2' /\ measure S2 < measure S1)%nat)
/\ match_states S2 S2'.
Proof.
+
+Ltac NOTIN :=
+ match goal with
+ | [ H1: In ?x ?l, H2: list_disjoint ?l _ |- ~In ?x _ ] =>
+ red; intro; elim (H2 x x); auto; fail
+ | [ H1: In ?x ?l, H2: list_disjoint _ ?l |- ~In ?x _ ] =>
+ red; intro; elim (H2 x x); auto; fail
+ end.
+
induction 1; intros; inv MS.
-(* expr *)
- assert (tr_expr le dest r sl a tmps).
- inv H9. contradiction. auto.
+- (* expr *)
+ assert (tr_expr (prog_comp_env cu) le dest r sl a tmps).
+ { inv TR. contradiction. auto. }
exploit tr_simple_rvalue; eauto. destruct dest.
- (* for val *)
++ (* for val *)
intros [SL1 [TY1 EV1]]. subst sl.
econstructor; split.
right; split. apply star_refl. destruct r; simpl; (contradiction || lia).
econstructor; eauto.
instantiate (1 := tmps). apply tr_top_val_val; auto.
- (* for effects *)
++ (* for effects *)
intros SL1. subst sl.
econstructor; split.
right; split. apply star_refl. destruct r; simpl; (contradiction || lia).
econstructor; eauto.
instantiate (1 := tmps). apply tr_top_base. constructor.
- (* for set *)
- inv H10.
-(* rval volatile *)
- exploit tr_top_leftcontext; eauto. clear H11.
++ (* for set *)
+ inv MK.
+- (* rval volatile *)
+ exploit tr_top_leftcontext; eauto. clear TR.
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.
+ exploit is_bitfield_access_sound; eauto. intros EQ; subst bf0.
econstructor; split.
left. eapply plus_two. constructor. eapply step_make_set; eauto. traceEq.
econstructor; eauto.
change (final dst' (Etempvar t0 (Csyntax.typeof l)) ++ sl2) with (nil ++ (final dst' (Etempvar t0 (Csyntax.typeof l)) ++ sl2)).
apply S. apply tr_val_gen. auto.
- intros. constructor. rewrite H5; auto. apply PTree.gss.
- intros. apply PTree.gso. red; intros; subst; elim H5; auto.
+ intros. constructor. rewrite H7; auto. apply PTree.gss.
+ intros. apply PTree.gso. red; intros; subst; elim H7; auto.
auto.
-(* seqand true *)
- exploit tr_top_leftcontext; eauto. clear H9.
+- (* seqand true *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
- (* for val *)
++ (* for val *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1488,7 +1595,7 @@ Proof.
eapply match_exprstates; eauto.
apply S. apply tr_paren_val with (a1 := a2); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1499,7 +1606,7 @@ Proof.
eapply match_exprstates; eauto.
apply S. apply tr_paren_effects with (a1 := a2); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
- (* for set *)
++ (* for set *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1510,11 +1617,11 @@ Proof.
eapply match_exprstates; eauto.
apply S. apply tr_paren_set with (a1 := a2) (t := sd_temp sd); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
-(* seqand false *)
- exploit tr_top_leftcontext; eauto. clear H9.
+- (* seqand false *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
- (* for val *)
++ (* for val *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1526,7 +1633,7 @@ Proof.
intros. constructor. rewrite H2. apply PTree.gss. auto.
intros. apply PTree.gso. congruence.
auto.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1536,7 +1643,7 @@ Proof.
eapply match_exprstates; eauto.
change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
auto. auto.
- (* for set *)
++ (* for set *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1546,11 +1653,11 @@ Proof.
rewrite <- Kseqlist_app.
eapply match_exprstates; eauto.
apply S. econstructor; eauto. intros. constructor. auto. auto.
-(* seqor true *)
- exploit tr_top_leftcontext; eauto. clear H9.
+- (* seqor true *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
- (* for val *)
++ (* for val *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1562,7 +1669,7 @@ Proof.
intros. constructor. rewrite H2. apply PTree.gss. auto.
intros. apply PTree.gso. congruence.
auto.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1572,7 +1679,7 @@ Proof.
eapply match_exprstates; eauto.
change sl2 with (nil ++ sl2). apply S. econstructor; eauto.
auto. auto.
- (* for set *)
++ (* for set *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1582,11 +1689,11 @@ Proof.
rewrite <- Kseqlist_app.
eapply match_exprstates; eauto.
apply S. econstructor; eauto. intros. constructor. auto. auto.
-(* seqand false *)
- exploit tr_top_leftcontext; eauto. clear H9.
+- (* seqand false *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
- (* for val *)
++ (* for val *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1597,7 +1704,7 @@ Proof.
eapply match_exprstates; eauto.
apply S. apply tr_paren_val with (a1 := a2); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1608,7 +1715,7 @@ Proof.
eapply match_exprstates; eauto.
apply S. apply tr_paren_effects with (a1 := a2); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
- (* for set *)
++ (* for set *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist.
econstructor; split.
@@ -1619,11 +1726,11 @@ Proof.
eapply match_exprstates; eauto.
apply S. apply tr_paren_set with (a1 := a2) (t := sd_temp sd); auto.
apply tr_expr_monotone with tmp2; eauto. auto. auto.
-(* condition *)
- exploit tr_top_leftcontext; eauto. clear H9.
+- (* condition *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
- (* for value *)
++ (* for value *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist. destruct b.
econstructor; split.
@@ -1640,7 +1747,7 @@ Proof.
rewrite <- Kseqlist_app.
eapply match_exprstates; eauto.
apply S. econstructor; eauto. apply tr_expr_monotone with tmp3; eauto. auto. auto.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist. destruct b.
econstructor; split.
@@ -1649,7 +1756,7 @@ Proof.
apply push_seq.
reflexivity. traceEq.
rewrite <- Kseqlist_app.
- econstructor. eauto. apply S.
+ econstructor; eauto. apply S.
econstructor; eauto. apply tr_expr_monotone with tmp2; eauto.
econstructor; eauto.
auto. auto.
@@ -1659,11 +1766,11 @@ Proof.
apply push_seq.
reflexivity. traceEq.
rewrite <- Kseqlist_app.
- econstructor. eauto. apply S.
+ econstructor; eauto. apply S.
econstructor; eauto. apply tr_expr_monotone with tmp3; eauto.
econstructor; eauto.
- auto. auto.
- (* for set *)
+ auto.
++ (* for set *)
exploit tr_simple_rvalue; eauto. intros [SL [TY EV]].
subst sl0; simpl Kseqlist. destruct b.
econstructor; split.
@@ -1672,40 +1779,42 @@ Proof.
apply push_seq.
reflexivity. traceEq.
rewrite <- Kseqlist_app.
- econstructor. eauto. apply S.
+ econstructor; eauto. apply S.
econstructor; eauto. apply tr_expr_monotone with tmp2; eauto.
econstructor; eauto.
- auto. 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 S.
econstructor; eauto. apply tr_expr_monotone with tmp3; eauto.
econstructor; eauto.
- auto. auto.
-(* assign *)
- exploit tr_top_leftcontext; eauto. clear H12.
+ auto.
+- (* assign *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H4.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
+ assert (bf0 = bf) by (eapply is_bitfield_access_sound; eauto).
subst; simpl Kseqlist.
econstructor; split.
left. eapply plus_left. constructor.
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 *)
+ econstructor; eauto. change sl2 with (nil ++ sl2). apply S.
+ constructor. auto. auto.
++ (* for value *)
exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
exploit tr_simple_lvalue. eauto.
- eapply tr_expr_invariant with (le' := PTree.set t0 v' le). eauto.
+ eapply tr_expr_invariant with (le' := PTree.set t0 v1 le). eauto.
intros. apply PTree.gso. intuition congruence.
intros [SL1 [TY1 EV1]].
+ assert (bf0 = bf) by (eapply is_bitfield_access_sound; eauto).
subst; simpl Kseqlist.
econstructor; split.
left. eapply plus_left. constructor.
@@ -1714,22 +1823,24 @@ Proof.
apply star_one. eapply step_make_assign; eauto.
constructor. apply PTree.gss. simpl. eapply cast_idempotent; eauto.
reflexivity. reflexivity. traceEq.
- econstructor. auto. apply S.
- apply tr_val_gen. auto. intros. constructor.
- rewrite H4; auto. apply PTree.gss.
+ econstructor; eauto. apply S.
+ apply tr_val_gen. rewrite typeof_make_assign_value; auto.
+ intros. eapply make_assign_value_sound; eauto.
+ constructor. rewrite H4; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
- auto. auto.
-(* assignop *)
- exploit tr_top_leftcontext; eauto. clear H15.
+ auto.
+- (* assignop *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H6.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
intros. apply INV. NOTIN. intros [? [? EV1']].
exploit tr_simple_rvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
intros. apply INV. NOTIN. simpl. intros [SL2 [TY2 EV2]].
+ assert (bf0 = bf) by (eapply is_bitfield_access_sound; eauto).
subst; simpl Kseqlist.
econstructor; split.
left. eapply star_plus_trans. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
@@ -1737,9 +1848,9 @@ Proof.
econstructor. eexact EV3. eexact EV2.
rewrite TY3; rewrite <- TY1; rewrite <- TY2; rewrite comp_env_preserved; auto.
reflexivity. traceEq.
- econstructor. auto. change sl2 with (nil ++ sl2). apply S.
- constructor. auto. auto. auto.
- (* for value *)
+ econstructor; eauto. change sl2 with (nil ++ sl2). apply S.
+ constructor. auto. auto.
++ (* for value *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
@@ -1750,6 +1861,7 @@ Proof.
eapply tr_expr_invariant with (le := le) (le' := PTree.set t v4 le'). eauto.
intros. rewrite PTree.gso. apply INV. NOTIN. intuition congruence.
intros [? [? EV1'']].
+ assert (bf0 = bf) by (eapply is_bitfield_access_sound; eauto).
subst; simpl Kseqlist.
econstructor; split.
left. rewrite app_ass. rewrite Kseqlist_app.
@@ -1761,18 +1873,19 @@ Proof.
econstructor. eapply step_make_assign; eauto.
constructor. apply PTree.gss. simpl. eapply cast_idempotent; eauto.
reflexivity. traceEq.
- econstructor. auto. apply S.
- apply tr_val_gen. auto. intros. constructor.
- rewrite H10; auto. apply PTree.gss.
+ econstructor; eauto. apply S.
+ apply tr_val_gen. rewrite typeof_make_assign_value; auto.
+ intros. eapply make_assign_value_sound; eauto.
+ constructor. rewrite H10; auto. apply PTree.gss.
intros. rewrite PTree.gso. apply INV.
red; intros; elim H10; auto.
intuition congruence.
- auto. auto.
-(* assignop stuck *)
- exploit tr_top_leftcontext; eauto. clear H12.
+ auto.
+- (* assignop stuck *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H4.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
@@ -1781,7 +1894,7 @@ Proof.
right; split. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
simpl. lia.
constructor.
- (* for value *)
++ (* for value *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_rvalue; eauto. intros [SL2 [TY2 EV2]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
@@ -1790,15 +1903,16 @@ Proof.
right; split. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC.
simpl. lia.
constructor.
-(* postincr *)
- exploit tr_top_leftcontext; eauto. clear H14.
+- (* postincr *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H5.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
exploit tr_simple_lvalue. eauto. eapply tr_expr_invariant with (le := le) (le' := le'). eauto.
intros. apply INV. NOTIN. intros [? [? EV1']].
+ assert (bf0 = bf) by (eapply is_bitfield_access_sound; eauto).
subst; simpl Kseqlist.
econstructor; split.
left. rewrite app_ass; rewrite Kseqlist_app.
@@ -1810,14 +1924,15 @@ Proof.
econstructor. eauto. constructor. rewrite TY3; rewrite <- TY1; rewrite comp_env_preserved. simpl; eauto.
destruct id; auto.
reflexivity. traceEq.
- econstructor. auto. change sl2 with (nil ++ sl2). apply S.
- constructor. auto. auto. auto.
- (* for value *)
+ econstructor; eauto. change sl2 with (nil ++ sl2). apply S.
+ constructor. auto. auto.
++ (* for value *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_lvalue. eauto.
eapply tr_expr_invariant with (le' := PTree.set t v1 le). eauto.
intros. apply PTree.gso. intuition congruence.
intros [SL2 [TY2 EV2]].
+ assert (bf0 = bf) by (eapply is_bitfield_access_sound; eauto).
subst; simpl Kseqlist.
econstructor; split.
left. eapply plus_four. constructor.
@@ -1831,16 +1946,16 @@ Proof.
rewrite comp_env_preserved; simpl; eauto.
destruct id; auto.
traceEq.
- econstructor. auto. apply S.
+ econstructor; eauto. apply S.
apply tr_val_gen. auto. intros. econstructor; eauto.
rewrite H5; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
- auto. auto.
-(* postincr stuck *)
- exploit tr_top_leftcontext; eauto. clear H11.
+ auto.
+- (* postincr stuck *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H3.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
exploit step_tr_rvalof; eauto. intros [le' [EXEC [EV3 [TY3 INV]]]].
subst. simpl Kseqlist.
@@ -1848,15 +1963,16 @@ Proof.
right; split. rewrite app_ass; rewrite Kseqlist_app. eexact EXEC.
simpl; lia.
constructor.
- (* for value *)
++ (* for value *)
exploit tr_simple_lvalue; eauto. intros [SL1 [TY1 EV1]].
+ assert (bf0 = bf) by (eapply is_bitfield_access_sound; eauto).
subst. simpl Kseqlist.
econstructor; split.
left. eapply plus_two. constructor. eapply step_make_set; eauto.
traceEq.
constructor.
-(* comma *)
- exploit tr_top_leftcontext; eauto. clear H9.
+- (* comma *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H1.
exploit tr_simple_rvalue; eauto. simpl; intro SL1.
@@ -1867,11 +1983,11 @@ Proof.
econstructor; eauto. apply S.
eapply tr_expr_monotone; eauto.
auto. auto.
-(* paren *)
- exploit tr_top_leftcontext; eauto. clear H9.
+- (* paren *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
- (* for value *)
++ (* for value *)
exploit tr_simple_rvalue; eauto. intros [b [SL1 [TY1 EV1]]].
subst sl1; simpl Kseqlist.
econstructor; split.
@@ -1882,14 +1998,14 @@ Proof.
constructor. auto. intros. constructor. rewrite H2; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
auto.
- (* for effects *)
++ (* for effects *)
econstructor; split.
right; split. apply star_refl. simpl. apply plus_lt_compat_r.
apply (leftcontext_size _ _ _ H). simpl. lia.
econstructor; eauto.
exploit tr_simple_rvalue; eauto. simpl. intros A. subst sl1.
apply S. constructor; auto. auto. auto.
- (* for set *)
++ (* for set *)
exploit tr_simple_rvalue; eauto. simpl. intros [b [SL1 [TY1 EV1]]]. subst sl1.
simpl Kseqlist.
econstructor; split.
@@ -1901,46 +2017,46 @@ Proof.
intros. apply PTree.gso. congruence.
auto.
-(* call *)
- exploit tr_top_leftcontext; eauto. clear H12.
+- (* call *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H5.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_rvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_exprlist; eauto. intros [SL2 EV2].
subst. simpl Kseqlist.
- exploit functions_translated; eauto. intros [tfd [J K]].
+ exploit functions_translated; eauto. intros (cu' & tfd & J & K & L).
econstructor; split.
left. eapply plus_left. constructor. apply star_one.
econstructor; eauto. rewrite <- TY1; eauto.
exploit type_of_fundef_preserved; eauto. congruence.
traceEq.
- constructor; auto. econstructor; eauto.
+ econstructor. eexact L. eauto. econstructor. eexact LINK. auto. auto.
intros. change sl2 with (nil ++ sl2). apply S.
- constructor. auto. auto.
- (* for value *)
+ constructor. auto. auto. auto.
++ (* for value *)
exploit tr_simple_rvalue; eauto. intros [SL1 [TY1 EV1]].
exploit tr_simple_exprlist; eauto. intros [SL2 EV2].
subst. simpl Kseqlist.
- exploit functions_translated; eauto. intros [tfd [J K]].
+ exploit functions_translated; eauto. intros (cu' & tfd & J & K & L).
econstructor; split.
left. eapply plus_left. constructor. apply star_one.
econstructor; eauto. rewrite <- TY1; eauto.
exploit type_of_fundef_preserved; eauto. congruence.
traceEq.
- constructor; auto. econstructor; eauto.
+ econstructor. eexact L. eauto. econstructor. eexact LINK. auto. auto.
intros. apply S.
destruct dst'; constructor.
auto. intros. constructor. rewrite H5; auto. apply PTree.gss.
auto. intros. constructor. rewrite H5; auto. apply PTree.gss.
intros. apply PTree.gso. intuition congruence.
- auto.
+ auto. auto.
-(* builtin *)
- exploit tr_top_leftcontext; eauto. clear H9.
+- (* builtin *)
+ exploit tr_top_leftcontext; eauto. clear TR.
intros [dst' [sl1 [sl2 [a' [tmp' [P [Q [R S]]]]]]]].
inv P. inv H2.
- (* for effects *)
++ (* for effects *)
exploit tr_simple_exprlist; eauto. intros [SL EV].
subst. simpl Kseqlist.
econstructor; split.
@@ -1950,7 +2066,7 @@ Proof.
traceEq.
econstructor; eauto.
change sl2 with (nil ++ sl2). apply S. constructor. simpl; auto. auto.
- (* for value *)
++ (* for value *)
exploit tr_simple_exprlist; eauto. intros [SL EV].
subst. simpl Kseqlist.
econstructor; split.
@@ -1968,8 +2084,8 @@ Qed.
(** Forward simulation for statements. *)
Lemma tr_top_val_for_val_inv:
- forall e le m v ty sl a tmps,
- tr_top tge e le m For_val (Csyntax.Eval v ty) sl a tmps ->
+ forall ce e le m v ty sl a tmps,
+ tr_top ce tge e le m For_val (Csyntax.Eval v ty) sl a tmps ->
sl = nil /\ typeof a = ty /\ eval_expr tge e le m a v.
Proof.
intros. inv H. auto. inv H0. auto.
@@ -2011,264 +2127,263 @@ Lemma sstep_simulation:
/\ match_states S2 S2'.
Proof.
induction 1; intros; inv MS.
-(* do 1 *)
- inv H6. inv H0.
+- (* do 1 *)
+ inv TR. inv H0.
econstructor; split.
right; split. apply push_seq.
simpl. lia.
econstructor; eauto. constructor. auto.
-(* do 2 *)
- inv H7. inv H6. inv H.
+- (* do 2 *)
+ inv MK. inv TR. inv H.
econstructor; split.
right; split. apply star_refl. simpl. lia.
econstructor; eauto. constructor.
-(* seq *)
- inv H6. econstructor; split. left. apply plus_one. constructor.
+- (* seq *)
+ inv TR. econstructor; split. left. apply plus_one. constructor.
econstructor; eauto. constructor; auto.
-(* skip seq *)
- inv H6; inv H7. econstructor; split.
+- (* skip seq *)
+ inv TR; inv MK. econstructor; split.
left. apply plus_one; constructor.
econstructor; eauto.
-(* continue seq *)
- inv H6; inv H7. econstructor; split.
+- (* continue seq *)
+ inv TR; inv MK. econstructor; split.
left. apply plus_one; constructor.
econstructor; eauto. constructor.
-(* break seq *)
- inv H6; inv H7. econstructor; split.
+- (* break seq *)
+ inv TR; inv MK. econstructor; split.
left. apply plus_one; constructor.
econstructor; eauto. constructor.
-(* ifthenelse *)
- inv H6.
-(* ifthenelse empty *)
+- (* ifthenelse *)
+ inv TR.
++ (* ifthenelse empty *)
inv H3. econstructor; split.
left. eapply plus_left. constructor. apply push_seq.
econstructor; eauto.
econstructor; eauto.
econstructor; eauto.
-(* ifthenelse non empty *)
++ (* ifthenelse non empty *)
inv H2. econstructor; split.
left. eapply plus_left. constructor. apply push_seq. traceEq.
econstructor; eauto. econstructor; eauto.
-(* ifthenelse *)
- inv H8.
-(* ifthenelse empty *)
+- (* ifthenelse *)
+ inv MK.
++ (* ifthenelse empty *)
exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split; simpl.
right. destruct b; econstructor; eauto.
eapply star_left. apply step_skip_seq. econstructor. traceEq.
eapply star_left. apply step_skip_seq. econstructor. traceEq.
destruct b; econstructor; eauto. econstructor; eauto. econstructor; eauto.
- (* ifthenelse non empty *)
++ (* ifthenelse non empty *)
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.
-(* while *)
- inv H6. inv H1. econstructor; split.
+- (* while *)
+ inv TR. inv H1. econstructor; split.
left. eapply plus_left. constructor.
eapply star_left. constructor.
apply push_seq.
reflexivity. traceEq. rewrite Kseqlist_app.
- econstructor; eauto. simpl. econstructor; eauto. econstructor; eauto.
-(* while false *)
- inv H8.
+ econstructor; eauto. simpl. econstructor; eauto. econstructor; eauto.
+- (* while false *)
+ inv MK.
exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (v1 := v) (b := false); auto.
eapply star_two. constructor. apply step_break_loop1.
reflexivity. reflexivity. traceEq.
- constructor; auto. constructor.
-(* while true *)
- inv H8.
+ econstructor; eauto. constructor.
+- (* while true *)
+ inv MK.
exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
constructor.
reflexivity. traceEq.
- constructor; auto. constructor; auto.
-(* skip-or-continue while *)
- assert (ts = Sskip \/ ts = Scontinue). destruct H; subst s0; inv H7; auto.
- inv H8.
+ econstructor; eauto. constructor; auto.
+- (* skip-or-continue while *)
+ assert (ts = Sskip \/ ts = Scontinue). { destruct H; subst s0; inv TR; auto. }
+ inv MK.
econstructor; split.
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; eauto. constructor; auto.
+- (* break while *)
+ inv TR. inv MK.
econstructor; split.
left. apply plus_one. apply step_break_loop1.
- constructor; auto. constructor.
+ econstructor; eauto. constructor.
-(* dowhile *)
- inv H6.
+- (* dowhile *)
+ inv TR.
econstructor; split.
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; eauto. constructor; auto.
+- (* skip_or_continue dowhile *)
+ assert (ts = Sskip \/ ts = Scontinue). { destruct H; subst s0; inv TR; auto. }
+ inv MK. inv H5.
econstructor; split.
left. eapply plus_left. apply step_skip_or_continue_loop1. auto.
apply push_seq.
traceEq.
rewrite Kseqlist_app.
econstructor; eauto. simpl. econstructor; auto. econstructor; eauto.
-(* dowhile false *)
- inv H8.
+- (* dowhile false *)
+ inv MK.
exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := false); auto.
constructor.
reflexivity. traceEq.
- constructor; auto. constructor.
-(* dowhile true *)
- inv H8.
+ econstructor; eauto. constructor.
+- (* dowhile true *)
+ inv MK.
exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
constructor.
reflexivity. traceEq.
- constructor; auto. constructor; auto.
-(* break dowhile *)
- inv H6. inv H7.
+ econstructor; eauto. constructor; auto.
+- (* break dowhile *)
+ inv TR. inv MK.
econstructor; split.
left. apply plus_one. apply step_break_loop1.
- constructor; auto. constructor.
+ econstructor; eauto. constructor.
-(* for start *)
- inv H7. congruence.
+- (* for start *)
+ inv TR. congruence.
econstructor; split.
left; apply plus_one. constructor.
econstructor; eauto. constructor; auto. econstructor; eauto.
-(* for *)
- inv H6; try congruence. inv H2.
+- (* for *)
+ inv TR; try congruence. inv H2.
econstructor; split.
left. eapply plus_left. apply step_loop.
eapply star_left. constructor. apply push_seq.
reflexivity. traceEq.
rewrite Kseqlist_app. econstructor; eauto. simpl. constructor; auto. econstructor; eauto.
-(* for false *)
- inv H8. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+- (* for false *)
+ inv MK. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_trans. apply step_makeif with (v1 := v) (b := false); auto.
eapply star_two. constructor. apply step_break_loop1.
reflexivity. reflexivity. traceEq.
- constructor; auto. constructor.
-(* for true *)
- inv H8. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+ econstructor; eauto. constructor.
+- (* for true *)
+ inv MK. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. simpl. eapply plus_left. constructor.
eapply star_right. apply step_makeif with (v1 := v) (b := true); auto.
constructor.
reflexivity. traceEq.
- constructor; auto. constructor; auto.
-(* skip_or_continue for3 *)
- assert (ts = Sskip \/ ts = Scontinue). destruct H; subst x; inv H7; auto.
- inv H8.
+ econstructor; eauto. constructor; auto.
+- (* skip_or_continue for3 *)
+ assert (ts = Sskip \/ ts = Scontinue). { destruct H; subst x; inv TR; auto. }
+ inv MK.
econstructor; split.
left. apply plus_one. apply step_skip_or_continue_loop1. auto.
econstructor; eauto. econstructor; auto.
-(* break for3 *)
- inv H6. inv H7.
+- (* break for3 *)
+ inv TR. inv MK.
econstructor; split.
left. apply plus_one. apply step_break_loop1.
econstructor; eauto. constructor.
-(* skip for4 *)
- inv H6. inv H7.
+- (* skip for4 *)
+ inv TR. inv MK.
econstructor; split.
left. apply plus_one. constructor.
econstructor; eauto. constructor; auto.
-
-(* return none *)
- inv H7. econstructor; split.
+- (* return none *)
+ inv TR. econstructor; split.
left. apply plus_one. econstructor; eauto. rewrite blocks_of_env_preserved; eauto.
- constructor. apply match_cont_call; auto.
-(* return some 1 *)
- inv H6. inv H0. econstructor; split.
+ econstructor. intros; eapply match_cont_call_cont; eauto.
+- (* return some 1 *)
+ inv TR. inv H0. econstructor; split.
left; eapply plus_left. constructor. apply push_seq. traceEq.
econstructor; eauto. constructor. auto.
-(* return some 2 *)
- inv H9. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+- (* return some 2 *)
+ inv MK. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left. eapply plus_two. constructor. econstructor. eauto.
erewrite function_return_preserved; eauto. rewrite blocks_of_env_preserved; eauto.
eauto. traceEq.
- constructor. apply match_cont_call; auto.
-(* skip return *)
- inv H8.
- assert (is_call_cont tk). inv H9; simpl in *; auto.
+ econstructor. intros; eapply match_cont_call_cont; eauto.
+- (* skip return *)
+ inv TR.
+ assert (is_call_cont tk). { inv MK; simpl in *; auto. }
econstructor; split.
left. apply plus_one. apply step_skip_call; eauto. rewrite blocks_of_env_preserved; eauto.
- constructor. auto.
+ econstructor. intros; eapply match_cont_is_call_cont; eauto.
-(* switch *)
- inv H6. inv H1.
+- (* switch *)
+ inv TR. inv H1.
econstructor; split.
left; eapply plus_left. constructor. apply push_seq. traceEq.
econstructor; eauto. constructor; auto.
-(* expr switch *)
- inv H8. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
+- (* expr switch *)
+ inv MK. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst.
econstructor; split.
left; eapply plus_two. constructor. econstructor; eauto. traceEq.
econstructor; eauto.
apply tr_seq_of_labeled_statement. apply tr_select_switch. auto.
constructor; auto.
-(* skip-or-break switch *)
- assert (ts = Sskip \/ ts = Sbreak). destruct H; subst x; inv H7; auto.
- inv H8.
+- (* skip-or-break switch *)
+ assert (ts = Sskip \/ ts = Sbreak). { destruct H; subst x; inv TR; auto. }
+ inv MK.
econstructor; split.
left; apply plus_one. apply step_skip_break_switch. auto.
- constructor; auto. constructor.
+ econstructor; eauto. constructor.
-(* continue switch *)
- inv H6. inv H7.
+- (* continue switch *)
+ inv TR. inv MK.
econstructor; split.
left; apply plus_one. apply step_continue_switch.
- constructor; auto. constructor.
+ econstructor; eauto. constructor.
-(* label *)
- inv H6. econstructor; split.
+- (* label *)
+ inv TR. econstructor; split.
left; apply plus_one. constructor.
- constructor; auto.
+ econstructor; eauto.
-(* goto *)
- inv H7.
- inversion H6; subst.
- exploit tr_find_label. eauto. apply match_cont_call. eauto.
+- (* goto *)
+ inv TR.
+ inversion TRF; subst.
+ exploit tr_find_label. eauto. eapply match_cont_call_cont; eauto.
instantiate (1 := lbl). rewrite H.
intros [ts' [tk' [P [Q R]]]].
econstructor; split.
left. apply plus_one. econstructor; eauto.
econstructor; eauto.
-(* internal function *)
- inv H7. inversion H3; subst.
+- (* internal function *)
+ inv TR. inversion H3; subst.
econstructor; split.
left; apply plus_one. eapply step_internal_function. econstructor.
rewrite H6; rewrite H7; auto.
rewrite H6; rewrite H7. eapply alloc_variables_preserved; eauto.
rewrite H6. eapply bind_parameters_preserved; eauto.
eauto.
- constructor; auto.
+ econstructor; eauto.
-(* external function *)
- inv H5.
+- (* external function *)
+ inv TR.
econstructor; split.
left; apply plus_one. econstructor; eauto.
eapply external_call_symbols_preserved; eauto. apply senv_preserved.
- constructor; auto.
+ econstructor; eauto.
-(* return *)
- inv H3.
+- (* return *)
+ specialize (MK (PTree.empty _)). inv MK.
econstructor; split.
left; apply plus_one. constructor.
econstructor; eauto.
@@ -2295,7 +2410,7 @@ Lemma transl_initial_states:
exists S', Clight.initial_state tprog S' /\ match_states S S'.
Proof.
intros. inv H.
- exploit function_ptr_translated; eauto. intros [tf [FIND TR]].
+ exploit function_ptr_translated; eauto. intros (cu & tf & FIND & TR & L).
econstructor; split.
econstructor.
eapply (Genv.init_mem_match (proj1 TRANSL)); eauto.
@@ -2303,15 +2418,15 @@ Proof.
rewrite symbols_preserved. eauto.
destruct TRANSL. destruct H as (A & B & C). simpl in B. auto.
eexact FIND.
- rewrite <- H3. apply type_of_fundef_preserved. auto.
- constructor. auto. constructor.
+ rewrite <- H3. eapply type_of_fundef_preserved; eauto.
+ econstructor; eauto. intros; constructor.
Qed.
Lemma transl_final_states:
forall S S' r,
match_states S S' -> Csem.final_state S r -> Clight.final_state S' r.
Proof.
- intros. inv H0. inv H. inv H4. constructor.
+ intros. inv H0. inv H. specialize (MK (PTree.empty _)). inv MK. constructor.
Qed.
Theorem transl_program_correct:
@@ -2331,13 +2446,13 @@ End PRESERVATION.
Instance TransfSimplExprLink : TransfLink match_prog.
Proof.
- red; intros. eapply Ctypes.link_match_program; eauto.
+ red; intros. eapply Ctypes.link_match_program_gen; eauto.
- intros.
Local Transparent Linker_fundef.
simpl in *; unfold link_fundef in *. inv H3; inv H4; try discriminate.
- destruct ef; inv H2. exists (Internal tf); split; auto. constructor; auto.
- destruct ef; inv H2. exists (Internal tf); split; auto. constructor; auto.
+ destruct ef; inv H2. exists (Internal tf); split; auto. left; constructor; auto.
+ destruct ef; inv H2. exists (Internal tf); split; auto. right; constructor; auto.
destruct (external_function_eq ef ef0 && typelist_eq targs targs0 &&
type_eq tres tres0 && calling_convention_eq cconv cconv0); inv H2.
- exists (External ef targs tres cconv); split; auto. constructor.
+ exists (External ef targs tres cconv); split; auto. left; constructor.
Qed.
diff --git a/cfrontend/SimplExprspec.v b/cfrontend/SimplExprspec.v
index 98425311..b689bdeb 100644
--- a/cfrontend/SimplExprspec.v
+++ b/cfrontend/SimplExprspec.v
@@ -18,6 +18,8 @@ Require Import Ctypes Cop Csyntax Clight SimplExpr.
Section SPEC.
+Variable ce: composite_env.
+
Local Open Scope gensym_monad_scope.
(** * Relational specification of the translation. *)
@@ -40,13 +42,28 @@ Definition final (dst: destination) (a: expr) : list statement :=
| For_set sd => do_set sd a
end.
+Definition tr_is_bitfield_access (l: expr) (bf: bitfield) : Prop :=
+ match l with
+ | Efield r f _ =>
+ exists co ofs,
+ match typeof r with
+ | Tstruct id _ =>
+ ce!id = Some co /\ field_offset ce f (co_members co) = OK (ofs, bf)
+ | Tunion id _ =>
+ ce!id = Some co /\ union_field_offset ce f (co_members co) = OK (ofs, bf)
+ | _ => False
+ end
+ | _ => bf = Full
+ end.
+
Inductive tr_rvalof: type -> expr -> list statement -> expr -> list ident -> Prop :=
| tr_rvalof_nonvol: forall ty a tmp,
type_is_volatile ty = false ->
tr_rvalof ty a nil a tmp
- | tr_rvalof_vol: forall ty a t tmp,
+ | tr_rvalof_vol: forall ty a t bf tmp,
type_is_volatile ty = true -> In t tmp ->
- tr_rvalof ty a (make_set t a :: nil) (Etempvar t ty) tmp.
+ tr_is_bitfield_access a bf ->
+ tr_rvalof ty a (make_set bf t a :: nil) (Etempvar t ty) tmp.
Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement -> expr -> list ident -> Prop :=
| tr_var: forall le dst id ty tmp,
@@ -200,15 +217,16 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
tr_expr le (For_set sd) (Csyntax.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,
+ | tr_assign_effects: forall le e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 bf any tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_expr le For_val e2 sl2 a2 tmp2 ->
list_disjoint tmp1 tmp2 ->
incl tmp1 tmp -> incl tmp2 tmp ->
+ tr_is_bitfield_access a1 bf ->
tr_expr le For_effects (Csyntax.Eassign e1 e2 ty)
- (sl1 ++ sl2 ++ make_assign a1 a2 :: nil)
+ (sl1 ++ sl2 ++ make_assign bf 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_assign_val: forall le dst e1 e2 ty sl1 a1 tmp1 sl2 a2 tmp2 t tmp ty1 ty2 bf,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_expr le For_val e2 sl2 a2 tmp2 ->
incl tmp1 tmp -> incl tmp2 tmp ->
@@ -216,23 +234,25 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
In t tmp -> ~In t tmp1 -> ~In t tmp2 ->
ty1 = Csyntax.typeof e1 ->
ty2 = Csyntax.typeof e2 ->
+ tr_is_bitfield_access a1 bf ->
tr_expr le dst (Csyntax.Eassign e1 e2 ty)
(sl1 ++ sl2 ++
Sset t (Ecast a2 ty1) ::
- make_assign a1 (Etempvar t ty1) ::
- final dst (Etempvar t ty1))
- (Etempvar t ty1) tmp
- | tr_assignop_effects: forall le op e1 e2 tyres ty ty1 sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 any tmp,
+ make_assign bf a1 (Etempvar t ty1) ::
+ final dst (make_assign_value bf (Etempvar t ty1)))
+ (make_assign_value bf (Etempvar t ty1)) tmp
+ | tr_assignop_effects: forall le op e1 e2 tyres ty ty1 sl1 a1 tmp1 sl2 a2 tmp2 bf sl3 a3 tmp3 any tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_expr le For_val e2 sl2 a2 tmp2 ->
ty1 = Csyntax.typeof e1 ->
tr_rvalof ty1 a1 sl3 a3 tmp3 ->
list_disjoint tmp1 tmp2 -> list_disjoint tmp1 tmp3 -> list_disjoint tmp2 tmp3 ->
incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp ->
+ tr_is_bitfield_access a1 bf ->
tr_expr le For_effects (Csyntax.Eassignop op e1 e2 tyres ty)
- (sl1 ++ sl2 ++ sl3 ++ make_assign a1 (Ebinop op a3 a2 tyres) :: nil)
+ (sl1 ++ sl2 ++ sl3 ++ make_assign bf 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_assignop_val: forall le dst op e1 e2 tyres ty sl1 a1 tmp1 sl2 a2 tmp2 sl3 a3 tmp3 t bf tmp ty1,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_expr le For_val e2 sl2 a2 tmp2 ->
tr_rvalof ty1 a1 sl3 a3 tmp3 ->
@@ -240,28 +260,31 @@ Inductive tr_expr: temp_env -> destination -> Csyntax.expr -> list statement ->
incl tmp1 tmp -> incl tmp2 tmp -> incl tmp3 tmp ->
In t tmp -> ~In t tmp1 -> ~In t tmp2 -> ~In t tmp3 ->
ty1 = Csyntax.typeof e1 ->
+ tr_is_bitfield_access a1 bf ->
tr_expr le dst (Csyntax.Eassignop op e1 e2 tyres ty)
(sl1 ++ sl2 ++ sl3 ++
Sset t (Ecast (Ebinop op a3 a2 tyres) ty1) ::
- make_assign a1 (Etempvar t ty1) ::
- final dst (Etempvar t ty1))
- (Etempvar t ty1) tmp
- | tr_postincr_effects: forall le id e1 ty ty1 sl1 a1 tmp1 sl2 a2 tmp2 any tmp,
+ make_assign bf a1 (Etempvar t ty1) ::
+ final dst (make_assign_value bf (Etempvar t ty1)))
+ (make_assign_value bf (Etempvar t ty1)) tmp
+ | tr_postincr_effects: forall le id e1 ty ty1 sl1 a1 tmp1 sl2 a2 tmp2 bf any tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
tr_rvalof ty1 a1 sl2 a2 tmp2 ->
ty1 = Csyntax.typeof e1 ->
incl tmp1 tmp -> incl tmp2 tmp ->
list_disjoint tmp1 tmp2 ->
+ tr_is_bitfield_access a1 bf ->
tr_expr le For_effects (Csyntax.Epostincr id e1 ty)
- (sl1 ++ sl2 ++ make_assign a1 (transl_incrdecr id a2 ty1) :: nil)
+ (sl1 ++ sl2 ++ make_assign bf 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_postincr_val: forall le dst id e1 ty sl1 a1 tmp1 bf t ty1 tmp,
tr_expr le For_val e1 sl1 a1 tmp1 ->
incl tmp1 tmp -> In t tmp -> ~In t tmp1 ->
ty1 = Csyntax.typeof e1 ->
+ tr_is_bitfield_access a1 bf ->
tr_expr le dst (Csyntax.Epostincr id e1 ty)
- (sl1 ++ make_set t a1 ::
- make_assign a1 (transl_incrdecr id (Etempvar t ty1) ty1) ::
+ (sl1 ++ make_set bf t a1 ::
+ make_assign bf 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,
@@ -746,14 +769,31 @@ Proof.
intros. apply tr_expr_monotone with tmps; auto. apply add_dest_incl.
Qed.
+Lemma is_bitfield_access_meets_spec: forall l g bf g' I,
+ is_bitfield_access ce l g = Res bf g' I ->
+ tr_is_bitfield_access l bf.
+Proof.
+ unfold is_bitfield_access; intros; red. destruct l; try (monadInv H; auto).
+ assert (AUX: forall fn id,
+ is_bitfield_access_aux ce fn id i g = Res bf g' I ->
+ exists co ofs,
+ ce!id = Some co /\ fn ce i (co_members co) = OK (ofs, bf)).
+ { unfold is_bitfield_access_aux; intros.
+ destruct ce!id as [co|]; try discriminate.
+ destruct (fn ce i (co_members co)) as [[ofs1 bf1]|] eqn:FN; inv H0.
+ exists co, ofs1; auto. }
+ destruct (typeof l); try discriminate; apply AUX; auto.
+Qed.
+
Lemma transl_valof_meets_spec:
forall ty a g sl b g' I,
- transl_valof ty a g = Res (sl, b) g' I ->
+ transl_valof ce ty a g = Res (sl, b) g' I ->
exists tmps, tr_rvalof ty a sl b tmps /\ contained tmps g g'.
Proof.
unfold transl_valof; intros.
destruct (type_is_volatile ty) eqn:?; monadInv H.
- exists (x :: nil); split; eauto with gensym. econstructor; eauto with coqlib.
+ exists (x :: nil); split; eauto with gensym.
+ econstructor; eauto using is_bitfield_access_meets_spec with coqlib.
exists (@nil ident); split; eauto with gensym. constructor; auto.
Qed.
@@ -763,12 +803,12 @@ 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 ->
+ transl_expr ce dst r g = Res (sl, a) g' I ->
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 ->
+ transl_exprlist ce 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; simpl add_dest; intros.
@@ -920,9 +960,10 @@ Opaque makeif.
- (* assign *)
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [Csyntax D]].
- destruct dst; monadInv EQ2; simpl add_dest in *.
+ apply is_bitfield_access_meets_spec in EQ0.
+ destruct dst; monadInv EQ3; simpl add_dest in *.
+ (* for value *)
- exists (x1 :: tmp1 ++ tmp2); split.
+ exists (x2 :: tmp1 ++ tmp2); split.
intros. eapply tr_assign_val with (dst := For_val); eauto with gensym.
apply contained_cons. eauto with gensym.
apply contained_app; eauto with gensym.
@@ -931,7 +972,7 @@ Opaque makeif.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
+ (* for set *)
- exists (x1 :: tmp1 ++ tmp2); split.
+ exists (x2 :: tmp1 ++ tmp2); split.
repeat rewrite app_ass. simpl.
intros. eapply tr_assign_val with (dst := For_set sd); eauto with gensym.
apply contained_cons. eauto with gensym.
@@ -940,37 +981,39 @@ Opaque makeif.
monadInv H1. exploit H; eauto. intros [tmp1 [A B]].
exploit H0; eauto. intros [tmp2 [Csyntax D]].
exploit transl_valof_meets_spec; eauto. intros [tmp3 [E F]].
- destruct dst; monadInv EQ3; simpl add_dest in *.
+ apply is_bitfield_access_meets_spec in EQ2.
+ destruct dst; monadInv EQ4; simpl add_dest in *.
+ (* for value *)
- exists (x2 :: tmp1 ++ tmp2 ++ tmp3); split.
- intros. eapply tr_assignop_val with (dst := For_val); eauto with gensym.
- apply contained_cons. eauto with gensym.
- apply contained_app; eauto with gensym.
+ exists (x3 :: tmp1 ++ tmp2 ++ tmp3); split.
+ intros. eapply tr_assignop_val with (dst := For_val); eauto 6 with gensym.
+ apply contained_cons. eauto 6 with gensym.
+ apply contained_app; eauto 6 with gensym.
+ (* for effects *)
exists (tmp1 ++ tmp2 ++ tmp3); split.
econstructor; eauto with gensym.
apply contained_app; eauto with gensym.
+ (* for set *)
- exists (x2 :: tmp1 ++ tmp2 ++ tmp3); split.
+ exists (x3 :: tmp1 ++ tmp2 ++ tmp3); split.
repeat rewrite app_ass. simpl.
- intros. eapply tr_assignop_val with (dst := For_set sd); eauto with gensym.
- apply contained_cons. eauto with gensym.
- apply contained_app; eauto with gensym.
+ intros. eapply tr_assignop_val with (dst := For_set sd); eauto 6 with gensym.
+ apply contained_cons. eauto 6 with gensym.
+ apply contained_app; eauto 6 with gensym.
- (* postincr *)
monadInv H0. exploit H; eauto. intros [tmp1 [A B]].
- destruct dst; monadInv EQ0; simpl add_dest in *.
+ apply is_bitfield_access_meets_spec in EQ1.
+ destruct dst; monadInv EQ2; simpl add_dest in *.
+ (* for value *)
- exists (x0 :: tmp1); split.
+ exists (x1 :: tmp1); split.
econstructor; eauto with gensym.
apply contained_cons; eauto with gensym.
+ (* for effects *)
- exploit transl_valof_meets_spec; eauto. intros [tmp2 [Csyntax D]].
+ exploit transl_valof_meets_spec; eauto. intros [tmp2 [C D]].
exists (tmp1 ++ tmp2); split.
econstructor; eauto with gensym.
eauto with gensym.
+ (* for set *)
repeat rewrite app_ass; simpl.
- exists (x0 :: tmp1); split.
+ exists (x1 :: tmp1); split.
econstructor; eauto with gensym.
apply contained_cons; eauto with gensym.
- (* comma *)
@@ -1032,7 +1075,7 @@ Qed.
Lemma transl_expr_meets_spec:
forall r dst g sl a g' I,
- transl_expr dst r g = Res (sl, a) g' I ->
+ transl_expr ce 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.
@@ -1042,7 +1085,7 @@ Qed.
Lemma transl_expression_meets_spec:
forall r g s a g' I,
- transl_expression r g = Res (s, a) g' I ->
+ transl_expression ce r g = Res (s, a) g' I ->
tr_expression r s a.
Proof.
intros. monadInv H. exploit transl_expr_meets_spec; eauto.
@@ -1051,7 +1094,7 @@ Qed.
Lemma transl_expr_stmt_meets_spec:
forall r g s g' I,
- transl_expr_stmt r g = Res s g' I ->
+ transl_expr_stmt ce r g = Res s g' I ->
tr_expr_stmt r s.
Proof.
intros. monadInv H. exploit transl_expr_meets_spec; eauto.
@@ -1060,7 +1103,7 @@ Qed.
Lemma transl_if_meets_spec:
forall r s1 s2 g s g' I,
- transl_if r s1 s2 g = Res s g' I ->
+ transl_if ce r s1 s2 g = Res s g' I ->
tr_if r s1 s2 s.
Proof.
intros. monadInv H. exploit transl_expr_meets_spec; eauto.
@@ -1068,9 +1111,9 @@ Proof.
Qed.
Lemma transl_stmt_meets_spec:
- forall s g ts g' I, transl_stmt s g = Res ts g' I -> tr_stmt s ts
+ forall s g ts g' I, transl_stmt ce s g = Res ts g' I -> tr_stmt s ts
with transl_lblstmt_meets_spec:
- forall s g ts g' I, transl_lblstmt s g = Res ts g' I -> tr_lblstmts s ts.
+ forall s g ts g' I, transl_lblstmt ce s g = Res ts g' I -> tr_lblstmts s ts.
Proof.
generalize transl_expression_meets_spec transl_expr_stmt_meets_spec transl_if_meets_spec; intros T1 T2 T3.
Opaque transl_expression transl_expr_stmt.
@@ -1099,32 +1142,32 @@ Inductive tr_function: Csyntax.function -> Clight.function -> Prop :=
fn_vars tf = Csyntax.fn_vars f ->
tr_function f tf.
-Inductive tr_fundef: Csyntax.fundef -> Clight.fundef -> Prop :=
- | tr_internal: forall f tf,
- tr_function f tf ->
- tr_fundef (Internal f) (Internal tf)
- | tr_external: forall ef targs tres cconv,
- tr_fundef (External ef targs tres cconv) (External ef targs tres cconv).
-
Lemma transl_function_spec:
forall f tf,
- transl_function f = OK tf ->
+ transl_function ce f = OK tf ->
tr_function f tf.
Proof.
unfold transl_function; intros.
- destruct (transl_stmt (Csyntax.fn_body f) (initial_generator tt)) eqn:T; inv H.
+ destruct (transl_stmt ce (Csyntax.fn_body f) (initial_generator tt)) eqn:T; inv H.
constructor; auto. simpl. eapply transl_stmt_meets_spec; eauto.
Qed.
+End SPEC.
+
+Inductive tr_fundef (p: Csyntax.program): Csyntax.fundef -> Clight.fundef -> Prop :=
+ | tr_internal: forall f tf,
+ tr_function p.(prog_comp_env) f tf ->
+ tr_fundef p (Internal f) (Internal tf)
+ | tr_external: forall ef targs tres cconv,
+ tr_fundef p (External ef targs tres cconv) (External ef targs tres cconv).
+
Lemma transl_fundef_spec:
- forall fd tfd,
- transl_fundef fd = OK tfd ->
- tr_fundef fd tfd.
+ forall p fd tfd,
+ transl_fundef p.(prog_comp_env) fd = OK tfd ->
+ tr_fundef p fd tfd.
Proof.
unfold transl_fundef; intros.
destruct fd; Errors.monadInv H.
+ constructor. eapply transl_function_spec; eauto.
+ constructor.
Qed.
-
-End SPEC.
diff --git a/cfrontend/SimplLocalsproof.v b/cfrontend/SimplLocalsproof.v
index 988988a1..e4b759c4 100644
--- a/cfrontend/SimplLocalsproof.v
+++ b/cfrontend/SimplLocalsproof.v
@@ -391,7 +391,7 @@ Lemma match_envs_assign_lifted:
e!id = Some(b, ty) ->
val_casted v ty ->
Val.inject f v tv ->
- assign_loc ge ty m b Ptrofs.zero v m' ->
+ assign_loc ge ty m b Ptrofs.zero Full v m' ->
VSet.mem id cenv = true ->
match_envs f cenv e le m' lo hi te (PTree.set id tv tle) tlo thi.
Proof.
@@ -1004,13 +1004,13 @@ Proof.
Qed.
Lemma assign_loc_inject:
- forall f ty m loc ofs v m' tm loc' ofs' v',
- assign_loc ge ty m loc ofs v m' ->
+ forall f ty m loc ofs bf v m' tm loc' ofs' v',
+ assign_loc ge ty m loc ofs bf v m' ->
Val.inject f (Vptr loc ofs) (Vptr loc' ofs') ->
Val.inject f v v' ->
Mem.inject f m tm ->
exists tm',
- assign_loc tge ty tm loc' ofs' v' tm'
+ assign_loc tge ty tm loc' ofs' bf v' tm'
/\ Mem.inject f m' tm'
/\ (forall b chunk v,
f b = None -> Mem.load chunk m b 0 = Some v -> Mem.load chunk m' b 0 = Some v).
@@ -1078,15 +1078,25 @@ Proof.
split. auto.
intros. rewrite <- H0. eapply Mem.load_storebytes_other; eauto.
left. congruence.
+- (* bitfield *)
+ inv H3.
+ exploit Mem.loadv_inject; eauto. intros (vc' & LD' & INJ). inv INJ.
+ exploit Mem.storev_mapped_inject; eauto. intros [tm' [A B]].
+ inv H1.
+ exists tm'; split. eapply assign_loc_bitfield; eauto. econstructor; eauto.
+ split. auto.
+ intros. rewrite <- H3. eapply Mem.load_store_other; eauto.
+ left. inv H0. congruence.
Qed.
Lemma assign_loc_nextblock:
- forall ge ty m b ofs v m',
- assign_loc ge ty m b ofs v m' -> Mem.nextblock m' = Mem.nextblock m.
+ forall ge ty m b ofs bf v m',
+ assign_loc ge ty m b ofs bf v m' -> Mem.nextblock m' = Mem.nextblock m.
Proof.
induction 1.
simpl in H0. eapply Mem.nextblock_store; eauto.
eapply Mem.nextblock_storebytes; eauto.
+ inv H. eapply Mem.nextblock_store; eauto.
Qed.
Theorem store_params_correct:
@@ -1168,7 +1178,7 @@ Local Opaque Conventions1.parameter_needs_normalization.
reflexivity. reflexivity.
eexact U.
traceEq.
- rewrite (assign_loc_nextblock _ _ _ _ _ _ _ A) in Z. auto.
+ rewrite (assign_loc_nextblock _ _ _ _ _ _ _ _ A) in Z. auto.
Qed.
Lemma bind_parameters_nextblock:
@@ -1400,19 +1410,22 @@ Proof.
Qed.
Lemma deref_loc_inject:
- forall ty loc ofs v loc' ofs',
- deref_loc ty m loc ofs v ->
+ forall ty loc ofs bf v loc' ofs',
+ deref_loc ty m loc ofs bf v ->
Val.inject f (Vptr loc ofs) (Vptr loc' ofs') ->
- exists tv, deref_loc ty tm loc' ofs' tv /\ Val.inject f v tv.
+ exists tv, deref_loc ty tm loc' ofs' bf tv /\ Val.inject f v tv.
Proof.
intros. inv H.
- (* by value *)
+- (* by value *)
exploit Mem.loadv_inject; eauto. intros [tv [A B]].
exists tv; split; auto. eapply deref_loc_value; eauto.
- (* by reference *)
+- (* by reference *)
exists (Vptr loc' ofs'); split; auto. eapply deref_loc_reference; eauto.
- (* by copy *)
+- (* by copy *)
exists (Vptr loc' ofs'); split; auto. eapply deref_loc_copy; eauto.
+- (* bitfield *)
+ inv H1. exploit Mem.loadv_inject; eauto. intros [tc [A B]]. inv B.
+ econstructor; split. eapply deref_loc_bitfield. econstructor; eauto. constructor.
Qed.
Lemma eval_simpl_expr:
@@ -1422,11 +1435,11 @@ Lemma eval_simpl_expr:
exists tv, eval_expr tge te tle tm (simpl_expr cenv a) tv /\ Val.inject f v tv
with eval_simpl_lvalue:
- forall a b ofs,
- eval_lvalue ge e le m a b ofs ->
+ forall a b ofs bf,
+ eval_lvalue ge e le m a b ofs bf ->
compat_cenv (addr_taken_expr a) cenv ->
match a with Evar id ty => VSet.mem id cenv = false | _ => True end ->
- exists b', exists ofs', eval_lvalue tge te tle tm (simpl_expr cenv a) b' ofs' /\ Val.inject f (Vptr b ofs) (Vptr b' ofs').
+ exists b', exists ofs', eval_lvalue tge te tle tm (simpl_expr cenv a) b' ofs' bf /\ Val.inject f (Vptr b ofs) (Vptr b' ofs').
Proof.
destruct 1; simpl; intros.
@@ -1472,7 +1485,7 @@ Proof.
subst a. simpl. rewrite OPT.
exploit me_vars; eauto. instantiate (1 := id). intros MV.
inv H; inv MV; try congruence.
- rewrite ENV in H6; inv H6.
+ rewrite ENV in H7; inv H7.
inv H0; try congruence.
assert (chunk0 = chunk). simpl in H. congruence. subst chunk0.
assert (v0 = v). unfold Mem.loadv in H2. rewrite Ptrofs.unsigned_zero in H2. congruence. subst v0.
@@ -1516,7 +1529,8 @@ Proof.
exploit eval_simpl_expr; eauto. intros [tv [A B]].
inversion B. subst.
econstructor; econstructor; split.
- eapply eval_Efield_union; eauto. rewrite typeof_simpl_expr; eauto. auto.
+ eapply eval_Efield_union; eauto. rewrite typeof_simpl_expr; eauto.
+ econstructor; eauto. repeat rewrite Ptrofs.add_assoc. decEq. apply Ptrofs.add_commut.
Qed.
Lemma eval_simpl_exprlist:
@@ -1607,18 +1621,20 @@ Qed.
(** Invariance by assignment to location "above" *)
Lemma match_cont_assign_loc:
- forall f cenv k tk m bound tbound ty loc ofs v m',
+ forall f cenv k tk m bound tbound ty loc ofs bf v m',
match_cont f cenv k tk m bound tbound ->
- assign_loc ge ty m loc ofs v m' ->
+ assign_loc ge ty m loc ofs bf v m' ->
Ple bound loc ->
match_cont f cenv k tk m' bound tbound.
Proof.
intros. eapply match_cont_invariant; eauto.
intros. rewrite <- H4. inv H0.
- (* scalar *)
+- (* scalar *)
simpl in H6. eapply Mem.load_store_other; eauto. left. unfold block; extlia.
- (* block copy *)
+- (* block copy *)
eapply Mem.load_storebytes_other; eauto. left. unfold block; extlia.
+- (* bitfield *)
+ inv H5. eapply Mem.load_store_other; eauto. left. unfold block; extlia.
Qed.
(** Invariance by external calls *)