From b279716c76c387c6c5eec96388c0c35629858b4c Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Wed, 26 Nov 2014 14:46:07 +0100 Subject: Introduce symbol environments (type Senv.t) as a restricted view on global environments (type Genv.t). Use symbol environments instead of global environments for external functions (module Events). --- cfrontend/Cexec.v | 24 ++++++++++++------------ cfrontend/Cshmgenproof.v | 36 ++++++------------------------------ cfrontend/SimplExprproof.v | 4 ++-- 3 files changed, 20 insertions(+), 44 deletions(-) (limited to 'cfrontend') diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v index 80748df1..5c062158 100644 --- a/cfrontend/Cexec.v +++ b/cfrontend/Cexec.v @@ -149,7 +149,7 @@ Lemma eventval_of_val_complete: forall ev t v, eventval_match ge ev t v -> eventval_of_val v t = Some ev. Proof. induction 1; simpl; auto. - rewrite (Genv.find_invert_symbol _ _ H0). rewrite H. auto. + rewrite (Genv.find_invert_symbol _ _ H0). simpl in H; rewrite H. auto. Qed. Lemma list_eventval_of_val_sound: @@ -181,14 +181,14 @@ Qed. Lemma val_of_eventval_complete: forall ev t v, eventval_match ge ev t v -> val_of_eventval ev t = Some v. Proof. - induction 1; simpl; auto. rewrite H, H0; auto. + induction 1; simpl; auto. simpl in *. rewrite H, H0; auto. Qed. (** Volatile memory accesses. *) Definition do_volatile_load (w: world) (chunk: memory_chunk) (m: mem) (b: block) (ofs: int) : option (world * trace * val) := - if block_is_volatile ge b then + if Genv.block_is_volatile ge b then do id <- Genv.invert_symbol ge b; match nextworld_vload w chunk id ofs with | None => None @@ -202,7 +202,7 @@ Definition do_volatile_load (w: world) (chunk: memory_chunk) (m: mem) (b: block) Definition do_volatile_store (w: world) (chunk: memory_chunk) (m: mem) (b: block) (ofs: int) (v: val) : option (world * trace * mem) := - if block_is_volatile ge b then + 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; @@ -239,7 +239,7 @@ Lemma do_volatile_load_complete: volatile_load ge chunk m b ofs t v -> possible_trace w t w' -> do_volatile_load w chunk m b ofs = Some(w', t, v). Proof. - unfold do_volatile_load; intros. inv H. + unfold do_volatile_load; intros. inv H; simpl in *. rewrite H1. rewrite (Genv.find_invert_symbol _ _ H2). inv H0. inv H8. inv H6. rewrite H9. rewrite (val_of_eventval_complete _ _ _ H3). auto. rewrite H1. rewrite H2. inv H0. auto. @@ -262,7 +262,7 @@ Lemma do_volatile_store_complete: volatile_store ge chunk m b ofs v t m' -> possible_trace w t w' -> do_volatile_store w chunk m b ofs v = Some(w', t, m'). Proof. - unfold do_volatile_store; intros. inv H. + unfold do_volatile_store; intros. inv H; simpl in *. rewrite H1. rewrite (Genv.find_invert_symbol _ _ H2). rewrite (eventval_of_val_complete _ _ _ H3). inv H0. inv H8. inv H6. rewrite H9. auto. @@ -387,7 +387,7 @@ Qed. (** External calls *) Variable do_external_function: - ident -> signature -> genv -> world -> list val -> mem -> option (world * trace * val * mem). + ident -> signature -> Senv.t -> world -> list val -> mem -> option (world * trace * val * mem). Hypothesis do_external_function_sound: forall id sg ge vargs m t vres m' w w', @@ -401,7 +401,7 @@ Hypothesis do_external_function_complete: do_external_function id sg ge w vargs m = Some(w', t, vres, m'). Variable do_inline_assembly: - ident -> genv -> world -> list val -> mem -> option (world * trace * val * mem). + ident -> Senv.t -> world -> list val -> mem -> option (world * trace * val * mem). Hypothesis do_inline_assembly_sound: forall txt ge vargs m t vres m' w w', @@ -573,11 +573,11 @@ Proof with try congruence. (* EF_vstore *) auto. (* EF_vload_global *) - rewrite volatile_load_global_charact. + rewrite volatile_load_global_charact; simpl. unfold do_ef_volatile_load_global. destruct (Genv.find_symbol ge)... intros. exploit VLOAD; eauto. intros [A B]. split; auto. exists b; auto. (* EF_vstore_global *) - rewrite volatile_store_global_charact. + rewrite volatile_store_global_charact; simpl. unfold do_ef_volatile_store_global. destruct (Genv.find_symbol ge)... intros. exploit VSTORE; eauto. intros [A B]. split; auto. exists b; auto. (* EF_malloc *) @@ -633,10 +633,10 @@ Proof. (* EF_vstore *) auto. (* EF_vload_global *) - rewrite volatile_load_global_charact in H. destruct H as [b [P Q]]. + rewrite volatile_load_global_charact in H; simpl in H. destruct H as [b [P Q]]. unfold do_ef_volatile_load_global. rewrite P. auto. (* EF_vstore *) - rewrite volatile_store_global_charact in H. destruct H as [b [P Q]]. + rewrite volatile_store_global_charact in H; simpl in H. destruct H as [b [P Q]]. unfold do_ef_volatile_store_global. rewrite P. auto. (* EF_malloc *) inv H; unfold do_ef_malloc. diff --git a/cfrontend/Cshmgenproof.v b/cfrontend/Cshmgenproof.v index 9cb112b0..dc47df04 100644 --- a/cfrontend/Cshmgenproof.v +++ b/cfrontend/Cshmgenproof.v @@ -759,29 +759,9 @@ Lemma function_ptr_translated: exists tf, Genv.find_funct_ptr tge b = Some tf /\ transl_fundef f = OK tf. Proof (Genv.find_funct_ptr_transf_partial2 transl_fundef transl_globvar _ TRANSL). -Lemma var_info_translated: - forall b v, - Genv.find_var_info ge b = Some v -> - exists tv, Genv.find_var_info tge b = Some tv /\ transf_globvar transl_globvar v = OK tv. -Proof (Genv.find_var_info_transf_partial2 transl_fundef transl_globvar _ TRANSL). - -Lemma var_info_rev_translated: - forall b tv, - Genv.find_var_info tge b = Some tv -> - exists v, Genv.find_var_info ge b = Some v /\ transf_globvar transl_globvar v = OK tv. -Proof (Genv.find_var_info_rev_transf_partial2 transl_fundef transl_globvar _ TRANSL). - Lemma block_is_volatile_preserved: - forall b, block_is_volatile tge b = block_is_volatile ge b. -Proof. - intros. unfold block_is_volatile. - destruct (Genv.find_var_info ge b) eqn:?. - exploit var_info_translated; eauto. intros [tv [A B]]. rewrite A. - unfold transf_globvar in B. monadInv B. auto. - destruct (Genv.find_var_info tge b) eqn:?. - exploit var_info_rev_translated; eauto. intros [tv [A B]]. congruence. - auto. -Qed. + forall b, Genv.block_is_volatile tge b = Genv.block_is_volatile ge b. +Proof (Genv.block_is_volatile_transf_partial2 transl_fundef transl_globvar _ TRANSL). (** * Matching between environments *) @@ -1288,10 +1268,8 @@ Proof. econstructor; split. apply plus_one. econstructor. eapply transl_arglist_correct; eauto. - eapply external_call_symbols_preserved_2; eauto. - exact symbols_preserved. exact public_preserved. - eexact (Genv.find_var_info_transf_partial2 transl_fundef transl_globvar _ TRANSL). - eexact (Genv.find_var_info_rev_transf_partial2 transl_fundef transl_globvar _ TRANSL). + eapply external_call_symbols_preserved_gen with (ge1 := ge). + exact symbols_preserved. exact public_preserved. exact block_is_volatile_preserved. eauto. eapply match_states_skip; eauto. (* seq *) @@ -1469,10 +1447,8 @@ Proof. exploit match_cont_is_call_cont; eauto. intros [A B]. econstructor; split. apply plus_one. constructor. eauto. - eapply external_call_symbols_preserved_2; eauto. - exact symbols_preserved. exact public_preserved. - eexact (Genv.find_var_info_transf_partial2 transl_fundef transl_globvar _ TRANSL). - eexact (Genv.find_var_info_rev_transf_partial2 transl_fundef transl_globvar _ TRANSL). + eapply external_call_symbols_preserved_gen with (ge1 := ge). + exact symbols_preserved. exact public_preserved. exact block_is_volatile_preserved. eauto. econstructor; eauto. (* returnstate *) diff --git a/cfrontend/SimplExprproof.v b/cfrontend/SimplExprproof.v index 250f2b26..6d8b5c86 100644 --- a/cfrontend/SimplExprproof.v +++ b/cfrontend/SimplExprproof.v @@ -91,9 +91,9 @@ Proof. Qed. Lemma block_is_volatile_preserved: - forall b, block_is_volatile tge b = block_is_volatile ge b. + forall b, Genv.block_is_volatile tge b = Genv.block_is_volatile ge b. Proof. - intros. unfold block_is_volatile. rewrite varinfo_preserved. auto. + intros. unfold Genv.block_is_volatile. rewrite varinfo_preserved. auto. Qed. Lemma type_of_fundef_preserved: -- cgit From e89f1e606bc8c9c425628392adc9c69cec666b5e Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Mon, 22 Dec 2014 19:34:45 +0100 Subject: Represent struct and union types by name instead of by structure. --- cfrontend/C2C.ml | 251 +++++------- cfrontend/Cexec.v | 80 ++-- cfrontend/Clight.v | 142 ++++--- cfrontend/ClightBigstep.v | 18 +- cfrontend/Cop.v | 82 ++-- cfrontend/Csem.v | 96 ++--- cfrontend/Cshmgen.v | 179 +++++---- cfrontend/Cshmgenproof.v | 154 ++++---- cfrontend/Cstrategy.v | 97 ++--- cfrontend/Csyntax.v | 45 ++- cfrontend/Ctypes.v | 883 ++++++++++++++++++++++++++---------------- cfrontend/Initializers.v | 115 +++--- cfrontend/Initializersproof.v | 162 ++++---- cfrontend/PrintClight.ml | 71 +--- cfrontend/PrintCsyntax.ml | 141 +------ cfrontend/SimplExpr.v | 9 +- cfrontend/SimplExprproof.v | 77 ++-- cfrontend/SimplExprspec.v | 4 +- cfrontend/SimplLocals.v | 18 +- cfrontend/SimplLocalsproof.v | 202 +++++----- 20 files changed, 1518 insertions(+), 1308 deletions(-) (limited to 'cfrontend') diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index 4d5d6c07..ea278ac1 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -100,6 +100,10 @@ let atom_location a = with Not_found -> Cutil.no_loc +(** The current environment of composite definitions *) + +let comp_env : composite_env ref = ref Maps.PTree.empty + (** Hooks -- overriden in machine-dependent CPragmas module *) let process_pragma_hook = ref (fun (s: string) -> false) @@ -243,12 +247,12 @@ let make_builtin_memcpy args = match args with | Econs(dst, Econs(src, Econs(sz, Econs(al, Enil)))) -> let sz1 = - match Initializers.constval sz with + match Initializers.constval !comp_env sz with | Errors.OK(Vint n) -> n | _ -> error "ill-formed __builtin_memcpy_aligned (3rd argument must be a constant)"; Integers.Int.zero in let al1 = - match Initializers.constval al with + match Initializers.constval !comp_env al with | Errors.OK(Vint n) -> n | _ -> error "ill-formed __builtin_memcpy_aligned (4th argument must be a constant)"; Integers.Int.one in @@ -270,7 +274,7 @@ let va_list_ptr e = let make_builtin_va_arg env ty e = let (helper, ty_ret) = match ty with - | Tint _ | Tpointer _ | Tcomp_ptr _ -> + | Tint _ | Tpointer _ -> ("__compcert_va_int32", Tint(I32, Unsigned, noattr)) | Tlong _ -> ("__compcert_va_int64", Tlong(Unsigned, noattr)) @@ -303,27 +307,6 @@ let convertAttr a = let n = Cutil.alignas_attribute a in if n > 0 then Some (N.of_int (log2 n)) else None } -let mergeAttr a1 a2 = - { attr_volatile = a1.attr_volatile || a2.attr_volatile; - attr_alignas = - match a1.attr_alignas, a2.attr_alignas with - | None, aa -> aa - | aa, None -> aa - | Some n1, Some n2 -> Some (if N.le n1 n2 then n1 else n2) } - -let mergeTypAttr ty a2 = - match ty with - | Tvoid -> ty - | Tint(sz, sg, a1) -> Tint(sz, sg, mergeAttr a1 a2) - | Tfloat(sz, a1) -> Tfloat(sz, mergeAttr a1 a2) - | Tlong(sg, a1) -> Tlong(sg, mergeAttr a1 a2) - | Tpointer(ty', a1) -> Tpointer(ty', mergeAttr a1 a2) - | Tarray(ty', sz, a1) -> Tarray(ty', sz, mergeAttr a1 a2) - | Tfunction(targs, tres, cc) -> ty - | Tstruct(id, fld, a1) -> Tstruct(id, fld, mergeAttr a1 a2) - | Tunion(id, fld, a1) -> Tunion(id, fld, mergeAttr a1 a2) - | Tcomp_ptr(id, a1) -> Tcomp_ptr(id, mergeAttr a1 a2) - let convertCallconv va attr = let sr = Cutil.find_custom_attributes ["structreturn"; "__structreturn"] attr in @@ -353,93 +336,48 @@ let convertFkind = function if not !Clflags.option_flongdouble then unsupported "'long double' type"; F64 -(** A cache for structs and unions already converted *) - -let compositeCache : (C.ident, coq_type) Hashtbl.t = Hashtbl.create 77 - -let convertTyp env t = - - let rec convertTyp seen t = - match Cutil.unroll env t with - | C.TVoid a -> Tvoid - | C.TInt(C.ILongLong, a) -> - Tlong(Signed, convertAttr a) - | C.TInt(C.IULongLong, a) -> - Tlong(Unsigned, convertAttr a) - | C.TInt(ik, a) -> - let (sg, sz) = convertIkind ik in Tint(sz, sg, convertAttr a) - | C.TFloat(fk, a) -> - Tfloat(convertFkind fk, convertAttr a) - | C.TPtr(ty, a) -> - begin match Cutil.unroll env ty with - | C.TStruct(id, _) when List.mem id seen -> - Tcomp_ptr(intern_string ("struct " ^ id.name), convertAttr a) - | C.TUnion(id, _) when List.mem id seen -> - Tcomp_ptr(intern_string ("union " ^ id.name), convertAttr a) - | _ -> - Tpointer(convertTyp seen ty, convertAttr a) - end - | C.TArray(ty, None, a) -> - (* Cparser verified that the type ty[] occurs only in - contexts that are safe for Clight, so just treat as ty[0]. *) - (* warning "array type of unspecified size"; *) - Tarray(convertTyp seen ty, coqint_of_camlint 0l, convertAttr a) - | C.TArray(ty, Some sz, a) -> - Tarray(convertTyp seen ty, convertInt sz, convertAttr a) - | C.TFun(tres, targs, va, a) -> - if Cutil.is_composite_type env tres then - unsupported "return type is a struct or union (consider adding option -fstruct-return)"; - Tfunction(begin match targs with - | None -> Tnil - | Some tl -> convertParams seen tl - end, - convertTyp seen tres, - convertCallconv va a) - | C.TNamed _ -> - assert false - | C.TStruct(id, a) -> - let a' = convertAttr a in - begin try - merge_attributes (Hashtbl.find compositeCache id) a' - with Not_found -> - let flds = - try - convertFields (id :: seen) (Env.find_struct env id) - with Env.Error e -> - error (Env.error_message e); Fnil in - Tstruct(intern_string("struct " ^ id.name), flds, a') - end - | C.TUnion(id, a) -> - let a' = convertAttr a in - begin try - merge_attributes (Hashtbl.find compositeCache id) a' - with Not_found -> - let flds = - try - convertFields (id :: seen) (Env.find_union env id) - with Env.Error e -> - error (Env.error_message e); Fnil in - Tunion(intern_string("union " ^ id.name), flds, a') - end - | C.TEnum(id, a) -> - let (sg, sz) = convertIkind Cutil.enum_ikind in - Tint(sz, sg, convertAttr a) - - and convertParams seen = function +let rec convertTyp env t = + match t with + | C.TVoid a -> Tvoid + | C.TInt(C.ILongLong, a) -> + Tlong(Signed, convertAttr a) + | C.TInt(C.IULongLong, a) -> + Tlong(Unsigned, convertAttr a) + | C.TInt(ik, a) -> + let (sg, sz) = convertIkind ik in Tint(sz, sg, convertAttr a) + | C.TFloat(fk, a) -> + Tfloat(convertFkind fk, convertAttr a) + | C.TPtr(ty, a) -> + Tpointer(convertTyp env ty, convertAttr a) + | C.TArray(ty, None, a) -> + (* Cparser verified that the type ty[] occurs only in + contexts that are safe for Clight, so just treat as ty[0]. *) + (* warning "array type of unspecified size"; *) + Tarray(convertTyp env ty, coqint_of_camlint 0l, convertAttr a) + | C.TArray(ty, Some sz, a) -> + Tarray(convertTyp env ty, convertInt sz, convertAttr a) + | C.TFun(tres, targs, va, a) -> + if Cutil.is_composite_type env tres then + unsupported "return type is a struct or union (consider adding option -fstruct-return)"; + Tfunction(begin match targs with + | None -> Tnil + | Some tl -> convertParams env tl + end, + convertTyp env tres, + convertCallconv va a) + | C.TNamed _ -> + convertTyp env (Cutil.unroll env t) + | C.TStruct(id, a) -> + Tstruct(intern_string id.name, convertAttr a) + | C.TUnion(id, a) -> + Tunion(intern_string id.name, convertAttr a) + | C.TEnum(id, a) -> + let (sg, sz) = convertIkind Cutil.enum_ikind in + Tint(sz, sg, convertAttr a) + +and convertParams env = function | [] -> Tnil - | (id, ty) :: rem -> - Tcons(convertTyp seen ty, convertParams seen rem) - - and convertFields seen ci = - convertFieldList seen ci.Env.ci_members - - and convertFieldList seen = function - | [] -> Fnil - | f :: fl -> - Fcons(intern_string f.fld_name, convertTyp seen f.fld_typ, - convertFieldList seen fl) - - in convertTyp [] t + | (id, ty) :: rem -> Tcons(convertTyp env ty, convertParams env rem) let rec convertTypArgs env tl el = match tl, el with @@ -450,12 +388,16 @@ let rec convertTypArgs env tl el = | (id, t1) :: tl, e1 :: el -> Tcons(convertTyp env t1, convertTypArgs env tl el) -let cacheCompositeDef env su id attr flds = - let ty = - match su with - | C.Struct -> C.TStruct(id, attr) - | C.Union -> C.TUnion(id, attr) in - Hashtbl.add compositeCache id (convertTyp env ty) +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 convertCompositedef env su id attr members = + Composite(intern_string id.name, + begin match su with C.Struct -> Struct | C.Union -> Union end, + List.map (convertField env) members, + convertAttr attr) let rec projFunType env ty = match Cutil.unroll env ty with @@ -958,7 +900,8 @@ and convertInitList env il = | i :: il' -> Init_cons(convertInit env i, convertInitList env il') let convertInitializer env ty i = - match Initializers.transl_init (convertTyp env ty) (convertInit env i) + match Initializers.transl_init + !comp_env (convertTyp env ty) (convertInit env i) with | Errors.OK init -> init | Errors.Error msg -> @@ -970,8 +913,8 @@ let convertInitializer env ty i = let convertGlobvar loc env (sto, id, ty, optinit) = let id' = intern_string id.name in let ty' = convertTyp env ty in - let sz = Ctypes.sizeof ty' in - let al = Ctypes.alignof ty' in + let sz = Ctypes.sizeof !comp_env ty' in + let al = Ctypes.alignof !comp_env ty' in let attr = Cutil.attributes_of_type env ty in let init' = match optinit with @@ -998,16 +941,6 @@ let convertGlobvar loc env (sto, id, ty, optinit) = (id', Gvar {gvar_info = ty'; gvar_init = init'; gvar_readonly = readonly; gvar_volatile = volatile}) -(** Sanity checks on composite declarations. *) - -let checkComposite env si id attr flds = - let checkField f = - if f.fld_bitfield <> None then - unsupported "bit field in struct or union (consider adding option -fbitfields)" in - List.iter checkField flds; - if Cutil.find_custom_attributes ["packed";"__packed__"] attr <> [] then - unsupported "packed struct (consider adding option -fpacked-struct)" - (** Convert a list of global declarations. Result is a list of CompCert C global declarations (functions + variables). *) @@ -1031,21 +964,31 @@ let rec convertGlobdecls env res gl = end | C.Gfundef fd -> convertGlobdecls env (convertFundef g.gloc env fd :: res) gl' - | C.Gcompositedecl _ | C.Gtypedef _ | C.Genumdef _ -> - (* typedefs are unrolled, structs are expanded inline, and - enum tags are folded. So we just skip their declarations. *) - convertGlobdecls env res gl' - | C.Gcompositedef(su, id, attr, flds) -> - (* sanity checks on fields *) - checkComposite env su id attr flds; - (* convert it to a CompCert C type and cache this type *) - cacheCompositeDef env su id attr flds; + | C.Gcompositedecl _ | C.Gcompositedef _ | C.Gtypedef _ | C.Genumdef _ -> + (* Composites are treated in a separate pass, + typedefs are unrolled, and enum tags are folded. + So we just skip their declarations. *) convertGlobdecls env res gl' | C.Gpragma s -> if not (!process_pragma_hook s) then warning ("'#pragma " ^ s ^ "' directive ignored"); convertGlobdecls env res gl' +(** Convert struct and union declarations. + Result is a list of CompCert C composite declarations. *) + +let rec convertCompositedefs env res gl = + match gl with + | [] -> List.rev res + | g :: gl' -> + updateLoc g.gloc; + match g.gdesc with + | C.Gcompositedef(su, id, a, m) -> + convertCompositedefs env + (convertCompositedef env su id a m :: res) gl' + | _ -> + convertCompositedefs env res gl' + (** Build environment of typedefs, structs, unions and enums *) let rec translEnv env = function @@ -1130,17 +1073,29 @@ let convertProgram p = stringNum := 0; Hashtbl.clear decl_atom; Hashtbl.clear stringTable; - Hashtbl.clear compositeCache; - let p = Builtins.declarations() @ p in + let p = cleanupGlobals (Builtins.declarations() @ p) in try - let gl1 = convertGlobdecls (translEnv Env.empty p) [] (cleanupGlobals p) in - let gl2 = globals_for_strings gl1 in - let p' = { AST.prog_defs = gl2; - AST.prog_public = public_globals gl2; - AST.prog_main = intern_string "main" } in - if !numErrors > 0 - then None - else Some p' + let env = translEnv Env.empty p in + let typs = convertCompositedefs env [] p in + match build_composite_env typs with + | Errors.Error msg -> + error (sprintf "Incorrect struct or union definition: %s" + (string_of_errmsg msg)); + None + | Errors.OK ce -> + comp_env := ce; + let gl1 = convertGlobdecls env [] p in + let gl2 = globals_for_strings gl1 in + comp_env := Maps.PTree.empty; + let p' = + { prog_defs = gl2; + prog_public = public_globals gl2; + prog_main = intern_string "main"; + prog_types = typs; + prog_comp_env = ce } in + if !numErrors > 0 + then None + else Some p' with Env.Error msg -> error (Env.error_message msg); None diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v index 5c062158..952d148d 100644 --- a/cfrontend/Cexec.v +++ b/cfrontend/Cexec.v @@ -284,31 +284,31 @@ Definition do_deref_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: int) : o end. Definition assign_copy_ok (ty: type) (b: block) (ofs: int) (b': block) (ofs': int) : Prop := - (alignof_blockcopy ty | Int.unsigned ofs') /\ (alignof_blockcopy ty | Int.unsigned ofs) /\ + (alignof_blockcopy ge ty | Int.unsigned ofs') /\ (alignof_blockcopy ge ty | Int.unsigned ofs) /\ (b' <> b \/ Int.unsigned ofs' = Int.unsigned ofs - \/ Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs - \/ Int.unsigned ofs + sizeof ty <= Int.unsigned ofs'). + \/ Int.unsigned ofs' + sizeof ge ty <= Int.unsigned ofs + \/ Int.unsigned ofs + sizeof ge ty <= Int.unsigned ofs'). Remark check_assign_copy: forall (ty: type) (b: block) (ofs: int) (b': block) (ofs': int), { assign_copy_ok ty b ofs b' ofs' } + {~ assign_copy_ok ty b ofs b' ofs' }. Proof with try (right; intuition omega). intros. unfold assign_copy_ok. - assert (alignof_blockcopy ty > 0) by apply alignof_blockcopy_pos. - destruct (Zdivide_dec (alignof_blockcopy ty) (Int.unsigned ofs')); auto... - destruct (Zdivide_dec (alignof_blockcopy ty) (Int.unsigned ofs)); auto... + assert (alignof_blockcopy ge ty > 0) by apply alignof_blockcopy_pos. + destruct (Zdivide_dec (alignof_blockcopy ge ty) (Int.unsigned ofs')); auto... + destruct (Zdivide_dec (alignof_blockcopy ge ty) (Int.unsigned ofs)); auto... assert (Y: {b' <> b \/ Int.unsigned ofs' = Int.unsigned ofs \/ - Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs \/ - Int.unsigned ofs + sizeof ty <= Int.unsigned ofs'} + + Int.unsigned ofs' + sizeof ge ty <= Int.unsigned ofs \/ + Int.unsigned ofs + sizeof ge ty <= Int.unsigned ofs'} + {~(b' <> b \/ Int.unsigned ofs' = Int.unsigned ofs \/ - Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs \/ - Int.unsigned ofs + sizeof ty <= Int.unsigned ofs')}). + Int.unsigned ofs' + sizeof ge ty <= Int.unsigned ofs \/ + Int.unsigned ofs + sizeof ge ty <= Int.unsigned ofs')}). destruct (eq_block b' b); auto. destruct (zeq (Int.unsigned ofs') (Int.unsigned ofs)); auto. - destruct (zle (Int.unsigned ofs' + sizeof ty) (Int.unsigned ofs)); auto. - destruct (zle (Int.unsigned ofs + sizeof ty) (Int.unsigned ofs')); auto. + destruct (zle (Int.unsigned ofs' + sizeof ge ty) (Int.unsigned ofs)); auto. + destruct (zle (Int.unsigned ofs + sizeof ge ty) (Int.unsigned ofs')); auto. right; intuition omega. destruct Y... left; intuition omega. Defined. @@ -324,7 +324,7 @@ Definition do_assign_loc (w: world) (ty: type) (m: mem) (b: block) (ofs: int) (v match v with | Vptr b' ofs' => if check_assign_copy ty b ofs b' ofs' then - do bytes <- Mem.loadbytes m b' (Int.unsigned ofs') (sizeof ty); + do bytes <- Mem.loadbytes m b' (Int.unsigned ofs') (sizeof ge ty); do m' <- Mem.storebytes m b (Int.unsigned ofs) bytes; Some(w, E0, m') else None @@ -746,12 +746,14 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match is_val r with | Some(Vptr b ofs, ty') => match ty' with - | Tstruct id fList _ => - match field_offset f fList with + | Tstruct id _ => + do co <- ge.(genv_cenv)!id; + match field_offset ge f (co_members co) with | Error _ => stuck | OK delta => topred (Lred (Eloc b (Int.add ofs (Int.repr delta)) ty) m) end - | Tunion id fList _ => + | Tunion id _ => + do co <- ge.(genv_cenv)!id; topred (Lred (Eloc b ofs ty) m) | _ => stuck end @@ -787,7 +789,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := | RV, Ebinop op r1 r2 ty => match is_val r1, is_val r2 with | Some(v1, ty1), Some(v2, ty2) => - do v <- sem_binary_operation op v1 ty1 v2 ty2 m; + do v <- sem_binary_operation ge op v1 ty1 v2 ty2 m; topred (Rred (Eval v ty) m E0) | _, _ => incontext2 (fun x => Ebinop op x r2 ty) (step_expr RV r1 m) @@ -828,9 +830,9 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := incontext (fun x => Econdition x r2 r3 ty) (step_expr RV r1 m) end | RV, Esizeof ty' ty => - topred (Rred (Eval (Vint (Int.repr (sizeof ty'))) ty) m E0) + topred (Rred (Eval (Vint (Int.repr (sizeof ge ty'))) ty) m E0) | RV, Ealignof ty' ty => - topred (Rred (Eval (Vint (Int.repr (alignof ty'))) ty) m E0) + topred (Rred (Eval (Vint (Int.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) => @@ -988,8 +990,8 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop := | Efield (Eval v ty1) f ty => exists b, exists ofs, v = Vptr b ofs /\ match ty1 with - | Tstruct _ fList _ => exists delta, field_offset f fList = Errors.OK delta - | Tunion _ _ _ => True + | Tstruct 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 | _ => False end | Eval v ty => False @@ -998,7 +1000,7 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop := | Eunop op (Eval v1 ty1) ty => exists v, sem_unary_operation op v1 ty1 = Some v | Ebinop op (Eval v1 ty1) (Eval v2 ty2) ty => - exists v, sem_binary_operation op v1 ty1 v2 ty2 m = Some v + exists v, sem_binary_operation ge op v1 ty1 v2 ty2 m = Some v | Ecast (Eval v1 ty1) ty => exists v, sem_cast v1 ty1 ty = Some v | Eseqand (Eval v1 ty1) r2 ty => @@ -1043,8 +1045,8 @@ Proof. exists b; auto. exists b; auto. exists b; exists ofs; auto. - exists b; exists ofs; split; auto. exists delta; auto. - exists b; exists ofs; auto. + exists b; exists ofs; split; auto. exists co, delta; auto. + exists b; exists ofs; split; auto. exists co; auto. Qed. Lemma rred_invert: @@ -1385,10 +1387,12 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence; destruct v... destruct ty'... (* top struct *) - destruct (field_offset f f0) as [delta|] eqn:?... - apply topred_ok; auto. apply red_field_struct; auto. + destruct (ge.(genv_cenv)!i0) as [co|] eqn:?... + destruct (field_offset ge f (co_members co)) as [delta|] eqn:?... + apply topred_ok; auto. eapply red_field_struct; eauto. (* top union *) - apply topred_ok; auto. apply red_field_union; auto. + destruct (ge.(genv_cenv)!i0) as [co|] eqn:?... + apply topred_ok; auto. eapply red_field_union; eauto. (* in depth *) eapply incontext_ok; eauto. (* Evalof *) @@ -1425,7 +1429,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence; destruct (is_val a2) as [[v2 ty2] | ] eqn:?. rewrite (is_val_inv _ _ _ Heqo). rewrite (is_val_inv _ _ _ Heqo0). (* top *) - destruct (sem_binary_operation op v1 ty1 v2 ty2 m) as [v|] eqn:?... + destruct (sem_binary_operation ge op v1 ty1 v2 ty2 m) as [v|] eqn:?... apply topred_ok; auto. split. apply red_binop; auto. exists w; constructor. (* depth *) eapply incontext2_ok; eauto. @@ -1589,9 +1593,9 @@ Proof. (* deref *) auto. (* field struct *) - rewrite H; auto. + rewrite H, H0; auto. (* field union *) - auto. + rewrite H; auto. Qed. Lemma rred_topred: @@ -1895,21 +1899,21 @@ Fixpoint do_alloc_variables (e: env) (m: mem) (l: list (ident * type)) {struct l match l with | nil => (e,m) | (id, ty) :: l' => - let (m1,b1) := Mem.alloc m 0 (sizeof ty) in + let (m1,b1) := Mem.alloc m 0 (sizeof ge ty) in do_alloc_variables (PTree.set id (b1, ty) e) m1 l' end. Lemma do_alloc_variables_sound: - forall l e m, alloc_variables e m l (fst (do_alloc_variables e m l)) (snd (do_alloc_variables e m l)). + forall l e m, alloc_variables ge e m l (fst (do_alloc_variables e m l)) (snd (do_alloc_variables e m l)). Proof. induction l; intros; simpl. constructor. - destruct a as [id ty]. destruct (Mem.alloc m 0 (sizeof ty)) as [m1 b1] eqn:?; simpl. + destruct a as [id ty]. destruct (Mem.alloc m 0 (sizeof ge ty)) as [m1 b1] eqn:?; simpl. econstructor; eauto. Qed. Lemma do_alloc_variables_complete: - forall e1 m1 l e2 m2, alloc_variables e1 m1 l e2 m2 -> + forall e1 m1 l e2 m2, alloc_variables ge e1 m1 l e2 m2 -> do_alloc_variables e1 m1 l = (e2, m2). Proof. induction 1; simpl. @@ -1985,7 +1989,7 @@ Definition do_step (w: world) (s: state) : list (trace * state) := if b then ret (State f s (Kfor3 a2 a3 s k) e m) else ret (State f Sskip k e m) | Kreturn k => do v' <- sem_cast v ty f.(fn_return); - do m' <- Mem.free_list m (blocks_of_env e); + do m' <- Mem.free_list m (blocks_of_env ge e); ret (Returnstate v' (call_cont k) m') | Kswitch1 sl k => do n <- sem_switch_arg v ty; @@ -2024,11 +2028,11 @@ Definition do_step (w: world) (s: state) : list (trace * state) := | State f Sskip (Kfor4 a2 a3 s k) e m => ret (State f (Sfor Sskip a2 a3 s) k e m) | State f (Sreturn None) k e m => - do m' <- Mem.free_list m (blocks_of_env e); + do m' <- Mem.free_list m (blocks_of_env ge e); ret (Returnstate Vundef (call_cont k) m') | State f (Sreturn (Some x)) k e m => ret (ExprState f x (Kreturn k) e m) | State f Sskip ((Kstop | Kcall _ _ _ _ _) as k) e m => - do m' <- Mem.free_list m (blocks_of_env e); + do m' <- Mem.free_list m (blocks_of_env ge e); ret (Returnstate Vundef k m') | State f (Sswitch x sl) k e m => ret (ExprState f x (Kswitch1 sl k) e m) @@ -2209,7 +2213,7 @@ End EXEC. Local Open Scope option_monad_scope. Definition do_initial_state (p: program): option (genv * state) := - let ge := Genv.globalenv p in + let ge := globalenv p in do m0 <- Genv.init_mem p; do b <- Genv.find_symbol ge p.(prog_main); do f <- Genv.find_funct_ptr ge b; diff --git a/cfrontend/Clight.v b/cfrontend/Clight.v index 40206d38..7a45b453 100644 --- a/cfrontend/Clight.v +++ b/cfrontend/Clight.v @@ -57,12 +57,9 @@ Inductive expr : Type := | Eunop: unary_operation -> expr -> type -> expr (**r unary operation *) | Ebinop: binary_operation -> expr -> expr -> type -> expr (**r binary operation *) | Ecast: expr -> type -> expr (**r type cast ([(ty) e]) *) - | Efield: expr -> ident -> type -> expr. (**r access to a member of a struct or union *) - -(** [sizeof] and [alignof] are derived forms. *) - -Definition Esizeof (ty' ty: type) : expr := Econst_int (Int.repr(sizeof ty')) ty. -Definition Ealignof (ty' ty: type) : expr := Econst_int (Int.repr(alignof ty')) ty. + | Efield: expr -> ident -> type -> expr (**r access to a member of a struct or union *) + | Esizeof: type -> type -> expr (**r size of a type *) + | Ealignof: type -> type -> expr. (**r alignment of a type *) (** Extract the type part of a type-annotated Clight expression. *) @@ -80,6 +77,8 @@ Definition typeof (e: expr) : type := | Ebinop _ _ _ ty => ty | Ecast _ ty => ty | Efield _ _ ty => ty + | Esizeof _ ty => ty + | Ealignof _ ty => ty end. (** ** Statements *) @@ -164,19 +163,57 @@ Definition type_of_fundef (f: fundef) : type := (** ** Programs *) -(** A program is a collection of named functions, plus a collection - of named global variables, carrying their types and optional initialization - data. See module [AST] for more details. *) +(** A program is composed of: +- a list of definitions of functions and global variables; +- the names of functions and global variables that are public (not static); +- the name of the function that acts as entry point ("main" function). +- a list of definitions for structure and union names; +- the corresponding composite environment; +*) + +Record program : Type := { + prog_defs: list (ident * globdef fundef type); + prog_public: list ident; + prog_main: ident; + prog_types: list composite_definition; + prog_comp_env: composite_env; + prog_comp_env_eq: build_composite_env prog_types = OK prog_comp_env +}. -Definition program : Type := AST.program fundef type. +Definition program_of_program (p: program) : AST.program fundef type := + {| AST.prog_defs := p.(prog_defs); + AST.prog_public := p.(prog_public); + AST.prog_main := p.(prog_main) |}. + +Coercion program_of_program: program >-> AST.program. + +Program Definition make_program (types: list composite_definition) + (defs: list (ident * globdef fundef type)) + (public: list ident) + (main: ident): res program := + match build_composite_env types with + | OK env => + OK {| prog_defs := defs; + prog_public := public; + prog_main := main; + prog_types := types; + prog_comp_env := env; + prog_comp_env_eq := _ |} + | Error msg => + Error msg + end. (** * Operational semantics *) (** The semantics uses two environments. The global environment maps names of functions and global variables to memory block references, - and function pointers to their definitions. (See module [Globalenvs].) *) + and function pointers to their definitions. (See module [Globalenvs].) + It also contains a composite environment, used by type-dependent operations. *) -Definition genv := Genv.t fundef type. +Record genv := { genv_genv :> Genv.t fundef type; genv_cenv :> composite_env }. + +Definition globalenv (p: program) := + {| genv_genv := Genv.globalenv p; genv_cenv := p.(prog_comp_env) |}. (** The local environment maps local variables to block references and types. The current value of the variable is stored in the @@ -214,22 +251,26 @@ Inductive deref_loc (ty: type) (m: mem) (b: block) (ofs: int) : val -> Prop := This is allowed only if [ty] indicates an access by value or by copy. [m'] is the updated memory state. *) -Inductive assign_loc (ty: type) (m: mem) (b: block) (ofs: int): +Inductive assign_loc (ce: composite_env) (ty: type) (m: mem) (b: block) (ofs: int): val -> mem -> Prop := | assign_loc_value: forall v chunk m', access_mode ty = By_value chunk -> Mem.storev chunk m (Vptr b ofs) v = Some m' -> - assign_loc ty m b ofs v m' + assign_loc ce ty m b ofs v m' | assign_loc_copy: forall b' ofs' bytes m', access_mode ty = By_copy -> - (sizeof ty > 0 -> (alignof_blockcopy ty | Int.unsigned ofs')) -> - (sizeof ty > 0 -> (alignof_blockcopy ty | Int.unsigned ofs)) -> + (sizeof ce ty > 0 -> (alignof_blockcopy ce ty | Int.unsigned ofs')) -> + (sizeof ce ty > 0 -> (alignof_blockcopy ce ty | Int.unsigned ofs)) -> b' <> b \/ Int.unsigned ofs' = Int.unsigned ofs - \/ Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs - \/ Int.unsigned ofs + sizeof ty <= Int.unsigned ofs' -> - Mem.loadbytes m b' (Int.unsigned ofs') (sizeof ty) = Some bytes -> + \/ Int.unsigned ofs' + sizeof ce ty <= Int.unsigned ofs + \/ Int.unsigned ofs + sizeof ce ty <= Int.unsigned ofs' -> + Mem.loadbytes m b' (Int.unsigned ofs') (sizeof ce ty) = Some bytes -> Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' -> - assign_loc ty m b ofs (Vptr b' ofs') m'. + assign_loc ce ty m b ofs (Vptr b' ofs') m'. + +Section SEMANTICS. + +Variable ge: genv. (** Allocation of function-local variables. [alloc_variables e1 m1 vars e2 m2] allocates one memory block @@ -246,7 +287,7 @@ Inductive alloc_variables: env -> mem -> alloc_variables e m nil e m | alloc_variables_cons: forall e m id ty vars m1 b1 m2 e2, - Mem.alloc m 0 (sizeof ty) = (m1, b1) -> + Mem.alloc m 0 (sizeof ge ty) = (m1, b1) -> alloc_variables (PTree.set id (b1, ty) e) m1 vars e2 m2 -> alloc_variables e m ((id, ty) :: vars) e2 m2. @@ -264,7 +305,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 ty m b Int.zero v1 m1 -> + assign_loc ge ty m b Int.zero v1 m1 -> bind_parameters e m1 params vl m2 -> bind_parameters e m ((id, ty) :: params) (v1 :: vl) m2. @@ -289,7 +330,7 @@ Fixpoint bind_parameter_temps (formals: list (ident * type)) (args: list val) (** Return the list of blocks in the codomain of [e], with low and high bounds. *) Definition block_of_binding (id_b_ty: ident * (block * type)) := - match id_b_ty with (id, (b, ty)) => (b, 0, sizeof ty) end. + match id_b_ty with (id, (b, ty)) => (b, 0, sizeof ge ty) end. Definition blocks_of_env (e: env) : list (block * Z * Z) := List.map block_of_binding (PTree.elements e). @@ -333,10 +374,6 @@ Fixpoint seq_of_labeled_statement (sl: labeled_statements) : statement := | LScons _ s sl' => Ssequence s (seq_of_labeled_statement sl') end. -Section SEMANTICS. - -Variable ge: genv. - (** ** Evaluation of expressions *) Section EXPR. @@ -371,12 +408,16 @@ Inductive eval_expr: expr -> val -> Prop := | eval_Ebinop: forall op a1 a2 ty v1 v2 v, eval_expr a1 v1 -> eval_expr a2 v2 -> - sem_binary_operation op v1 (typeof a1) v2 (typeof a2) m = Some v -> + sem_binary_operation ge op v1 (typeof a1) v2 (typeof a2) m = Some v -> eval_expr (Ebinop op a1 a2 ty) v | eval_Ecast: forall a ty v1 v, eval_expr a v1 -> sem_cast v1 (typeof a) ty = Some v -> eval_expr (Ecast a ty) v + | eval_Esizeof: forall ty1 ty, + eval_expr (Esizeof ty1 ty) (Vint (Int.repr (sizeof ge ty1))) + | eval_Ealignof: forall ty1 ty, + eval_expr (Ealignof ty1 ty) (Vint (Int.repr (alignof ge ty1))) | eval_Elvalue: forall a loc ofs v, eval_lvalue a loc ofs -> deref_loc (typeof a) m loc ofs v -> @@ -397,14 +438,16 @@ with eval_lvalue: expr -> block -> int -> Prop := | eval_Ederef: forall a ty l ofs, eval_expr a (Vptr l ofs) -> eval_lvalue (Ederef a ty) l ofs - | eval_Efield_struct: forall a i ty l ofs id fList att delta, + | eval_Efield_struct: forall a i ty l ofs id co att delta, eval_expr a (Vptr l ofs) -> - typeof a = Tstruct id fList att -> - field_offset i fList = OK delta -> + 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 (Int.add ofs (Int.repr delta)) - | eval_Efield_union: forall a i ty l ofs id fList att, + | eval_Efield_union: forall a i ty l ofs id co att, eval_expr a (Vptr l ofs) -> - typeof a = Tunion id fList att -> + typeof a = Tunion id att -> + ge.(genv_cenv)!id = Some co -> eval_lvalue (Efield a i ty) l ofs. Scheme eval_expr_ind2 := Minimality for eval_expr Sort Prop @@ -538,7 +581,7 @@ Inductive step: state -> trace -> state -> Prop := eval_lvalue e le m a1 loc ofs -> eval_expr e le m a2 v2 -> sem_cast v2 (typeof a2) (typeof a1) = Some v -> - assign_loc (typeof a1) m loc ofs v m' -> + assign_loc ge (typeof a1) m loc ofs v m' -> step (State f (Sassign a1 a2) k e le m) E0 (State f Sskip k e le m') @@ -676,45 +719,48 @@ End SEMANTICS. (** The two semantics for function parameters. First, parameters as local variables. *) -Inductive function_entry1 (f: function) (vargs: list val) (m: mem) (e: env) (le: temp_env) (m': mem) : Prop := +Inductive function_entry1 (ge: genv) (f: function) (vargs: list val) (m: mem) (e: env) (le: temp_env) (m': mem) : Prop := | function_entry1_intro: forall m1, list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) -> - alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> - bind_parameters e m1 f.(fn_params) vargs m' -> + alloc_variables ge empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> + bind_parameters ge e m1 f.(fn_params) vargs m' -> le = create_undef_temps f.(fn_temps) -> - function_entry1 f vargs m e le m'. + function_entry1 ge f vargs m e le m'. -Definition step1 (ge: genv) := step ge function_entry1. +Definition step1 (ge: genv) := step ge (function_entry1 ge). (** Second, parameters as temporaries. *) -Inductive function_entry2 (f: function) (vargs: list val) (m: mem) (e: env) (le: temp_env) (m': mem) : Prop := +Inductive function_entry2 (ge: genv) (f: function) (vargs: list val) (m: mem) (e: env) (le: temp_env) (m': mem) : Prop := | function_entry2_intro: list_norepet (var_names f.(fn_vars)) -> list_norepet (var_names f.(fn_params)) -> list_disjoint (var_names f.(fn_params)) (var_names f.(fn_temps)) -> - alloc_variables empty_env m f.(fn_vars) e m' -> + alloc_variables ge empty_env m f.(fn_vars) e m' -> bind_parameter_temps f.(fn_params) vargs (create_undef_temps f.(fn_temps)) = Some le -> - function_entry2 f vargs m e le m'. + function_entry2 ge f vargs m e le m'. -Definition step2 (ge: genv) := step ge function_entry2. +Definition step2 (ge: genv) := step ge (function_entry2 ge). (** Wrapping up these definitions in two small-step semantics. *) Definition semantics1 (p: program) := - Semantics step1 (initial_state p) final_state (Genv.globalenv p). + let ge := globalenv p in + Semantics_gen step1 (initial_state p) final_state ge ge. Definition semantics2 (p: program) := - Semantics step2 (initial_state p) final_state (Genv.globalenv p). + let ge := globalenv p in + Semantics_gen step2 (initial_state p) final_state ge ge. (** This semantics is receptive to changes in events. *) Lemma semantics_receptive: forall (p: program), receptive (semantics1 p). Proof. - intros. constructor; simpl; intros. + intros. unfold semantics1. + set (ge := globalenv p). constructor; simpl; intros. (* receptiveness *) - assert (t1 = E0 -> exists s2, step1 (Genv.globalenv p) s t2 s2). + assert (t1 = E0 -> exists s2, step1 ge s t2 s2). intros. subst. inv H0. exists s1; auto. inversion H; subst; auto. (* builtin *) @@ -724,7 +770,7 @@ Proof. exploit external_call_receptive; eauto. intros [vres2 [m2 EC2]]. exists (Returnstate vres2 k m2). econstructor; eauto. (* trace length *) - red; intros. inv H; simpl; try omega. + red; simpl; intros. inv H; simpl; try omega. eapply external_call_trace_length; eauto. eapply external_call_trace_length; eauto. Qed. diff --git a/cfrontend/ClightBigstep.v b/cfrontend/ClightBigstep.v index d61e4eef..5b092db7 100644 --- a/cfrontend/ClightBigstep.v +++ b/cfrontend/ClightBigstep.v @@ -82,7 +82,7 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env -> eval_lvalue ge e le m a1 loc ofs -> eval_expr ge e le m a2 v2 -> sem_cast v2 (typeof a2) (typeof a1) = Some v -> - assign_loc (typeof a1) m loc ofs v m' -> + assign_loc ge (typeof a1) m loc ofs v m' -> exec_stmt e le m (Sassign a1 a2) E0 le m' Out_normal | exec_Sset: forall e le m id a v, @@ -164,12 +164,12 @@ Inductive exec_stmt: env -> temp_env -> mem -> statement -> trace -> temp_env -> with eval_funcall: mem -> fundef -> list val -> trace -> mem -> val -> Prop := | eval_funcall_internal: forall le m f vargs t e m1 m2 m3 out vres m4, - alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> + alloc_variables ge empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) -> - bind_parameters e m1 f.(fn_params) vargs m2 -> + bind_parameters ge e m1 f.(fn_params) vargs m2 -> exec_stmt e (create_undef_temps f.(fn_temps)) m2 f.(fn_body) t le m3 out -> outcome_result_value out f.(fn_return) vres -> - Mem.free_list m3 (blocks_of_env e) = Some m4 -> + Mem.free_list m3 (blocks_of_env ge e) = Some m4 -> eval_funcall m (Internal f) vargs t m4 vres | eval_funcall_external: forall m ef targs tres cconv vargs t vres m', external_call ef ge vargs m t vres m' -> @@ -232,9 +232,9 @@ CoInductive execinf_stmt: env -> temp_env -> mem -> statement -> traceinf -> Pro with evalinf_funcall: mem -> fundef -> list val -> traceinf -> Prop := | evalinf_funcall_internal: forall m f vargs t e m1 m2, - alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> + alloc_variables ge empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) -> - bind_parameters e m1 f.(fn_params) vargs m2 -> + bind_parameters ge e m1 f.(fn_params) vargs m2 -> execinf_stmt e (create_undef_temps f.(fn_temps)) m2 f.(fn_body) t -> evalinf_funcall m (Internal f) vargs t. @@ -244,7 +244,7 @@ End BIGSTEP. Inductive bigstep_program_terminates (p: program): trace -> int -> Prop := | bigstep_program_terminates_intro: forall b f m0 m1 t r, - let ge := Genv.globalenv p in + let ge := globalenv p in Genv.init_mem p = Some m0 -> Genv.find_symbol ge p.(prog_main) = Some b -> Genv.find_funct_ptr ge b = Some f -> @@ -254,7 +254,7 @@ Inductive bigstep_program_terminates (p: program): trace -> int -> Prop := Inductive bigstep_program_diverges (p: program): traceinf -> Prop := | bigstep_program_diverges_intro: forall b f m0 t, - let ge := Genv.globalenv p in + let ge := globalenv p in Genv.init_mem p = Some m0 -> Genv.find_symbol ge p.(prog_main) = Some b -> Genv.find_funct_ptr ge b = Some f -> @@ -270,7 +270,7 @@ Definition bigstep_semantics (p: program) := Section BIGSTEP_TO_TRANSITIONS. Variable prog: program. -Let ge : genv := Genv.globalenv prog. +Let ge : genv := globalenv prog. Inductive outcome_state_match (e: env) (le: temp_env) (m: mem) (f: function) (k: cont): outcome -> state -> Prop := diff --git a/cfrontend/Cop.v b/cfrontend/Cop.v index a5d4c662..4e572277 100644 --- a/cfrontend/Cop.v +++ b/cfrontend/Cop.v @@ -93,17 +93,17 @@ Inductive classify_cast_cases : Type := | cast_case_s2bool (**r single -> bool *) | cast_case_l2bool (**r long -> bool *) | cast_case_p2bool (**r pointer -> bool *) - | cast_case_struct (id1: ident) (fld1: fieldlist) (id2: ident) (fld2: fieldlist) (**r struct -> struct *) - | cast_case_union (id1: ident) (fld1: fieldlist) (id2: ident) (fld2: fieldlist) (**r union -> union *) + | cast_case_struct (id1 id2: ident) (**r struct -> struct *) + | cast_case_union (id1 id2: ident) (**r union -> union *) | cast_case_void (**r any -> void *) | cast_case_default. Definition classify_cast (tfrom tto: type) : classify_cast_cases := match tto, tfrom with - | Tint I32 si2 _, (Tint _ _ _ | Tpointer _ _ | Tcomp_ptr _ _ | Tarray _ _ _ | Tfunction _ _ _) => cast_case_neutral + | Tint I32 si2 _, (Tint _ _ _ | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _) => cast_case_neutral | Tint IBool _ _, Tfloat F64 _ => cast_case_f2bool | Tint IBool _ _, Tfloat F32 _ => cast_case_s2bool - | Tint IBool _ _, (Tpointer _ _ | Tcomp_ptr _ _ | Tarray _ _ _ | Tfunction _ _ _) => cast_case_p2bool + | Tint IBool _ _, (Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _) => cast_case_p2bool | Tint sz2 si2 _, Tint sz1 si1 _ => cast_case_i2i sz2 si2 | Tint sz2 si2 _, Tfloat F64 _ => cast_case_f2i sz2 si2 | Tint sz2 si2 _, Tfloat F32 _ => cast_case_s2i sz2 si2 @@ -113,7 +113,7 @@ Definition classify_cast (tfrom tto: type) : classify_cast_cases := | Tfloat F32 _, Tfloat F64 _ => cast_case_f2s | Tfloat F64 _, Tint sz1 si1 _ => cast_case_i2f si1 | Tfloat F32 _, Tint sz1 si1 _ => cast_case_i2s si1 - | (Tpointer _ _ | Tcomp_ptr _ _), (Tint _ _ _ | Tpointer _ _ | Tcomp_ptr _ _ | Tarray _ _ _ | Tfunction _ _ _) => cast_case_neutral + | Tpointer _ _, (Tint _ _ _ | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _) => cast_case_neutral | Tlong _ _, Tlong _ _ => cast_case_l2l | Tlong _ _, Tint sz1 si1 _ => cast_case_i2l si1 | Tint IBool _ _, Tlong _ _ => cast_case_l2bool @@ -122,10 +122,10 @@ Definition classify_cast (tfrom tto: type) : classify_cast_cases := | Tlong si2 _, Tfloat F32 _ => cast_case_s2l si2 | Tfloat F64 _, Tlong si1 _ => cast_case_l2f si1 | Tfloat F32 _, Tlong si1 _ => cast_case_l2s si1 - | (Tpointer _ _ | Tcomp_ptr _ _), Tlong _ _ => cast_case_l2i I32 Unsigned - | Tlong si2 _, (Tpointer _ _ | Tcomp_ptr _ _ | Tarray _ _ _ | Tfunction _ _ _) => cast_case_i2l si2 - | Tstruct id2 fld2 _, Tstruct id1 fld1 _ => cast_case_struct id1 fld1 id2 fld2 - | Tunion id2 fld2 _, Tunion id1 fld1 _ => cast_case_union id1 fld1 id2 fld2 + | Tpointer _ _, Tlong _ _ => cast_case_l2i I32 Unsigned + | Tlong si2 _, (Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _) => cast_case_i2l si2 + | Tstruct id2 _, Tstruct id1 _ => cast_case_struct id1 id2 + | Tunion id2 _, Tunion id1 _ => cast_case_union id1 id2 | Tvoid, _ => cast_case_void | _, _ => cast_case_default end. @@ -325,16 +325,16 @@ Definition sem_cast (v: val) (t1 t2: type) : option val := end | _ => None end - | cast_case_struct id1 fld1 id2 fld2 => + | cast_case_struct id1 id2 => match v with | Vptr b ofs => - if ident_eq id1 id2 && fieldlist_eq fld1 fld2 then Some v else None + if ident_eq id1 id2 then Some v else None | _ => None end - | cast_case_union id1 fld1 id2 fld2 => + | cast_case_union id1 id2 => match v with | Vptr b ofs => - if ident_eq id1 id2 && fieldlist_eq fld1 fld2 then Some v else None + if ident_eq id1 id2 then Some v else None | _ => None end | cast_case_void => @@ -359,7 +359,7 @@ Inductive classify_bool_cases : Type := Definition classify_bool (ty: type) : classify_bool_cases := match typeconv ty with | Tint _ _ _ => bool_case_i - | Tpointer _ _ | Tcomp_ptr _ _ => bool_case_p + | Tpointer _ _ => bool_case_p | Tfloat F64 _ => bool_case_f | Tfloat F32 _ => bool_case_s | Tlong _ _ => bool_case_l @@ -402,7 +402,6 @@ Definition bool_val (v: val) (t: type) : option bool := | bool_default => None end. - (** ** Unary operators *) (** *** Boolean negation *) @@ -639,32 +638,32 @@ Definition classify_add (ty1: type) (ty2: type) := | _, _ => add_default end. -Definition sem_add (v1:val) (t1:type) (v2: val) (t2:type) : option val := +Definition sem_add (cenv: composite_env) (v1:val) (t1:type) (v2: val) (t2:type) : option val := match classify_add t1 t2 with | add_case_pi ty => (**r pointer plus integer *) match v1,v2 with | Vptr b1 ofs1, Vint n2 => - Some (Vptr b1 (Int.add ofs1 (Int.mul (Int.repr (sizeof ty)) n2))) + Some (Vptr b1 (Int.add ofs1 (Int.mul (Int.repr (sizeof cenv ty)) n2))) | _, _ => None end | add_case_ip ty => (**r integer plus pointer *) match v1,v2 with | Vint n1, Vptr b2 ofs2 => - Some (Vptr b2 (Int.add ofs2 (Int.mul (Int.repr (sizeof ty)) n1))) + Some (Vptr b2 (Int.add ofs2 (Int.mul (Int.repr (sizeof cenv ty)) n1))) | _, _ => None end | add_case_pl ty => (**r pointer plus long *) match v1,v2 with | Vptr b1 ofs1, Vlong n2 => let n2 := Int.repr (Int64.unsigned n2) in - Some (Vptr b1 (Int.add ofs1 (Int.mul (Int.repr (sizeof ty)) n2))) + Some (Vptr b1 (Int.add ofs1 (Int.mul (Int.repr (sizeof cenv ty)) n2))) | _, _ => None end | add_case_lp ty => (**r long plus pointer *) match v1,v2 with | Vlong n1, Vptr b2 ofs2 => let n1 := Int.repr (Int64.unsigned n1) in - Some (Vptr b2 (Int.add ofs2 (Int.mul (Int.repr (sizeof ty)) n1))) + Some (Vptr b2 (Int.add ofs2 (Int.mul (Int.repr (sizeof cenv ty)) n1))) | _, _ => None end | add_default => @@ -692,27 +691,27 @@ Definition classify_sub (ty1: type) (ty2: type) := | _, _ => sub_default end. -Definition sem_sub (v1:val) (t1:type) (v2: val) (t2:type) : option val := +Definition sem_sub (cenv: composite_env) (v1:val) (t1:type) (v2: val) (t2:type) : option val := match classify_sub t1 t2 with | sub_case_pi ty => (**r pointer minus integer *) match v1,v2 with | Vptr b1 ofs1, Vint n2 => - Some (Vptr b1 (Int.sub ofs1 (Int.mul (Int.repr (sizeof ty)) n2))) + Some (Vptr b1 (Int.sub ofs1 (Int.mul (Int.repr (sizeof cenv ty)) n2))) | _, _ => None end | sub_case_pl ty => (**r pointer minus long *) match v1,v2 with | Vptr b1 ofs1, Vlong n2 => let n2 := Int.repr (Int64.unsigned n2) in - Some (Vptr b1 (Int.sub ofs1 (Int.mul (Int.repr (sizeof ty)) n2))) + Some (Vptr b1 (Int.sub ofs1 (Int.mul (Int.repr (sizeof cenv ty)) n2))) | _, _ => None end | sub_case_pp ty => (**r pointer minus pointer *) match v1,v2 with | Vptr b1 ofs1, Vptr b2 ofs2 => if eq_block b1 b2 then - if Int.eq (Int.repr (sizeof ty)) Int.zero then None - else Some (Vint (Int.divu (Int.sub ofs1 ofs2) (Int.repr (sizeof ty)))) + if Int.eq (Int.repr (sizeof cenv ty)) Int.zero then None + else Some (Vint (Int.divu (Int.sub ofs1 ofs2) (Int.repr (sizeof cenv ty)))) else None | _, _ => None end @@ -983,12 +982,13 @@ Definition sem_unary_operation end. Definition sem_binary_operation + (cenv: composite_env) (op: binary_operation) (v1: val) (t1: type) (v2: val) (t2:type) (m: mem): option val := match op with - | Oadd => sem_add v1 t1 v2 t2 - | Osub => sem_sub v1 t1 v2 t2 + | Oadd => sem_add cenv v1 t1 v2 t2 + | Osub => sem_sub cenv v1 t1 v2 t2 | Omul => sem_mul v1 t1 v2 t2 | Omod => sem_mod v1 t1 v2 t2 | Odiv => sem_div v1 t1 v2 t2 @@ -1005,10 +1005,10 @@ Definition sem_binary_operation | Oge => sem_cmp Cge v1 t1 v2 t2 m end. -Definition sem_incrdecr (id: incr_or_decr) (v: val) (ty: type) := +Definition sem_incrdecr (cenv: composite_env) (id: incr_or_decr) (v: val) (ty: type) := match id with - | Incr => sem_add v ty (Vint Int.one) type_int32s - | Decr => sem_sub v ty (Vint Int.one) type_int32s + | Incr => sem_add cenv v ty (Vint Int.one) type_int32s + | Decr => sem_sub cenv v ty (Vint Int.one) type_int32s end. Definition incrdecr_type (ty: type) := @@ -1086,8 +1086,8 @@ Proof. - destruct (cast_single_int si2 f0); inv H1; TrivialInject. - destruct (cast_float_long si2 f0); inv H1; TrivialInject. - destruct (cast_single_long si2 f0); inv H1; TrivialInject. -- destruct (ident_eq id1 id2 && fieldlist_eq fld1 fld2); inv H2; TrivialInject. econstructor; eauto. -- destruct (ident_eq id1 id2 && fieldlist_eq fld1 fld2); inv H2; TrivialInject. econstructor; eauto. +- destruct (ident_eq id1 id2); inv H2; TrivialInject. econstructor; eauto. +- destruct (ident_eq id1 id2); inv H2; TrivialInject. econstructor; eauto. - econstructor; eauto. Qed. @@ -1191,10 +1191,10 @@ Proof. Qed. Lemma sem_binary_operation_inj: - forall op v1 ty1 v2 ty2 v tv1 tv2, - sem_binary_operation op v1 ty1 v2 ty2 m = Some v -> + forall cenv op v1 ty1 v2 ty2 v tv1 tv2, + sem_binary_operation cenv op v1 ty1 v2 ty2 m = Some v -> val_inject f v1 tv1 -> val_inject f v2 tv2 -> - exists tv, sem_binary_operation op tv1 ty1 tv2 ty2 m' = Some tv /\ val_inject f v tv. + exists tv, sem_binary_operation cenv op tv1 ty1 tv2 ty2 m' = Some tv /\ val_inject f v tv. Proof. unfold sem_binary_operation; intros; destruct op. - (* add *) @@ -1215,7 +1215,7 @@ Proof. + inv H0; inv H1; inv H. TrivialInject. destruct (eq_block b1 b0); try discriminate. subst b1. rewrite H0 in H2; inv H2. rewrite dec_eq_true. - destruct (Int.eq (Int.repr (sizeof ty)) Int.zero); inv H3. + destruct (Int.eq (Int.repr (sizeof cenv ty)) Int.zero); inv H3. rewrite Int.sub_shifted. TrivialInject. + inv H0; inv H1; inv H. TrivialInject. econstructor. eauto. rewrite Int.sub_add_l. auto. @@ -1278,11 +1278,11 @@ Qed. End GENERIC_INJECTION. Lemma sem_binary_operation_inject: - forall f m m' op v1 ty1 v2 ty2 v tv1 tv2, - sem_binary_operation op v1 ty1 v2 ty2 m = Some v -> + forall f m m' cenv op v1 ty1 v2 ty2 v tv1 tv2, + sem_binary_operation cenv op v1 ty1 v2 ty2 m = Some v -> val_inject f v1 tv1 -> val_inject f v2 tv2 -> Mem.inject f m m' -> - exists tv, sem_binary_operation op tv1 ty1 tv2 ty2 m' = Some tv /\ val_inject f v tv. + exists tv, sem_binary_operation cenv op tv1 ty1 tv2 ty2 m' = Some tv /\ val_inject f v tv. Proof. intros. eapply sem_binary_operation_inj; eauto. intros; eapply Mem.valid_pointer_inject_val; eauto. @@ -1309,7 +1309,7 @@ Proof. assert (A: classify_bool t = match t with | Tint _ _ _ => bool_case_i - | Tpointer _ _ | Tcomp_ptr _ _ | Tarray _ _ _ | Tfunction _ _ _ => bool_case_p + | Tpointer _ _ | Tarray _ _ _ | Tfunction _ _ _ => bool_case_p | Tfloat F64 _ => bool_case_f | Tfloat F32 _ => bool_case_s | Tlong _ _ => bool_case_l @@ -1332,7 +1332,6 @@ Proof. destruct (Int.eq i Int.zero); auto. destruct (Int.eq i Int.zero); auto. destruct (Int.eq i Int.zero); auto. - destruct (Int.eq i0 Int.zero); auto. Qed. (** Relation between Boolean value and Boolean negation. *) @@ -1528,5 +1527,4 @@ End ArithConv. - diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v index 06cea006..e6e3a321 100644 --- a/cfrontend/Csem.v +++ b/cfrontend/Csem.v @@ -34,9 +34,13 @@ Require Import Smallstep. (** The semantics uses two environments. The global environment maps names of functions and global variables to memory block references, - and function pointers to their definitions. (See module [Globalenvs].) *) + and function pointers to their definitions. (See module [Globalenvs].) + It also contains a composite environment, used by type-dependent operations. *) -Definition genv := Genv.t fundef type. +Record genv := { genv_genv :> Genv.t fundef type; genv_cenv :> composite_env }. + +Definition globalenv (p: program) := + {| genv_genv := Genv.globalenv p; genv_cenv := p.(prog_comp_env) |}. (** The local environment maps local variables to block references and types. The current value of the variable is stored in the associated memory @@ -46,6 +50,11 @@ Definition env := PTree.t (block * type). (* map variable -> location & type *) Definition empty_env: env := (PTree.empty (block * type)). + +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]. If the type [ty] indicates an access by value, the corresponding @@ -54,22 +63,22 @@ Definition empty_env: env := (PTree.empty (block * type)). returned, and [t] the trace of observables (nonempty if this is a volatile access). *) -Inductive deref_loc {F V: Type} (ge: Genv.t F V) (ty: type) (m: mem) (b: block) (ofs: int) : trace -> val -> Prop := +Inductive deref_loc (ty: type) (m: mem) (b: block) (ofs: int) : trace -> val -> Prop := | deref_loc_value: forall chunk v, access_mode ty = By_value chunk -> type_is_volatile ty = false -> Mem.loadv chunk m (Vptr b ofs) = Some v -> - deref_loc ge ty m b ofs E0 v + deref_loc ty m b ofs E0 v | deref_loc_volatile: forall chunk t v, access_mode ty = By_value chunk -> type_is_volatile ty = true -> volatile_load ge chunk m b ofs t v -> - deref_loc ge ty m b ofs t v + deref_loc ty m b ofs t v | deref_loc_reference: access_mode ty = By_reference -> - deref_loc ge ty m b ofs E0 (Vptr b ofs) + deref_loc ty m b ofs E0 (Vptr b ofs) | deref_loc_copy: access_mode ty = By_copy -> - deref_loc ge ty m b ofs E0 (Vptr b ofs). + deref_loc ty m b ofs E0 (Vptr b ofs). (** Symmetrically, [assign_loc ty m b ofs v t m'] returns the memory state after storing the value [v] in the datum @@ -78,27 +87,27 @@ Inductive deref_loc {F V: Type} (ge: Genv.t F V) (ty: type) (m: mem) (b: block) [m'] is the updated memory state and [t] the trace of observables (nonempty if this is a volatile store). *) -Inductive assign_loc {F V: Type} (ge: Genv.t F V) (ty: type) (m: mem) (b: block) (ofs: int): +Inductive assign_loc (ty: type) (m: mem) (b: block) (ofs: int): val -> trace -> mem -> Prop := | assign_loc_value: forall v chunk m', access_mode ty = By_value chunk -> type_is_volatile ty = false -> Mem.storev chunk m (Vptr b ofs) v = Some m' -> - assign_loc ge ty m b ofs v E0 m' + assign_loc ty m b ofs v E0 m' | assign_loc_volatile: forall v chunk t m', access_mode ty = By_value chunk -> type_is_volatile ty = true -> volatile_store ge chunk m b ofs v t m' -> - assign_loc ge ty m b ofs v t m' + assign_loc ty m b ofs v t m' | assign_loc_copy: forall b' ofs' bytes m', access_mode ty = By_copy -> - (alignof_blockcopy ty | Int.unsigned ofs') -> - (alignof_blockcopy ty | Int.unsigned ofs) -> + (alignof_blockcopy ge ty | Int.unsigned ofs') -> + (alignof_blockcopy ge ty | Int.unsigned ofs) -> b' <> b \/ Int.unsigned ofs' = Int.unsigned ofs - \/ Int.unsigned ofs' + sizeof ty <= Int.unsigned ofs - \/ Int.unsigned ofs + sizeof ty <= Int.unsigned ofs' -> - Mem.loadbytes m b' (Int.unsigned ofs') (sizeof ty) = Some bytes -> + \/ Int.unsigned ofs' + sizeof ge ty <= Int.unsigned ofs + \/ Int.unsigned ofs + sizeof ge ty <= Int.unsigned ofs' -> + Mem.loadbytes m b' (Int.unsigned ofs') (sizeof ge ty) = Some bytes -> Mem.storebytes m b (Int.unsigned ofs) bytes = Some m' -> - assign_loc ge ty m b ofs (Vptr b' ofs') E0 m'. + assign_loc ty m b ofs (Vptr b' ofs') E0 m'. (** Allocation of function-local variables. [alloc_variables e1 m1 vars e2 m2] allocates one memory block @@ -115,7 +124,7 @@ Inductive alloc_variables: env -> mem -> alloc_variables e m nil e m | alloc_variables_cons: forall e m id ty vars m1 b1 m2 e2, - Mem.alloc m 0 (sizeof ty) = (m1, b1) -> + Mem.alloc m 0 (sizeof ge ty) = (m1, b1) -> alloc_variables (PTree.set id (b1, ty) e) m1 vars e2 m2 -> alloc_variables e m ((id, ty) :: vars) e2 m2. @@ -124,23 +133,23 @@ Inductive alloc_variables: env -> mem -> in the memory blocks corresponding to the variables [params]. [m1] is the initial memory state and [m2] the final memory state. *) -Inductive bind_parameters {F V: Type} (ge: Genv.t F V) (e: env): +Inductive bind_parameters (e: env): mem -> list (ident * type) -> list val -> mem -> Prop := | bind_parameters_nil: forall m, - bind_parameters ge e m nil nil m + bind_parameters e m nil nil m | bind_parameters_cons: forall m id ty params v1 vl b m1 m2, PTree.get id e = Some(b, ty) -> - assign_loc ge ty m b Int.zero v1 E0 m1 -> - bind_parameters ge e m1 params vl m2 -> - bind_parameters ge e m ((id, ty) :: params) (v1 :: vl) m2. + assign_loc ty m b Int.zero v1 E0 m1 -> + bind_parameters e m1 params vl m2 -> + bind_parameters e m ((id, ty) :: params) (v1 :: vl) m2. (** Return the list of blocks in the codomain of [e], with low and high bounds. *) Definition block_of_binding (id_b_ty: ident * (block * type)) := - match id_b_ty with (id, (b, ty)) => (b, 0, sizeof ty) end. + match id_b_ty with (id, (b, ty)) => (b, 0, sizeof ge ty) end. Definition blocks_of_env (e: env) : list (block * Z * Z) := List.map block_of_binding (PTree.elements e). @@ -185,10 +194,6 @@ Inductive cast_arguments: exprlist -> typelist -> list val -> Prop := sem_cast v ty targ1 = Some v1 -> cast_arguments el targs vl -> cast_arguments (Econs (Eval v ty) el) (Tcons targ1 targs) (v1 :: vl). -Section SEMANTICS. - -Variable ge: genv. - (** ** Reduction semantics for expressions *) Section EXPR. @@ -215,19 +220,21 @@ Inductive lred: expr -> mem -> expr -> mem -> Prop := | red_deref: forall b ofs ty1 ty m, lred (Ederef (Eval (Vptr b ofs) ty1) ty) m (Eloc b ofs ty) m - | red_field_struct: forall b ofs id fList a f ty m delta, - field_offset f fList = OK delta -> - lred (Efield (Eval (Vptr b ofs) (Tstruct id fList a)) f ty) m + | red_field_struct: forall b ofs id co a f ty m delta, + ge.(genv_cenv)!id = Some co -> + field_offset ge f (co_members co) = OK delta -> + lred (Efield (Eval (Vptr b ofs) (Tstruct id a)) f ty) m (Eloc b (Int.add ofs (Int.repr delta)) ty) m - | red_field_union: forall b ofs id fList a f ty m, - lred (Efield (Eval (Vptr b ofs) (Tunion id fList a)) f ty) m + | red_field_union: forall b ofs id co a f ty m, + ge.(genv_cenv)!id = Some co -> + lred (Efield (Eval (Vptr b ofs) (Tunion id a)) f ty) m (Eloc b ofs 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 ge ty m b ofs t v -> + deref_loc ty m b ofs t v -> rred (Evalof (Eloc b ofs ty) ty) m t (Eval v ty) m | red_addrof: forall b ofs ty1 ty m, @@ -238,7 +245,7 @@ Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop := rred (Eunop op (Eval v1 ty1) ty) m E0 (Eval v ty) m | red_binop: forall op v1 ty1 v2 ty2 ty m v, - sem_binary_operation op v1 ty1 v2 ty2 m = Some v -> + sem_binary_operation ge op v1 ty1 v2 ty2 m = Some v -> rred (Ebinop op (Eval v1 ty1) (Eval v2 ty2) ty) m E0 (Eval v ty) m | red_cast: forall ty v1 ty1 m v, @@ -267,22 +274,22 @@ Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop := E0 (Eparen (if b then r1 else r2) ty ty) m | red_sizeof: forall ty1 ty m, rred (Esizeof ty1 ty) m - E0 (Eval (Vint (Int.repr (sizeof ty1))) ty) m + E0 (Eval (Vint (Int.repr (sizeof ge ty1))) ty) m | red_alignof: forall ty1 ty m, rred (Ealignof ty1 ty) m - E0 (Eval (Vint (Int.repr (alignof ty1))) ty) m + E0 (Eval (Vint (Int.repr (alignof ge ty1))) ty) m | red_assign: forall b ofs ty1 v2 ty2 m v t m', sem_cast v2 ty2 ty1 = Some v -> - assign_loc ge ty1 m b ofs v t m' -> + 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 ge ty1 m b ofs 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) (Ebinop op (Eval v1 ty1) (Eval v2 ty2) tyres) ty1) m | red_postincr: forall id b ofs ty m t v1 op, - deref_loc ge ty m b ofs t v1 -> + deref_loc ty m b ofs t v1 -> op = match id with Incr => Oadd | Decr => Osub end -> rred (Epostincr id (Eloc b ofs ty) ty) m t (Ecomma (Eassign (Eloc b ofs ty) @@ -739,7 +746,7 @@ Inductive sstep: state -> trace -> state -> Prop := | step_internal_function: forall f vargs k m e m1 m2, list_norepet (var_names (fn_params f) ++ var_names (fn_vars f)) -> alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> - bind_parameters ge e m1 f.(fn_params) vargs m2 -> + bind_parameters e m1 f.(fn_params) vargs m2 -> sstep (Callstate (Internal f) vargs k m) E0 (State f f.(fn_body) k e m2) @@ -766,7 +773,7 @@ End SEMANTICS. Inductive initial_state (p: program): state -> Prop := | initial_state_intro: forall b f m0, - let ge := Genv.globalenv p in + let ge := globalenv p in Genv.init_mem p = Some m0 -> Genv.find_symbol ge p.(prog_main) = Some b -> Genv.find_funct_ptr ge b = Some f -> @@ -782,19 +789,20 @@ Inductive final_state: state -> int -> Prop := (** Wrapping up these definitions in a small-step semantics. *) Definition semantics (p: program) := - Semantics step (initial_state p) final_state (Genv.globalenv p). + Semantics_gen step (initial_state p) final_state (globalenv p) (globalenv p). (** This semantics has the single-event property. *) Lemma semantics_single_events: forall p, single_events (semantics p). Proof. - intros; red; intros. destruct H. - set (ge := globalenv (semantics p)) in *. + 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 omega. inv H3; simpl; try omega. assert (ASSIGN: forall chunk m b ofs t v m', assign_loc ge chunk m b ofs v t m' -> (length t <= 1)%nat). intros. inv H0; simpl; try omega. inv H3; simpl; try omega. + destruct H. inv H; simpl; try omega. inv H0; eauto; simpl; try omega. eapply external_call_trace_length; eauto. inv H; simpl; try omega. eapply external_call_trace_length; eauto. diff --git a/cfrontend/Cshmgen.v b/cfrontend/Cshmgen.v index 3b23a547..cb83731a 100644 --- a/cfrontend/Cshmgen.v +++ b/cfrontend/Cshmgen.v @@ -21,6 +21,7 @@ *) Require Import Coqlib. +Require Import Maps. Require Import Errors. Require Import Integers. Require Import Floats. @@ -160,8 +161,8 @@ Definition make_cast (from to: type) (e: expr) := | cast_case_s2bool => OK (Ebinop (Ocmpfs Cne) e (make_singleconst Float32.zero)) | cast_case_l2bool => OK (Ebinop (Ocmpl Cne) e (make_longconst Int64.zero)) | cast_case_p2bool => OK (Ebinop (Ocmpu Cne) e (make_intconst Int.zero)) - | cast_case_struct id1 fld1 id2 fld2 => OK e - | cast_case_union id1 fld1 id2 fld2 => OK e + | cast_case_struct id1 id2 => OK e + | cast_case_union id1 id2 => OK e | cast_case_void => OK e | cast_case_default => Error (msg "Cshmgen.make_cast") end. @@ -234,34 +235,34 @@ Definition make_binarith (iop iopu fop sop lop lopu: binary_operation) | bin_default => Error (msg "Cshmgen.make_binarith") end. -Definition make_add (e1: expr) (ty1: type) (e2: expr) (ty2: type) := +Definition make_add (ce: composite_env) (e1: expr) (ty1: type) (e2: expr) (ty2: type) := match classify_add ty1 ty2 with | add_case_pi ty => - let n := make_intconst (Int.repr (Ctypes.sizeof ty)) in + let n := make_intconst (Int.repr (Ctypes.sizeof ce ty)) in OK (Ebinop Oadd e1 (Ebinop Omul n e2)) | add_case_ip ty => - let n := make_intconst (Int.repr (Ctypes.sizeof ty)) in + let n := make_intconst (Int.repr (Ctypes.sizeof ce ty)) in OK (Ebinop Oadd e2 (Ebinop Omul n e1)) | add_case_pl ty => - let n := make_intconst (Int.repr (Ctypes.sizeof ty)) in + let n := make_intconst (Int.repr (Ctypes.sizeof ce ty)) in OK (Ebinop Oadd e1 (Ebinop Omul n (Eunop Ointoflong e2))) | add_case_lp ty => - let n := make_intconst (Int.repr (Ctypes.sizeof ty)) in + let n := make_intconst (Int.repr (Ctypes.sizeof ce ty)) in OK (Ebinop Oadd e2 (Ebinop Omul n (Eunop Ointoflong e1))) | add_default => make_binarith Oadd Oadd Oaddf Oaddfs Oaddl Oaddl e1 ty1 e2 ty2 end. -Definition make_sub (e1: expr) (ty1: type) (e2: expr) (ty2: type) := +Definition make_sub (ce: composite_env) (e1: expr) (ty1: type) (e2: expr) (ty2: type) := match classify_sub ty1 ty2 with | sub_case_pi ty => - let n := make_intconst (Int.repr (Ctypes.sizeof ty)) in + let n := make_intconst (Int.repr (Ctypes.sizeof ce ty)) in OK (Ebinop Osub e1 (Ebinop Omul n e2)) | sub_case_pp ty => - let n := make_intconst (Int.repr (Ctypes.sizeof ty)) in + let n := make_intconst (Int.repr (Ctypes.sizeof ce ty)) in OK (Ebinop Odivu (Ebinop Osub e1 e2) n) | sub_case_pl ty => - let n := make_intconst (Int.repr (Ctypes.sizeof ty)) in + let n := make_intconst (Int.repr (Ctypes.sizeof ce ty)) in OK (Ebinop Osub e1 (Ebinop Omul n (Eunop Ointoflong e2))) | sub_default => make_binarith Osub Osub Osubf Osubfs Osubl Osubl e1 ty1 e2 ty2 @@ -349,8 +350,8 @@ Definition make_load (addr: expr) (ty_res: type) := (** [make_memcpy dst src ty] returns a [memcpy] builtin appropriate for by-copy assignment of a value of Clight type [ty]. *) -Definition make_memcpy (dst src: expr) (ty: type) := - Sbuiltin None (EF_memcpy (Ctypes.sizeof ty) (Ctypes.alignof_blockcopy ty)) +Definition make_memcpy (ce: composite_env) (dst src: expr) (ty: type) := + Sbuiltin None (EF_memcpy (Ctypes.sizeof ce ty) (Ctypes.alignof_blockcopy ce ty)) (dst :: src :: nil). (** [make_store addr ty rhs] stores the value of the @@ -358,10 +359,10 @@ Definition make_memcpy (dst src: expr) (ty: type) := Csharpminor expression [addr]. [ty] is the type of the memory location. *) -Definition make_store (addr: expr) (ty: type) (rhs: expr) := +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 => OK (make_memcpy addr rhs ty) + | By_copy => OK (make_memcpy ce addr rhs ty) | _ => Error (msg "Cshmgen.make_store") end. @@ -375,12 +376,13 @@ Definition transl_unop (op: Cop.unary_operation) (a: expr) (ta: type) : res expr | Cop.Oabsfloat => make_absfloat a ta end. -Definition transl_binop (op: Cop.binary_operation) +Definition transl_binop (ce: composite_env) + (op: Cop.binary_operation) (a: expr) (ta: type) (b: expr) (tb: type) : res expr := match op with - | Cop.Oadd => make_add a ta b tb - | Cop.Osub => make_sub a ta b tb + | Cop.Oadd => make_add ce a ta b tb + | Cop.Osub => make_sub ce a ta b tb | Cop.Omul => make_mul a ta b tb | Cop.Odiv => make_div a ta b tb | Cop.Omod => make_mod a ta b tb @@ -397,13 +399,31 @@ Definition transl_binop (op: Cop.binary_operation) | Cop.Oge => make_cmp Cge a ta b tb end. +(** ** 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 (Ebinop Oadd a (make_intconst (Int.repr ofs))) + end + | Tunion id _ => + OK a + | _ => + Error(msg "Cshmgen.make_field_access") + end. + (** * Translation of expressions *) (** [transl_expr a] returns the Csharpminor code that computes the value of expression [a]. The computation is performed in the error monad (see module [Errors]) to enable error reporting. *) -Fixpoint transl_expr (a: Clight.expr) {struct a} : res expr := +Fixpoint transl_expr (ce: composite_env) (a: Clight.expr) {struct a} : res expr := match a with | Clight.Econst_int n _ => OK(make_intconst n) @@ -418,34 +438,28 @@ Fixpoint transl_expr (a: Clight.expr) {struct a} : res expr := | Clight.Etempvar id ty => OK(Evar id) | Clight.Ederef b ty => - do tb <- transl_expr b; + do tb <- transl_expr ce b; make_load tb ty | Clight.Eaddrof b _ => - transl_lvalue b + transl_lvalue ce b | Clight.Eunop op b _ => - do tb <- transl_expr b; + do tb <- transl_expr ce b; transl_unop op tb (typeof b) | Clight.Ebinop op b c _ => - do tb <- transl_expr b; - do tc <- transl_expr c; - transl_binop op tb (typeof b) tc (typeof c) + do tb <- transl_expr ce b; + do tc <- transl_expr ce c; + transl_binop ce op tb (typeof b) tc (typeof c) | Clight.Ecast b ty => - do tb <- transl_expr b; + do tb <- transl_expr ce b; make_cast (typeof b) ty tb - | Clight.Efield b i ty => - match typeof b with - | Tstruct _ fld _ => - do tb <- transl_expr b; - do ofs <- field_offset i fld; - make_load - (Ebinop Oadd tb (make_intconst (Int.repr ofs))) - ty - | Tunion _ fld _ => - do tb <- transl_expr b; - make_load tb ty - | _ => - Error(msg "Cshmgen.transl_expr(field)") - end + | 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 + | Clight.Esizeof ty' ty => + OK(make_intconst (Int.repr (sizeof ce ty'))) + | Clight.Ealignof ty' ty => + OK(make_intconst (Int.repr (alignof ce ty'))) end (** [transl_lvalue a] returns the Csharpminor code that evaluates @@ -453,23 +467,15 @@ Fixpoint transl_expr (a: Clight.expr) {struct a} : res expr := where the value of [a] is stored. *) -with transl_lvalue (a: Clight.expr) {struct a} : res expr := +with transl_lvalue (ce: composite_env) (a: Clight.expr) {struct a} : res expr := match a with | Clight.Evar id _ => OK (Eaddrof id) | Clight.Ederef b _ => - transl_expr b + transl_expr ce b | Clight.Efield b i ty => - match typeof b with - | Tstruct _ fld _ => - do tb <- transl_expr b; - do ofs <- field_offset i fld; - OK (Ebinop Oadd tb (make_intconst (Int.repr ofs))) - | Tunion _ fld _ => - transl_expr b - | _ => - Error(msg "Cshmgen.transl_lvalue(field)") - end + do tb <- transl_expr ce b; + make_field_access ce (typeof b) i tb | _ => Error(msg "Cshmgen.transl_lvalue") end. @@ -479,20 +485,20 @@ with transl_lvalue (a: Clight.expr) {struct a} : res expr := casted to the corresponding types in [tyl]. Used for function applications. *) -Fixpoint transl_arglist (al: list Clight.expr) (tyl: typelist) +Fixpoint transl_arglist (ce: composite_env) (al: list Clight.expr) (tyl: typelist) {struct al}: res (list expr) := match al, tyl with | nil, Tnil => OK nil | a1 :: a2, Tcons ty1 ty2 => - do ta1 <- transl_expr a1; + do ta1 <- transl_expr ce a1; do ta1' <- make_cast (typeof a1) ty1 ta1; - do ta2 <- transl_arglist a2 ty2; + do ta2 <- transl_arglist ce a2 ty2; OK (ta1' :: ta2) | a1 :: a2, Tnil => (* Tolerance for calls to K&R or variadic functions *) - do ta1 <- transl_expr a1; + do ta1 <- transl_expr ce a1; do ta1' <- make_cast (typeof a1) (default_argument_conversion (typeof a1)) ta1; - do ta2 <- transl_arglist a2 Tnil; + do ta2 <- transl_arglist ce a2 Tnil; OK (ta1' :: ta2) | _, _ => Error(msg "Cshmgen.transl_arglist: arity mismatch") @@ -536,24 +542,24 @@ loop s1 s2 ---> block { // break in s1 and s2 branches here *) -Fixpoint transl_statement (tyret: type) (nbrk ncnt: nat) +Fixpoint transl_statement (ce: composite_env) (tyret: type) (nbrk ncnt: nat) (s: Clight.statement) {struct s} : res stmt := match s with | Clight.Sskip => OK Sskip | Clight.Sassign b c => - do tb <- transl_lvalue b; - do tc <- transl_expr c; + do tb <- transl_lvalue ce b; + do tc <- transl_expr ce c; do tc' <- make_cast (typeof c) (typeof b) tc; - make_store tb (typeof b) tc' + make_store ce tb (typeof b) tc' | Clight.Sset x b => - do tb <- transl_expr b; + do tb <- transl_expr ce b; OK(Sset x tb) | Clight.Scall x b cl => match classify_fun (typeof b) with | fun_case_f args res cconv => - do tb <- transl_expr b; - do tcl <- transl_arglist cl args; + do tb <- transl_expr ce b; + do tcl <- transl_arglist ce cl args; OK(Scall x {| sig_args := typlist_of_arglist cl args; sig_res := opttyp_of_type res; sig_cc := cconv |} @@ -561,80 +567,80 @@ Fixpoint transl_statement (tyret: type) (nbrk ncnt: nat) | _ => Error(msg "Cshmgen.transl_stmt(call)") end | Clight.Sbuiltin x ef tyargs bl => - do tbl <- transl_arglist bl tyargs; + do tbl <- transl_arglist ce bl tyargs; OK(Sbuiltin x ef tbl) | Clight.Ssequence s1 s2 => - do ts1 <- transl_statement tyret nbrk ncnt s1; - do ts2 <- transl_statement tyret nbrk ncnt s2; + do ts1 <- transl_statement ce tyret nbrk ncnt s1; + do ts2 <- transl_statement ce tyret nbrk ncnt s2; OK (Sseq ts1 ts2) | Clight.Sifthenelse e s1 s2 => - do te <- transl_expr e; - do ts1 <- transl_statement tyret nbrk ncnt s1; - do ts2 <- transl_statement tyret nbrk ncnt s2; + do te <- transl_expr ce e; + do ts1 <- transl_statement ce tyret nbrk ncnt s1; + do ts2 <- transl_statement ce tyret nbrk ncnt s2; OK (Sifthenelse (make_boolean te (typeof e)) ts1 ts2) | Clight.Sloop s1 s2 => - do ts1 <- transl_statement tyret 1%nat 0%nat s1; - do ts2 <- transl_statement tyret 0%nat (S ncnt) s2; + do ts1 <- transl_statement ce tyret 1%nat 0%nat s1; + do ts2 <- transl_statement ce tyret 0%nat (S ncnt) s2; OK (Sblock (Sloop (Sseq (Sblock ts1) ts2))) | Clight.Sbreak => OK (Sexit nbrk) | Clight.Scontinue => OK (Sexit ncnt) | Clight.Sreturn (Some e) => - do te <- transl_expr e; + do te <- transl_expr ce e; do te' <- make_cast (typeof e) tyret te; OK (Sreturn (Some te')) | Clight.Sreturn None => OK (Sreturn None) | Clight.Sswitch a sl => - do ta <- transl_expr a; - do tsl <- transl_lbl_stmt tyret 0%nat (S ncnt) sl; + do ta <- transl_expr ce a; + do tsl <- transl_lbl_stmt ce tyret 0%nat (S ncnt) sl; match classify_switch (typeof a) with | switch_case_i => OK (Sblock (Sswitch false ta tsl)) | switch_case_l => OK (Sblock (Sswitch true ta tsl)) | switch_default => Error(msg "Cshmgen.transl_stmt(switch)") end | Clight.Slabel lbl s => - do ts <- transl_statement tyret nbrk ncnt s; + do ts <- transl_statement ce tyret nbrk ncnt s; OK (Slabel lbl ts) | Clight.Sgoto lbl => OK (Sgoto lbl) end -with transl_lbl_stmt (tyret: type) (nbrk ncnt: nat) +with transl_lbl_stmt (ce: composite_env) (tyret: type) (nbrk ncnt: nat) (sl: Clight.labeled_statements) {struct sl}: res lbl_stmt := match sl with | Clight.LSnil => OK LSnil | Clight.LScons n s sl' => - do ts <- transl_statement tyret nbrk ncnt s; - do tsl' <- transl_lbl_stmt tyret nbrk ncnt sl'; + do ts <- transl_statement ce tyret nbrk ncnt s; + do tsl' <- transl_lbl_stmt ce tyret nbrk ncnt sl'; OK (LScons n ts tsl') end. (*** Translation of functions *) -Definition transl_var (v: ident * type) := (fst v, sizeof (snd v)). +Definition transl_var (ce: composite_env) (v: ident * type) := (fst v, sizeof ce (snd v)). Definition signature_of_function (f: Clight.function) := {| sig_args := map typ_of_type (map snd (Clight.fn_params f)); sig_res := opttyp_of_type (Clight.fn_return f); sig_cc := Clight.fn_callconv f |}. -Definition transl_function (f: Clight.function) : res function := - do tbody <- transl_statement f.(Clight.fn_return) 1%nat 0%nat (Clight.fn_body f); +Definition transl_function (ce: composite_env) (f: Clight.function) : res function := + do tbody <- transl_statement ce f.(Clight.fn_return) 1%nat 0%nat (Clight.fn_body f); OK (mkfunction (signature_of_function f) (map fst (Clight.fn_params f)) - (map transl_var (Clight.fn_vars f)) + (map (transl_var ce) (Clight.fn_vars f)) (map fst (Clight.fn_temps f)) tbody). -Definition transl_fundef (f: Clight.fundef) : res fundef := +Definition transl_fundef (ce: composite_env) (f: Clight.fundef) : res fundef := match f with | Clight.Internal g => - do tg <- transl_function g; OK(AST.Internal tg) + do tg <- transl_function ce g; OK(AST.Internal tg) | Clight.External ef args res cconv => if signature_eq (ef_sig ef) (signature_of_type args res cconv) then OK(AST.External ef) @@ -646,4 +652,5 @@ Definition transl_fundef (f: Clight.fundef) : res fundef := Definition transl_globvar (ty: type) := OK tt. Definition transl_program (p: Clight.program) : res program := - transform_partial_program2 transl_fundef transl_globvar p. + transform_partial_program2 (transl_fundef p.(prog_comp_env)) transl_globvar p. + diff --git a/cfrontend/Cshmgenproof.v b/cfrontend/Cshmgenproof.v index dc47df04..7f61c657 100644 --- a/cfrontend/Cshmgenproof.v +++ b/cfrontend/Cshmgenproof.v @@ -40,8 +40,8 @@ Proof. Qed. Lemma transl_fundef_sig1: - forall f tf args res cc, - transl_fundef f = OK tf -> + forall ce f tf args res cc, + transl_fundef ce f = OK tf -> classify_fun (type_of_fundef f) = fun_case_f args res cc -> funsig tf = signature_of_type args res cc. Proof. @@ -54,8 +54,8 @@ Proof. Qed. Lemma transl_fundef_sig2: - forall f tf args res cc, - transl_fundef f = OK tf -> + forall ce f tf args res cc, + transl_fundef ce f = OK tf -> type_of_fundef f = Tfunction args res cc -> funsig tf = signature_of_type args res cc. Proof. @@ -70,8 +70,8 @@ Qed. Lemma transl_expr_lvalue: forall ge e le m a loc ofs ta, Clight.eval_lvalue ge e le m a loc ofs -> - transl_expr a = OK ta -> - (exists tb, transl_lvalue a = OK tb /\ make_load tb (typeof a) = OK ta). + transl_expr ge a = OK ta -> + (exists tb, transl_lvalue ge a = OK tb /\ make_load tb (typeof a) = OK ta). Proof. intros until ta; intros EVAL TR. inv EVAL; simpl in TR. (* var local *) @@ -81,39 +81,36 @@ Proof. (* deref *) monadInv TR. exists x; auto. (* field struct *) - rewrite H0 in TR. monadInv TR. - econstructor; split. simpl. rewrite H0. - rewrite EQ; rewrite EQ1; simpl; eauto. auto. + monadInv TR. exists x0; split; auto. simpl; rewrite EQ; auto. (* field union *) - rewrite H0 in TR. monadInv TR. - econstructor; split. simpl. rewrite H0. rewrite EQ; simpl; eauto. auto. + monadInv TR. exists x0; split; auto. simpl; rewrite EQ; auto. Qed. (** Properties of labeled statements *) Lemma transl_lbl_stmt_1: - forall tyret nbrk ncnt n sl tsl, - transl_lbl_stmt tyret nbrk ncnt sl = OK tsl -> - transl_lbl_stmt tyret nbrk ncnt (Clight.select_switch n sl) = OK (select_switch n tsl). + forall ce tyret nbrk ncnt n sl tsl, + transl_lbl_stmt ce tyret nbrk ncnt sl = OK tsl -> + transl_lbl_stmt ce tyret nbrk ncnt (Clight.select_switch n sl) = OK (select_switch n tsl). Proof. intros until n. assert (DFL: forall sl tsl, - transl_lbl_stmt tyret nbrk ncnt sl = OK tsl -> - transl_lbl_stmt tyret nbrk ncnt (Clight.select_switch_default sl) = OK (select_switch_default tsl)). + transl_lbl_stmt ce tyret nbrk ncnt sl = OK tsl -> + transl_lbl_stmt ce tyret nbrk ncnt (Clight.select_switch_default sl) = OK (select_switch_default tsl)). { induction sl; simpl; intros. inv H; auto. monadInv H. simpl. destruct o; eauto. simpl; rewrite EQ; simpl; rewrite EQ1; auto. } assert (CASE: forall sl tsl, - transl_lbl_stmt tyret nbrk ncnt sl = OK tsl -> + transl_lbl_stmt ce tyret nbrk ncnt sl = OK tsl -> match Clight.select_switch_case n sl with | None => select_switch_case n tsl = None | Some sl' => exists tsl', select_switch_case n tsl = Some tsl' - /\ transl_lbl_stmt tyret nbrk ncnt sl' = OK tsl' + /\ transl_lbl_stmt ce tyret nbrk ncnt sl' = OK tsl' end). { induction sl; simpl; intros. @@ -130,9 +127,9 @@ Proof. Qed. Lemma transl_lbl_stmt_2: - forall tyret nbrk ncnt sl tsl, - transl_lbl_stmt tyret nbrk ncnt sl = OK tsl -> - transl_statement tyret nbrk ncnt (seq_of_labeled_statement sl) = OK (seq_of_lbl_stmt tsl). + forall ce tyret nbrk ncnt sl tsl, + transl_lbl_stmt ce tyret nbrk ncnt sl = OK tsl -> + transl_statement ce tyret nbrk ncnt (seq_of_labeled_statement sl) = OK (seq_of_lbl_stmt tsl). Proof. induction sl; intros. monadInv H. auto. @@ -143,6 +140,7 @@ Qed. Section CONSTRUCTORS. +Variable ce: composite_env. Variable ge: genv. Lemma make_intconst_correct: @@ -305,9 +303,9 @@ Proof. simpl. unfold Val.cmpu, Val.cmpu_bool, Int.cmpu. destruct (Int.eq i Int.zero); auto. (* struct *) - destruct (ident_eq id1 id2 && fieldlist_eq fld1 fld2); inv H2; auto. + destruct (ident_eq id1 id2); inv H2; auto. (* union *) - destruct (ident_eq id1 id2 && fieldlist_eq fld1 fld2); inv H2; auto. + destruct (ident_eq id1 id2); inv H2; auto. Qed. Lemma make_boolean_correct: @@ -467,7 +465,7 @@ End MAKE_BIN. Hint Extern 2 (@eq (option val) _ _) => (simpl; reflexivity) : cshm. -Lemma make_add_correct: binary_constructor_correct make_add sem_add. +Lemma make_add_correct: binary_constructor_correct (make_add ce) (sem_add ce). Proof. red; unfold make_add, sem_add; intros until m; intros SEM MAKE EV1 EV2; @@ -479,14 +477,14 @@ Proof. - eapply make_binarith_correct; eauto; intros; auto. Qed. -Lemma make_sub_correct: binary_constructor_correct make_sub sem_sub. +Lemma make_sub_correct: binary_constructor_correct (make_sub ce) (sem_sub ce). Proof. red; unfold make_sub, sem_sub; intros until m; intros SEM MAKE EV1 EV2; destruct (classify_sub tya tyb); inv MAKE. - destruct va; try discriminate; destruct vb; inv SEM; eauto with cshm. - destruct va; try discriminate; destruct vb; inv SEM. - destruct (eq_block b0 b1); try discriminate. destruct (Int.eq (Int.repr (sizeof ty)) Int.zero) eqn:E; inv H0. + destruct (eq_block b0 b1); try discriminate. destruct (Int.eq (Int.repr (sizeof ce ty)) Int.zero) eqn:E; inv H0. econstructor; eauto with cshm. rewrite dec_eq_true. simpl. rewrite E; auto. - destruct va; try discriminate; destruct vb; inv SEM; eauto with cshm. - eapply make_binarith_correct; eauto; intros; auto. @@ -648,8 +646,8 @@ Qed. Lemma transl_binop_correct: forall op a tya b tyb c va vb v e le m, - transl_binop op a tya b tyb = OK c -> - sem_binary_operation op va tya vb tyb m = Some v -> + transl_binop ce op a tya b tyb = OK c -> + sem_binary_operation ce op va tya vb tyb m = Some v -> eval_expr ge e le m a va -> eval_expr ge e le m b vb -> eval_expr ge e le m c v. @@ -691,12 +689,12 @@ Proof. Qed. Lemma make_memcpy_correct: - forall f dst src ty k e le m b ofs v m', + forall ce f dst src ty k e le m b ofs v m', eval_expr ge e le m dst (Vptr b ofs) -> eval_expr ge e le m src v -> - assign_loc ty m b ofs v m' -> + assign_loc ce ty m b ofs v m' -> access_mode ty = By_copy -> - step ge (State f (make_memcpy dst src ty) k e le m) E0 (State f Sskip k e le m'). + step ge (State f (make_memcpy ce dst src ty) k e le m) E0 (State f Sskip k e le m'). Proof. intros. inv H1; try congruence. unfold make_memcpy. change le with (set_optvar None Vundef le) at 2. @@ -710,10 +708,10 @@ Qed. Lemma make_store_correct: forall addr ty rhs code e le m b ofs v m' f k, - make_store addr ty rhs = OK code -> + make_store ce addr ty rhs = OK code -> eval_expr ge e le m addr (Vptr b ofs) -> eval_expr ge e le m rhs v -> - assign_loc ty m b ofs v m' -> + assign_loc ce ty m b ofs v m' -> step ge (State f code k e le m) E0 (State f Sskip k e le m'). Proof. unfold make_store. intros until k; intros MKSTORE EV1 EV2 ASSIGN. @@ -736,32 +734,32 @@ Variable prog: Clight.program. Variable tprog: Csharpminor.program. Hypothesis TRANSL: transl_program prog = OK tprog. -Let ge := Genv.globalenv prog. +Let ge := globalenv prog. Let tge := Genv.globalenv tprog. Lemma symbols_preserved: forall s, Genv.find_symbol tge s = Genv.find_symbol ge s. -Proof (Genv.find_symbol_transf_partial2 transl_fundef transl_globvar _ TRANSL). +Proof (Genv.find_symbol_transf_partial2 (transl_fundef ge) transl_globvar _ TRANSL). Lemma public_preserved: forall s, Genv.public_symbol tge s = Genv.public_symbol ge s. -Proof (Genv.public_symbol_transf_partial2 transl_fundef transl_globvar _ TRANSL). +Proof (Genv.public_symbol_transf_partial2 (transl_fundef ge) transl_globvar _ TRANSL). Lemma functions_translated: forall v f, Genv.find_funct ge v = Some f -> - exists tf, Genv.find_funct tge v = Some tf /\ transl_fundef f = OK tf. -Proof (Genv.find_funct_transf_partial2 transl_fundef transl_globvar _ TRANSL). + exists tf, Genv.find_funct tge v = Some tf /\ transl_fundef ge f = OK tf. +Proof (Genv.find_funct_transf_partial2 (transl_fundef ge) transl_globvar _ 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 /\ transl_fundef f = OK tf. -Proof (Genv.find_funct_ptr_transf_partial2 transl_fundef transl_globvar _ TRANSL). + exists tf, Genv.find_funct_ptr tge b = Some tf /\ transl_fundef ge f = OK tf. +Proof (Genv.find_funct_ptr_transf_partial2 (transl_fundef ge) transl_globvar _ TRANSL). Lemma block_is_volatile_preserved: forall b, Genv.block_is_volatile tge b = Genv.block_is_volatile ge b. -Proof (Genv.block_is_volatile_transf_partial2 transl_fundef transl_globvar _ TRANSL). +Proof (Genv.block_is_volatile_transf_partial2 (transl_fundef ge) transl_globvar _ TRANSL). (** * Matching between environments *) @@ -772,7 +770,7 @@ Record match_env (e: Clight.env) (te: Csharpminor.env) : Prop := mk_match_env { me_local: forall id b ty, - e!id = Some (b, ty) -> te!id = Some(b, sizeof ty); + e!id = Some (b, ty) -> te!id = Some(b, sizeof ge ty); me_local_inv: forall id b sz, te!id = Some (b, sz) -> exists ty, e!id = Some(b, ty) @@ -791,18 +789,18 @@ Qed. Lemma match_env_same_blocks: forall e te, match_env e te -> - blocks_of_env te = Clight.blocks_of_env e. + blocks_of_env te = Clight.blocks_of_env ge e. Proof. intros. set (R := fun (x: (block * type)) (y: (block * Z)) => match x, y with - | (b1, ty), (b2, sz) => b2 = b1 /\ sz = sizeof ty + | (b1, ty), (b2, sz) => b2 = b1 /\ sz = sizeof ge ty end). assert (list_forall2 (fun i_x i_y => fst i_x = fst i_y /\ R (snd i_x) (snd i_y)) (PTree.elements e) (PTree.elements te)). apply PTree.elements_canonical_order. - intros id [b ty] GET. exists (b, sizeof ty); split. eapply me_local; eauto. red; auto. + intros id [b ty] GET. exists (b, sizeof ge ty); split. eapply me_local; eauto. red; auto. intros id [b sz] GET. exploit me_local_inv; eauto. intros [ty EQ]. exploit me_local; eauto. intros EQ1. exists (b, ty); split. auto. red; split; congruence. @@ -818,7 +816,7 @@ Qed. Lemma match_env_free_blocks: forall e te m m', match_env e te -> - Mem.free_list m (Clight.blocks_of_env e) = Some m' -> + Mem.free_list m (Clight.blocks_of_env ge e) = Some m' -> Mem.free_list m (blocks_of_env te) = Some m'. Proof. intros. rewrite (match_env_same_blocks _ _ H). auto. @@ -839,16 +837,16 @@ Qed. Lemma match_env_alloc_variables: forall e1 m1 vars e2 m2, - Clight.alloc_variables e1 m1 vars e2 m2 -> + Clight.alloc_variables ge e1 m1 vars e2 m2 -> forall te1, match_env e1 te1 -> exists te2, - Csharpminor.alloc_variables te1 m1 (map transl_var vars) te2 m2 + Csharpminor.alloc_variables te1 m1 (map (transl_var ge) vars) te2 m2 /\ match_env e2 te2. Proof. induction 1; intros; simpl. exists te1; split. constructor. auto. - exploit (IHalloc_variables (PTree.set id (b1, sizeof ty) te1)). + exploit (IHalloc_variables (PTree.set id (b1, sizeof ge ty) te1)). constructor. (* me_local *) intros until ty0. repeat rewrite PTree.gsspec. @@ -913,11 +911,11 @@ Hypothesis MENV: match_env e te. Lemma transl_expr_lvalue_correct: (forall a v, Clight.eval_expr ge e le m a v -> - forall ta (TR: transl_expr a = OK ta) , + forall ta (TR: transl_expr ge 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 a = OK ta), + forall ta (TR: transl_lvalue ge a = OK ta), Csharpminor.eval_expr tge te le m ta (Vptr b ofs)). Proof. apply eval_expr_lvalue_ind; intros; try (monadInv TR). @@ -939,6 +937,10 @@ Proof. eapply transl_binop_correct; eauto. (* cast *) eapply make_cast_correct; eauto. +(* sizeof *) + apply make_intconst_correct. +(* alignof *) + apply make_intconst_correct. (* rvalue out of lvalue *) exploit transl_expr_lvalue; eauto. intros [tb [TRLVAL MKLOAD]]. eapply make_load_correct; eauto. @@ -952,32 +954,34 @@ Proof. (* deref *) simpl in TR. eauto. (* field struct *) - simpl in TR. rewrite H1 in TR. monadInv TR. + change (prog_comp_env prog) with (genv_cenv ge) in EQ0. + unfold make_field_access in EQ0; rewrite H1, H2 in EQ0; monadInv EQ0. eapply eval_Ebinop; eauto. apply make_intconst_correct. simpl. congruence. (* field union *) - simpl in TR. rewrite H1 in TR. eauto. + unfold make_field_access in EQ0; rewrite H1 in EQ0; monadInv EQ0. + auto. Qed. Lemma transl_expr_correct: forall a v, Clight.eval_expr ge e le m a v -> - forall ta, transl_expr a = OK ta -> + forall ta, transl_expr ge a = OK ta -> Csharpminor.eval_expr tge te le m ta v. 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 a = OK ta -> + forall ta, transl_lvalue ge a = OK ta -> Csharpminor.eval_expr tge te le m ta (Vptr b ofs). Proof (proj2 transl_expr_lvalue_correct). Lemma transl_arglist_correct: forall al tyl vl, Clight.eval_exprlist ge e le m al tyl vl -> - forall tal, transl_arglist al tyl = OK tal -> + forall tal, transl_arglist ge al tyl = OK tal -> Csharpminor.eval_exprlist tge te le m tal vl. Proof. induction 1; intros. @@ -1035,21 +1039,21 @@ Inductive match_cont: type -> nat -> nat -> Clight.cont -> Csharpminor.cont -> P | match_Kstop: forall tyret nbrk ncnt, match_cont tyret nbrk ncnt Clight.Kstop Kstop | match_Kseq: forall tyret nbrk ncnt s k ts tk, - transl_statement tyret nbrk ncnt s = OK ts -> + transl_statement ge tyret nbrk ncnt s = OK ts -> match_cont tyret nbrk ncnt k tk -> match_cont tyret nbrk ncnt (Clight.Kseq s k) (Kseq ts tk) | match_Kloop1: forall tyret s1 s2 k ts1 ts2 nbrk ncnt tk, - transl_statement tyret 1%nat 0%nat s1 = OK ts1 -> - transl_statement tyret 0%nat (S ncnt) s2 = OK ts2 -> + transl_statement ge tyret 1%nat 0%nat s1 = OK ts1 -> + transl_statement ge tyret 0%nat (S ncnt) s2 = OK ts2 -> match_cont tyret nbrk ncnt k tk -> match_cont tyret 1%nat 0%nat (Clight.Kloop1 s1 s2 k) (Kblock (Kseq ts2 (Kseq (Sloop (Sseq (Sblock ts1) ts2)) (Kblock tk)))) | match_Kloop2: forall tyret s1 s2 k ts1 ts2 nbrk ncnt tk, - transl_statement tyret 1%nat 0%nat s1 = OK ts1 -> - transl_statement tyret 0%nat (S ncnt) s2 = OK ts2 -> + transl_statement ge tyret 1%nat 0%nat s1 = OK ts1 -> + transl_statement ge tyret 0%nat (S ncnt) s2 = OK ts2 -> match_cont tyret nbrk ncnt k tk -> match_cont tyret 0%nat (S ncnt) (Clight.Kloop2 s1 s2 k) @@ -1060,7 +1064,7 @@ Inductive match_cont: type -> nat -> nat -> Clight.cont -> Csharpminor.cont -> P (Clight.Kswitch k) (Kblock tk) | match_Kcall_some: forall tyret nbrk ncnt nbrk' ncnt' f e k id tf te le tk, - transl_function f = OK tf -> + transl_function ge f = OK tf -> match_env e te -> match_cont (Clight.fn_return f) nbrk' ncnt' k tk -> match_cont tyret nbrk ncnt @@ -1070,8 +1074,8 @@ Inductive match_cont: type -> nat -> nat -> Clight.cont -> Csharpminor.cont -> P Inductive match_states: Clight.state -> Csharpminor.state -> Prop := | match_state: forall f nbrk ncnt s k e le m tf ts tk te ts' tk' - (TRF: transl_function f = OK tf) - (TR: transl_statement (Clight.fn_return f) nbrk ncnt s = OK ts) + (TRF: transl_function ge f = OK tf) + (TR: transl_statement ge (Clight.fn_return f) nbrk ncnt s = OK ts) (MTR: match_transl ts tk ts' tk') (MENV: match_env e te) (MK: match_cont (Clight.fn_return f) nbrk ncnt k tk), @@ -1079,7 +1083,7 @@ Inductive match_states: Clight.state -> Csharpminor.state -> Prop := (State tf ts' tk' te le m) | match_callstate: forall fd args k m tfd tk targs tres cconv - (TR: transl_fundef fd = OK tfd) + (TR: transl_fundef ge fd = OK tfd) (MK: match_cont Tvoid 0%nat 0%nat k tk) (ISCC: Clight.is_call_cont k) (TY: type_of_fundef fd = Tfunction targs tres cconv), @@ -1093,7 +1097,7 @@ Inductive match_states: Clight.state -> Csharpminor.state -> Prop := Remark match_states_skip: forall f e le te nbrk ncnt k tf tk m, - transl_function f = OK tf -> + transl_function ge f = OK tf -> match_env e te -> match_cont (Clight.fn_return f) nbrk ncnt k tk -> match_states (Clight.State f Clight.Sskip k e le m) (State tf Sskip tk te le m). @@ -1109,27 +1113,27 @@ Variable tyret: type. Lemma transl_find_label: forall s nbrk ncnt k ts tk - (TR: transl_statement tyret nbrk ncnt s = OK ts) + (TR: transl_statement ge tyret nbrk ncnt s = OK ts) (MC: match_cont tyret nbrk ncnt k tk), match Clight.find_label lbl s k with | None => find_label lbl ts tk = None | Some (s', k') => exists ts', exists tk', exists nbrk', exists ncnt', find_label lbl ts tk = Some (ts', tk') - /\ transl_statement tyret nbrk' ncnt' s' = OK ts' + /\ transl_statement ge tyret nbrk' ncnt' s' = OK ts' /\ match_cont tyret nbrk' ncnt' k' tk' end with transl_find_label_ls: forall ls nbrk ncnt k tls tk - (TR: transl_lbl_stmt tyret nbrk ncnt ls = OK tls) + (TR: transl_lbl_stmt ge tyret nbrk ncnt ls = OK tls) (MC: match_cont tyret nbrk ncnt k tk), match Clight.find_label_ls lbl ls k with | None => find_label_ls lbl tls tk = None | Some (s', k') => exists ts', exists tk', exists nbrk', exists ncnt', find_label_ls lbl tls tk = Some (ts', tk') - /\ transl_statement tyret nbrk' ncnt' s' = OK ts' + /\ transl_statement ge tyret nbrk' ncnt' s' = OK ts' /\ match_cont tyret nbrk' ncnt' k' tk' end. @@ -1340,7 +1344,8 @@ Proof. monadInv TR. inv MTR. inv MK. econstructor; split. apply plus_one. constructor. - econstructor; eauto. + econstructor; eauto. +Local Opaque ge. simpl. rewrite H5; simpl. rewrite H7; simpl. eauto. constructor. @@ -1464,9 +1469,10 @@ Lemma transl_initial_states: Proof. intros. inv H. exploit function_ptr_translated; eauto. intros [tf [A B]]. - assert (C: Genv.find_symbol tge (prog_main tprog) = Some b). - rewrite symbols_preserved. replace (prog_main tprog) with (prog_main prog). - auto. symmetry. unfold transl_program in TRANSL. + assert (C: Genv.find_symbol tge (AST.prog_main tprog) = Some b). + rewrite symbols_preserved. replace (AST.prog_main tprog) with (prog_main prog). + auto. symmetry. unfold transl_program in TRANSL. + change (prog_main prog) with (AST.prog_main (program_of_program prog)). eapply transform_partial_program2_main; eauto. assert (funsig tf = signature_of_type Tnil type_int32s cc_default). eapply transl_fundef_sig2; eauto. diff --git a/cfrontend/Cstrategy.v b/cfrontend/Cstrategy.v index f8c3963e..3b0eb84f 100644 --- a/cfrontend/Cstrategy.v +++ b/cfrontend/Cstrategy.v @@ -97,13 +97,16 @@ Inductive eval_simple_lvalue: expr -> block -> int -> Prop := | esl_deref: forall r ty b ofs, eval_simple_rvalue r (Vptr b ofs) -> eval_simple_lvalue (Ederef r ty) b ofs - | esl_field_struct: forall r f ty b ofs id fList a delta, + | esl_field_struct: forall r f ty b ofs id co a delta, eval_simple_rvalue r (Vptr b ofs) -> - typeof r = Tstruct id fList a -> field_offset f fList = OK delta -> + typeof r = Tstruct id a -> + ge.(genv_cenv)!id = Some co -> + field_offset ge f (co_members co) = OK delta -> eval_simple_lvalue (Efield r f ty) b (Int.add ofs (Int.repr delta)) - | esl_field_union: forall r f ty b ofs id fList a, + | esl_field_union: forall r f ty b ofs id co a, eval_simple_rvalue r (Vptr b ofs) -> - typeof r = Tunion id fList a -> + typeof r = Tunion id a -> + ge.(genv_cenv)!id = Some co -> eval_simple_lvalue (Efield r f ty) b ofs with eval_simple_rvalue: expr -> val -> Prop := @@ -123,16 +126,16 @@ with eval_simple_rvalue: expr -> val -> Prop := eval_simple_rvalue (Eunop op r1 ty) v | esr_binop: forall op r1 r2 ty v1 v2 v, eval_simple_rvalue r1 v1 -> eval_simple_rvalue r2 v2 -> - sem_binary_operation op v1 (typeof r1) v2 (typeof r2) m = Some v -> + sem_binary_operation ge op v1 (typeof r1) v2 (typeof r2) m = Some v -> eval_simple_rvalue (Ebinop op r1 r2 ty) v | esr_cast: forall ty r1 v1 v, eval_simple_rvalue r1 v1 -> sem_cast v1 (typeof r1) ty = Some v -> eval_simple_rvalue (Ecast r1 ty) v | esr_sizeof: forall ty1 ty, - eval_simple_rvalue (Esizeof ty1 ty) (Vint (Int.repr (sizeof ty1))) + eval_simple_rvalue (Esizeof ty1 ty) (Vint (Int.repr (sizeof ge ty1))) | esr_alignof: forall ty1 ty, - eval_simple_rvalue (Ealignof ty1 ty) (Vint (Int.repr (alignof ty1))). + eval_simple_rvalue (Ealignof ty1 ty) (Vint (Int.repr (alignof ge ty1))). Inductive eval_simple_list: exprlist -> typelist -> list val -> Prop := | esrl_nil: @@ -291,7 +294,7 @@ Inductive estep: state -> trace -> state -> Prop := eval_simple_lvalue e m l b ofs -> deref_loc ge (typeof l) m b ofs t1 v1 -> eval_simple_rvalue e m r v2 -> - sem_binary_operation op v1 (typeof l) v2 (typeof r) m = Some v3 -> + sem_binary_operation ge op v1 (typeof l) v2 (typeof r) m = Some v3 -> sem_cast v3 tyres (typeof l) = Some v4 -> assign_loc ge (typeof l) m b ofs v4 t2 m' -> ty = typeof l -> @@ -304,7 +307,7 @@ Inductive estep: state -> trace -> state -> Prop := eval_simple_lvalue e m l b ofs -> deref_loc ge (typeof l) m b ofs t v1 -> eval_simple_rvalue e m r v2 -> - match sem_binary_operation op v1 (typeof l) v2 (typeof r) m with + 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) with @@ -320,7 +323,7 @@ Inductive estep: state -> trace -> state -> Prop := leftcontext RV RV C -> eval_simple_lvalue e m l b ofs -> deref_loc ge ty m b ofs t1 v1 -> - sem_incrdecr id v1 ty = Some v2 -> + sem_incrdecr ge id v1 ty = Some v2 -> sem_cast v2 (incrdecr_type ty) ty = Some v3 -> assign_loc ge ty m b ofs v3 t2 m' -> ty = typeof l -> @@ -332,7 +335,7 @@ Inductive estep: state -> trace -> state -> Prop := leftcontext RV RV C -> eval_simple_lvalue e m l b ofs -> deref_loc ge ty m b ofs t v1 -> - match sem_incrdecr id v1 ty with + match sem_incrdecr ge id v1 ty with | None => True | Some v2 => match sem_cast v2 (incrdecr_type ty) ty with @@ -525,8 +528,8 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop := | Efield (Eval v ty1) f ty => exists b, exists ofs, v = Vptr b ofs /\ match ty1 with - | Tstruct _ fList _ => exists delta, field_offset f fList = Errors.OK delta - | Tunion _ _ _ => True + | Tstruct 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 | _ => False end | Eval v ty => False @@ -535,7 +538,7 @@ Definition invert_expr_prop (a: expr) (m: mem) : Prop := | Eunop op (Eval v1 ty1) ty => exists v, sem_unary_operation op v1 ty1 = Some v | Ebinop op (Eval v1 ty1) (Eval v2 ty2) ty => - exists v, sem_binary_operation op v1 ty1 v2 ty2 m = Some v + exists v, sem_binary_operation ge op v1 ty1 v2 ty2 m = Some v | Ecast (Eval v1 ty1) ty => exists v, sem_cast v1 ty1 ty = Some v | Eseqand (Eval v1 ty1) r2 ty => @@ -581,8 +584,8 @@ Proof. exists b; auto. exists b; auto. exists b; exists ofs; auto. - exists b; exists ofs; split; auto. exists delta; auto. - exists b; exists ofs; auto. + exists b; exists ofs; split; auto. exists co, delta; auto. + exists b; exists ofs; split; auto. exists co; auto. Qed. Lemma rred_invert: @@ -705,11 +708,11 @@ Lemma eval_simple_steps: (forall a v, eval_simple_rvalue e m a v -> 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)) + E0 (ExprState f (C (Eval v (typeof a))) k e m)) /\ (forall a b ofs, eval_simple_lvalue e m a b ofs -> forall C, context LV RV C -> 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 (typeof a))) k e m)). Proof. Ltac Steps REC C' := eapply star_trans; [apply (REC C'); eauto | idtac | simpl; reflexivity]. @@ -816,8 +819,8 @@ Ltac StepR REC C' a := 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 [delta OFS]. exists b; exists (Int.add ofs (Int.repr delta)); econstructor; eauto. - exists b; exists ofs; econstructor; eauto. + destruct TY as (co & delta & CE & OFS). exists b; exists (Int.add ofs (Int.repr delta)); econstructor; eauto. + destruct TY as (co & CE). exists b; exists ofs; econstructor; 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. @@ -1197,7 +1200,7 @@ Proof. eapply eval_simple_rvalue_steps with (C := fun x => C(Eassignop op (Eloc b ofs (typeof l)) x tyres (typeof l))); eauto. eapply plus_left. left; apply step_rred; auto. econstructor; eauto. - destruct (sem_binary_operation op v1 (typeof l) v2 (typeof r) m) as [v3|] eqn:?. + destruct (sem_binary_operation ge op v1 (typeof l) v2 (typeof r) m) as [v3|] eqn:?. eapply star_left. left; apply step_rred with (C := fun x => C(Eassign (Eloc b ofs (typeof l)) x (typeof l))); eauto. econstructor; eauto. apply star_one. @@ -1230,10 +1233,10 @@ Proof. eapply plus_left. left; apply step_rred; auto. econstructor; eauto. set (op := match id with Incr => Oadd | Decr => Osub end). - assert (SEM: sem_binary_operation op v1 (typeof l) (Vint Int.one) type_int32s m = - sem_incrdecr id v1 (typeof l)). + assert (SEM: sem_binary_operation ge op v1 (typeof l) (Vint Int.one) type_int32s m = + sem_incrdecr ge id v1 (typeof l)). destruct id; auto. - destruct (sem_incrdecr id v1 (typeof l)) as [v2|]. + destruct (sem_incrdecr ge id v1 (typeof l)) as [v2|]. eapply star_left. left; apply step_rred with (C := fun x => C (Ecomma (Eassign (Eloc b ofs (typeof l)) x (typeof l)) (Eval v1 (typeof l)) (typeof l))); eauto. econstructor; eauto. @@ -1325,7 +1328,7 @@ Proof. exploit (simple_can_eval_rval f k e m b2 (fun x => C(Eassignop op (Eloc b ofs (typeof b1)) x tyres ty))); eauto. intros [v [E2 S2]]. exploit safe_inv. eexact S2. eauto. simpl. intros [t1 [v1 [A B]]]. - destruct (sem_binary_operation op v1 (typeof b1) v (typeof b2) m) as [v3|] eqn:?. + destruct (sem_binary_operation ge op v1 (typeof b1) v (typeof b2) m) as [v3|] eqn:?. destruct (sem_cast v3 tyres (typeof b1)) as [v4|] eqn:?. destruct (classic (exists t2, exists m', assign_loc ge (typeof b1) m b ofs v4 t2 m')). destruct H2 as [t2 [m' D]]. @@ -1340,7 +1343,7 @@ Proof. exploit (simple_can_eval_lval f k e m b (fun x => C(Epostincr id x ty))); eauto. intros [b1 [ofs [E1 S1]]]. exploit safe_inv. eexact S1. eauto. simpl. intros [t [v1 [A B]]]. - destruct (sem_incrdecr id v1 ty) as [v2|] eqn:?. + destruct (sem_incrdecr ge id v1 ty) as [v2|] eqn:?. destruct (sem_cast v2 (incrdecr_type ty) ty) as [v3|] eqn:?. destruct (classic (exists t2, exists m', assign_loc ge ty m b1 ofs v3 t2 m')). destruct H0 as [t2 [m' D]]. @@ -1431,12 +1434,13 @@ End STRATEGY. (** The semantics that follows the strategy. *) Definition semantics (p: program) := - Semantics step (initial_state p) final_state (Genv.globalenv p). + let ge := globalenv p in + Semantics_gen step (initial_state p) final_state ge ge. (** This semantics is receptive to changes in events. *) Remark deref_loc_trace: - forall F V (ge: Genv.t F V) ty m b ofs t v, + forall ge ty m b ofs t v, deref_loc ge ty m b ofs t v -> match t with nil => True | ev :: nil => True | _ => False end. Proof. @@ -1444,7 +1448,7 @@ Proof. Qed. Remark deref_loc_receptive: - forall F V (ge: Genv.t F V) ty m b ofs ev1 t1 v ev2, + forall ge ty m b ofs ev1 t1 v ev2, deref_loc ge ty m b ofs (ev1 :: t1) v -> match_traces ge (ev1 :: nil) (ev2 :: nil) -> t1 = nil /\ exists v', deref_loc ge ty m b ofs (ev2 :: nil) v'. @@ -1456,7 +1460,7 @@ Proof. Qed. Remark assign_loc_trace: - forall F V (ge: Genv.t F V) ty m b ofs t v m', + forall ge ty m b ofs t v m', assign_loc ge ty m b ofs v t m' -> match t with nil => True | ev :: nil => output_event ev | _ => False end. Proof. @@ -1464,7 +1468,7 @@ Proof. Qed. Remark assign_loc_receptive: - forall F V (ge: Genv.t F V) ty m b ofs ev1 t1 v m' ev2, + forall ge ty m b ofs ev1 t1 v m' ev2, assign_loc ge ty m b ofs v (ev1 :: t1) m' -> match_traces ge (ev1 :: nil) (ev2 :: nil) -> ev1 :: t1 = ev2 :: nil. @@ -1479,6 +1483,7 @@ Lemma semantics_strongly_receptive: Proof. intros. constructor; simpl; intros. (* receptiveness *) + set (ge := globalenv p) in *. inversion H; subst. inv H1. (* valof volatile *) @@ -1492,9 +1497,9 @@ Proof. subst t2. exploit assign_loc_receptive; eauto. intros EQ; rewrite EQ in H. econstructor; econstructor; eauto. inv H10. exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t0. - destruct (sem_binary_operation op v1' (typeof l) v2 (typeof r) m) as [v3'|] eqn:?. + destruct (sem_binary_operation ge op v1' (typeof l) v2 (typeof r) m) as [v3'|] eqn:?. destruct (sem_cast v3' tyres (typeof l)) as [v4'|] eqn:?. - destruct (classic (exists t2', exists m'', assign_loc (Genv.globalenv p) (typeof l) m b ofs v4' t2' m'')). + destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v4' t2' m'')). destruct H1 as [t2' [m'' P]]. econstructor; econstructor. left; eapply step_assignop with (v1 := v1'); eauto. simpl; reflexivity. econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto. @@ -1505,9 +1510,9 @@ Proof. rewrite Heqo; auto. (* assignop stuck *) exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t1. - destruct (sem_binary_operation op v1' (typeof l) v2 (typeof r) m) as [v3'|] eqn:?. + destruct (sem_binary_operation ge op v1' (typeof l) v2 (typeof r) m) as [v3'|] eqn:?. destruct (sem_cast v3' tyres (typeof l)) as [v4'|] eqn:?. - destruct (classic (exists t2', exists m'', assign_loc (Genv.globalenv p) (typeof l) m b ofs v4' t2' m'')). + destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v4' t2' m'')). destruct H1 as [t2' [m'' P]]. econstructor; econstructor. left; eapply step_assignop with (v1 := v1'); eauto. simpl; reflexivity. econstructor; econstructor. left; eapply step_assignop_stuck with (v1 := v1'); eauto. @@ -1521,9 +1526,9 @@ Proof. subst t2. exploit assign_loc_receptive; eauto. intros EQ; rewrite EQ in H. econstructor; econstructor; eauto. inv H9. exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t0. - destruct (sem_incrdecr id v1' (typeof l)) as [v2'|] eqn:?. + destruct (sem_incrdecr ge id v1' (typeof l)) as [v2'|] eqn:?. destruct (sem_cast v2' (incrdecr_type (typeof l)) (typeof l)) as [v3'|] eqn:?. - destruct (classic (exists t2', exists m'', assign_loc (Genv.globalenv p) (typeof l) m b ofs v3' t2' m'')). + destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v3' t2' m'')). destruct H1 as [t2' [m'' P]]. econstructor; econstructor. left; eapply step_postincr with (v1 := v1'); eauto. simpl; reflexivity. econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto. @@ -1534,9 +1539,9 @@ Proof. rewrite Heqo; auto. (* postincr stuck *) exploit deref_loc_receptive; eauto. intros [EQ [v1' A]]. subst t1. - destruct (sem_incrdecr id v1' (typeof l)) as [v2'|] eqn:?. + destruct (sem_incrdecr ge id v1' (typeof l)) as [v2'|] eqn:?. destruct (sem_cast v2' (incrdecr_type (typeof l)) (typeof l)) as [v3'|] eqn:?. - destruct (classic (exists t2', exists m'', assign_loc (Genv.globalenv p) (typeof l) m b ofs v3' t2' m'')). + destruct (classic (exists t2', exists m'', assign_loc ge (typeof l) m b ofs v3' t2' m'')). destruct H1 as [t2' [m'' P]]. econstructor; econstructor. left; eapply step_postincr with (v1 := v1'); eauto. simpl; reflexivity. econstructor; econstructor. left; eapply step_postincr_stuck with (v1 := v1'); eauto. @@ -1732,7 +1737,7 @@ with eval_expr: env -> mem -> kind -> expr -> trace -> mem -> expr -> Prop := eval_simple_lvalue ge e m2 l' b ofs -> deref_loc ge (typeof l) m2 b ofs t3 v1 -> eval_simple_rvalue ge e m2 r' v2 -> - sem_binary_operation op v1 (typeof l) v2 (typeof r) m2 = Some v3 -> + sem_binary_operation ge op v1 (typeof l) v2 (typeof r) m2 = Some v3 -> sem_cast v3 tyres (typeof l) = Some v4 -> assign_loc ge (typeof l) m2 b ofs v4 t4 m3 -> ty = typeof l -> @@ -1741,7 +1746,7 @@ with eval_expr: env -> mem -> kind -> expr -> trace -> mem -> expr -> Prop := 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 -> - sem_incrdecr id v1 ty = Some v2 -> + sem_incrdecr ge id v1 ty = Some v2 -> sem_cast v2 (incrdecr_type ty) ty = Some v3 -> assign_loc ge ty m1 b ofs v3 t3 m2 -> ty = typeof l -> @@ -1893,11 +1898,11 @@ with exec_stmt: env -> mem -> statement -> trace -> mem -> outcome -> Prop := with eval_funcall: mem -> fundef -> list val -> trace -> mem -> val -> Prop := | eval_funcall_internal: forall m f vargs t e m1 m2 m3 out vres m4, list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) -> - alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> + alloc_variables ge empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> bind_parameters ge e m1 f.(fn_params) vargs m2 -> exec_stmt e m2 f.(fn_body) t m3 out -> outcome_result_value out f.(fn_return) vres -> - Mem.free_list m3 (blocks_of_env e) = Some m4 -> + Mem.free_list m3 (blocks_of_env ge e) = Some m4 -> eval_funcall m (Internal f) vargs t m4 vres | eval_funcall_external: forall m ef targs tres cconv vargs t vres m', external_call ef ge vargs m t vres m' -> @@ -2115,7 +2120,7 @@ with execinf_stmt: env -> mem -> statement -> traceinf -> Prop := with evalinf_funcall: mem -> fundef -> list val -> traceinf -> Prop := | evalinf_funcall_internal: forall m f vargs t e m1 m2, list_norepet (var_names f.(fn_params) ++ var_names f.(fn_vars)) -> - alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> + alloc_variables ge empty_env m (f.(fn_params) ++ f.(fn_vars)) e m1 -> bind_parameters ge e m1 f.(fn_params) vargs m2 -> execinf_stmt e m2 f.(fn_body) t -> evalinf_funcall m (Internal f) vargs t. @@ -3019,7 +3024,7 @@ End BIGSTEP. Inductive bigstep_program_terminates (p: program): trace -> int -> Prop := | bigstep_program_terminates_intro: forall b f m0 m1 t r, - let ge := Genv.globalenv p in + let ge := globalenv p in Genv.init_mem p = Some m0 -> Genv.find_symbol ge p.(prog_main) = Some b -> Genv.find_funct_ptr ge b = Some f -> @@ -3029,7 +3034,7 @@ Inductive bigstep_program_terminates (p: program): trace -> int -> Prop := Inductive bigstep_program_diverges (p: program): traceinf -> Prop := | bigstep_program_diverges_intro: forall b f m0 t, - let ge := Genv.globalenv p in + let ge := globalenv p in Genv.init_mem p = Some m0 -> Genv.find_symbol ge p.(prog_main) = Some b -> Genv.find_funct_ptr ge b = Some f -> diff --git a/cfrontend/Csyntax.v b/cfrontend/Csyntax.v index fa74f11c..8ea4e077 100644 --- a/cfrontend/Csyntax.v +++ b/cfrontend/Csyntax.v @@ -20,6 +20,7 @@ Require Import Integers. Require Import Floats. Require Import Values. Require Import AST. +Require Import Errors. Require Import Ctypes. Require Import Cop. @@ -207,8 +208,44 @@ Definition type_of_fundef (f: fundef) : type := (** ** Programs *) -(** A program is a collection of named functions, plus a collection - of named global variables, carrying their types and optional initialization - data. See module [AST] for more details. *) +(** A program is composed of: +- a list of definitions of functions and global variables; +- the names of functions and global variables that are public (not static); +- the name of the function that acts as entry point ("main" function). +- a list of definitions for structure and union names; +- the corresponding composite environment; +*) + +Record program : Type := { + prog_defs: list (ident * globdef fundef type); + prog_public: list ident; + prog_main: ident; + prog_types: list composite_definition; + prog_comp_env: composite_env; + prog_comp_env_eq: build_composite_env prog_types = OK prog_comp_env +}. + +Definition program_of_program (p: program) : AST.program fundef type := + {| AST.prog_defs := p.(prog_defs); + AST.prog_public := p.(prog_public); + AST.prog_main := p.(prog_main) |}. + +Coercion program_of_program: program >-> AST.program. + +Program Definition make_program (types: list composite_definition) + (defs: list (ident * globdef fundef type)) + (public: list ident) + (main: ident): res program := + match build_composite_env types with + | OK env => + OK {| prog_defs := defs; + prog_public := public; + prog_main := main; + prog_types := types; + prog_comp_env := env; + prog_comp_env_eq := _ |} + | Error msg => + Error msg + end. + -Definition program : Type := AST.program fundef type. diff --git a/cfrontend/Ctypes.v b/cfrontend/Ctypes.v index c437a6bc..091c5276 100644 --- a/cfrontend/Ctypes.v +++ b/cfrontend/Ctypes.v @@ -16,6 +16,7 @@ (** Type expressions for the Compcert C and Clight languages *) Require Import Coqlib. +Require Import Maps. Require Import AST. Require Import Errors. Require Archi. @@ -63,32 +64,6 @@ Definition noattr := {| attr_volatile := false; attr_alignas := None |}. of the function arguments (list [targs]), and the type of the function result ([tres]). Variadic functions and old-style unprototyped functions are not supported. -- In C, struct and union types are named and compared by name. - This enables the definition of recursive struct types such as -<< - struct s1 { int n; struct * s1 next; }; ->> - Note that recursion within types must go through a pointer type. - For instance, the following is not allowed in C. -<< - struct s2 { int n; struct s2 next; }; ->> - In Compcert C, struct and union types [Tstruct id fields] and - [Tunion id fields] are compared by structure: the [fields] - argument gives the names and types of the members. The identifier - [id] is a local name which can be used in conjuction with the - [Tcomp_ptr] constructor to express recursive types. [Tcomp_ptr id] - stands for a pointer type to the nearest enclosing [Tstruct] - or [Tunion] type named [id]. For instance. the structure [s1] - defined above in C is expressed by -<< - Tstruct "s1" (Fcons "n" (Tint I32 Signed) - (Fcons "next" (Tcomp_ptr "s1") - Fnil)) ->> - Note that the incorrect structure [s2] above cannot be expressed at - all, since [Tcomp_ptr] lets us refer to a pointer to an enclosing - structure or union, but not to the structure or union directly. *) Inductive type : Type := @@ -99,35 +74,27 @@ Inductive type : Type := | Tpointer: type -> attr -> type (**r pointer types ([*ty]) *) | Tarray: type -> Z -> attr -> type (**r array types ([ty[len]]) *) | Tfunction: typelist -> type -> calling_convention -> type (**r function types *) - | Tstruct: ident -> fieldlist -> attr -> type (**r struct types *) - | Tunion: ident -> fieldlist -> attr -> type (**r union types *) - | Tcomp_ptr: ident -> attr -> type (**r pointer to named struct or union *) - + | Tstruct: ident -> attr -> type (**r struct types *) + | Tunion: ident -> attr -> type (**r union types *) with typelist : Type := | Tnil: typelist - | Tcons: type -> typelist -> typelist - -with fieldlist : Type := - | Fnil: fieldlist - | Fcons: ident -> type -> fieldlist -> fieldlist. + | Tcons: type -> typelist -> typelist. Lemma type_eq: forall (ty1 ty2: type), {ty1=ty2} + {ty1<>ty2} -with typelist_eq: forall (tyl1 tyl2: typelist), {tyl1=tyl2} + {tyl1<>tyl2} -with fieldlist_eq: forall (fld1 fld2: fieldlist), {fld1=fld2} + {fld1<>fld2}. +with typelist_eq: forall (tyl1 tyl2: typelist), {tyl1=tyl2} + {tyl1<>tyl2}. Proof. assert (forall (x y: intsize), {x=y} + {x<>y}) by decide equality. 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. intros E1 E2 E3. + generalize ident_eq zeq bool_dec ident_eq; intros. decide equality. decide equality. decide equality. - generalize ident_eq. intros E1. decide equality. Defined. -Opaque type_eq typelist_eq fieldlist_eq. +Opaque type_eq typelist_eq. (** Extract the attributes of a type. *) @@ -140,9 +107,8 @@ Definition attr_of_type (ty: type) := | Tpointer elt a => a | Tarray elt sz a => a | Tfunction args res cc => noattr - | Tstruct id fld a => a - | Tunion id fld a => a - | Tcomp_ptr id a => a + | Tstruct id a => a + | Tunion id a => a end. (** Change the top-level attributes of a type *) @@ -156,9 +122,8 @@ Definition change_attributes (f: attr -> attr) (ty: type) : type := | Tpointer elt a => Tpointer elt (f a) | Tarray elt sz a => Tarray elt sz (f a) | Tfunction args res cc => ty - | Tstruct id fld a => Tstruct id fld (f a) - | Tunion id fld a => Tunion id fld (f a) - | Tcomp_ptr id a => Tcomp_ptr id (f a) + | Tstruct id a => Tstruct id (f a) + | Tunion id a => Tunion id (f a) end. (** Erase the top-level attributes of a type *) @@ -181,6 +146,39 @@ Definition attr_union (a1 a2: attr) : attr := Definition merge_attributes (ty: type) (a: attr) : type := change_attributes (attr_union a) ty. +(** 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 composite_definition : Type := + Composite (id: ident) (su: struct_or_union) (m: members) (a: attr). + +(** For type-checking, compilation and semantics purposes, the composite + definitions are collected in the following [composite_env] environment. + The [composite] record contains additional information compared with + the [composite_definition], such as size and alignment information. *) + +Record composite : Type := { + co_su: struct_or_union; + co_members: members; + co_attr: attr; + co_sizeof: Z; + co_alignof: Z; + co_sizeof_pos: co_sizeof >= 0; + co_alignof_two_p: exists n, co_alignof = two_power_nat n; + co_sizeof_alignof: (co_alignof | co_sizeof) +}. + +Definition composite_env : Type := PTree.t composite. + +(** * Operations over types *) + +(** ** Conversions *) + Definition type_int32s := Tint I32 Signed noattr. Definition type_bool := Tint IBool Signed noattr. @@ -208,15 +206,51 @@ Definition default_argument_conversion (ty: type) : type := | _ => remove_attributes ty end. -(** * Operations over types *) +(** ** Complete types *) -(** Alignment of a type, in bytes. *) +(** A type is complete if it fully describes an object. + All struct and union names appearing in the type must be defined, + unless they occur under a pointer or function type. [void] and + function types are incomplete types. *) -Fixpoint alignof (t: type) : Z := - match attr_alignas (attr_of_type t) with +Fixpoint complete_type (env: composite_env) (t: type) : bool := + match t with + | Tvoid => false + | Tint _ _ _ => true + | Tlong _ _ => true + | Tfloat _ _ => true + | Tpointer _ _ => true + | Tarray t' _ _ => complete_type env t' + | Tfunction _ _ _ => false + | Tstruct id _ | Tunion id _ => + match env!id with Some co => true | None => false end + end. + +Definition complete_or_function_type (env: composite_env) (t: type) : bool := + match t with + | Tfunction _ _ _ => true + | _ => complete_type env t + end. + +(** ** Alignment of a type *) + +(** Adjust the natural alignment [al] based on the attributes [a] attached + to the type. If an "alignas" attribute is given, use it as alignment + in preference to [al]. *) + +Definition align_attr (a: attr) (al: Z) : Z := + match attr_alignas a with | Some l => two_p (Z.of_N l) - | None => - match t with + | None => al + end. + +(** In the ISO C standard, alignment is defined only for complete + types. However, it is convenient that [alignof] is a total + function. For incomplete types, it returns 1. *) + +Fixpoint alignof (env: composite_env) (t: type) : Z := + align_attr (attr_of_type t) + (match t with | Tvoid => 1 | Tint I8 _ _ => 1 | Tint I16 _ _ => 2 @@ -226,42 +260,26 @@ Fixpoint alignof (t: type) : Z := | Tfloat F32 _ => 4 | Tfloat F64 _ => Archi.align_float64 | Tpointer _ _ => 4 - | Tarray t' _ _ => alignof t' + | Tarray t' _ _ => alignof env t' | Tfunction _ _ _ => 1 - | Tstruct _ fld _ => alignof_fields fld - | Tunion _ fld _ => alignof_fields fld - | Tcomp_ptr _ _ => 4 - end - end - -with alignof_fields (f: fieldlist) : Z := - match f with - | Fnil => 1 - | Fcons id t f' => Zmax (alignof t) (alignof_fields f') - end. - -Scheme type_ind2 := Induction for type Sort Prop - with fieldlist_ind2 := Induction for fieldlist Sort Prop. + | Tstruct id _ | Tunion id _ => + match env!id with Some co => co_alignof co | None => 1 end + end). + +Remark align_attr_two_p: + forall al a, + (exists n, al = two_power_nat n) -> + (exists n, align_attr a al = two_power_nat n). +Proof. + intros. unfold align_attr. destruct (attr_alignas a). + exists (N.to_nat n). rewrite two_power_nat_two_p. rewrite N_nat_Z. auto. + auto. +Qed. Lemma alignof_two_p: - forall t, exists n, alignof t = two_power_nat n -with alignof_fields_two_p: - forall f, exists n, alignof_fields f = two_power_nat n. + forall env t, exists n, alignof env t = two_power_nat n. Proof. - assert (X: forall t a, - (exists n, a = two_power_nat n) -> - exists n, - match attr_alignas (attr_of_type t) with - | Some l => two_p (Z.of_N l) - | None => a - end = two_power_nat n). - { - intros. - destruct (attr_alignas (attr_of_type t)). - exists (N.to_nat n). rewrite two_power_nat_two_p. rewrite N_nat_Z. auto. - auto. - } - induction t; apply X; simpl; auto. + induction t; apply align_attr_two_p; simpl. exists 0%nat; auto. destruct i. exists 0%nat; auto. @@ -273,28 +291,28 @@ Proof. exists 2%nat; auto. (exists 2%nat; reflexivity) || (exists 3%nat; reflexivity). exists 2%nat; auto. + apply IHt. exists 0%nat; auto. - exists 2%nat; auto. - induction f; simpl. - exists 0%nat; auto. - apply Z.max_case; auto. + destruct (env!i). apply co_alignof_two_p. exists 0%nat; auto. + destruct (env!i). apply co_alignof_two_p. exists 0%nat; auto. Qed. Lemma alignof_pos: - forall t, alignof t > 0. + forall env t, alignof env t > 0. Proof. - intros. destruct (alignof_two_p t) as [n EQ]. rewrite EQ. apply two_power_nat_pos. + intros. destruct (alignof_two_p env t) as [n EQ]. rewrite EQ. apply two_power_nat_pos. Qed. -Lemma alignof_fields_pos: - forall f, alignof_fields f > 0. -Proof. - intros. destruct (alignof_fields_two_p f) as [n EQ]. rewrite EQ. apply two_power_nat_pos. -Qed. +(** ** Size of a type *) -(** Size of a type, in bytes. *) +(** In the ISO C standard, size is defined only for complete + types. However, it is convenient that [sizeof] is a total + function. For [void] and function types, we follow GCC and define + their size to be 1. For undefined structures and unions, the size is + arbitrarily taken to be 0. +*) -Fixpoint sizeof (t: type) : Z := +Fixpoint sizeof (env: composite_env) (t: type) : Z := match t with | Tvoid => 1 | Tint I8 _ _ => 1 @@ -305,199 +323,217 @@ Fixpoint sizeof (t: type) : Z := | Tfloat F32 _ => 4 | Tfloat F64 _ => 8 | Tpointer _ _ => 4 - | Tarray t' n _ => sizeof t' * Zmax 0 n + | Tarray t' n _ => sizeof env t' * Z.max 0 n | Tfunction _ _ _ => 1 - | Tstruct _ fld _ => align (sizeof_struct fld 0) (alignof t) - | Tunion _ fld _ => align (sizeof_union fld) (alignof t) - | Tcomp_ptr _ _ => 4 - end - -with sizeof_struct (fld: fieldlist) (pos: Z) {struct fld} : Z := - match fld with - | Fnil => pos - | Fcons id t fld' => sizeof_struct fld' (align pos (alignof t) + sizeof t) - end - -with sizeof_union (fld: fieldlist) : Z := - match fld with - | Fnil => 0 - | Fcons id t fld' => Zmax (sizeof t) (sizeof_union fld') + | Tstruct id _ | Tunion id _ => + match env!id with Some co => co_sizeof co | None => 0 end end. Lemma sizeof_pos: - forall t, sizeof t >= 0 -with sizeof_struct_incr: - forall fld pos, pos <= sizeof_struct fld pos. + forall env t, sizeof env t >= 0. Proof. -- Local Opaque alignof. - assert (X: forall n t, n >= 0 -> align n (alignof t) >= 0). - { - intros. generalize (align_le n (alignof t) (alignof_pos t)). omega. - } - induction t; simpl; try xomega. + induction t; simpl; try omega. destruct i; omega. destruct f; omega. - change 0 with (0 * Z.max 0 z) at 2. apply Zmult_ge_compat_r. auto. xomega. - apply X. apply Zle_ge. apply sizeof_struct_incr. - apply X. induction f; simpl; xomega. -- induction fld; intros; simpl. - omega. - eapply Zle_trans. 2: apply IHfld. - apply Zle_trans with (align pos (alignof t)). - apply align_le. apply alignof_pos. - generalize (sizeof_pos t); omega. + change 0 with (0 * Z.max 0 z) at 2. apply Zmult_ge_compat_r. auto. xomega. + destruct (env!i). apply co_sizeof_pos. omega. + destruct (env!i). apply co_sizeof_pos. omega. Qed. +(** The size of a type is an integral multiple of its alignment, + unless the alignment was artificially increased with the [__Alignas] + attribute. *) + Fixpoint naturally_aligned (t: type) : Prop := + attr_alignas (attr_of_type t) = None /\ match t with - | Tint _ _ a | Tlong _ a | Tfloat _ a | Tpointer _ a | Tcomp_ptr _ a => - attr_alignas a = None - | Tarray t' _ a => - attr_alignas a = None /\ naturally_aligned t' - | Tvoid | Tfunction _ _ _ | Tstruct _ _ _ | Tunion _ _ _ => - True + | Tarray t' _ _ => naturally_aligned t' + | _ => True end. Lemma sizeof_alignof_compat: - forall t, naturally_aligned t -> (alignof t | sizeof t). + forall env t, naturally_aligned t -> (alignof env t | sizeof env t). Proof. -Local Transparent alignof. - induction t; simpl; intros. + induction t; intros [A B]; unfold alignof, align_attr; rewrite A; simpl. - apply Zdivide_refl. -- rewrite H. destruct i; apply Zdivide_refl. -- rewrite H. exists (8 / Archi.align_int64); reflexivity. -- rewrite H. destruct f. apply Zdivide_refl. exists (8 / Archi.align_float64); reflexivity. -- rewrite H; apply Zdivide_refl. -- destruct H. rewrite H. apply Z.divide_mul_l; auto. +- destruct i; apply Zdivide_refl. +- exists (8 / Archi.align_int64); reflexivity. +- destruct f. apply Zdivide_refl. exists (8 / Archi.align_float64); reflexivity. - apply Zdivide_refl. -- change (alignof (Tstruct i f a) | align (sizeof_struct f 0) (alignof (Tstruct i f a))). - apply align_divides. apply alignof_pos. -- change (alignof (Tunion i f a) | align (sizeof_union f) (alignof (Tunion i f a))). - apply align_divides. apply alignof_pos. -- rewrite H; apply Zdivide_refl. +- apply Z.divide_mul_l; auto. +- apply Zdivide_refl. +- destruct (env!i). apply co_sizeof_alignof. apply Zdivide_0. +- destruct (env!i). apply co_sizeof_alignof. apply Zdivide_0. Qed. -(** Byte offset for a field in a struct or union. - Field are laid out consecutively, and padding is inserted - to align each field to the natural alignment for its type. *) +(** ** Size and alignment for composite definitions *) + +(** The alignment for a structure or union is the max of the alignment + of its members. *) -Open Local Scope string_scope. +Fixpoint alignof_composite (env: composite_env) (m: members) : Z := + match m with + | nil => 1 + | (id, t) :: m' => Z.max (alignof env t) (alignof_composite env m') + 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. *) + +Fixpoint sizeof_struct (env: composite_env) (cur: Z) (m: members) : Z := + match m with + | nil => cur + | (id, t) :: m' => sizeof_struct env (align cur (alignof env t) + sizeof env t) m' + end. + +(** The size of an union is the max of the sizes of its members. *) -Fixpoint field_offset_rec (id: ident) (fld: fieldlist) (pos: Z) - {struct fld} : res Z := +Fixpoint sizeof_union (env: composite_env) (m: members) : Z := + match m with + | nil => 0 + | (id, t) :: m' => Z.max (sizeof env t) (sizeof_union env m') + end. + +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. +- exists 0%nat; auto. +- apply Z.max_case; auto. apply alignof_two_p. +Qed. + +Lemma alignof_composite_pos: + forall env m a, align_attr a (alignof_composite env m) > 0. +Proof. + intros. + exploit align_attr_two_p. apply (alignof_composite_two_p env m). + instantiate (1 := a). intros [n EQ]. + rewrite EQ; apply two_power_nat_pos. +Qed. + +Lemma sizeof_struct_incr: + forall env m cur, cur <= sizeof_struct env cur m. +Proof. + induction m as [|[id t]]; simpl; intros. +- omega. +- apply Zle_trans with (align cur (alignof env t)). + apply align_le. apply alignof_pos. + apply Zle_trans with (align cur (alignof env t) + sizeof env t). + generalize (sizeof_pos env t); omega. + apply IHm. +Qed. + +Lemma sizeof_union_pos: + forall env m, 0 <= sizeof_union env m. +Proof. + induction m as [|[id t]]; simpl; xomega. +Qed. + +(** ** Byte offset 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 - | Fnil => Error (MSG "Unknown field " :: CTX id :: nil) - | Fcons id' t fld' => + | nil => Error (MSG "Unknown field " :: CTX id :: nil) + | (id', t) :: fld' => if ident_eq id id' - then OK (align pos (alignof t)) - else field_offset_rec id fld' (align pos (alignof t) + sizeof t) + then OK (align pos (alignof env t)) + else field_offset_rec env id fld' (align pos (alignof env t) + sizeof env t) end. -Definition field_offset (id: ident) (fld: fieldlist) : res Z := - field_offset_rec id fld 0. +Definition field_offset (env: composite_env) (id: ident) (fld: members) : res Z := + field_offset_rec env id fld 0. -Fixpoint field_type (id: ident) (fld: fieldlist) {struct fld} : res type := +Fixpoint field_type (id: ident) (fld: members) {struct fld} : res type := match fld with - | Fnil => Error (MSG "Unknown field " :: CTX id :: nil) - | Fcons id' t fld' => if ident_eq id id' then OK t else field_type id fld' + | nil => Error (MSG "Unknown field " :: CTX id :: nil) + | (id', t) :: fld' => if ident_eq id id' then OK t else field_type id fld' end. (** Some sanity checks about field offsets. First, field offsets are within the range of acceptable offsets. *) Remark field_offset_rec_in_range: - forall id ofs ty fld pos, - field_offset_rec id fld pos = OK ofs -> field_type id fld = OK ty -> - pos <= ofs /\ ofs + sizeof ty <= sizeof_struct fld pos. + 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. Proof. - intros until ty. induction fld; simpl. - congruence. - destruct (ident_eq id i); intros. - inv H. inv H0. split. apply align_le. apply alignof_pos. apply sizeof_struct_incr. + intros until ty. induction fld as [|[i t]]; 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 Zle_trans; eauto. apply Zle_trans with (align pos (alignof t)). - apply align_le. apply alignof_pos. generalize (sizeof_pos t). omega. + eapply Zle_trans; eauto. apply Zle_trans with (align pos (alignof env t)). + apply align_le. apply alignof_pos. generalize (sizeof_pos env t). omega. Qed. Lemma field_offset_in_range: - forall sid fld a fid ofs ty, - field_offset fid fld = OK ofs -> field_type fid fld = OK ty -> - 0 <= ofs /\ ofs + sizeof ty <= sizeof (Tstruct sid fld a). + 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. Proof. - intros. exploit field_offset_rec_in_range; eauto. intros [A B]. - split. auto. -Local Opaque alignof. - simpl. eapply Zle_trans; eauto. - apply align_le. apply alignof_pos. + intros. eapply field_offset_rec_in_range; eauto. Qed. (** Second, two distinct fields do not overlap *) Lemma field_offset_no_overlap: - forall id1 ofs1 ty1 id2 ofs2 ty2 fld, - field_offset id1 fld = OK ofs1 -> field_type id1 fld = OK ty1 -> - field_offset id2 fld = OK ofs2 -> field_type id2 fld = OK ty2 -> + 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 -> id1 <> id2 -> - ofs1 + sizeof ty1 <= ofs2 \/ ofs2 + sizeof ty2 <= ofs1. + ofs1 + sizeof env ty1 <= ofs2 \/ ofs2 + sizeof env ty2 <= ofs1. Proof. - intros until ty2. intros fld0 A B C D NEQ. - assert (forall fld pos, - field_offset_rec id1 fld pos = OK ofs1 -> field_type id1 fld = OK ty1 -> - field_offset_rec id2 fld pos = OK ofs2 -> field_type id2 fld = OK ty2 -> - ofs1 + sizeof ty1 <= ofs2 \/ ofs2 + sizeof ty2 <= ofs1). - induction fld; intro pos; simpl. congruence. - destruct (ident_eq id1 i); destruct (ident_eq id2 i). - congruence. - subst i. intros. inv H; inv H0. + intros until fld. unfold field_offset. generalize 0 as pos. + induction fld as [|[i t]]; simpl; intros. +- discriminate. +- destruct (ident_eq id1 i); destruct (ident_eq id2 i). ++ congruence. ++ subst i. inv H; inv H0. exploit field_offset_rec_in_range. eexact H1. eauto. tauto. - subst i. intros. inv H1; inv H2. ++ subst i. inv H1; inv H2. exploit field_offset_rec_in_range. eexact H. eauto. tauto. - intros. eapply IHfld; eauto. - - apply H with fld0 0; auto. ++ eapply IHfld; eauto. Qed. (** Third, if a struct is a prefix of another, the offsets of common fields are the same. *) -Fixpoint fieldlist_app (fld1 fld2: fieldlist) {struct fld1} : fieldlist := - match fld1 with - | Fnil => fld2 - | Fcons id ty fld => Fcons id ty (fieldlist_app fld fld2) - end. - Lemma field_offset_prefix: - forall id ofs fld2 fld1, - field_offset id fld1 = OK ofs -> - field_offset id (fieldlist_app fld1 fld2) = OK ofs. + forall env id ofs fld2 fld1, + field_offset env id fld1 = OK ofs -> + field_offset env id (fld1 ++ fld2) = OK ofs. Proof. - intros until fld2. - assert (forall fld1 pos, - field_offset_rec id fld1 pos = OK ofs -> - field_offset_rec id (fieldlist_app fld1 fld2) pos = OK ofs). - induction fld1; intros pos; simpl. congruence. - destruct (ident_eq id i); auto. - intros. unfold field_offset; auto. + intros until fld1. unfold field_offset. generalize 0 as pos. + induction fld1 as [|[i t]]; simpl; intros. +- discriminate. +- destruct (ident_eq id i); auto. Qed. (** Fourth, the position of each field respects its alignment. *) Lemma field_offset_aligned: - forall id fld ofs ty, - field_offset id fld = OK ofs -> field_type id fld = OK ty -> - (alignof ty | ofs). + forall env id fld ofs ty, + field_offset env id fld = OK ofs -> field_type id fld = OK ty -> + (alignof env ty | ofs). Proof. - assert (forall id ofs ty fld pos, - field_offset_rec id fld pos = OK ofs -> field_type id fld = OK ty -> - (alignof ty | ofs)). - induction fld; simpl; intros. - discriminate. - destruct (ident_eq id i). inv H; inv H0. - apply align_divides. apply alignof_pos. - eapply IHfld; eauto. - intros. eapply H with (pos := 0); eauto. + intros until ty. unfold field_offset. generalize 0 as pos. revert fld. + induction fld as [|[i t]]; simpl; intros. +- discriminate. +- destruct (ident_eq id i). ++ inv H; inv H0. apply align_divides. apply alignof_pos. ++ eauto. Qed. +(** ** Access modes *) + (** The [access_mode] function describes how a l-value of the given type must be accessed: - [By_value ch]: access by value, i.e. by loading from the address @@ -530,9 +566,8 @@ Definition access_mode (ty: type) : mode := | Tpointer _ _ => By_value Mint32 | Tarray _ _ _ => By_reference | Tfunction _ _ _ => By_reference - | Tstruct _ _ _ => By_copy - | Tunion _ _ _ => By_copy - | Tcomp_ptr _ _ => By_nothing + | Tstruct _ _ => By_copy + | Tunion _ _ => By_copy end. (** For the purposes of the semantics and the compiler, a type denotes @@ -545,87 +580,13 @@ Definition type_is_volatile (ty: type) : bool := | _ => false end. -(** Unroll the type of a structure or union field, substituting - [Tcomp_ptr] by a pointer to the structure. *) - -Section UNROLL_COMPOSITE. - -Variable cid: ident. -Variable comp: type. - -Fixpoint unroll_composite (ty: type) : type := - match ty with - | Tvoid => ty - | Tint _ _ _ => ty - | Tlong _ _ => ty - | Tfloat _ _ => ty - | Tpointer t1 a => Tpointer (unroll_composite t1) a - | Tarray t1 sz a => Tarray (unroll_composite t1) sz a - | Tfunction t1 t2 a => Tfunction (unroll_composite_list t1) (unroll_composite t2) a - | Tstruct id fld a => if ident_eq id cid then ty else Tstruct id (unroll_composite_fields fld) a - | Tunion id fld a => if ident_eq id cid then ty else Tunion id (unroll_composite_fields fld) a - | Tcomp_ptr id a => if ident_eq id cid then Tpointer comp a else ty - end - -with unroll_composite_list (tl: typelist) : typelist := - match tl with - | Tnil => Tnil - | Tcons t1 tl' => Tcons (unroll_composite t1) (unroll_composite_list tl') - end - -with unroll_composite_fields (fld: fieldlist) : fieldlist := - match fld with - | Fnil => Fnil - | Fcons id ty fld' => Fcons id (unroll_composite ty) (unroll_composite_fields fld') - end. - -Lemma attr_of_type_unroll_composite: - forall ty, attr_of_type (unroll_composite ty) = attr_of_type ty. -Proof. - intros. destruct ty; simpl; auto; destruct (ident_eq i cid); auto. -Qed. - -Lemma alignof_unroll_composite: - forall ty, alignof (unroll_composite ty) = alignof ty. -Proof. -Local Transparent alignof. - apply (type_ind2 (fun ty => alignof (unroll_composite ty) = alignof ty) - (fun fld => alignof_fields (unroll_composite_fields fld) = alignof_fields fld)); - simpl; intros; auto. - rewrite H; auto. - destruct (ident_eq i cid); auto. simpl. rewrite H; auto. - destruct (ident_eq i cid); auto. simpl. rewrite H; auto. - destruct (ident_eq i cid); auto. congruence. -Qed. - -Lemma sizeof_unroll_composite: - forall ty, sizeof (unroll_composite ty) = sizeof ty. -Proof. -Local Opaque alignof. - apply (type_ind2 (fun ty => sizeof (unroll_composite ty) = sizeof ty) - (fun fld => - sizeof_union (unroll_composite_fields fld) = sizeof_union fld - /\ forall pos, - sizeof_struct (unroll_composite_fields fld) pos = sizeof_struct fld pos)); - simpl; intros; auto. -- rewrite H. auto. -- rewrite <- (alignof_unroll_composite (Tstruct i f a)). simpl. - destruct H. destruct (ident_eq i cid). auto. simpl. rewrite H0. auto. -- rewrite <- (alignof_unroll_composite (Tunion i f a)). simpl. - destruct H. destruct (ident_eq i cid). auto. simpl. rewrite H. auto. -- destruct (ident_eq i cid); auto. -- destruct H0. split. - + congruence. - + intros. rewrite H1. rewrite H. rewrite alignof_unroll_composite. auto. -Qed. - -End UNROLL_COMPOSITE. +(** ** Alignment for block copy operations *) (** A variant of [alignof] for use in block copy operations. Block copy operations do not support alignments greater than 8, and require the size to be an integral multiple of the alignment. *) -Fixpoint alignof_blockcopy (t: type) : Z := +Fixpoint alignof_blockcopy (env: composite_env) (t: type) : Z := match t with | Tvoid => 1 | Tint I8 _ _ => 1 @@ -636,20 +597,22 @@ Fixpoint alignof_blockcopy (t: type) : Z := | Tfloat F32 _ => 4 | Tfloat F64 _ => 8 | Tpointer _ _ => 4 - | Tarray t' _ _ => alignof_blockcopy t' + | Tarray t' _ _ => alignof_blockcopy env t' | Tfunction _ _ _ => 1 - | Tstruct _ fld _ => Zmin 8 (alignof t) - | Tunion _ fld _ => Zmin 8 (alignof t) - | Tcomp_ptr _ _ => 4 + | Tstruct id _ | Tunion id _ => + match env!id with + | Some co => Z.min 8 (co_alignof co) + | None => 1 + end end. Lemma alignof_blockcopy_1248: - forall ty, let a := alignof_blockcopy ty in a = 1 \/ a = 2 \/ a = 4 \/ a = 8. + forall env ty, let a := alignof_blockcopy env ty in a = 1 \/ a = 2 \/ a = 4 \/ a = 8. Proof. - assert (X: forall ty, let a := Zmin 8 (alignof ty) in + assert (X: forall co, let a := Zmin 8 (co_alignof co) in a = 1 \/ a = 2 \/ a = 4 \/ a = 8). { - intros. destruct (alignof_two_p ty) as [n EQ]. unfold a; rewrite EQ. + intros. destruct (co_alignof_two_p co) as [n EQ]. unfold a; rewrite EQ. destruct n; auto. destruct n; auto. destruct n; auto. @@ -660,45 +623,47 @@ Proof. induction ty; simpl; auto. destruct i; auto. destruct f; auto. + destruct (env!i); auto. + destruct (env!i); auto. Qed. Lemma alignof_blockcopy_pos: - forall ty, alignof_blockcopy ty > 0. + forall env ty, alignof_blockcopy env ty > 0. Proof. - intros. generalize (alignof_blockcopy_1248 ty). simpl. intuition omega. + intros. generalize (alignof_blockcopy_1248 env ty). simpl. intuition omega. Qed. Lemma sizeof_alignof_blockcopy_compat: - forall ty, (alignof_blockcopy ty | sizeof ty). + forall env ty, (alignof_blockcopy env ty | sizeof env ty). Proof. - assert (X: forall ty sz, (alignof ty | sz) -> (Zmin 8 (alignof ty) | sz)). + assert (X: forall co, (Z.min 8 (co_alignof co) | co_sizeof co)). { - intros. destruct (alignof_two_p ty) as [n EQ]. rewrite EQ in *. - destruct n; auto. - destruct n; auto. - destruct n; auto. - eapply Zdivide_trans; eauto. + intros. apply Zdivide_trans with (co_alignof co). 2: apply co_sizeof_alignof. + destruct (co_alignof_two_p co) as [n EQ]. rewrite EQ. + destruct n. apply Zdivide_refl. + destruct n. apply Zdivide_refl. + destruct n. apply Zdivide_refl. apply Z.min_case. - replace (two_power_nat (S(S(S n)))) with (two_p (3 + Z.of_nat n)). - rewrite two_p_is_exp by omega. change (two_p 3) with 8. - exists (two_p (Z.of_nat n)). ring. + exists (two_p (Z.of_nat n)). + change 8 with (two_p 3). + rewrite <- two_p_is_exp by omega. rewrite two_power_nat_two_p. rewrite !inj_S. f_equal. omega. apply Zdivide_refl. } - Local Opaque alignof. induction ty; simpl. apply Zdivide_refl. - destruct i; apply Zdivide_refl. apply Zdivide_refl. - destruct f; apply Zdivide_refl. apply Zdivide_refl. - apply Z.divide_mul_l. auto. apply Zdivide_refl. - apply X. apply align_divides. apply alignof_pos. - apply X. apply align_divides. apply alignof_pos. apply Zdivide_refl. + apply Z.divide_mul_l. auto. + apply Zdivide_refl. + destruct (env!i). apply X. apply Zdivide_0. + destruct (env!i). apply X. apply Zdivide_0. Qed. +(** ** C types and back-end types *) + (** Extracting a type list from a function parameter declaration. *) Fixpoint type_of_params (params: list (ident * type)) : typelist := @@ -735,3 +700,259 @@ Fixpoint typlist_of_typelist (tl: typelist) : list AST.typ := Definition signature_of_type (args: typelist) (res: type) (cc: calling_convention): signature := mksignature (typlist_of_typelist args) (opttyp_of_type res) cc. +(** * Construction of the composite environment *) + +Definition sizeof_composite (env: composite_env) (su: struct_or_union) (m: members) : Z := + match su with + | Struct => sizeof_struct env 0 m + | Union => sizeof_union env m + end. + +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. +Qed. + +Fixpoint complete_members (env: composite_env) (m: members) : bool := + match m with + | nil => true + | (id, t) :: m' => complete_type env t && complete_members env m' + end. + +(** Convert a composite definition to its internal representation. + The size and alignment of the composite are determined at this time. + The alignment takes into account the [__Alignas] attributes + associated with the definition. The size is rounded up to a multiple + of the alignment. + + The conversion fails if a type of a member is not complete. This rules + out incorrect recursive definitions such as +<< + struct s { int x; struct s next; } +>> + Here, when we process the definition of [struct s], the identifier [s] + is not bound yet in the composite environment, hence field [next] + has an incomplete type. However, recursions that go through a pointer type + are correctly handled: +<< + struct s { int x; struct s * next; } +>> + Here, [next] has a pointer type, which is always complete, even though + [s] is not yet bound to a composite. +*) + +Program Definition composite_of_def + (env: composite_env) (id: ident) (su: struct_or_union) (m: members) (a: attr) + : res composite := + match env!id, complete_members env m return _ with + | Some _, _ => + Error (MSG "Multiple definitions of struct or union " :: CTX id :: nil) + | None, false => + Error (MSG "Incomplete struct or union " :: CTX id :: nil) + | None, true => + let al := align_attr a (alignof_composite env m) in + OK {| co_su := su; + co_members := m; + co_attr := a; + co_sizeof := align (sizeof_composite env su m) al; + co_alignof := al; + co_sizeof_pos := _; + co_alignof_two_p := _; + co_sizeof_alignof := _ |} + end. +Next Obligation. + apply Zle_ge. eapply Zle_trans. eapply sizeof_composite_pos. + apply align_le; apply alignof_composite_pos. +Qed. +Next Obligation. + apply align_attr_two_p. apply alignof_composite_two_p. +Qed. +Next Obligation. + apply align_divides. apply alignof_composite_pos. +Qed. + +(** The composite environment for a program is obtained by entering + its composite definitions in sequence. The definitions are assumed + to be listed in dependency order: the definition of a composite + 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 + | Composite id su m a :: defs => + do co <- composite_of_def env id su m a; + add_composite_definitions (PTree.set id co env) defs + end. + +Definition build_composite_env (defs: list composite_definition) := + add_composite_definitions (PTree.empty _) defs. + +(** Stability properties for alignments and sizes. If the type is + complete in a composite environment [env], its size and alignment + are unchanged if we add more definitions to [env]. *) + +Section STABILITY. + +Variables env env': composite_env. +Hypothesis extends: forall id co, env!id = Some co -> env'!id = Some co. + +Lemma alignof_stable: + forall t, complete_type env t = true -> alignof env' t = alignof env t. +Proof. + induction t; simpl; intros; f_equal; auto. + destruct (env!i) as [co|] eqn:E; try discriminate. + erewrite extends by eauto. auto. + destruct (env!i) as [co|] eqn:E; try discriminate. + erewrite extends by eauto. auto. +Qed. + +Lemma sizeof_stable: + forall t, complete_type env t = true -> sizeof env' t = sizeof env t. +Proof. + induction t; simpl; intros; auto. + rewrite IHt by auto. auto. + destruct (env!i) as [co|] eqn:E; try discriminate. + erewrite extends by eauto. auto. + destruct (env!i) as [co|] eqn:E; try discriminate. + erewrite extends by eauto. auto. +Qed. + +Lemma complete_type_stable: + forall t, complete_type env t = true -> complete_type env' t = true. +Proof. + induction t; simpl; intros; auto. + destruct (env!i) as [co|] eqn:E; try discriminate. + erewrite extends by eauto. auto. + destruct (env!i) as [co|] eqn:E; try discriminate. + erewrite extends by eauto. auto. +Qed. + +Lemma alignof_composite_stable: + forall m, complete_members env m = true -> alignof_composite env' m = alignof_composite env m. +Proof. + induction m as [|[id t]]; simpl; intros. + auto. + InvBooleans. rewrite alignof_stable by auto. rewrite IHm 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. +Proof. + induction m as [|[id t]]; simpl; intros. + auto. + InvBooleans. rewrite alignof_stable by auto. rewrite sizeof_stable by auto. + rewrite IHm by auto. auto. +Qed. + +Lemma sizeof_union_stable: + forall m, complete_members env m = true -> sizeof_union env' m = sizeof_union env m. +Proof. + induction m as [|[id t]]; simpl; intros. + auto. + InvBooleans. rewrite sizeof_stable by auto. rewrite IHm 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. +Proof. + intros. destruct su; simpl. + apply sizeof_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. +Proof. + induction m as [|[id t]]; simpl; intros. + auto. + InvBooleans. rewrite complete_type_stable by auto. rewrite IHm by auto. auto. +Qed. + +End STABILITY. + +Lemma add_composite_definitions_incr: + forall id co defs env1 env2, + add_composite_definitions env1 defs = OK env2 -> + env1!id = Some co -> env2!id = Some co. +Proof. + induction defs; simpl; intros. +- inv H; auto. +- destruct a; monadInv H. + eapply IHdefs; eauto. rewrite PTree.gso; auto. + red; intros; subst id0. unfold composite_of_def in EQ. rewrite H0 in EQ; discriminate. +Qed. + +(** It follows that the sizes and alignments contained in the composite + environment produced by [build_composite_env] are consistent with + the sizes and alignments of the members of the composite types. *) + +Definition composite_consistent (env: composite_env) (co: composite) : Prop := + complete_members env (co_members co) = true + /\ co_alignof co = align_attr (co_attr co) (alignof_composite env (co_members co)) + /\ co_sizeof co = align (sizeof_composite env (co_su co) (co_members co)) (co_alignof co). + +Definition composite_env_consistent (env: composite_env) : Prop := + forall id co, env!id = Some co -> composite_consistent env co. + +Theorem build_composite_env_consistent: + forall defs env, build_composite_env defs = OK env -> composite_env_consistent env. +Proof. + cut (forall defs env0 env, + add_composite_definitions env0 defs = OK env -> + composite_env_consistent env0 -> + composite_env_consistent env). + intros. eapply H; eauto. red; intros. rewrite PTree.gempty in H1; discriminate. + induction defs as [|d1 defs]; simpl; intros. +- inv H; auto. +- destruct d1; monadInv H. + eapply IHdefs; eauto. + set (env1 := PTree.set id x env0) in *. + unfold composite_of_def in EQ. + destruct (env0!id) eqn:E; try discriminate. + destruct (complete_members env0 m) eqn:C; inversion EQ; clear EQ. + assert (forall id1 co1, env0!id1 = Some co1 -> env1!id1 = Some co1). + { intros. unfold env1. rewrite PTree.gso; auto. congruence. } + red; intros. unfold env1 in H2; rewrite PTree.gsspec in H2; destruct (peq id0 id). ++ subst id0. inversion H2; clear H2. subst co. + red; rewrite <- H1; simpl. + assert (A: alignof_composite env1 m = alignof_composite env0 m) + by (apply alignof_composite_stable; assumption). + rewrite A. + split. eapply complete_members_stable; eauto. + split. auto. + f_equal. symmetry. apply sizeof_composite_stable; assumption. ++ exploit H0; eauto. intros (P & Q & R). + assert (A: alignof_composite env1 (co_members co) = alignof_composite env0 (co_members co)) + by (apply alignof_composite_stable; assumption). + red. rewrite A. + split. eapply complete_members_stable; eauto. + split. auto. + rewrite R. f_equal. symmetry. apply sizeof_composite_stable; assumption. +Qed. + +(** Moreover, every composite definition is reflected in the composite environment. *) + +Theorem build_composite_env_charact: + forall id su m a defs env, + build_composite_env defs = OK env -> + In (Composite id su m a) defs -> + exists co, env!id = Some co /\ co_members co = m /\ co_attr co = a /\ co_su co = su. +Proof. + intros until defs. unfold build_composite_env. generalize (PTree.empty composite) as env0. + revert defs. induction defs as [|d1 defs]; simpl; intros. +- contradiction. +- destruct d1; monadInv H. + destruct H0; [idtac|eapply IHdefs;eauto]. inv H. + unfold composite_of_def in EQ. + destruct (env0!id) eqn:E; try discriminate. + destruct (complete_members env0 m) eqn:C; simplify_eq EQ. clear EQ; intros EQ. + exists x. + split. eapply add_composite_definitions_incr; eauto. apply PTree.gss. + subst x; auto. +Qed. diff --git a/cfrontend/Initializers.v b/cfrontend/Initializers.v index 6f193cd9..025960d7 100644 --- a/cfrontend/Initializers.v +++ b/cfrontend/Initializers.v @@ -13,6 +13,7 @@ (** Compile-time evaluation of initializers for global C variables. *) Require Import Coqlib. +Require Import Maps. Require Import Errors. Require Import Integers. Require Import Floats. @@ -51,7 +52,13 @@ Definition do_cast (v: val) (t1 t2: type) : res val := | None => Error(msg "undefined cast") end. -Fixpoint constval (a: expr) : res val := +Definition lookup_composite (ce: composite_env) (id: ident) : res composite := + match ce!id with + | Some co => OK co + | None => Error (MSG "Undefined struct or union " :: CTX id :: nil) + end. + +Fixpoint constval (ce: composite_env) (a: expr) : res val := match a with | Eval v ty => match v with @@ -60,74 +67,75 @@ Fixpoint constval (a: expr) : res val := end | Evalof l ty => match access_mode ty with - | By_reference | By_copy => constval l + | By_reference | By_copy => constval ce l | _ => Error(msg "dereferencing of an l-value") end | Eaddrof l ty => - constval l + constval ce l | Eunop op r1 ty => - do v1 <- constval r1; + do v1 <- constval ce r1; match sem_unary_operation op v1 (typeof r1) with | Some v => OK v | None => Error(msg "undefined unary operation") end | Ebinop op r1 r2 ty => - do v1 <- constval r1; - do v2 <- constval r2; - match sem_binary_operation op v1 (typeof r1) v2 (typeof r2) Mem.empty with + do v1 <- constval ce r1; + do v2 <- constval ce r2; + match sem_binary_operation ce op v1 (typeof r1) v2 (typeof r2) Mem.empty with | Some v => OK v | None => Error(msg "undefined binary operation") end | Ecast r ty => - do v1 <- constval r; do_cast v1 (typeof r) ty + do v1 <- constval ce r; do_cast v1 (typeof r) ty | Esizeof ty1 ty => - OK (Vint (Int.repr (sizeof ty1))) + OK (Vint (Int.repr (sizeof ce ty1))) | Ealignof ty1 ty => - OK (Vint (Int.repr (alignof ty1))) + OK (Vint (Int.repr (alignof ce ty1))) | Eseqand r1 r2 ty => - do v1 <- constval r1; - do v2 <- constval r2; + do v1 <- constval ce r1; + do v2 <- constval ce r2; match bool_val v1 (typeof r1) with | Some true => do_cast v2 (typeof r2) type_bool | Some false => OK (Vint Int.zero) | None => Error(msg "undefined && operation") end | Eseqor r1 r2 ty => - do v1 <- constval r1; - do v2 <- constval r2; + do v1 <- constval ce r1; + do v2 <- constval ce r2; match bool_val v1 (typeof r1) with | Some false => do_cast v2 (typeof r2) type_bool | Some true => OK (Vint Int.one) | None => Error(msg "undefined || operation") end | Econdition r1 r2 r3 ty => - do v1 <- constval r1; - do v2 <- constval r2; - do v3 <- constval r3; + do v1 <- constval ce r1; + do v2 <- constval ce r2; + do v3 <- constval ce r3; match bool_val v1 (typeof r1) with | Some true => do_cast v2 (typeof r2) ty | Some false => do_cast v3 (typeof r3) ty | None => Error(msg "condition is undefined") end | Ecomma r1 r2 ty => - do v1 <- constval r1; constval r2 + do v1 <- constval ce r1; constval ce r2 | Evar x ty => OK(Vptr x Int.zero) | Ederef r ty => - constval r + constval ce r | Efield l f ty => match typeof l with - | Tstruct id fList _ => - do delta <- field_offset f fList; - do v <- constval l; + | Tstruct id _ => + do co <- lookup_composite ce id; + do delta <- field_offset ce f (co_members co); + do v <- constval ce l; OK (Val.add v (Vint (Int.repr delta))) - | Tunion id fList _ => - constval l + | Tunion id _ => + constval ce l | _ => Error(msg "ill-typed field access") end | Eparen r tycast ty => - do v <- constval r; do_cast v (typeof r) tycast + do v <- constval ce r; do_cast v (typeof r) tycast | _ => Error(msg "not a compile-time constant") end. @@ -146,21 +154,19 @@ with initializer_list := (** Translate an initializing expression [a] for a scalar variable of type [ty]. Return the corresponding initialization datum. *) -Definition transl_init_single (ty: type) (a: expr) : res init_data := - do v1 <- constval a; +Definition transl_init_single (ce: composite_env) (ty: type) (a: expr) : res init_data := + do v1 <- constval ce a; do v2 <- do_cast v1 (typeof a) ty; match v2, ty with | Vint n, Tint (I8|IBool) sg _ => OK(Init_int8 n) | Vint n, Tint I16 sg _ => OK(Init_int16 n) | Vint n, Tint I32 sg _ => OK(Init_int32 n) | Vint n, Tpointer _ _ => OK(Init_int32 n) - | Vint n, Tcomp_ptr _ _ => OK(Init_int32 n) | Vlong n, Tlong _ _ => OK(Init_int64 n) | Vsingle f, Tfloat F32 _ => OK(Init_float32 f) | Vfloat f, Tfloat F64 _ => OK(Init_float64 f) | Vptr id ofs, Tint I32 sg _ => OK(Init_addrof id ofs) | Vptr id ofs, Tpointer _ _ => OK(Init_addrof id ofs) - | Vptr id ofs, Tcomp_ptr _ _ => OK(Init_addrof id ofs) | Vundef, _ => Error(msg "undefined operation in initializer") | _, _ => Error (msg "type mismatch in initializer") end. @@ -171,46 +177,55 @@ Definition transl_init_single (ty: type) (a: expr) : res init_data := Definition padding (frm to: Z) : list init_data := if zlt frm to then Init_space (to - frm) :: nil else nil. -Fixpoint transl_init (ty: type) (i: initializer) +Fixpoint transl_init (ce: composite_env) (ty: type) (i: initializer) {struct i} : res (list init_data) := match i, ty with | Init_single a, _ => - do d <- transl_init_single ty a; OK (d :: nil) + do d <- transl_init_single ce ty a; OK (d :: nil) | Init_array il, Tarray tyelt nelt _ => - transl_init_array tyelt il (Zmax 0 nelt) - | Init_struct il, Tstruct id fl _ => - transl_init_struct id ty fl il 0 - | Init_union f i1, Tunion id fl _ => - do ty1 <- field_type f fl; - do d <- transl_init ty1 i1; - OK (d ++ padding (sizeof ty1) (sizeof ty)) + transl_init_array ce tyelt il (Zmax 0 nelt) + | 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 + | 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 d <- transl_init ce ty1 i1; + OK (d ++ padding (sizeof ce ty1) (sizeof ce ty)) + end | _, _ => Error (msg "wrong type for compound initializer") end -with transl_init_array (ty: type) (il: initializer_list) (sz: Z) +with transl_init_array (ce: composite_env) (ty: type) (il: initializer_list) (sz: Z) {struct il} : res (list init_data) := match il with | Init_nil => if zeq sz 0 then OK nil - else if zle 0 sz then OK (Init_space (sz * sizeof ty) :: nil) + else if zle 0 sz then OK (Init_space (sz * sizeof ce ty) :: nil) else Error (msg "wrong number of elements in array initializer") | Init_cons i1 il' => - do d1 <- transl_init ty i1; - do d2 <- transl_init_array ty il' (sz - 1); + do d1 <- transl_init ce ty i1; + do d2 <- transl_init_array ce ty il' (sz - 1); OK (d1 ++ d2) end -with transl_init_struct (id: ident) (ty: type) - (fl: fieldlist) (il: initializer_list) (pos: Z) +with transl_init_struct (ce: composite_env) (ty: type) + (fl: members) (il: initializer_list) (pos: Z) {struct il} : res (list init_data) := match il, fl with - | Init_nil, Fnil => - OK (padding pos (sizeof ty)) - | Init_cons i1 il', Fcons _ ty1 fl' => - let pos1 := align pos (alignof ty1) in - do d1 <- transl_init ty1 i1; - do d2 <- transl_init_struct id ty fl' il' (pos1 + sizeof ty1); + | Init_nil, nil => + OK (padding pos (sizeof ce ty)) + | Init_cons i1 il', (_, ty1) :: fl' => + let pos1 := align pos (alignof ce ty1) in + do d1 <- transl_init ce ty1 i1; + do d2 <- transl_init_struct ce ty fl' il' (pos1 + sizeof ce ty1); OK (padding pos pos1 ++ d1 ++ d2) | _, _ => Error (msg "wrong number of elements in struct initializer") diff --git a/cfrontend/Initializersproof.v b/cfrontend/Initializersproof.v index 73fd90b7..225dc23e 100644 --- a/cfrontend/Initializersproof.v +++ b/cfrontend/Initializersproof.v @@ -90,13 +90,13 @@ Inductive eval_simple_lvalue: expr -> block -> int -> Prop := | esl_deref: forall r ty b ofs, eval_simple_rvalue r (Vptr b ofs) -> eval_simple_lvalue (Ederef r ty) b ofs - | esl_field_struct: forall r f ty b ofs id fList a delta, + | esl_field_struct: forall r f ty b ofs id co a delta, eval_simple_rvalue r (Vptr b ofs) -> - typeof r = Tstruct id fList a -> field_offset f fList = OK delta -> + typeof r = Tstruct id a -> ge.(genv_cenv)!id = Some co -> field_offset ge f (co_members co) = OK delta -> eval_simple_lvalue (Efield r f ty) b (Int.add ofs (Int.repr delta)) - | esl_field_union: forall r f ty b ofs id fList a, + | esl_field_union: forall r f ty b ofs id a, eval_simple_rvalue r (Vptr b ofs) -> - typeof r = Tunion id fList a -> + typeof r = Tunion id a -> eval_simple_lvalue (Efield r f ty) b ofs with eval_simple_rvalue: expr -> val -> Prop := @@ -116,16 +116,16 @@ with eval_simple_rvalue: expr -> val -> Prop := eval_simple_rvalue (Eunop op r1 ty) v | esr_binop: forall op r1 r2 ty v1 v2 v, eval_simple_rvalue r1 v1 -> eval_simple_rvalue r2 v2 -> - sem_binary_operation op v1 (typeof r1) v2 (typeof r2) m = Some v -> + sem_binary_operation ge op v1 (typeof r1) v2 (typeof r2) m = Some v -> eval_simple_rvalue (Ebinop op r1 r2 ty) v | esr_cast: forall ty r1 v1 v, eval_simple_rvalue r1 v1 -> sem_cast v1 (typeof r1) ty = Some v -> eval_simple_rvalue (Ecast r1 ty) v | esr_sizeof: forall ty1 ty, - eval_simple_rvalue (Esizeof ty1 ty) (Vint (Int.repr (sizeof ty1))) + eval_simple_rvalue (Esizeof ty1 ty) (Vint (Int.repr (sizeof ge ty1))) | esr_alignof: forall ty1 ty, - eval_simple_rvalue (Ealignof ty1 ty) (Vint (Int.repr (alignof ty1))) + eval_simple_rvalue (Ealignof ty1 ty) (Vint (Int.repr (alignof ge ty1))) | esr_seqand_true: forall r1 r2 ty v1 v2 v3, eval_simple_rvalue r1 v1 -> bool_val v1 (typeof r1) = Some true -> eval_simple_rvalue r2 v2 -> @@ -372,13 +372,13 @@ Lemma constval_rvalue: forall m a v, eval_simple_rvalue empty_env m a v -> forall v', - constval a = OK v' -> + 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 v', - constval a = OK v' -> + constval ge a = OK v' -> val_inject inj v' (Vptr b ofs). Proof. (* rvalue *) @@ -394,7 +394,7 @@ Proof. exploit sem_unary_operation_inject. eexact E. eauto. intros [v' [A B]]. congruence. (* binop *) - destruct (sem_binary_operation op x (typeof r1) x0 (typeof r2) Mem.empty) as [v1'|] eqn:E; inv EQ2. + destruct (sem_binary_operation ge op x (typeof r1) x0 (typeof r2) Mem.empty) as [v1'|] eqn:E; inv EQ2. exploit (sem_binary_operation_inj inj Mem.empty m). intros. rewrite mem_empty_not_valid_pointer in H3; discriminate. intros. rewrite mem_empty_not_weak_valid_pointer in H3; discriminate. @@ -443,8 +443,9 @@ Proof. (* deref *) eauto. (* field struct *) - rewrite H0 in CV. monadInv CV. exploit constval_rvalue; eauto. intro MV. inv MV. - simpl. replace x with delta by congruence. econstructor; eauto. + rewrite H0 in CV. monadInv CV. unfold lookup_composite in EQ; rewrite H1 in EQ; monadInv EQ. + exploit constval_rvalue; eauto. intro MV. inv MV. + simpl. replace x0 with delta by congruence. econstructor; eauto. rewrite ! Int.add_assoc. f_equal. apply Int.add_commut. simpl. auto. (* field union *) @@ -452,7 +453,7 @@ Proof. Qed. Lemma constval_simple: - forall a v, constval a = OK v -> simple a. + 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. @@ -466,7 +467,7 @@ Qed. Theorem constval_steps: forall f r m v v' ty m', star step ge (ExprState f r Kstop empty_env m) E0 (ExprState f (Eval v' ty) Kstop empty_env m') -> - constval r = OK v -> + constval ge r = OK v -> m' = m /\ ty = typeof r /\ val_inject inj v v'. Proof. intros. exploit eval_simple_steps; eauto. eapply constval_simple; eauto. @@ -479,7 +480,7 @@ Qed. Theorem transl_init_single_steps: forall ty a data f m v1 ty1 m' v chunk b ofs m'', - transl_init_single ty a = OK data -> + 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 = Some v -> access_mode ty = By_value chunk -> @@ -522,15 +523,14 @@ Qed. Lemma transl_init_single_size: forall ty a data, - transl_init_single ty a = OK data -> - Genv.init_data_size data = sizeof ty. + transl_init_single ge ty a = OK data -> + Genv.init_data_size data = sizeof ge ty. Proof. intros. monadInv H. destruct x0. - monadInv EQ2. - destruct ty; try discriminate. destruct i0; inv EQ2; auto. inv EQ2; auto. - inv EQ2; auto. - destruct ty; inv EQ2; auto. - destruct ty; try discriminate. destruct f0; inv EQ2; auto. @@ -539,7 +539,6 @@ Proof. - destruct ty; try discriminate. destruct i0; inv EQ2; auto. inv EQ2; auto. - inv EQ2; auto. Qed. Notation idlsize := Genv.init_data_list_size. @@ -561,30 +560,32 @@ Proof. Qed. Remark union_field_size: - forall f ty fl, field_type f fl = OK ty -> sizeof ty <= sizeof_union fl. + forall f ty fl, field_type f fl = OK ty -> sizeof ge ty <= sizeof_union ge fl. Proof. - induction fl; simpl; intros. + induction fl as [|[i t]]; simpl; intros. - inv H. - destruct (ident_eq f i). + inv H. xomega. + specialize (IHfl H). xomega. Qed. +Hypothesis ce_consistent: composite_env_consistent ge. + Lemma transl_init_size: forall i ty data, - transl_init ty i = OK data -> - idlsize data = sizeof ty + transl_init ge ty i = OK data -> + idlsize data = sizeof ge ty with transl_init_list_size: forall il, (forall ty sz data, - transl_init_array ty il sz = OK data -> - idlsize data = sizeof ty * sz) + transl_init_array ge ty il sz = OK data -> + idlsize data = sizeof ge ty * sz) /\ - (forall id ty fl pos data, - transl_init_struct id ty fl il pos = OK data -> - sizeof_struct fl pos <= sizeof ty -> - idlsize data + pos = sizeof ty). + (forall ty fl pos data, + transl_init_struct ge ty fl il pos = OK data -> + sizeof_struct ge pos fl <= sizeof ge ty -> + idlsize data + pos = sizeof ge ty). Proof. - induction i; intros. @@ -595,27 +596,35 @@ Proof. simpl. eapply (proj1 (transl_init_list_size il)); eauto. + (* struct *) simpl in H. destruct ty; try discriminate. + monadInv H. destruct (co_su x) eqn:?; try discriminate. replace (idlsize data) with (idlsize data + 0) by omega. eapply (proj2 (transl_init_list_size il)). eauto. -Local Opaque alignof. - simpl. apply align_le. apply alignof_pos. + unfold lookup_composite in EQ. simpl. destruct (ge.(genv_cenv)!i) as [co|] eqn:?; inv EQ. + exploit ce_consistent; eauto. intros (A & B & C). + rewrite C. unfold sizeof_composite. rewrite Heqs. apply align_le. + destruct (co_alignof_two_p x) as [n EQ]. rewrite EQ. apply two_power_nat_pos. + (* union *) simpl in H. destruct ty; try discriminate. - set (sz := sizeof (Tunion i0 f0 a)) in *. - monadInv H. rewrite idlsize_app. rewrite (IHi _ _ EQ1). - rewrite padding_size. omega. unfold sz. simpl. - apply Zle_trans with (sizeof_union f0). eapply union_field_size; eauto. - apply align_le. apply alignof_pos. + monadInv H. destruct (co_su x) eqn:?; try discriminate. + monadInv EQ0. + rewrite idlsize_app. rewrite (IHi _ _ EQ0). + unfold lookup_composite in EQ. simpl. destruct (ge.(genv_cenv)!i0) as [co|] eqn:?; inv EQ. + rewrite padding_size. omega. + apply Zle_trans with (sizeof_union ge (co_members x)). + eapply union_field_size; eauto. + exploit ce_consistent; eauto. intros (A & B & C). + rewrite C. unfold sizeof_composite. rewrite Heqs. apply align_le. + destruct (co_alignof_two_p x) as [n EQ]. rewrite EQ. apply two_power_nat_pos. - induction il. + (* base cases *) - simpl. intuition. + simpl. intuition auto. * (* arrays *) destruct (zeq sz 0). inv H. simpl; ring. destruct (zle 0 sz); inv H. simpl. rewrite Z.mul_comm. - assert (0 <= sizeof ty * sz). - { apply Zmult_gt_0_le_0_compat. omega. generalize (sizeof_pos ty); omega. } + assert (0 <= sizeof ge ty * sz). + { apply Zmult_gt_0_le_0_compat. omega. generalize (sizeof_pos ge ty); omega. } zify; omega. * (* structs *) destruct fl; inv H. @@ -629,12 +638,12 @@ Local Opaque alignof. rewrite (A _ _ _ EQ1). ring. * (* structs *) - intros. simpl in H. destruct fl; monadInv H. + intros. simpl in H. destruct fl as [|[i1 t1]]; monadInv H. rewrite ! idlsize_app. simpl in H0. rewrite padding_size. rewrite (transl_init_size _ _ _ EQ). - rewrite <- (B _ _ _ _ _ EQ1). omega. + rewrite <- (B _ _ _ _ EQ1). omega. auto. apply align_le. apply alignof_pos. Qed. @@ -642,11 +651,11 @@ Qed. Definition dummy_function := mkfunction Tvoid cc_default nil nil Sskip. -Fixpoint fields_of_struct (id: ident) (ty: type) (fl: fieldlist) (pos: Z) : list (Z * type) := +Fixpoint fields_of_struct (fl: members) (pos: Z) : list (Z * type) := match fl with - | Fnil => nil - | Fcons id1 ty1 fl' => - (align pos (alignof ty1), ty1) :: fields_of_struct id ty fl' (align pos (alignof ty1) + sizeof ty1) + | nil => nil + | (id1, ty1) :: fl' => + (align pos (alignof ge ty1), ty1) :: fields_of_struct fl' (align pos (alignof ge ty1) + sizeof ge ty1) end. Inductive exec_init: mem -> block -> Z -> type -> initializer -> mem -> Prop := @@ -660,13 +669,15 @@ Inductive exec_init: mem -> block -> Z -> type -> initializer -> mem -> Prop := | 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 fl a il m', - exec_init_list m b ofs (fields_of_struct id (Tstruct id fl a) fl 0) il m' -> - exec_init m b ofs (Tstruct id fl a) (Init_struct il) m' - | exec_init_union: forall m b ofs id fl a f i ty m', - field_type f fl = OK ty -> + | exec_init_struct: forall m b ofs id a il co 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', + 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 fl a) (Init_union f i) m' + exec_init m b ofs (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, @@ -674,7 +685,7 @@ with exec_init_array: mem -> block -> Z -> type -> Z -> initializer_list -> mem 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_array m' b (ofs + sizeof ty) ty (sz - 1) il 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 := @@ -718,15 +729,15 @@ Qed. Lemma transl_init_sound_gen: (forall m b ofs ty i m', exec_init m b ofs ty i m' -> - forall data, transl_init ty i = OK data -> + forall data, transl_init ge ty i = OK 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, transl_init_array ty il sz = OK data -> + forall data, transl_init_array ge ty il sz = OK 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 id ty fl data pos, - l = fields_of_struct id ty fl pos -> - transl_init_struct id ty fl il pos = OK data -> + forall ty fl data pos, + l = fields_of_struct fl pos -> + transl_init_struct ge ty fl il pos = OK data -> Genv.store_init_data_list ge m b (ofs + pos) data = Some m'). Proof. Local Opaque sizeof. @@ -737,9 +748,11 @@ Local Opaque sizeof. replace (Z.max 0 sz) with sz in H1. eauto. assert (sz >= 0) by (eapply exec_init_array_length; eauto). xomega. - (* struct *) + unfold lookup_composite in H3. rewrite H in H3. simpl in H3. rewrite H0 in H3. replace ofs with (ofs + 0) by omega. eauto. - (* union *) - rewrite H in H2. monadInv H2. inv EQ. + unfold lookup_composite in H4. rewrite H in H4. simpl in H4. rewrite H0 in H4. + monadInv H4. assert (x = ty) by congruence. subst x. eapply store_init_data_list_app. eauto. apply store_init_data_list_padding. @@ -753,15 +766,15 @@ Local Opaque sizeof. eauto. rewrite (transl_init_size _ _ _ EQ). eauto. - (* struct, empty *) - destruct fl; simpl in H; inv H. - inv H0. apply store_init_data_list_padding. + destruct fl as [|[i t]]; simpl in H0; inv H0. + apply store_init_data_list_padding. - (* struct, nonempty *) - destruct fl; simpl in H3; inv H3. - monadInv H4. + destruct fl as [|[i t]]; simpl in H4; monadInv H4. + simpl in H3; inv H3. eapply store_init_data_list_app. apply store_init_data_list_padding. rewrite padding_size. - replace (ofs + pos0 + (align pos0 (alignof t) - pos0)) - with (ofs + align pos0 (alignof t)) by omega. + replace (ofs + pos0 + (align pos0 (alignof ge t) - pos0)) + with (ofs + align pos0 (alignof ge t)) by omega. eapply store_init_data_list_app. eauto. rewrite (transl_init_size _ _ _ EQ). @@ -769,15 +782,18 @@ Local Opaque sizeof. apply align_le. apply alignof_pos. Qed. +End SOUNDNESS. + Theorem transl_init_sound: - forall m b ty i m' data, - exec_init m b 0 ty i m' -> - transl_init ty i = OK data -> - Genv.store_init_data_list ge m b 0 data = Some m'. + forall p m b ty i m' data, + exec_init (globalenv p) m b 0 ty i m' -> + transl_init (prog_comp_env p) ty i = OK data -> + Genv.store_init_data_list (globalenv p) m b 0 data = Some m'. Proof. - intros. eapply (proj1 transl_init_sound_gen); eauto. + intros. + set (ge := globalenv p) in *. + change (prog_comp_env p) with (genv_cenv ge) in H0. + destruct (transl_init_sound_gen ge) as (A & B & C). + eapply build_composite_env_consistent. apply prog_comp_env_eq. + eapply A; eauto. Qed. - -End SOUNDNESS. - - diff --git a/cfrontend/PrintClight.ml b/cfrontend/PrintClight.ml index c5a6e6e1..ebd06c54 100644 --- a/cfrontend/PrintClight.ml +++ b/cfrontend/PrintClight.ml @@ -45,6 +45,8 @@ let rec precedence = function | Econst_float _ -> (16, NA) | Econst_single _ -> (16, NA) | Econst_long _ -> (16, NA) + | Esizeof _ -> (15, RtoL) + | Ealignof _ -> (15, RtoL) | Eunop _ -> (15, RtoL) | Eaddrof _ -> (15, RtoL) | Ecast _ -> (14, RtoL) @@ -100,6 +102,10 @@ let rec expr p (prec, e) = expr (prec1, a1) (name_binop op) expr (prec2, a2) | Ecast(a1, ty) -> fprintf p "(%s) %a" (name_type ty) expr (prec', a1) + | Esizeof(ty, _) -> + fprintf p "sizeof(%s)" (name_type ty) + | Ealignof(ty, _) -> + fprintf p "__alignof__(%s)" (name_type ty) end; if prec' < prec then fprintf p ")@]" else fprintf p "@]" @@ -265,71 +271,10 @@ let print_globdef p (id, gd) = | Gfun f -> print_fundef p id f | Gvar v -> print_globvar p id v (* from PrintCsyntax *) -(* Collect struct and union types *) - -let rec collect_expr e = - collect_type (typeof e); - match e with - | Econst_int _ -> () - | Econst_float _ -> () - | Econst_single _ -> () - | Econst_long _ -> () - | Evar _ -> () - | Etempvar _ -> () - | Ederef(r, _) -> collect_expr r - | Efield(l, _, _) -> collect_expr l - | Eaddrof(l, _) -> collect_expr l - | Eunop(_, r, _) -> collect_expr r - | Ebinop(_, r1, r2, _) -> collect_expr r1; collect_expr r2 - | Ecast(r, _) -> collect_expr r - -let rec collect_exprlist = function - | [] -> () - | r1 :: rl -> collect_expr r1; collect_exprlist rl - -let rec collect_stmt = function - | Sskip -> () - | Sassign(e1, e2) -> collect_expr e1; collect_expr e2 - | Sset(id, e2) -> collect_expr e2 - | Scall(optid, e1, el) -> collect_expr e1; collect_exprlist el - | Sbuiltin(optid, ef, tyargs, el) -> collect_exprlist el - | Ssequence(s1, s2) -> collect_stmt s1; collect_stmt s2 - | Sifthenelse(e, s1, s2) -> collect_expr e; collect_stmt s1; collect_stmt s2 - | Sloop(s1, s2) -> collect_stmt s1; collect_stmt s2 - | Sbreak -> () - | Scontinue -> () - | Sswitch(e, cases) -> collect_expr e; collect_cases cases - | Sreturn None -> () - | Sreturn (Some e) -> collect_expr e - | Slabel(lbl, s) -> collect_stmt s - | Sgoto lbl -> () - -and collect_cases = function - | LSnil -> () - | LScons(lbl, s, rem) -> collect_stmt s; collect_cases rem - -let collect_function f = - collect_type f.fn_return; - List.iter (fun (id, ty) -> collect_type ty) f.fn_params; - List.iter (fun (id, ty) -> collect_type ty) f.fn_vars; - List.iter (fun (id, ty) -> collect_type ty) f.fn_temps; - collect_stmt f.fn_body - -let collect_globdef (id, gd) = - match gd with - | Gfun(External(_, args, res, _)) -> collect_type_list args; collect_type res - | Gfun(Internal f) -> collect_function f - | Gvar v -> collect_type v.gvar_info - -let collect_program p = - List.iter collect_globdef p.prog_defs - let print_program p prog = - struct_unions := StructUnion.empty; - collect_program prog; fprintf p "@["; - StructUnion.iter (declare_struct_or_union p) !struct_unions; - StructUnion.iter (print_struct_or_union p) !struct_unions; + List.iter (declare_composite p) prog.prog_types; + List.iter (define_composite p) prog.prog_types; List.iter (print_globdef p) prog.prog_defs; fprintf p "@]@." diff --git a/cfrontend/PrintCsyntax.ml b/cfrontend/PrintCsyntax.ml index e1b53af8..8a4d60a5 100644 --- a/cfrontend/PrintCsyntax.ml +++ b/cfrontend/PrintCsyntax.ml @@ -70,12 +70,6 @@ let name_longtype sg = | Signed -> "long long" | Unsigned -> "unsigned long long" -(* Collecting the names and fields of structs and unions *) - -module StructUnion = Map.Make(String) - -let struct_unions = ref StructUnion.empty - (* Declarator (identifier + type) *) let attributes a = @@ -132,12 +126,10 @@ let rec name_cdecl id ty = add_args true args; Buffer.add_char b ')'; name_cdecl (Buffer.contents b) res - | Tstruct(name, fld, a) -> - extern_atom name ^ attributes a ^ name_optid id - | Tunion(name, fld, a) -> - extern_atom name ^ attributes a ^ name_optid id - | Tcomp_ptr(name, a) -> - extern_atom name ^ " *" ^ attributes a ^ id + | Tstruct(name, a) -> + "struct " ^ extern_atom name ^ attributes a ^ name_optid id + | Tunion(name, a) -> + "union " ^ extern_atom name ^ attributes a ^ name_optid id (* Type *) @@ -466,7 +458,7 @@ let print_globvar p id v = fprintf p "@[%s = " (name_cdecl name2 v.gvar_info); begin match v.gvar_info, v.gvar_init with - | (Tint _ | Tlong _ | Tfloat _ | Tpointer _ | Tfunction _ | Tcomp_ptr _), + | (Tint _ | Tlong _ | Tfloat _ | Tpointer _ | Tfunction _), [i1] -> print_init p i1 | _, il -> @@ -482,119 +474,24 @@ let print_globdef p (id, gd) = | Gfun f -> print_fundef p id f | Gvar v -> print_globvar p id v -(* Collect struct and union types *) - -let rec collect_type = function - | Tvoid -> () - | Tint _ -> () - | Tfloat _ -> () - | Tlong _ -> () - | Tpointer(t, _) -> collect_type t - | Tarray(t, _, _) -> collect_type t - | Tfunction(args, res, _) -> collect_type_list args; collect_type res - | Tstruct(id, fld, _) | Tunion(id, fld, _) -> - let s = extern_atom id in - if not (StructUnion.mem s !struct_unions) then begin - struct_unions := StructUnion.add s fld !struct_unions; - collect_fields fld - end - | Tcomp_ptr _ -> () - -and collect_type_list = function - | Tnil -> () - | Tcons(hd, tl) -> collect_type hd; collect_type_list tl - -and collect_fields = function - | Fnil -> () - | Fcons(id, hd, tl) -> collect_type hd; collect_fields tl - -let rec collect_expr e = - collect_type (typeof e); - match e with - | Eloc _ -> assert false - | Evar _ -> () - | Ederef(r, _) -> collect_expr r - | Efield(l, _, _) -> collect_expr l - | Eval _ -> () - | Evalof(l, _) -> collect_expr l - | Eaddrof(l, _) -> collect_expr l - | Eunop(_, r, _) -> collect_expr r - | Ebinop(_, r1, r2, _) -> collect_expr r1; collect_expr r2 - | Ecast(r, _) -> collect_expr r - | Eseqand(r1, r2, _) -> collect_expr r1; collect_expr r2 - | Eseqor(r1, r2, _) -> collect_expr r1; collect_expr r2 - | Econdition(r1, r2, r3, _) -> - collect_expr r1; collect_expr r2; collect_expr r3 - | Esizeof(ty, _) -> collect_type ty - | Ealignof(ty, _) -> collect_type ty - | Eassign(l, r, _) -> collect_expr l; collect_expr r - | Eassignop(_, l, r, _, _) -> collect_expr l; collect_expr r - | Epostincr(_, l, _) -> collect_expr l - | Ecomma(r1, r2, _) -> collect_expr r1; collect_expr r2 - | Ecall(r1, rl, _) -> collect_expr r1; collect_exprlist rl - | Ebuiltin(_, _, rl, _) -> collect_exprlist rl - | Eparen _ -> assert false - -and collect_exprlist = function - | Enil -> () - | Econs(r1, rl) -> collect_expr r1; collect_exprlist rl - -let rec collect_stmt = function - | Sskip -> () - | Sdo e -> collect_expr e - | Ssequence(s1, s2) -> collect_stmt s1; collect_stmt s2 - | Sifthenelse(e, s1, s2) -> collect_expr e; collect_stmt s1; collect_stmt s2 - | Swhile(e, s) -> collect_expr e; collect_stmt s - | Sdowhile(e, s) -> collect_stmt s; collect_expr e - | Sfor(s_init, e, s_iter, s_body) -> - collect_stmt s_init; collect_expr e; - collect_stmt s_iter; collect_stmt s_body - | Sbreak -> () - | Scontinue -> () - | Sswitch(e, cases) -> collect_expr e; collect_cases cases - | Sreturn None -> () - | Sreturn (Some e) -> collect_expr e - | Slabel(lbl, s) -> collect_stmt s - | Sgoto lbl -> () - -and collect_cases = function - | LSnil -> () - | LScons(lbl, s, rem) -> collect_stmt s; collect_cases rem - -let collect_function f = - collect_type f.fn_return; - List.iter (fun (id, ty) -> collect_type ty) f.fn_params; - List.iter (fun (id, ty) -> collect_type ty) f.fn_vars; - collect_stmt f.fn_body - -let collect_globdef (id, gd) = - match gd with - | Gfun(External(_, args, res, _)) -> collect_type_list args; collect_type res - | Gfun(Internal f) -> collect_function f - | Gvar v -> collect_type v.gvar_info - -let collect_program p = - List.iter collect_globdef p.prog_defs - -let declare_struct_or_union p name fld = - fprintf p "%s;@ @ " name - -let print_struct_or_union p name fld = - fprintf p "@[%s {" name; - let rec print_fields = function - | Fnil -> () - | Fcons(id, ty, rem) -> - fprintf p "@ %s;" (name_cdecl (extern_atom id) ty); - print_fields rem in - print_fields fld; +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 define_composite p (Composite(id, su, m, a)) = + fprintf p "@[%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; fprintf p "@;<0 -2>};@]@ @ " let print_program p prog = - struct_unions := StructUnion.empty; - collect_program prog; fprintf p "@["; - StructUnion.iter (declare_struct_or_union p) !struct_unions; - StructUnion.iter (print_struct_or_union p) !struct_unions; + List.iter (declare_composite p) prog.prog_types; + List.iter (define_composite p) prog.prog_types; List.iter (print_globdef p) prog.prog_defs; fprintf p "@]@." diff --git a/cfrontend/SimplExpr.v b/cfrontend/SimplExpr.v index 089797f2..36fe07ae 100644 --- a/cfrontend/SimplExpr.v +++ b/cfrontend/SimplExpr.v @@ -540,5 +540,10 @@ Fixpoint transl_globdefs end. Definition transl_program (p: Csyntax.program) : res program := - do gl' <- transl_globdefs p.(prog_defs) (initial_generator tt); - OK (mkprogram gl' p.(prog_public) p.(prog_main)). + do gl' <- transl_globdefs (Csyntax.prog_defs p) (initial_generator tt); + OK {| prog_defs := gl'; + prog_public := Csyntax.prog_public p; + prog_main := Csyntax.prog_main p; + prog_types := Csyntax.prog_types p; + prog_comp_env := Csyntax.prog_comp_env p; + prog_comp_env_eq := Csyntax.prog_comp_env_eq p |}. diff --git a/cfrontend/SimplExprproof.v b/cfrontend/SimplExprproof.v index 6d8b5c86..74019061 100644 --- a/cfrontend/SimplExprproof.v +++ b/cfrontend/SimplExprproof.v @@ -37,11 +37,17 @@ Variable prog: Csyntax.program. Variable tprog: Clight.program. Hypothesis TRANSL: transl_program prog = OK tprog. -Let ge := Genv.globalenv prog. -Let tge := Genv.globalenv tprog. +Let ge := Csem.globalenv prog. +Let tge := Clight.globalenv tprog. (** Invariance properties. *) +Lemma comp_env_preserved: + Clight.genv_cenv tge = Csem.genv_cenv ge. +Proof. + monadInv TRANSL. unfold tge; rewrite <- H0; auto. +Qed. + Lemma symbols_preserved: forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s. Proof. @@ -87,7 +93,7 @@ Proof. - destruct (Genv.find_var_info tge b) as [v'|] eqn:V'; auto. exploit Genv.find_var_info_rev_match. eapply transl_program_spec; eauto. eassumption. simpl. destruct (plt b (Genv.genv_next (Genv.globalenv prog))); try tauto. - intros [v [A B]]. inv B. fold ge in A. congruence. + intros [v [A B]]. inv B. change (Genv.globalenv prog) with (Csem.genv_genv ge) in A. congruence. Qed. Lemma block_is_volatile_preserved: @@ -173,7 +179,7 @@ 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 ty m b ofs v m' + | 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' end. Proof. @@ -184,7 +190,8 @@ Proof. rewrite H0; rewrite H1. eapply volatile_store_preserved with (ge1 := ge); auto. exact symbols_preserved. exact public_preserved. exact block_is_volatile_preserved. (* By copy *) - rewrite H0. destruct (type_is_volatile ty); split; auto; eapply assign_loc_copy; eauto. + rewrite H0. rewrite <- comp_env_preserved in *. + destruct (type_is_volatile ty); split; auto; eapply assign_loc_copy; eauto. Qed. (** Evaluation of simple expressions and of their translation *) @@ -241,7 +248,7 @@ Opaque makeif. exploit H0; eauto. intros [A [B C]]. exploit H2; eauto. intros [D [E F]]. subst sl1 sl2; simpl. - assert (eval_expr tge e le m (Ebinop op a1 a2 ty) v). econstructor; eauto. congruence. + assert (eval_expr tge e le m (Ebinop op a1 a2 ty) v). econstructor; eauto. rewrite comp_env_preserved; congruence. destruct dst; auto. simpl; econstructor; eauto. (* cast *) exploit H0; eauto. intros [A [B C]]. @@ -249,11 +256,13 @@ Opaque makeif. assert (eval_expr tge e le m (Ecast a1 ty) v). econstructor; eauto. congruence. destruct dst; auto. simpl; econstructor; eauto. (* sizeof *) + rewrite <- comp_env_preserved. destruct dst. split; auto. split; auto. constructor. auto. exists (Esizeof ty1 ty). split. auto. split. auto. constructor. (* alignof *) + rewrite <- comp_env_preserved. destruct dst. split; auto. split; auto. constructor. auto. @@ -267,9 +276,11 @@ Opaque makeif. exploit H0; eauto. intros [A [B C]]. subst sl1. split; auto. split; auto. constructor; auto. (* field struct *) + rewrite <- comp_env_preserved in *. exploit H0; eauto. intros [A [B C]]. subst sl1. split; auto. split; auto. rewrite B in H1. eapply eval_Efield_struct; eauto. (* field union *) + rewrite <- comp_env_preserved in *. exploit H0; eauto. intros [A [B C]]. subst sl1. split; auto. split; auto. rewrite B in H1. eapply eval_Efield_union; eauto. Qed. @@ -1659,7 +1670,7 @@ Proof. left. eapply star_plus_trans. rewrite app_ass. rewrite Kseqlist_app. eexact EXEC. eapply plus_two. simpl. econstructor. eapply step_make_assign; eauto. econstructor. eexact EV3. eexact EV2. - rewrite TY3; rewrite <- TY1; rewrite <- TY2; eauto. + 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. @@ -1680,7 +1691,7 @@ Proof. eapply star_plus_trans. eexact EXEC. simpl. eapply plus_four. econstructor. econstructor. econstructor. eexact EV3. eexact EV2. - rewrite TY3; rewrite <- TY1; rewrite <- TY2. eauto. + rewrite TY3; rewrite <- TY1; rewrite <- TY2; rewrite comp_env_preserved; eauto. econstructor. eapply step_make_assign; eauto. constructor. apply PTree.gss. reflexivity. traceEq. @@ -1729,8 +1740,8 @@ Proof. eapply plus_two. simpl. constructor. eapply step_make_assign; eauto. unfold transl_incrdecr. destruct id; simpl in H2. - econstructor. eauto. constructor. simpl. rewrite TY3; rewrite <- TY1. eauto. - econstructor. eauto. constructor. simpl. rewrite TY3; rewrite <- TY1. eauto. + econstructor. eauto. constructor. rewrite TY3; rewrite <- TY1; rewrite comp_env_preserved. simpl; eauto. + econstructor. eauto. constructor. rewrite TY3; rewrite <- TY1; rewrite comp_env_preserved. simpl; eauto. destruct id; auto. reflexivity. traceEq. econstructor. auto. change sl2 with (nil ++ sl2). apply S. @@ -1748,8 +1759,10 @@ Proof. constructor. eapply step_make_assign; eauto. unfold transl_incrdecr. destruct id; simpl in H2. - econstructor. constructor. apply PTree.gss. constructor. simpl. eauto. - econstructor. constructor. apply PTree.gss. constructor. simpl. eauto. + econstructor. constructor. apply PTree.gss. constructor. + rewrite comp_env_preserved; simpl; eauto. + econstructor. constructor. apply PTree.gss. constructor. + rewrite comp_env_preserved; simpl; eauto. destruct id; auto. traceEq. econstructor. auto. apply S. @@ -1900,24 +1913,31 @@ Qed. Lemma alloc_variables_preserved: forall e m params e' m', - Csem.alloc_variables e m params e' m' -> - alloc_variables e m params e' m'. + Csem.alloc_variables ge e m params e' m' -> + alloc_variables tge e m params e' m'. Proof. - induction 1; econstructor; eauto. + induction 1; econstructor; eauto. rewrite comp_env_preserved; auto. Qed. Lemma bind_parameters_preserved: forall e m params args m', Csem.bind_parameters ge e m params args m' -> - bind_parameters e m params args m'. + bind_parameters tge e m params args m'. Proof. - induction 1; econstructor; eauto. - inv H0. - eapply assign_loc_value; eauto. - inv H4. eapply assign_loc_value; eauto. - eapply assign_loc_copy; eauto. + induction 1; econstructor; eauto. inv H0. +- eapply assign_loc_value; eauto. +- inv H4. eapply assign_loc_value; eauto. +- rewrite <- comp_env_preserved in *. eapply assign_loc_copy; eauto. Qed. +Lemma blocks_of_env_preserved: + forall e, blocks_of_env tge e = Csem.blocks_of_env ge e. +Proof. + intros; unfold blocks_of_env, Csem.blocks_of_env. + unfold block_of_binding, Csem.block_of_binding. + rewrite comp_env_preserved. auto. +Qed. + Lemma sstep_simulation: forall S1 t S2, Csem.sstep ge S1 t S2 -> forall S1' (MS: match_states S1 S1'), @@ -2090,9 +2110,10 @@ Proof. left. apply plus_one. constructor. econstructor; eauto. constructor; auto. + (* return none *) inv H7. econstructor; split. - left. apply plus_one. econstructor; eauto. + 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. @@ -2102,14 +2123,14 @@ Proof. inv H9. exploit tr_top_val_for_val_inv; eauto. intros [A [B C]]. subst. econstructor; split. left. eapply plus_two. constructor. econstructor. eauto. - erewrite function_return_preserved; eauto. + 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; split. - left. apply plus_one. apply step_skip_call; eauto. + left. apply plus_one. apply step_skip_call; eauto. rewrite blocks_of_env_preserved; eauto. constructor. auto. (* switch *) @@ -2198,14 +2219,14 @@ Lemma transl_initial_states: Csem.initial_state prog S -> exists S', Clight.initial_state tprog S' /\ match_states S S'. Proof. - intros. inv H. + intros. inv H. generalize TRANSL; intros TR; monadInv TR. rewrite H4. exploit transl_program_spec; eauto. intros MP. exploit function_ptr_translated; eauto. intros [tf [FIND TR]]. econstructor; split. econstructor. - exploit Genv.init_mem_match; eauto. - simpl. fold tge. rewrite symbols_preserved. - destruct MP as (A & B & C). rewrite B; eexact H1. + exploit Genv.init_mem_match; eauto. + change (Genv.globalenv tprog) with (genv_genv tge). rewrite symbols_preserved. + rewrite <- H4; simpl; eauto. eexact FIND. rewrite <- H3. apply type_of_fundef_preserved. auto. constructor. auto. constructor. diff --git a/cfrontend/SimplExprspec.v b/cfrontend/SimplExprspec.v index 83ddd1f4..9f9fb450 100644 --- a/cfrontend/SimplExprspec.v +++ b/cfrontend/SimplExprspec.v @@ -1144,10 +1144,10 @@ Qed. Theorem transl_program_spec: forall p tp, transl_program p = OK tp -> - match_program tr_fundef (fun v1 v2 => v1 = v2) nil p.(prog_main) p tp. + match_program tr_fundef (fun v1 v2 => v1 = v2) nil (Csyntax.prog_main p) p tp. Proof. unfold transl_program; intros. - destruct (transl_globdefs (prog_defs p) (initial_generator tt)) eqn:E; simpl in H; inv H. + destruct (transl_globdefs (Csyntax.prog_defs p) (initial_generator tt)) eqn:E; simpl in H; inv H. split; auto. exists l; split. eapply transl_globdefs_spec; eauto. rewrite <- app_nil_end; auto. Qed. diff --git a/cfrontend/SimplLocals.v b/cfrontend/SimplLocals.v index 9c529280..52ee8377 100644 --- a/cfrontend/SimplLocals.v +++ b/cfrontend/SimplLocals.v @@ -48,8 +48,8 @@ Definition make_cast (a: expr) (tto: type) : expr := | cast_case_f2f => a | cast_case_s2s => a | cast_case_l2l => a - | cast_case_struct _ _ _ _ => a - | cast_case_union _ _ _ _ => a + | cast_case_struct _ _ => a + | cast_case_union _ _ => a | cast_case_void => a | _ => Ecast a tto end. @@ -70,6 +70,8 @@ Fixpoint simpl_expr (cenv: compilenv) (a: expr) : expr := | Ebinop op a1 a2 ty => Ebinop op (simpl_expr cenv a1) (simpl_expr cenv a2) ty | Ecast a1 ty => Ecast (simpl_expr cenv a1) ty | Efield a1 fld ty => Efield (simpl_expr cenv a1) fld ty + | Esizeof _ _ => a + | Ealignof _ _ => a end. Definition simpl_exprlist (cenv: compilenv) (al: list expr) : list expr := @@ -170,6 +172,8 @@ Fixpoint addr_taken_expr (a: expr): VSet.t := | Ebinop op a1 a2 ty => VSet.union (addr_taken_expr a1) (addr_taken_expr a2) | Ecast a1 ty => addr_taken_expr a1 | Efield a1 fld ty => addr_taken_expr a1 + | Esizeof _ _ => VSet.empty + | Ealignof _ _ => VSet.empty end. Fixpoint addr_taken_exprlist (l: list expr) : VSet.t := @@ -247,6 +251,10 @@ Definition transf_fundef (fd: fundef) : res fundef := end. Definition transf_program (p: program) : res program := - AST.transform_partial_program transf_fundef p. - - + do p1 <- AST.transform_partial_program transf_fundef p; + OK {| prog_defs := AST.prog_defs p1; + prog_public := AST.prog_public p1; + prog_main := AST.prog_main p1; + prog_types := prog_types p; + prog_comp_env := prog_comp_env p; + prog_comp_env_eq := prog_comp_env_eq p |}. diff --git a/cfrontend/SimplLocalsproof.v b/cfrontend/SimplLocalsproof.v index 15d0af06..5cf59d94 100644 --- a/cfrontend/SimplLocalsproof.v +++ b/cfrontend/SimplLocalsproof.v @@ -39,25 +39,37 @@ Section PRESERVATION. Variable prog: program. Variable tprog: program. Hypothesis TRANSF: transf_program prog = OK tprog. -Let ge := Genv.globalenv prog. -Let tge := Genv.globalenv tprog. +Let ge := globalenv prog. +Let tge := globalenv tprog. + +Lemma comp_env_preserved: + genv_cenv tge = genv_cenv ge. +Proof. + monadInv TRANSF. unfold tge; rewrite <- H0; auto. +Qed. + +Lemma transf_programs: + AST.transform_partial_program transf_fundef (program_of_program prog) = OK (program_of_program tprog). +Proof. + monadInv TRANSF. rewrite EQ. destruct x; reflexivity. +Qed. Lemma symbols_preserved: forall (s: ident), Genv.find_symbol tge s = Genv.find_symbol ge s. Proof. - exact (Genv.find_symbol_transf_partial _ _ TRANSF). + exact (Genv.find_symbol_transf_partial _ _ transf_programs). Qed. Lemma public_preserved: forall (s: ident), Genv.public_symbol tge s = Genv.public_symbol ge s. Proof. - exact (Genv.public_symbol_transf_partial _ _ TRANSF). + exact (Genv.public_symbol_transf_partial _ _ transf_programs). Qed. Lemma varinfo_preserved: forall b, Genv.find_var_info tge b = Genv.find_var_info ge b. Proof. - exact (Genv.find_var_info_transf_partial _ _ TRANSF). + exact (Genv.find_var_info_transf_partial _ _ transf_programs). Qed. Lemma functions_translated: @@ -65,7 +77,7 @@ Lemma functions_translated: Genv.find_funct ge v = Some f -> exists tf, Genv.find_funct tge v = Some tf /\ transf_fundef f = OK tf. Proof. - exact (Genv.find_funct_transf_partial _ _ TRANSF). + exact (Genv.find_funct_transf_partial _ _ transf_programs). Qed. Lemma function_ptr_translated: @@ -73,7 +85,7 @@ Lemma function_ptr_translated: Genv.find_funct_ptr ge b = Some f -> exists tf, Genv.find_funct_ptr tge b = Some tf /\ transf_fundef f = OK tf. Proof. - exact (Genv.find_funct_ptr_transf_partial _ _ TRANSF). + exact (Genv.find_funct_ptr_transf_partial _ _ transf_programs). Qed. Lemma type_of_fundef_preserved: @@ -207,14 +219,10 @@ Inductive val_casted: val -> type -> Prop := val_casted (Vint n) (Tpointer ty attr) | val_casted_ptr_int: forall b ofs si attr, val_casted (Vptr b ofs) (Tint I32 si attr) - | val_casted_ptr_cptr: forall b ofs id attr, - val_casted (Vptr b ofs) (Tcomp_ptr id attr) - | val_casted_int_cptr: forall n id attr, - val_casted (Vint n) (Tcomp_ptr id attr) - | val_casted_struct: forall id fld attr b ofs, - val_casted (Vptr b ofs) (Tstruct id fld attr) - | val_casted_union: forall id fld attr b ofs, - val_casted (Vptr b ofs) (Tunion id fld attr) + | val_casted_struct: forall id attr b ofs, + val_casted (Vptr b ofs) (Tstruct id attr) + | val_casted_union: forall id attr b ofs, + val_casted (Vptr b ofs) (Tunion id attr) | val_casted_void: forall v, val_casted v Tvoid. @@ -254,8 +262,6 @@ Proof. constructor. constructor; auto. constructor. - constructor; auto. - constructor. constructor. simpl. destruct (Int.eq i0 Int.zero); auto. constructor. simpl. destruct (Int64.eq i Int64.zero); auto. constructor. simpl. destruct (Float32.cmp Ceq f Float32.zero); auto. @@ -266,8 +272,6 @@ Proof. constructor; auto. constructor. simpl. destruct (Int.eq i Int.zero); auto. constructor; auto. - constructor. simpl. destruct (Int.eq i0 Int.zero); auto. - constructor; auto. (* long *) destruct ty; try (destruct f); try discriminate. destruct v; inv H. constructor. @@ -277,7 +281,6 @@ Proof. destruct v; inv H. constructor. destruct v; inv H. constructor. destruct v; inv H. constructor. - destruct v; inv H. constructor. (* float *) destruct f; destruct ty; simpl in H; try (destruct f); try discriminate; destruct v; inv H; constructor. (* pointer *) @@ -287,12 +290,10 @@ Proof. discriminate. (* structs *) destruct ty; try discriminate; destruct v; try discriminate. - destruct (ident_eq i0 i && fieldlist_eq f0 f); inv H; constructor. + destruct (ident_eq i0 i); inv H; constructor. (* unions *) destruct ty; try discriminate; destruct v; try discriminate. - destruct (ident_eq i0 i && fieldlist_eq f0 f); inv H; constructor. -(* comp_ptr *) - destruct ty; simpl in H; try discriminate; destruct v; inv H; constructor. + destruct (ident_eq i0 i); inv H; constructor. Qed. Lemma val_casted_load_result: @@ -316,8 +317,6 @@ Proof. discriminate. discriminate. discriminate. - discriminate. - discriminate. Qed. Lemma cast_val_casted: @@ -374,9 +373,9 @@ Proof. destruct v1; inv H0; auto. destruct v1; inv H0; auto. destruct v1; try discriminate. - destruct (ident_eq id1 id2 && fieldlist_eq fld1 fld2); inv H0; auto. + destruct (ident_eq id1 id2); inv H0; auto. destruct v1; try discriminate. - destruct (ident_eq id1 id2 && fieldlist_eq fld1 fld2); inv H0; auto. + destruct (ident_eq id1 id2); inv H0; auto. inv H0; auto. Qed. @@ -388,7 +387,7 @@ Lemma match_envs_assign_lifted: e!id = Some(b, ty) -> val_casted v ty -> val_inject f v tv -> - assign_loc ty m b Int.zero v m' -> + assign_loc ge ty m b Int.zero 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. @@ -582,8 +581,8 @@ Hint Resolve compat_cenv_union_l compat_cenv_union_r compat_cenv_empty: compat. (** Allocation and initialization of parameters *) Lemma alloc_variables_nextblock: - forall e m vars e' m', - alloc_variables e m vars e' m' -> Ple (Mem.nextblock m) (Mem.nextblock m'). + forall ge e m vars e' m', + alloc_variables ge e m vars e' m' -> Ple (Mem.nextblock m) (Mem.nextblock m'). Proof. induction 1. apply Ple_refl. @@ -591,8 +590,8 @@ Proof. Qed. Lemma alloc_variables_range: - forall id b ty e m vars e' m', - alloc_variables e m vars e' m' -> + forall ge id b ty e m vars e' m', + alloc_variables ge e m vars e' m' -> e'!id = Some(b, ty) -> e!id = Some(b, ty) \/ Ple (Mem.nextblock m) b /\ Plt b (Mem.nextblock m'). Proof. induction 1; intros. @@ -600,15 +599,15 @@ Proof. exploit IHalloc_variables; eauto. rewrite PTree.gsspec. intros [A|A]. destruct (peq id id0). inv A. right. exploit Mem.alloc_result; eauto. exploit Mem.nextblock_alloc; eauto. - generalize (alloc_variables_nextblock _ _ _ _ _ H0). intros A B C. + generalize (alloc_variables_nextblock _ _ _ _ _ _ H0). intros A B C. subst b. split. apply Ple_refl. eapply Plt_le_trans; eauto. rewrite B. apply Plt_succ. auto. right. exploit Mem.nextblock_alloc; eauto. intros B. rewrite B in A. xomega. Qed. Lemma alloc_variables_injective: - forall id1 b1 ty1 id2 b2 ty2 e m vars e' m', - alloc_variables e m vars e' m' -> + forall ge id1 b1 ty1 id2 b2 ty2 e m vars e' m', + alloc_variables ge e m vars e' m' -> (e!id1 = Some(b1, ty1) -> e!id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2) -> (forall id b ty, e!id = Some(b, ty) -> Plt b (Mem.nextblock m)) -> (e'!id1 = Some(b1, ty1) -> e'!id2 = Some(b2, ty2) -> id1 <> id2 -> b1 <> b2). @@ -629,12 +628,12 @@ Qed. Lemma match_alloc_variables: forall cenv e m vars e' m', - alloc_variables e m vars e' m' -> + alloc_variables ge e m vars e' m' -> forall j tm te, list_norepet (var_names vars) -> Mem.inject j m tm -> exists j', exists te', exists tm', - alloc_variables te tm (remove_lifted cenv vars) te' tm' + alloc_variables tge te tm (remove_lifted cenv vars) te' tm' /\ Mem.inject j' m' tm' /\ inject_incr j j' /\ (forall b, Mem.valid_block m b -> j' b = j b) @@ -698,7 +697,7 @@ Proof. exploit IHalloc_variables; eauto. instantiate (1 := PTree.set id (tb1, ty) te). intros [j' [te' [tm' [J [K [L [M [N [Q [O P]]]]]]]]]]. exists j'; exists te'; exists tm'. - split. simpl. econstructor; eauto. + split. simpl. econstructor; eauto. rewrite comp_env_preserved; auto. split. auto. split. eapply inject_incr_trans; eauto. split. intros. transitivity (j1 b). apply M. eapply Mem.valid_block_alloc; eauto. @@ -737,7 +736,7 @@ Qed. Lemma alloc_variables_load: forall e m vars e' m', - alloc_variables e m vars e' m' -> + alloc_variables ge e m vars e' m' -> forall chunk b ofs v, Mem.load chunk m b ofs = Some v -> Mem.load chunk m' b ofs = Some v. @@ -749,10 +748,10 @@ Qed. Lemma sizeof_by_value: forall ty chunk, - access_mode ty = By_value chunk -> size_chunk chunk <= sizeof ty. + access_mode ty = By_value chunk -> size_chunk chunk <= sizeof ge ty. Proof. unfold access_mode; intros. - assert (size_chunk chunk = sizeof ty). + assert (size_chunk chunk = sizeof ge ty). { destruct ty; try destruct i; try destruct s; try destruct f; inv H; auto. } @@ -765,7 +764,7 @@ Definition env_initial_value (e: env) (m: mem) := Lemma alloc_variables_initial_value: forall e m vars e' m', - alloc_variables e m vars e' m' -> + alloc_variables ge e m vars e' m' -> env_initial_value e m -> env_initial_value e' m'. Proof. @@ -900,14 +899,14 @@ Qed. Theorem match_envs_alloc_variables: forall cenv m vars e m' temps j tm, - alloc_variables empty_env m vars e m' -> + alloc_variables ge empty_env m vars e m' -> list_norepet (var_names vars) -> Mem.inject j m tm -> (forall id ty, In (id, ty) vars -> VSet.mem id cenv = true -> exists chunk, access_mode ty = By_value chunk) -> (forall id, VSet.mem id cenv = true -> In id (var_names vars)) -> exists j', exists te, exists tm', - alloc_variables empty_env tm (remove_lifted cenv vars) te tm' + alloc_variables tge empty_env tm (remove_lifted cenv vars) te tm' /\ match_envs j' cenv e (create_undef_temps temps) m' (Mem.nextblock m) (Mem.nextblock m') te (create_undef_temps (add_lifted cenv vars temps)) (Mem.nextblock tm) (Mem.nextblock tm') /\ Mem.inject j' m' tm' @@ -996,12 +995,12 @@ Qed. Lemma assign_loc_inject: forall f ty m loc ofs v m' tm loc' ofs' v', - assign_loc ty m loc ofs v m' -> + assign_loc ge ty m loc ofs 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 ty tm loc' ofs' v' tm' + assign_loc tge ty tm loc' ofs' 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). @@ -1018,10 +1017,11 @@ Proof. rename b' into bsrc. rename ofs'0 into osrc. rename loc into bdst. rename ofs into odst. rename loc' into bdst'. rename b2 into bsrc'. - destruct (zeq (sizeof ty) 0). + rewrite <- comp_env_preserved in *. + destruct (zeq (sizeof tge ty) 0). + (* special case size = 0 *) assert (bytes = nil). - { exploit (Mem.loadbytes_empty m bsrc (Int.unsigned osrc) (sizeof ty)). + { exploit (Mem.loadbytes_empty m bsrc (Int.unsigned osrc) (sizeof tge ty)). omega. congruence. } subst. destruct (Mem.range_perm_storebytes tm bdst' (Int.unsigned (Int.add odst (Int.repr delta))) nil) @@ -1038,12 +1038,12 @@ Proof. left. congruence. + (* general case size > 0 *) exploit Mem.loadbytes_length; eauto. intros LEN. - assert (SZPOS: sizeof ty > 0). - { generalize (sizeof_pos ty); omega. } - assert (RPSRC: Mem.range_perm m bsrc (Int.unsigned osrc) (Int.unsigned osrc + sizeof ty) Cur Nonempty). + assert (SZPOS: sizeof tge ty > 0). + { generalize (sizeof_pos tge ty); omega. } + assert (RPSRC: Mem.range_perm m bsrc (Int.unsigned osrc) (Int.unsigned osrc + sizeof tge ty) Cur Nonempty). eapply Mem.range_perm_implies. eapply Mem.loadbytes_range_perm; eauto. auto with mem. - assert (RPDST: Mem.range_perm m bdst (Int.unsigned odst) (Int.unsigned odst + sizeof ty) Cur Nonempty). - replace (sizeof ty) with (Z_of_nat (length bytes)). + assert (RPDST: Mem.range_perm m bdst (Int.unsigned odst) (Int.unsigned odst + sizeof tge ty) Cur Nonempty). + replace (sizeof tge ty) with (Z_of_nat (length bytes)). eapply Mem.range_perm_implies. eapply Mem.storebytes_range_perm; eauto. auto with mem. rewrite LEN. apply nat_of_Z_eq. omega. assert (PSRC: Mem.perm m bsrc (Int.unsigned osrc) Cur Nonempty). @@ -1084,8 +1084,8 @@ Proof. Qed. Lemma assign_loc_nextblock: - forall ty m b ofs v m', - assign_loc ty m b ofs v m' -> Mem.nextblock m' = Mem.nextblock m. + forall ge ty m b ofs v m', + assign_loc ge ty m b ofs v m' -> Mem.nextblock m' = Mem.nextblock m. Proof. induction 1. simpl in H0. eapply Mem.nextblock_store; eauto. @@ -1094,7 +1094,7 @@ Qed. Theorem store_params_correct: forall j f k cenv le lo hi te tlo thi e m params args m', - bind_parameters e m params args m' -> + bind_parameters ge e m params args m' -> forall s tm tle1 tle2 targs, list_norepet (var_names params) -> list_forall2 val_casted args (map snd params) -> @@ -1105,7 +1105,7 @@ Theorem store_params_correct: (forall id, In id (var_names params) -> le!id = None) -> exists tle, exists tm', star step2 tge (State f (store_params cenv params s) k te tle tm) - E0 (State f s k te tle tm') + E0 (State f s k te tle tm') /\ bind_parameter_temps params targs tle2 = Some tle /\ Mem.inject j m' tm' /\ match_envs j cenv e le m' lo hi te tle tlo thi @@ -1157,12 +1157,12 @@ Proof. 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: - forall e m params args m', - bind_parameters e m params args m' -> Mem.nextblock m' = Mem.nextblock m. + forall ge e m params args m', + bind_parameters ge e m params args m' -> Mem.nextblock m' = Mem.nextblock m. Proof. induction 1. auto. @@ -1170,10 +1170,10 @@ Proof. Qed. Lemma bind_parameters_load: - forall e chunk b ofs, + forall ge e chunk b ofs, (forall id b' ty, e!id = Some(b', ty) -> b <> b') -> forall m params args m', - bind_parameters e m params args m' -> + bind_parameters ge e m params args m' -> Mem.load chunk m' b ofs = Mem.load chunk m b ofs. Proof. induction 2. @@ -1188,16 +1188,16 @@ Qed. (** Freeing of local variables *) Lemma free_blocks_of_env_perm_1: - forall m e m' id b ty ofs k p, - Mem.free_list m (blocks_of_env e) = Some m' -> + forall ce m e m' id b ty ofs k p, + Mem.free_list m (blocks_of_env ce e) = Some m' -> e!id = Some(b, ty) -> Mem.perm m' b ofs k p -> - 0 <= ofs < sizeof ty -> + 0 <= ofs < sizeof ce ty -> False. Proof. intros. exploit Mem.perm_free_list; eauto. intros [A B]. - apply B with 0 (sizeof ty); auto. - unfold blocks_of_env. change (b, 0, sizeof ty) with (block_of_binding (id, (b, ty))). + apply B with 0 (sizeof ce ty); auto. + unfold blocks_of_env. change (b, 0, sizeof ce ty) with (block_of_binding ce (id, (b, ty))). apply in_map. apply PTree.elements_correct. auto. Qed. @@ -1216,13 +1216,13 @@ Proof. Qed. Lemma free_blocks_of_env_perm_2: - forall m e m' id b ty, - Mem.free_list m (blocks_of_env e) = Some m' -> + forall ce m e m' id b ty, + Mem.free_list m (blocks_of_env ce e) = Some m' -> e!id = Some(b, ty) -> - Mem.range_perm m b 0 (sizeof ty) Cur Freeable. + Mem.range_perm m b 0 (sizeof ce ty) Cur Freeable. Proof. intros. eapply free_list_perm'; eauto. - unfold blocks_of_env. change (b, 0, sizeof ty) with (block_of_binding (id, (b, ty))). + unfold blocks_of_env. change (b, 0, sizeof ce ty) with (block_of_binding ce (id, (b, ty))). apply in_map. apply PTree.elements_correct. auto. Qed. @@ -1252,15 +1252,15 @@ Proof. Qed. Lemma blocks_of_env_no_overlap: - forall j cenv e le m lo hi te tle tlo thi tm, + forall (ge: genv) j cenv e le m lo hi te tle tlo thi tm, match_envs j cenv e le m lo hi te tle tlo thi -> Mem.inject j m tm -> (forall id b ty, - e!id = Some(b, ty) -> Mem.range_perm m b 0 (sizeof ty) Cur Freeable) -> + e!id = Some(b, ty) -> Mem.range_perm m b 0 (sizeof ge ty) Cur Freeable) -> forall l, list_norepet (List.map fst l) -> (forall id bty, In (id, bty) l -> te!id = Some bty) -> - freelist_no_overlap (List.map block_of_binding l). + freelist_no_overlap (List.map (block_of_binding ge) l). Proof. intros until tm; intros ME MINJ PERMS. induction l; simpl; intros. - auto. @@ -1272,8 +1272,8 @@ Proof. assert (TE': te!id' = Some(b', ty')) by eauto. exploit me_mapped. eauto. eexact TE. intros [b0 [INJ E]]. exploit me_mapped. eauto. eexact TE'. intros [b0' [INJ' E']]. - destruct (zle (sizeof ty) 0); auto. - destruct (zle (sizeof ty') 0); auto. + destruct (zle (sizeof ge0 ty) 0); auto. + destruct (zle (sizeof ge0 ty') 0); auto. assert (b0 <> b0'). { eapply me_inj; eauto. red; intros; subst; elim H3. change id' with (fst (id', (b', ty'))). apply List.in_map; auto. } @@ -1302,26 +1302,34 @@ Proof. eapply Mem.free_right_inject; eauto. Qed. +Lemma blocks_of_env_translated: + forall e, blocks_of_env tge e = blocks_of_env ge e. +Proof. + intros. unfold blocks_of_env, block_of_binding. + rewrite comp_env_preserved; auto. +Qed. + Theorem match_envs_free_blocks: forall j cenv e le m lo hi te tle tlo thi m' tm, match_envs j cenv e le m lo hi te tle tlo thi -> Mem.inject j m tm -> - Mem.free_list m (blocks_of_env e) = Some m' -> + Mem.free_list m (blocks_of_env ge e) = Some m' -> exists tm', - Mem.free_list tm (blocks_of_env te) = Some tm' + Mem.free_list tm (blocks_of_env tge te) = Some tm' /\ Mem.inject j m' tm'. Proof. intros. - assert (X: exists tm', Mem.free_list tm (blocks_of_env te) = Some tm'). +Local Opaque ge tge. + assert (X: exists tm', Mem.free_list tm (blocks_of_env tge te) = Some tm'). { - apply can_free_list. + rewrite blocks_of_env_translated. apply can_free_list. - (* permissions *) intros. unfold blocks_of_env in H2. exploit list_in_map_inv; eauto. intros [[id [b' ty]] [EQ IN]]. - simpl in EQ; inv EQ. + unfold block_of_binding in EQ; inv EQ. exploit me_mapped; eauto. eapply PTree.elements_complete; eauto. intros [b [A B]]. - change 0 with (0 + 0). replace (sizeof ty) with (sizeof ty + 0) by omega. + change 0 with (0 + 0). replace (sizeof ge ty) with (sizeof ge ty + 0) by omega. eapply Mem.range_perm_inject; eauto. eapply free_blocks_of_env_perm_2; eauto. - (* no overlap *) @@ -1335,10 +1343,10 @@ Proof. eapply free_list_right_inject; eauto. eapply Mem.free_list_left_inject; eauto. intros. unfold blocks_of_env in H3. exploit list_in_map_inv; eauto. - intros [[id [b' ty]] [EQ IN]]. simpl in EQ. inv EQ. + intros [[id [b' ty]] [EQ IN]]. unfold block_of_binding in EQ. inv EQ. exploit me_flat; eauto. apply PTree.elements_complete; eauto. intros [P Q]. subst delta. eapply free_blocks_of_env_perm_1 with (m := m); eauto. - omega. + rewrite <- comp_env_preserved. omega. Qed. (** Matching global environments *) @@ -1434,11 +1442,16 @@ Proof. exploit eval_simpl_expr. eexact H. eauto with compat. intros [tv1 [A B]]. exploit eval_simpl_expr. eexact H0. eauto with compat. intros [tv2 [C D]]. exploit sem_binary_operation_inject; eauto. intros [tv [E F]]. - exists tv; split; auto. econstructor; eauto. repeat rewrite typeof_simpl_expr; auto. + exists tv; split; auto. econstructor; eauto. + repeat rewrite typeof_simpl_expr; rewrite comp_env_preserved; auto. (* cast *) exploit eval_simpl_expr; eauto. intros [tv1 [A B]]. exploit sem_cast_inject; eauto. intros [tv2 [C D]]. - exists tv2; split; auto. econstructor. eauto. rewrite typeof_simpl_expr; auto. + exists tv2; split; auto. econstructor. eauto. rewrite typeof_simpl_expr; auto. +(* sizeof *) + econstructor; split. constructor. rewrite comp_env_preserved; auto. +(* alignof *) + econstructor; split. constructor. rewrite comp_env_preserved; auto. (* rval *) assert (EITHER: (exists id, exists ty, a = Evar id ty /\ VSet.mem id cenv = true) \/ (match a with Evar id _ => VSet.mem id cenv = false | _ => True end)). @@ -1481,12 +1494,14 @@ Proof. inversion B. subst. econstructor; econstructor; split; eauto. econstructor; eauto. (* field struct *) + rewrite <- comp_env_preserved in *. exploit eval_simpl_expr; eauto. intros [tv [A B]]. inversion B. subst. econstructor; econstructor; split. eapply eval_Efield_struct; eauto. rewrite typeof_simpl_expr; eauto. econstructor; eauto. repeat rewrite Int.add_assoc. decEq. apply Int.add_commut. (* field union *) + rewrite <- comp_env_preserved in *. exploit eval_simpl_expr; eauto. intros [tv [A B]]. inversion B. subst. econstructor; econstructor; split. @@ -1583,7 +1598,7 @@ Qed. Lemma match_cont_assign_loc: forall f cenv k tk m bound tbound ty loc ofs v m', match_cont f cenv k tk m bound tbound -> - assign_loc ty m loc ofs v m' -> + assign_loc ge ty m loc ofs v m' -> Ple bound loc -> match_cont f cenv k tk m' bound tbound. Proof. @@ -1687,8 +1702,8 @@ Lemma match_cont_free_env: match_cont f cenv k tk m lo tlo -> Ple hi (Mem.nextblock m) -> Ple thi (Mem.nextblock tm) -> - Mem.free_list m (blocks_of_env e) = Some m' -> - Mem.free_list tm (blocks_of_env te) = Some tm' -> + Mem.free_list m (blocks_of_env ge e) = Some m' -> + Mem.free_list tm (blocks_of_env tge te) = Some tm' -> match_cont f cenv k tk m' (Mem.nextblock m') (Mem.nextblock tm'). Proof. intros. apply match_cont_incr_bounds with lo tlo. @@ -2183,7 +2198,7 @@ Proof. red; intros; subst b'. xomega. eapply alloc_variables_load; eauto. apply compat_cenv_for. - rewrite (bind_parameters_nextblock _ _ _ _ _ H2). xomega. + rewrite (bind_parameters_nextblock _ _ _ _ _ _ H2). xomega. rewrite T; xomega. (* external function *) @@ -2216,8 +2231,9 @@ Proof. exploit function_ptr_translated; eauto. intros [tf [A B]]. econstructor; split. econstructor. - eapply Genv.init_mem_transf_partial; eauto. - rewrite (transform_partial_program_main _ _ TRANSF). + eapply Genv.init_mem_transf_partial. eexact transf_programs. eauto. + change (prog_main tprog) with (AST.prog_main tprog). + rewrite (transform_partial_program_main _ _ transf_programs). instantiate (1 := b). rewrite <- H1. apply symbols_preserved. eauto. rewrite <- H3; apply type_of_fundef_preserved; auto. -- cgit From 67b13ecb9cfd2511c1db62a6cc38cf796cfb2a14 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Wed, 31 Dec 2014 17:14:28 +0100 Subject: Add a type system for CompCert C and type-checking constructor functions. Use these constructor functions in C2C to rely less on the types produced by the unverified elaborator. --- cfrontend/C2C.ml | 144 ++-- cfrontend/Ctyping.v | 1999 +++++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 2079 insertions(+), 64 deletions(-) create mode 100644 cfrontend/Ctyping.v (limited to 'cfrontend') diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index ea278ac1..ef621a7c 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -127,6 +127,12 @@ let unsupported msg = let warning msg = eprintf "%aWarning: %s\n" Cutil.printloc !currentLocation msg +let string_of_errmsg msg = + let string_of_err = function + | Errors.MSG s -> camlstring_of_coqstring s + | Errors.CTX i -> extern_atom i + | Errors.POS i -> Z.to_string (Z.Zpos i) + in String.concat "" (List.map string_of_err msg) (** ** The builtin environment *) @@ -448,9 +454,9 @@ let convertFloat f kind = | Z.Z0 -> begin match kind with | FFloat -> - Vsingle (Float.to_single Float.zero) + Ctyping.econst_single (Float.to_single Float.zero) | FDouble | FLongDouble -> - Vfloat Float.zero + Ctyping.econst_float Float.zero end | Z.Zpos mant -> @@ -465,9 +471,9 @@ let convertFloat f kind = begin match kind with | FFloat -> - Vsingle (Float32.from_parsed base mant exp) + Ctyping.econst_single (Float32.from_parsed base mant exp) | FDouble | FLongDouble -> - Vfloat (Float.from_parsed base mant exp) + Ctyping.econst_float (Float.from_parsed base mant exp) end | Z.Zneg _ -> assert false @@ -476,53 +482,59 @@ let convertFloat f kind = let ezero = Eval(Vint(coqint_of_camlint 0l), type_int32s) +let ewrap = function + | Errors.OK e -> e + | Errors.Error msg -> + error ("retyping error: " ^ string_of_errmsg msg); ezero + let rec convertExpr env e = - let ty = convertTyp env e.etyp in + (*let ty = convertTyp env e.etyp in*) match e.edesc with | C.EVar _ | C.EUnop((C.Oderef|C.Odot _|C.Oarrow _), _) | C.EBinop(C.Oindex, _, _, _) -> let l = convertLvalue env e in - Evalof(l, ty) + ewrap (Ctyping.evalof l) - | C.EConst(C.CInt(i, (ILongLong|IULongLong), _)) -> - Eval(Vlong(coqint_of_camlint64 i), ty) | C.EConst(C.CInt(i, k, _)) -> - Eval(Vint(convertInt i), ty) + let sg = if Cutil.is_signed_ikind k then Signed else Unsigned in + if k = ILongLong || k = IULongLong + then Ctyping.econst_long (coqint_of_camlint64 i) sg + else Ctyping.econst_int (convertInt i) sg | C.EConst(C.CFloat(f, k)) -> if k = C.FLongDouble && not !Clflags.option_flongdouble then unsupported "'long double' floating-point literal"; - Eval(convertFloat f k, ty) + convertFloat f k | C.EConst(C.CStr s) -> let ty = typeStringLiteral s in Evalof(Evar(name_for_string_literal env s, ty), ty) | C.EConst(C.CWStr s) -> unsupported "wide string literal"; ezero | C.EConst(C.CEnum(id, i)) -> - Eval(Vint(convertInt i), ty) + Ctyping.econst_int (convertInt i) Signed | C.ESizeof ty1 -> - Esizeof(convertTyp env ty1, ty) + Ctyping.esizeof (convertTyp env ty1) | C.EAlignof ty1 -> - Ealignof(convertTyp env ty1, ty) + Ctyping.ealignof (convertTyp env ty1) | C.EUnop(C.Ominus, e1) -> - Eunop(Oneg, convertExpr env e1, ty) + ewrap (Ctyping.eunop Oneg (convertExpr env e1)) | C.EUnop(C.Oplus, e1) -> convertExpr env e1 | C.EUnop(C.Olognot, e1) -> - Eunop(Onotbool, convertExpr env e1, ty) + ewrap (Ctyping.eunop Onotbool (convertExpr env e1)) | C.EUnop(C.Onot, e1) -> - Eunop(Onotint, convertExpr env e1, ty) + ewrap (Ctyping.eunop Onotint (convertExpr env e1)) | C.EUnop(C.Oaddrof, e1) -> - Eaddrof(convertLvalue env e1, ty) + ewrap (Ctyping.eaddrof (convertLvalue env e1)) | C.EUnop(C.Opreincr, e1) -> - coq_Epreincr Incr (convertLvalue env e1) ty + ewrap (Ctyping.epreincr (convertLvalue env e1)) | C.EUnop(C.Opredecr, e1) -> - coq_Epreincr Decr (convertLvalue env e1) ty + ewrap (Ctyping.epredecr (convertLvalue env e1)) | C.EUnop(C.Opostincr, e1) -> - Epostincr(Incr, convertLvalue env e1, ty) + ewrap (Ctyping.epostincr (convertLvalue env e1)) | C.EUnop(C.Opostdecr, e1) -> - Epostincr(Decr, convertLvalue env e1, ty) + ewrap (Ctyping.epostdecr (convertLvalue env e1)) | C.EBinop((C.Oadd|C.Osub|C.Omul|C.Odiv|C.Omod|C.Oand|C.Oor|C.Oxor| C.Oshl|C.Oshr|C.Oeq|C.One|C.Olt|C.Ogt|C.Ole|C.Oge) as op, @@ -546,7 +558,7 @@ let rec convertExpr env e = | C.Ole -> Ole | C.Oge -> Oge | _ -> assert false in - Ebinop(op', convertExpr env e1, convertExpr env e2, ty) + ewrap (Ctyping.ebinop op' (convertExpr env e1) (convertExpr env e2)) | C.EBinop(C.Oassign, e1, e2, _) -> let e1' = convertLvalue env e1 in let e2' = convertExpr env e2 in @@ -554,12 +566,11 @@ let rec convertExpr env e = && List.mem AVolatile (Cutil.attributes_of_type env e1.etyp) then warning "assignment to a l-value of volatile composite type. \ The 'volatile' qualifier is ignored."; - Eassign(e1', e2', ty) + 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| C.Oshl_assign|C.Oshr_assign) as op, e1, e2, tyres) -> - let tyres = convertTyp env tyres in let op' = match op with | C.Oadd_assign -> Oadd @@ -575,18 +586,20 @@ let rec convertExpr env e = | _ -> assert false in let e1' = convertLvalue env e1 in let e2' = convertExpr env e2 in - Eassignop(op', e1', e2', tyres, ty) + ewrap (Ctyping.eassignop op' e1' e2') | C.EBinop(C.Ocomma, e1, e2, _) -> - Ecomma(convertExpr env e1, convertExpr env e2, ty) + ewrap (Ctyping.ecomma (convertExpr env e1) (convertExpr env e2)) | C.EBinop(C.Ologand, e1, e2, _) -> - Eseqand(convertExpr env e1, convertExpr env e2, ty) + ewrap (Ctyping.eseqand (convertExpr env e1) (convertExpr env e2)) | C.EBinop(C.Ologor, e1, e2, _) -> - Eseqor(convertExpr env e1, convertExpr env e2, ty) + ewrap (Ctyping.eseqor (convertExpr env e1) (convertExpr env e2)) | C.EConditional(e1, e2, e3) -> - Econdition(convertExpr env e1, convertExpr env e2, convertExpr env e3, ty) + ewrap (Ctyping.econdition' (convertExpr env e1) + (convertExpr env e2) (convertExpr env e3) + (convertTyp env e.etyp)) | C.ECast(ty1, e1) -> - Ecast(convertExpr env e1, convertTyp env ty1) + ewrap (Ctyping.ecast (convertTyp env ty1) (convertExpr env e1)) | C.ECompound(ty1, ie) -> unsupported "compound literals"; ezero @@ -597,7 +610,7 @@ let rec convertExpr env e = Ebuiltin( EF_annot(intern_string txt, List.map (fun t -> AA_arg t) (typlist_of_typelist targs1)), - targs1, convertExprList env args1, ty) + targs1, convertExprList env args1, convertTyp env e.etyp) | _ -> error "ill-formed __builtin_annot (first argument must be string literal)"; ezero @@ -609,7 +622,8 @@ let rec convertExpr env e = let targ = convertTyp env (Cutil.default_argument_conversion env arg.etyp) in Ebuiltin(EF_annot_val(intern_string txt, typ_of_type targ), - Tcons(targ, Tnil), convertExprList env [arg], ty) + Tcons(targ, Tnil), convertExprList env [arg], + convertTyp env e.etyp) | _ -> error "ill-formed __builtin_annot_intval (first argument must be string literal)"; ezero @@ -619,15 +633,15 @@ let rec convertExpr env e = make_builtin_memcpy (convertExprList env args) | C.ECall({edesc = C.EVar {name = "__builtin_fabs"}}, [arg]) -> - Eunop(Oabsfloat, convertExpr env arg, ty) + ewrap (Ctyping.eunop Oabsfloat (convertExpr env arg)) | C.ECall({edesc = C.EVar {name = "__builtin_va_start"}} as fn, [arg]) -> Ecall(convertExpr env fn, Econs(va_list_ptr(convertExpr env arg), Enil), - ty) + convertTyp env e.etyp) | C.ECall({edesc = C.EVar {name = "__builtin_va_arg"}}, [arg1; arg2]) -> - make_builtin_va_arg env ty (convertExpr env arg1) + make_builtin_va_arg env (convertTyp env e.etyp) (convertExpr env arg1) | C.ECall({edesc = C.EVar {name = "__builtin_va_end"}}, _) -> Ecast (ezero, Tvoid) @@ -643,12 +657,12 @@ let rec convertExpr env e = | C.ECall({edesc = C.EVar {name = "printf"}}, args) when !Clflags.option_interp -> - let targs = - convertTypArgs env [] args in + let targs = convertTypArgs env [] args + and tres = convertTyp env e.etyp in let sg = - signature_of_type targs ty {cc_vararg = true; cc_structret = false} in + signature_of_type targs tres {cc_vararg = true; cc_structret = false} in Ebuiltin(EF_external(intern_string "printf", sg), - targs, convertExprList env args, ty) + targs, convertExprList env args, tres) | C.ECall(fn, args) -> if not (supported_return_type env e.etyp) then @@ -662,26 +676,25 @@ let rec convertExpr env e = if va && not !Clflags.option_fvararg_calls then unsupported "call to variable-argument function (consider adding option -fvararg-calls)" end; - Ecall(convertExpr env fn, convertExprList env args, ty) + ewrap (Ctyping.ecall (convertExpr env fn) (convertExprList env args)) and convertLvalue env e = - let ty = convertTyp env e.etyp in match e.edesc with | C.EVar id -> - Evar(intern_string id.name, ty) + Evar(intern_string id.name, convertTyp env e.etyp) | C.EUnop(C.Oderef, e1) -> - Ederef(convertExpr env e1, ty) + ewrap (Ctyping.ederef (convertExpr env e1)) | C.EUnop(C.Odot id, e1) -> - Efield(convertExpr env e1, intern_string id, ty) + ewrap (Ctyping.efield !comp_env (convertExpr env e1) (intern_string id)) | C.EUnop(C.Oarrow id, e1) -> let e1' = convertExpr env e1 in - let ty1 = - match typeof e1' with - | Tpointer(t, _) | Tarray(t, _, _) -> t - | _ -> error ("wrong type for ->" ^ id ^ " access"); Tvoid in - Efield(Evalof(Ederef(e1', ty1), ty1), intern_string id, ty) + let e2' = ewrap (Ctyping.ederef e1') in + let e3' = ewrap (Ctyping.evalof e2') in + ewrap (Ctyping.efield !comp_env e3' (intern_string id)) | C.EBinop(C.Oindex, e1, e2, _) -> - coq_Eindex (convertExpr env e1) (convertExpr env e2) ty + let e1' = convertExpr env e1 and e2' = convertExpr env e2 in + let e3' = ewrap (Ctyping.ebinop Oadd e1' e2') in + ewrap (Ctyping.ederef e3') | _ -> error "illegal l-value"; ezero @@ -737,30 +750,39 @@ let add_lineno prev_loc this_loc s = (** Statements *) +let swrap = function + | Errors.OK s -> s + | Errors.Error msg -> + error ("retyping error: " ^ string_of_errmsg msg); Sskip + let rec convertStmt ploc env s = updateLoc s.sloc; match s.sdesc with | C.Sskip -> Sskip | C.Sdo e -> - add_lineno ploc s.sloc (Sdo(convertExpr env e)) + add_lineno ploc s.sloc (swrap (Ctyping.sdo (convertExpr env e))) | C.Sseq(s1, s2) -> Ssequence(convertStmt ploc env s1, convertStmt s1.sloc env s2) | C.Sif(e, s1, s2) -> let te = convertExpr env e in add_lineno ploc s.sloc - (Sifthenelse(te, convertStmt s.sloc env s1, convertStmt s.sloc env s2)) + (swrap (Ctyping.sifthenelse te + (convertStmt s.sloc env s1) (convertStmt s.sloc env s2))) | C.Swhile(e, s1) -> let te = convertExpr env e in - add_lineno ploc s.sloc (Swhile(te, convertStmt s.sloc env s1)) + add_lineno ploc s.sloc + (swrap (Ctyping.swhile te (convertStmt s.sloc env s1))) | C.Sdowhile(s1, e) -> let te = convertExpr env e in - add_lineno ploc s.sloc (Sdowhile(te, convertStmt s.sloc env s1)) + add_lineno ploc s.sloc + (swrap (Ctyping.sdowhile te (convertStmt s.sloc env s1))) | C.Sfor(s1, e, s2, s3) -> let te = convertExpr env e in add_lineno ploc s.sloc - (Sfor(convertStmt s.sloc env s1, te, - convertStmt s.sloc env s2, convertStmt s.sloc env s3)) + (swrap (Ctyping.sfor + (convertStmt s.sloc env s1) te + (convertStmt s.sloc env s2) (convertStmt s.sloc env s3))) | C.Sbreak -> Sbreak | C.Scontinue -> @@ -773,7 +795,8 @@ let rec convertStmt ploc env s = warning "ignored code at beginning of 'switch'"; let te = convertExpr env e in add_lineno ploc s.sloc - (Sswitch(te, convertSwitch s.sloc env (is_longlong env e.etyp) cases)) + (swrap (Ctyping.sswitch te + (convertSwitch s.sloc env (is_longlong env e.etyp) cases))) | C.Slabeled(C.Slabel lbl, s1) -> add_lineno ploc s.sloc (Slabel(intern_string lbl, convertStmt s.sloc env s1)) @@ -876,13 +899,6 @@ let convertFundecl env (sto, id, ty, optinit) = (** Initializers *) -let string_of_errmsg msg = - let string_of_err = function - | Errors.MSG s -> camlstring_of_coqstring s - | Errors.CTX i -> extern_atom i - | Errors.POS i -> Z.to_string (Z.Zpos i) - in String.concat "" (List.map string_of_err msg) - let rec convertInit env init = match init with | C.Init_single e -> diff --git a/cfrontend/Ctyping.v b/cfrontend/Ctyping.v new file mode 100644 index 00000000..43d34007 --- /dev/null +++ b/cfrontend/Ctyping.v @@ -0,0 +1,1999 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the GNU General Public License as published by *) +(* the Free Software Foundation, either version 2 of the License, or *) +(* (at your option) any later version. This file is also distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Typing rules and type-checking for the Compcert C language *) + +Require Import Coqlib. +Require Import String. +Require Import Maps. +Require Import Integers. +Require Import Floats. +Require Import Values. +Require Import Memory. +Require Import Globalenvs. +Require Import Events. +Require Import AST. +Require Import Ctypes. +Require Import Cop. +Require Import Csyntax. +Require Import Csem. +Require Import Errors. + +Local Open Scope error_monad_scope. + +Definition strict := false. +Opaque strict. + +(** * Operations over types *) + +(** The type of a member of a composite (struct or union). + The "volatile" attribute carried by the composite propagates + to the type of the member, but not the "alignas" attribute. *) + +Definition attr_add_volatile (vol: bool) (a: attr) := + {| attr_volatile := a.(attr_volatile) || vol; + attr_alignas := a.(attr_alignas) |}. + +Definition type_of_member (a: attr) (f: ident) (m: members) : res type := + do ty <- field_type f m; + OK (change_attributes (attr_add_volatile a.(attr_volatile)) ty). + +(** Type-checking of arithmetic and logical operators *) + +Definition type_unop (op: unary_operation) (ty: type) : res type := + match op with + | Onotbool => + match classify_bool ty with + | bool_default => Error (msg "operator !") + | _ => OK (Tint I32 Signed noattr) + end + | Onotint => + match classify_notint ty with + | notint_case_i sg => OK (Tint I32 sg noattr) + | notint_case_l sg => OK (Tlong sg noattr) + | notint_default => Error (msg "operator ~") + end + | Oneg => + match classify_neg ty with + | neg_case_i sg => OK (Tint I32 sg noattr) + | neg_case_f => OK (Tfloat F64 noattr) + | neg_case_s => OK (Tfloat F32 noattr) + | neg_case_l sg => OK (Tlong sg noattr) + | neg_default => Error (msg "operator prefix -") + end + | Oabsfloat => + match classify_neg ty with + | neg_default => Error (msg "operator __builtin_fabs") + | _ => OK (Tfloat F64 noattr) + end + end. + +Definition binarith_type (ty1 ty2: type) (m: string): res type := + match classify_binarith ty1 ty2 with + | bin_case_i sg => OK (Tint I32 sg noattr) + | bin_case_l sg => OK (Tlong sg noattr) + | bin_case_f => OK (Tfloat F64 noattr) + | bin_case_s => OK (Tfloat F32 noattr) + | bin_default => Error (msg m) + end. + +Definition binarith_int_type (ty1 ty2: type) (m: string): res type := + match classify_binarith ty1 ty2 with + | bin_case_i sg => OK (Tint I32 sg noattr) + | bin_case_l sg => OK (Tlong sg noattr) + | _ => Error (msg m) + end. + +Definition shift_op_type (ty1 ty2: type) (m: string): res type := + match classify_shift ty1 ty2 with + | shift_case_ii sg | shift_case_il sg => OK (Tint I32 sg noattr) + | shift_case_li sg | shift_case_ll sg => OK (Tlong sg noattr) + | shift_default => Error (msg m) + end. + +Definition comparison_type (ty1 ty2: type) (m: string): res type := + match classify_cmp ty1 ty2 with + | cmp_default => + match classify_binarith ty1 ty2 with + | bin_default => Error (msg m) + | _ => OK (Tint I32 Signed noattr) + end + | _ => OK (Tint I32 Signed noattr) + end. + +Definition type_binop (op: binary_operation) (ty1 ty2: type) : res type := + match op with + | Oadd => + match classify_add ty1 ty2 with + | add_case_pi ty | add_case_ip ty + | add_case_pl ty | add_case_lp ty => OK (Tpointer ty noattr) + | add_default => binarith_type ty1 ty2 "operator +" + end + | Osub => + match classify_sub ty1 ty2 with + | sub_case_pi ty | sub_case_pl ty => OK (Tpointer ty noattr) +(* + | sub_case_pp ty1 ty2 => + if type_eq (remove_attributes ty1) (remove_attributes ty2) + then OK (Tint I32 Signed noattr) + else Error (msg "operator - : incompatible pointer types") +*) + | sub_case_pp ty => OK (Tint I32 Signed noattr) + | sub_default => binarith_type ty1 ty2 "operator infix -" + end + | Omul => binarith_type ty1 ty2 "operator infix *" + | Odiv => binarith_type ty1 ty2 "operator /" + | Omod => binarith_int_type ty1 ty2 "operator %" + | Oand => binarith_int_type ty1 ty2 "operator &" + | Oor => binarith_int_type ty1 ty2 "operator |" + | Oxor => binarith_int_type ty1 ty2 "operator ^" + | Oshl => shift_op_type ty1 ty2 "operator <<" + | Oshr => shift_op_type ty1 ty2 "operator >>" + | Oeq => comparison_type ty1 ty2 "operator ==" + | One => comparison_type ty1 ty2 "operator !=" + | Olt => comparison_type ty1 ty2 "operator <" + | Ogt => comparison_type ty1 ty2 "operator >" + | Ole => comparison_type ty1 ty2 "operator <=" + | Oge => comparison_type ty1 ty2 "operator >=" + end. + +Definition type_deref (ty: type) : res type := + match ty with + | Tpointer tyelt _ => OK tyelt + | Tarray tyelt _ _ => OK tyelt + | Tfunction _ _ _ => OK ty + | _ => Error (msg "operator prefix *") + end. + +Definition is_void (ty: type) : bool := + match ty with Tvoid => true | _ => false end. + +Definition type_join (ty1 ty2: type) : res type := + match typeconv ty1, typeconv ty2 with + | (Tint _ _ _ | Tfloat _ _), (Tint _ _ _ | Tfloat _ _) => + binarith_type ty1 ty2 "conditional expression" + | Tpointer t1 a1, Tpointer t2 a2 => + OK (Tpointer (if is_void t1 || is_void t2 then Tvoid else t1) noattr) + | Tpointer t1 a1, Tint _ _ _ => + OK (Tpointer t1 noattr) + | Tint _ _ _, Tpointer t2 a2 => + OK (Tpointer t2 noattr) + | Tvoid, Tvoid => + OK Tvoid + | Tstruct id1 a1, Tstruct id2 a2 => + if ident_eq id1 id2 + then OK (Tstruct id1 noattr) + else Error (msg "conditional expression") + | Tunion id1 a1, Tunion id2 a2 => + if ident_eq id1 id2 + then OK (Tunion id1 noattr) + else Error (msg "conditional expression") + | _, _ => + Error (msg "conditional expression") + end. + +(** * Specification of the type system *) + +(** Type environments map identifiers to their types. *) + +Definition typenv := PTree.t type. + +Definition wt_cast (from to: type) : Prop := + classify_cast from to <> cast_case_default. + +Definition wt_bool (ty: type) : Prop := + classify_bool ty <> bool_default. + +Definition wt_int (n: int) (sz: intsize) (sg: signedness) : Prop := + match sz, sg with + | IBool, _ => Int.zero_ext 8 n = n + | I8, Unsigned => Int.zero_ext 8 n = n + | I8, Signed => Int.sign_ext 8 n = n + | I16, Unsigned => Int.zero_ext 16 n = n + | I16, Signed => Int.sign_ext 16 n = n + | I32, _ => True + end. + +Inductive wt_val : val -> type -> Prop := + | wt_val_int: forall n sz sg a, + wt_int n sz sg -> + wt_val (Vint n) (Tint sz sg a) + | wt_val_ptr_int: forall b ofs sg a, + wt_val (Vptr b ofs) (Tint I32 sg a) + | wt_val_long: forall n sg a, + wt_val (Vlong n) (Tlong sg a) + | wt_val_float: forall f a, + wt_val (Vfloat f) (Tfloat F64 a) + | wt_val_single: forall f a, + wt_val (Vsingle f) (Tfloat F32 a) + | wt_val_pointer: forall b ofs ty a, + wt_val (Vptr b ofs) (Tpointer ty a) + | wt_val_int_pointer: forall n ty a, + wt_val (Vint n) (Tpointer ty a) + | wt_val_array: forall b ofs ty sz a, + wt_val (Vptr b ofs) (Tarray ty sz a) + | wt_val_function: forall b ofs tyargs tyres cc, + wt_val (Vptr b ofs) (Tfunction tyargs tyres cc) + | wt_val_struct: forall b ofs id a, + wt_val (Vptr b ofs) (Tstruct id a) + | wt_val_union: forall b ofs id a, + wt_val (Vptr b ofs) (Tunion id a) + | wt_val_undef: forall ty, + wt_val Vundef ty + | wt_val_void: forall v, + wt_val v Tvoid. + +Inductive wt_arguments: exprlist -> typelist -> Prop := + | wt_arg_nil: + wt_arguments Enil Tnil + | wt_arg_cons: forall a al ty tyl, + wt_cast (typeof a) ty -> + wt_arguments al tyl -> + wt_arguments (Econs a al) (Tcons ty tyl) + | wt_arg_extra: forall a al, (**r tolerance for varargs *) + strict = false -> + wt_arguments (Econs a al) Tnil. + +Definition subtype (t1 t2: type) : Prop := + forall v, wt_val v t1 -> wt_val v t2. + +Section WT_EXPR_STMT. + +Variable ce: composite_env. +Variable e: typenv. + +Inductive wt_rvalue : expr -> Prop := + | wt_Eval: forall v ty, + wt_val v ty -> + wt_rvalue (Eval v ty) + | wt_Evalof: forall l, + wt_lvalue l -> + wt_rvalue (Evalof l (typeof l)) + | wt_Eaddrof: forall l, + wt_lvalue l -> + wt_rvalue (Eaddrof l (Tpointer (typeof l) noattr)) + | wt_Eunop: forall op r ty, + wt_rvalue r -> + type_unop op (typeof r) = OK ty -> + wt_rvalue (Eunop op r ty) + | wt_Ebinop: forall op r1 r2 ty, + wt_rvalue r1 -> wt_rvalue r2 -> + type_binop op (typeof r1) (typeof r2) = OK ty -> + wt_rvalue (Ebinop op r1 r2 ty) + | wt_Ecast: forall r ty, + wt_rvalue r -> wt_cast (typeof r) ty -> + wt_rvalue (Ecast r ty) + | wt_Eseqand: forall r1 r2, + wt_rvalue r1 -> wt_rvalue r2 -> + wt_bool (typeof r1) -> wt_bool (typeof r2) -> + wt_rvalue (Eseqand r1 r2 (Tint I32 Signed noattr)) + | wt_Eseqor: forall r1 r2, + wt_rvalue r1 -> wt_rvalue r2 -> + wt_bool (typeof r1) -> wt_bool (typeof r2) -> + wt_rvalue (Eseqor r1 r2 (Tint I32 Signed noattr)) + | wt_Econdition: forall r1 r2 r3 ty, + wt_rvalue r1 -> wt_rvalue r2 -> wt_rvalue r3 -> + wt_bool (typeof r1) -> + wt_cast (typeof r2) ty -> wt_cast (typeof r3) ty -> + wt_rvalue (Econdition r1 r2 r3 ty) + | wt_Esizeof: forall ty, + wt_rvalue (Esizeof ty (Tint I32 Unsigned noattr)) + | wt_Ealignof: forall ty, + wt_rvalue (Ealignof ty (Tint I32 Unsigned noattr)) + | wt_Eassign: forall l r, + wt_lvalue l -> wt_rvalue r -> wt_cast (typeof r) (typeof l) -> + wt_rvalue (Eassign l r (typeof l)) + | wt_Eassignop: forall op l r ty, + wt_lvalue l -> wt_rvalue r -> + type_binop op (typeof l) (typeof r) = OK ty -> + wt_cast ty (typeof l) -> + wt_rvalue (Eassignop op l r ty (typeof l)) + | wt_Epostincr: forall id l ty, + wt_lvalue l -> + type_binop (match id with Incr => Oadd | Decr => Osub end) + (typeof l) (Tint I32 Signed noattr) = OK ty -> + wt_cast (incrdecr_type (typeof l)) (typeof l) -> + wt_rvalue (Epostincr id l (typeof l)) + | wt_Ecomma: forall r1 r2, + wt_rvalue r1 -> wt_rvalue r2 -> + wt_rvalue (Ecomma r1 r2 (typeof r2)) + | wt_Ecall: forall r1 rargs tyargs tyres cconv, + wt_rvalue r1 -> wt_exprlist rargs -> + classify_fun (typeof r1) = fun_case_f tyargs tyres cconv -> + wt_arguments rargs tyargs -> + wt_rvalue (Ecall r1 rargs tyres) + | wt_Ebuiltin: forall ef tyargs rargs, + wt_exprlist rargs -> + wt_arguments rargs tyargs -> + (* This is specialized to builtins returning void, the only + case generated by C2C. *) + sig_res (ef_sig ef) = None -> + wt_rvalue (Ebuiltin ef tyargs rargs Tvoid) + | wt_Eparen: forall r tycast ty, + wt_rvalue r -> + wt_cast (typeof r) tycast -> subtype tycast ty -> + wt_rvalue (Eparen r tycast ty) + +with wt_lvalue : expr -> Prop := + | wt_Eloc: forall b ofs ty, + wt_lvalue (Eloc b ofs ty) + | wt_Evar: forall x ty, + e!x = Some ty -> + wt_lvalue (Evar x ty) + | wt_Ederef: forall r ty, + wt_rvalue r -> + type_deref (typeof r) = OK ty -> + wt_lvalue (Ederef r ty) + | wt_Efield: forall r f id a co ty, + wt_rvalue r -> + typeof r = Tstruct id a \/ typeof r = Tunion id a -> + ce!id = Some co -> + type_of_member a f co.(co_members) = OK ty -> + wt_lvalue (Efield r f ty) + +with wt_exprlist : exprlist -> Prop := + | wt_Enil: + wt_exprlist Enil + | wt_Econs: forall r1 rl, + wt_rvalue r1 -> wt_exprlist rl -> wt_exprlist (Econs r1 rl). + +Definition wt_expr_kind (k: kind) (a: expr) := + match k with + | RV => wt_rvalue a + | LV => wt_lvalue a + end. + +Definition expr_kind (a: expr) : kind := + match a with + | Eloc _ _ _ => LV + | Evar _ _ => LV + | Ederef _ _ => LV + | Efield _ _ _ => LV + | _ => RV + end. + +Definition wt_expr (a: expr) := + match expr_kind a with + | RV => wt_rvalue a + | LV => wt_lvalue a + end. + +Variable rt: type. + +Inductive wt_stmt: statement -> Prop := + | wt_Sskip: + wt_stmt Sskip + | wt_Sdo: forall r, + wt_rvalue r -> wt_stmt (Sdo r) + | wt_Ssequence: forall s1 s2, + wt_stmt s1 -> wt_stmt s2 -> wt_stmt (Ssequence s1 s2) + | wt_Sifthenelse: forall r s1 s2, + wt_rvalue r -> wt_stmt s1 -> wt_stmt s2 -> wt_bool (typeof r) -> + wt_stmt (Sifthenelse r s1 s2) + | wt_Swhile: forall r s, + wt_rvalue r -> wt_stmt s -> wt_bool (typeof r) -> + wt_stmt (Swhile r s) + | wt_Sdowhile: forall r s, + wt_rvalue r -> wt_stmt s -> wt_bool (typeof r) -> + wt_stmt (Sdowhile r s) + | wt_Sfor: forall s1 r s2 s3, + wt_rvalue r -> wt_stmt s1 -> wt_stmt s2 -> wt_stmt s3 -> + wt_bool (typeof r) -> + wt_stmt (Sfor s1 r s2 s3) + | wt_Sbreak: + wt_stmt Sbreak + | wt_Scontinue: + wt_stmt Scontinue + | wt_Sreturn_none: + wt_stmt (Sreturn None) + | wt_Sreturn_some: forall r, + wt_rvalue r -> + wt_cast (typeof r) rt -> + wt_stmt (Sreturn (Some r)) + | wt_Sswitch: forall r ls sg sz a, + wt_rvalue r -> + typeof r = Tint sz sg a \/ typeof r = Tlong sg a -> + wt_lblstmts ls -> + wt_stmt (Sswitch r ls) + | wt_Slabel: forall lbl s, + wt_stmt s -> wt_stmt (Slabel lbl s) + | wt_Sgoto: forall lbl, + wt_stmt (Sgoto lbl) + +with wt_lblstmts : labeled_statements -> Prop := + | wt_LSnil: + wt_lblstmts LSnil + | wt_LScons: forall case s ls, + wt_stmt s -> wt_lblstmts ls -> + wt_lblstmts (LScons case s ls). + +End WT_EXPR_STMT. + +Fixpoint bind_vars (e: typenv) (l: list (ident * type)) : typenv := + match l with + | nil => e + | (id, ty) :: l => bind_vars (PTree.set id ty e) l + end. + +Inductive wt_function (ce: composite_env) (e: typenv) : function -> Prop := + | wt_function_intro: forall f, + wt_stmt ce (bind_vars (bind_vars e f.(fn_params)) f.(fn_vars)) f.(fn_return) f.(fn_body) -> + wt_function ce e f. + +Fixpoint bind_globdef (e: typenv) (l: list (ident * globdef fundef type)) : typenv := + match l with + | nil => e + | (id, Gfun fd) :: l => bind_globdef (PTree.set id (type_of_fundef fd) e) l + | (id, Gvar v) :: l => bind_globdef (PTree.set id v.(gvar_info) e) l + end. + +Inductive wt_program : program -> Prop := + | wt_program_intro: forall p, + let e := bind_globdef (PTree.empty _) p.(prog_defs) in + (forall id f, In (id, Gfun (Internal f)) p.(prog_defs) -> + wt_function p.(prog_comp_env) e f) -> + wt_program p. + +Hint Constructors wt_val wt_rvalue wt_lvalue wt_stmt wt_lblstmts: ty. +Hint Extern 1 (wt_int _ _ _) => exact I: ty. +Hint Extern 1 (wt_int _ _ _) => reflexivity: ty. + +Ltac DestructCases := + match goal with + | [H: match match ?x with _ => _ end with _ => _ end = Some _ |- _ ] => destruct x; DestructCases + | [H: match match ?x with _ => _ end with _ => _ end = OK _ |- _ ] => destruct x; DestructCases + | [H: match ?x with _ => _ end = OK _ |- _ ] => destruct x; DestructCases + | [H: match ?x with _ => _ end = Some _ |- _ ] => destruct x; DestructCases + | [H: match ?x with _ => _ end = OK _ |- _ ] => destruct x; DestructCases + | [H: Some _ = Some _ |- _ ] => inv H; DestructCases + | [H: None = Some _ |- _ ] => discriminate + | [H: OK _ = OK _ |- _ ] => inv H; DestructCases + | [H: Error _ = OK _ |- _ ] => discriminate + | _ => idtac + end. + +Ltac DestructMatch := + match goal with + | [ |- match match ?x with _ => _ end with _ => _ end ] => destruct x; DestructMatch + | [ |- match ?x with _ => _ end ] => destruct x; DestructMatch + | _ => idtac + end. + +(** * Type checking *) + +Definition check_cast (t1 t2: type) : res unit := + match classify_cast t1 t2 with + | cast_case_default => Error (msg "illegal cast") + | _ => OK tt + end. + +Definition check_bool (t: type) : res unit := + match classify_bool t with + | bool_default => Error (msg "not a boolean") + | _ => OK tt + end. + +Definition check_literal (v: val) (t: type) : res unit := + match v, t with + | Vint n, Tint I32 sg a => OK tt + | Vint n, Tpointer t' a => OK tt + | Vlong n, Tlong sg a => OK tt + | Vsingle f, Tfloat F32 a => OK tt + | Vfloat f, Tfloat F64 a => OK tt + | _, _ => Error (msg "wrong literal") + end. + +Fixpoint check_arguments (el: exprlist) (tyl: typelist) : res unit := + match el, tyl with + | Enil, Tnil => OK tt + | Enil, _ => Error (msg "not enough arguments") + | _, Tnil => if strict then Error (msg "too many arguments") else OK tt + | Econs e1 el, Tcons ty1 tyl => do x <- check_cast (typeof e1) ty1; check_arguments el tyl + end. + +Definition check_rval (e: expr) : res unit := + match e with + | Eloc _ _ _ | Evar _ _ | Ederef _ _ | Efield _ _ _ => + Error (msg "not a r-value") + | _ => + OK tt + end. + +Definition check_lval (e: expr) : res unit := + match e with + | Eloc _ _ _ | Evar _ _ | Ederef _ _ | Efield _ _ _ => + OK tt + | _ => + Error (msg "not a l-value") + end. + +Fixpoint check_rvals (el: exprlist) : res unit := + match el with + | Enil => OK tt + | Econs e1 el => do x <- check_rval e1; check_rvals el + end. + +(** Type-checking of expressions is presented as smart constructors + that check type constraints and build the expression with the correct + type annotation. *) + +Definition evar (e: typenv) (x: ident) : res expr := + match e!x with + | Some ty => OK (Evar x ty) + | None => Error (MSG "unbound variable " :: CTX x :: nil) + end. + +Definition ederef (r: expr) : res expr := + do x1 <- check_rval r; + do ty <- type_deref (typeof r); + OK (Ederef r ty). + +Definition efield (ce: composite_env) (r: expr) (f: ident) : res expr := + do x1 <- check_rval r; + match typeof r with + | Tstruct id a | Tunion id a => + match ce!id with + | None => Error (MSG "unbound composite " :: CTX id :: nil) + | Some co => + do ty <- type_of_member a f co.(co_members); + OK (Efield r f ty) + end + | _ => + Error (MSG "argument of ." :: CTX f :: MSG " is not a struct or union" :: nil) + end. + +Definition econst_int (n: int) (sg: signedness) : expr := + (Eval (Vint n) (Tint I32 sg noattr)). + +Definition econst_ptr_int (n: int) (ty: type) : expr := + (Eval (Vint n) (Tpointer ty noattr)). + +Definition econst_long (n: int64) (sg: signedness) : expr := + (Eval (Vlong n) (Tlong sg noattr)). + +Definition econst_float (n: float) : expr := + (Eval (Vfloat n) (Tfloat F64 noattr)). + +Definition econst_single (n: float32) : expr := + (Eval (Vsingle n) (Tfloat F32 noattr)). + +Definition evalof (l: expr) : res expr := + do x <- check_lval l; OK (Evalof l (typeof l)). + +Definition eaddrof (l: expr) : res expr := + do x <- check_lval l; OK (Eaddrof l (Tpointer (typeof l) noattr)). + +Definition eunop (op: unary_operation) (r: expr) : res expr := + do x <- check_rval r; + do ty <- type_unop op (typeof r); + OK (Eunop op r ty). + +Definition ebinop (op: binary_operation) (r1 r2: expr) : res expr := + do x1 <- check_rval r1; do x2 <- check_rval r2; + do ty <- type_binop op (typeof r1) (typeof r2); + OK (Ebinop op r1 r2 ty). + +Definition ecast (ty: type) (r: expr) : res expr := + do x1 <- check_rval r; + do x2 <- check_cast (typeof r) ty; + OK (Ecast r ty). + +Definition eseqand (r1 r2: expr) : res expr := + do x1 <- check_rval r1; do x2 <- check_rval r2; + do y1 <- check_bool (typeof r1); do y2 <- check_bool (typeof r2); + OK (Eseqand r1 r2 type_int32s). + +Definition eseqor (r1 r2: expr) : res expr := + do x1 <- check_rval r1; do x2 <- check_rval r2; + do y1 <- check_bool (typeof r1); do y2 <- check_bool (typeof r2); + OK (Eseqor r1 r2 type_int32s). + +Definition econdition (r1 r2 r3: expr) : res expr := + do x1 <- check_rval r1; do x2 <- check_rval r2; do x3 <- check_rval r3; + do y1 <- check_bool (typeof r1); + do ty <- type_join (typeof r2) (typeof r3); + OK (Econdition r1 r2 r3 ty). + +Definition econdition' (r1 r2 r3: expr) (ty: type) : res expr := + do x1 <- check_rval r1; do x2 <- check_rval r2; do x3 <- check_rval r3; + do y1 <- check_bool (typeof r1); + do y2 <- check_cast (typeof r2) ty; + do y3 <- check_cast (typeof r3) ty; + OK (Econdition r1 r2 r3 ty). + +Definition esizeof (ty: type) : expr := + Esizeof ty (Tint I32 Unsigned noattr). + +Definition ealignof (ty: type) : expr := + Ealignof ty (Tint I32 Unsigned noattr). + +Definition eassign (l r: expr) : res expr := + do x1 <- check_lval l; do x2 <- check_rval r; + do y1 <- check_cast (typeof r) (typeof l); + OK (Eassign l r (typeof l)). + +Definition eassignop (op: binary_operation) (l r: expr) : res expr := + do x1 <- check_lval l; do x2 <- check_rval r; + do ty <- type_binop op (typeof l) (typeof r); + do y1 <- check_cast ty (typeof l); + OK (Eassignop op l r ty (typeof l)). + +Definition epostincrdecr (id: incr_or_decr) (l: expr) : res expr := + do x1 <- check_lval l; + do ty <- type_binop (match id with Incr => Oadd | Decr => Osub end) + (typeof l) type_int32s; + do y1 <- check_cast (incrdecr_type (typeof l)) (typeof l); + OK (Epostincr id l (typeof l)). + +Definition epostincr (l: expr) := epostincrdecr Incr l. +Definition epostdecr (l: expr) := epostincrdecr Decr l. + +Definition epreincr (l: expr) := eassignop Oadd l (Eval (Vint Int.one) type_int32s). +Definition epredecr (l: expr) := eassignop Osub l (Eval (Vint Int.one) type_int32s). + +Definition ecomma (r1 r2: expr) : res expr := + do x1 <- check_rval r1; do x2 <- check_rval r2; + OK (Ecomma r1 r2 (typeof r2)). + +Definition ecall (fn: expr) (args: exprlist) : res expr := + do x1 <- check_rval fn; do x2 <- check_rvals args; + match classify_fun (typeof fn) with + | fun_case_f tyargs tyres cconv => + do y1 <- check_arguments args tyargs; + OK (Ecall fn args tyres) + | _ => + Error (msg "call: not a function") + end. + +Definition ebuiltin (ef: external_function) (tyargs: typelist) (args: exprlist) (tyres: type) : res expr := + do x1 <- check_rvals args; + do x2 <- check_arguments args tyargs; + if type_eq tyres Tvoid + && opt_typ_eq (sig_res (ef_sig ef)) None + then OK (Ebuiltin ef tyargs args tyres) + else Error (msg "builtin: wrong type decoration"). + +Definition sdo (a: expr) : res statement := + do x <- check_rval a; OK (Sdo a). + +Definition sifthenelse (a: expr) (s1 s2: statement) : res statement := + do x <- check_rval a; do y <- check_bool (typeof a); OK (Sifthenelse a s1 s2). + +Definition swhile (a: expr) (s: statement) : res statement := + do x <- check_rval a; do y <- check_bool (typeof a); OK (Swhile a s). + +Definition sdowhile (a: expr) (s: statement) : res statement := + do x <- check_rval a; do y <- check_bool (typeof a); OK (Sdowhile a s). + +Definition sfor (s1: statement) (a: expr) (s2 s3: statement) : res statement := + do x <- check_rval a; do y <- check_bool (typeof a); OK (Sfor s1 a s2 s3). + +Definition sreturn (rt: type) (a: expr) : res statement := + do x <- check_rval a; do y <- check_cast (typeof a) rt; + OK (Sreturn (Some a)). + +Definition sswitch (a: expr) (sl: labeled_statements) : res statement := + do x <- check_rval a; + match typeof a with + | Tint _ _ _ | Tlong _ _ => OK (Sswitch a sl) + | _ => Error (msg "wrong type for argument of switch") + end. + +(** Using the smart constructors, we define a type checker that rebuilds + a correctly-type-annotated program. *) + +Fixpoint retype_expr (ce: composite_env) (e: typenv) (a: expr) : res expr := + match a with + | Eval (Vint n) (Tint _ sg _) => + OK (econst_int n sg) + | Eval (Vint n) (Tpointer ty _) => + OK (econst_ptr_int n ty) + | Eval (Vlong n) (Tlong sg _) => + OK (econst_long n sg) + | Eval (Vfloat n) _ => + OK (econst_float n) + | Eval (Vsingle n) _ => + OK (econst_single n) + | Eval _ _ => + Error (msg "bad literal") + | Evar x _ => + evar e x + | Efield r f _ => + do r' <- retype_expr ce e r; efield ce r' f + | Evalof l _ => + do l' <- retype_expr ce e l; evalof l' + | Ederef r _ => + do r' <- retype_expr ce e r; ederef r' + | Eaddrof l _ => + do l' <- retype_expr ce e l; eaddrof l' + | Eunop op r _ => + do r' <- retype_expr ce e r; eunop op r' + | Ebinop op r1 r2 _ => + do r1' <- retype_expr ce e r1; do r2' <- retype_expr ce e r2; ebinop op r1' r2' + | Ecast r ty => + do r' <- retype_expr ce e r; ecast ty r' + | Eseqand r1 r2 _ => + do r1' <- retype_expr ce e r1; do r2' <- retype_expr ce e r2; eseqand r1' r2' + | Eseqor r1 r2 _ => + do r1' <- retype_expr ce e r1; do r2' <- retype_expr ce e r2; eseqor r1' r2' + | Econdition r1 r2 r3 _ => + do r1' <- retype_expr ce e r1; do r2' <- retype_expr ce e r2; do r3' <- retype_expr ce e r3; econdition r1' r2' r3' + | Esizeof ty _ => + OK (esizeof ty) + | Ealignof ty _ => + OK (ealignof ty) + | Eassign l r _ => + do l' <- retype_expr ce e l; do r' <- retype_expr ce e r; eassign l' r' + | Eassignop op l r _ _ => + do l' <- retype_expr ce e l; do r' <- retype_expr ce e r; eassignop op l' r' + | Epostincr id l _ => + do l' <- retype_expr ce e l; epostincrdecr id l' + | Ecomma r1 r2 _ => + do r1' <- retype_expr ce e r1; do r2' <- retype_expr ce e r2; ecomma r1' r2' + | Ecall r1 rl _ => + 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 _ _ _ => + Error (msg "Eloc in source") + | Eparen _ _ _ => + Error (msg "Eparen in source") + end + +with retype_exprlist (ce: composite_env) (e: typenv) (al: exprlist) : res exprlist := + match al with + | Enil => OK Enil + | Econs a1 al => + do a1' <- retype_expr ce e a1; + do al' <- retype_exprlist ce e al; + do x1 <- check_rval a1'; + OK (Econs a1' al') + end. + +Fixpoint retype_stmt (ce: composite_env) (e: typenv) (rt: type) (s: statement) : res statement := + match s with + | Sskip => + OK Sskip + | Sdo a => + do a' <- retype_expr ce e a; sdo a' + | Ssequence s1 s2 => + do s1' <- retype_stmt ce e rt s1; do s2' <- retype_stmt ce e rt s2; OK (Ssequence s1' s2') + | Sifthenelse a s1 s2 => + do a' <- retype_expr ce e a; + do s1' <- retype_stmt ce e rt s1; do s2' <- retype_stmt ce e rt s2; + sifthenelse a' s1' s2' + | Swhile a s => + do a' <- retype_expr ce e a; + do s' <- retype_stmt ce e rt s; + swhile a' s' + | Sdowhile a s => + do a' <- retype_expr ce e a; + do s' <- retype_stmt ce e rt s; + sdowhile a' s' + | Sfor s1 a s2 s3 => + do a' <- retype_expr ce e a; + do s1' <- retype_stmt ce e rt s1; do s2' <- retype_stmt ce e rt s2; do s3' <- retype_stmt ce e rt s3; + sfor s1' a' s2' s3' + | Sbreak => + OK Sbreak + | Scontinue => + OK Scontinue + | Sreturn None => + OK (Sreturn None) + | Sreturn (Some a) => + do a' <- retype_expr ce e a; + sreturn rt a' + | Sswitch a sl => + do a' <- retype_expr ce e a; + do sl' <- retype_lblstmts ce e rt sl; + sswitch a' sl' + | Slabel lbl s => + do s' <- retype_stmt ce e rt s; OK (Slabel lbl s') + | Sgoto lbl => + OK (Sgoto lbl) + end + +with retype_lblstmts (ce: composite_env) (e: typenv) (rt: type) (sl: labeled_statements) : res labeled_statements := + match sl with + | LSnil => OK LSnil + | LScons case s sl => + do s' <- retype_stmt ce e rt s; do sl' <- retype_lblstmts ce e rt sl; + OK (LScons case s' sl') + end. + +Definition retype_function (ce: composite_env) (e: typenv) (f: function) : res function := + let e := bind_vars (bind_vars e f.(fn_params)) f.(fn_vars) in + do s <- retype_stmt ce e f.(fn_return) f.(fn_body); + OK (mkfunction f.(fn_return) + f.(fn_callconv) + f.(fn_params) + f.(fn_vars) + s). + +(** Soundness of the smart constructors. *) + +Lemma check_cast_sound: + forall t1 t2 x, check_cast t1 t2 = OK x -> wt_cast t1 t2. +Proof. + unfold check_cast, wt_cast; intros. + destruct (classify_cast t1 t2); congruence. +Qed. + +Lemma check_bool_sound: + forall t x, check_bool t = OK x -> wt_bool t. +Proof. + unfold check_bool, wt_bool; intros. + destruct (classify_bool t); congruence. +Qed. + +Hint Resolve check_cast_sound check_bool_sound: ty. + +Lemma check_arguments_sound: + forall el tl x, check_arguments el tl = OK x -> wt_arguments el tl. +Proof. + intros el tl; revert tl el. + induction tl; destruct el; simpl; intros; try discriminate. + constructor. + destruct strict eqn:S; try discriminate. constructor; auto. + monadInv H. constructor; eauto with ty. +Qed. + +Lemma check_rval_sound: + forall a x, check_rval a = OK x -> expr_kind a = RV. +Proof. + unfold check_rval; intros. destruct a; reflexivity || discriminate. +Qed. + +Lemma check_lval_sound: + forall a x, check_lval a = OK x -> expr_kind a = LV. +Proof. + unfold check_lval; intros. destruct a; reflexivity || discriminate. +Qed. + +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. + unfold wt_cast, binarith_type, classify_binarith; intros; DestructCases; + simpl; split; try congruence. destruct f; congruence. +Qed. + +Lemma typeconv_cast: + forall t1 t2, wt_cast (typeconv t1) t2 -> wt_cast t1 t2. +Proof. + unfold typeconv, wt_cast; intros. destruct t1; auto. + assert (classify_cast (Tint I32 Signed a) t2 <> cast_case_default -> + classify_cast (Tint i s a) t2 <> cast_case_default). + { + unfold classify_cast. destruct t2; try congruence. destruct f; congruence. + } + destruct i; auto. +Qed. + +Lemma type_join_cast: + forall t1 t2 t, + type_join t1 t2 = OK t -> wt_cast t1 t /\ wt_cast t2 t. +Proof. + intros. unfold type_join in H. + destruct (typeconv t1) eqn:T1; try discriminate; + destruct (typeconv t2) eqn:T2; inv H. +- unfold wt_cast; simpl; split; congruence. +- eapply binarith_type_cast; eauto. +- eapply binarith_type_cast; eauto. +- split; apply typeconv_cast; unfold wt_cast. + rewrite T1; simpl; congruence. rewrite T2; simpl; congruence. +- eapply binarith_type_cast; eauto. +- eapply binarith_type_cast; eauto. +- split; apply typeconv_cast; unfold wt_cast. + rewrite T1; simpl; congruence. rewrite T2; simpl; congruence. +- split; apply typeconv_cast; unfold wt_cast. + rewrite T1; simpl; congruence. rewrite T2; simpl; congruence. +- destruct (ident_eq i i0); inv H1. + split; apply typeconv_cast; unfold wt_cast. + rewrite T1; simpl; congruence. rewrite T2; simpl; congruence. +- destruct (ident_eq i i0); inv H1. + split; apply typeconv_cast; unfold wt_cast. + rewrite T1; simpl; congruence. rewrite T2; simpl; congruence. +Qed. + +Section SOUNDNESS_CONSTRUCTORS. + +Variable ce: composite_env. +Variable e: typenv. +Variable rt: type. + +Corollary check_rval_wt: + forall a x, wt_expr ce e a -> check_rval a = OK x -> wt_rvalue ce e a. +Proof. + unfold wt_expr; intros. erewrite check_rval_sound in H by eauto. auto. +Qed. + +Corollary check_lval_wt: + forall a x, wt_expr ce e a -> check_lval a = OK x -> wt_lvalue ce e a. +Proof. + unfold wt_expr; intros. erewrite check_lval_sound in H by eauto. auto. +Qed. + +Hint Resolve check_rval_wt check_lval_wt: ty. +Hint Extern 1 (wt_expr _ _ _) => (unfold wt_expr; simpl): ty. + +Lemma evar_sound: + forall x a, evar e x = OK a -> wt_expr ce e a. +Proof. + unfold evar; intros. destruct (e!x) as [ty|] eqn:E; inv H. eauto with ty. +Qed. + +Lemma ederef_sound: + forall r a, ederef r = OK a -> wt_expr ce e r -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma efield_sound: + forall r f a, efield ce r f = OK a -> wt_expr ce e r -> wt_expr ce e a. +Proof. + intros. monadInv H. + destruct (typeof r) eqn:TR; try discriminate; + destruct (ce!i) as [co|] eqn:CE; monadInv EQ0; eauto with ty. +Qed. + +Lemma econst_int_sound: + forall n sg, wt_expr ce e (econst_int n sg). +Proof. + unfold econst_int; auto with ty. +Qed. + +Lemma econst_ptr_int_sound: + forall n ty, wt_expr ce e (econst_ptr_int n ty). +Proof. + unfold econst_ptr_int; auto with ty. +Qed. + +Lemma econst_long_sound: + forall n sg, wt_expr ce e (econst_long n sg). +Proof. + unfold econst_long; auto with ty. +Qed. + +Lemma econst_float_sound: + forall n, wt_expr ce e (econst_float n). +Proof. + unfold econst_float; auto with ty. +Qed. + +Lemma econst_single_sound: + forall n, wt_expr ce e (econst_single n). +Proof. + unfold econst_single; auto with ty. +Qed. + +Lemma evalof_sound: + forall l a, evalof l = OK a -> wt_expr ce e l -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma eaddrof_sound: + forall l a, eaddrof l = OK a -> wt_expr ce e l -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma eunop_sound: + forall op r a, eunop op r = OK a -> wt_expr ce e r -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma ebinop_sound: + forall op r1 r2 a, ebinop op r1 r2 = OK a -> wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma ecast_sound: + forall ty r a, ecast ty r = OK a -> wt_expr ce e r -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma eseqand_sound: + forall r1 r2 a, eseqand r1 r2 = OK a -> wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto 10 with ty. +Qed. + +Lemma eseqor_sound: + forall r1 r2 a, eseqor r1 r2 = OK a -> wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto 10 with ty. +Qed. + +Lemma econdition_sound: + forall r1 r2 r3 a, econdition r1 r2 r3 = OK a -> + wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e r3 -> wt_expr ce e a. +Proof. + intros. monadInv H. apply type_join_cast in EQ3. destruct EQ3. eauto 10 with ty. +Qed. + +Lemma econdition'_sound: + forall r1 r2 r3 ty a, econdition' r1 r2 r3 ty = OK a -> + wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e r3 -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto 10 with ty. +Qed. + +Lemma esizeof_sound: + forall ty, wt_expr ce e (esizeof ty). +Proof. + unfold esizeof; auto with ty. +Qed. + +Lemma ealignof_sound: + forall ty, wt_expr ce e (ealignof ty). +Proof. + unfold ealignof; auto with ty. +Qed. + +Lemma eassign_sound: + forall l r a, eassign l r = OK a -> wt_expr ce e l -> wt_expr ce e r -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto 10 with ty. +Qed. + +Lemma eassignop_sound: + forall op l r a, eassignop op l r = OK a -> wt_expr ce e l -> wt_expr ce e r -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto 10 with ty. +Qed. + +Lemma epostincrdecr_sound: + forall id l a, epostincrdecr id l = OK a -> wt_expr ce e l -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto 10 with ty. +Qed. + +Lemma ecomma_sound: + forall r1 r2 a, ecomma r1 r2 = OK a -> wt_expr ce e r1 -> wt_expr ce e r2 -> wt_expr ce e a. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma ecall_sound: + forall fn args a, ecall fn args = OK a -> wt_expr ce e fn -> wt_exprlist ce e args -> wt_expr ce e a. +Proof. + intros. monadInv H. + destruct (classify_fun (typeof fn)) eqn:CF; monadInv EQ2. + econstructor; eauto with ty. eapply check_arguments_sound; eauto. +Qed. + +Lemma ebuiltin_sound: + forall ef tyargs args tyres a, + ebuiltin ef tyargs args tyres = OK a -> wt_exprlist ce e args -> wt_expr ce e a. +Proof. + intros. monadInv H. + destruct (type_eq tyres Tvoid); simpl in EQ2; try discriminate. + destruct (opt_typ_eq (sig_res (ef_sig ef)) None); inv EQ2. + econstructor; eauto. eapply check_arguments_sound; eauto. +Qed. + +Lemma sdo_sound: + forall a s, sdo a = OK s -> wt_expr ce e a -> wt_stmt ce e rt s. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma sifthenelse_sound: + forall a s1 s2 s, sifthenelse a s1 s2 = OK s -> + wt_expr ce e a -> wt_stmt ce e rt s1 -> wt_stmt ce e rt s2 -> wt_stmt ce e rt s. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma swhile_sound: + forall a s1 s, swhile a s1 = OK s -> + wt_expr ce e a -> wt_stmt ce e rt s1 -> wt_stmt ce e rt s. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma sdowhile_sound: + forall a s1 s, sdowhile a s1 = OK s -> + wt_expr ce e a -> wt_stmt ce e rt s1 -> wt_stmt ce e rt s. +Proof. + intros. monadInv H. eauto with ty. +Qed. + +Lemma sfor_sound: + forall s1 a s2 s3 s, sfor s1 a s2 s3 = OK s -> + wt_stmt ce e rt s1 -> wt_expr ce e a -> wt_stmt ce e rt s2 -> wt_stmt ce e rt s3 -> + wt_stmt ce e rt s. +Proof. + intros. monadInv H. eauto 10 with ty. +Qed. + +Lemma sreturn_sound: + forall a s, sreturn rt a = OK s -> wt_expr ce e a -> wt_stmt ce e rt s. +Proof. + intros. monadInv H; eauto with ty. +Qed. + +Lemma sswitch_sound: + forall a sl s, sswitch a sl = OK s -> + wt_expr ce e a -> wt_lblstmts ce e rt sl -> wt_stmt ce e rt s. +Proof. + intros. monadInv H. destruct (typeof a) eqn:TA; inv EQ0. + eauto with ty. + eapply wt_Sswitch with (sz := I32); eauto with ty. +Qed. + +Lemma retype_expr_sound: + forall a a', retype_expr ce e a = OK a' -> wt_expr ce e a' +with retype_exprlist_sound: + forall al al', retype_exprlist ce e al = OK al' -> wt_exprlist ce e al'. +Proof. +- destruct a; simpl; intros a' RT; try (monadInv RT). ++ destruct v; try discriminate. + destruct ty; inv RT. apply econst_int_sound. apply econst_ptr_int_sound. + destruct ty; inv RT. apply econst_long_sound. + inv RT. apply econst_float_sound. + inv RT. apply econst_single_sound. ++ eapply evar_sound; eauto. ++ eapply efield_sound; eauto. ++ eapply evalof_sound; eauto. ++ eapply ederef_sound; eauto. ++ eapply eaddrof_sound; eauto. ++ eapply eunop_sound; eauto. ++ eapply ebinop_sound; eauto. ++ eapply ecast_sound; eauto. ++ eapply eseqand_sound; eauto. ++ eapply eseqor_sound; eauto. ++ eapply econdition_sound; eauto. ++ apply esizeof_sound. ++ apply ealignof_sound. ++ eapply eassign_sound; eauto. ++ eapply eassignop_sound; eauto. ++ eapply epostincrdecr_sound; eauto. ++ eapply ecomma_sound; eauto. ++ eapply ecall_sound; eauto. ++ eapply ebuiltin_sound; eauto. +- destruct al; simpl; intros al' RT; monadInv RT. ++ constructor. ++ constructor; eauto with ty. +Qed. + +Lemma retype_stmt_sound: + forall s s', retype_stmt ce e rt s = OK s' -> wt_stmt ce e rt s' +with retype_lblstmts_sound: + forall sl sl', retype_lblstmts ce e rt sl = OK sl' -> wt_lblstmts ce e rt sl'. +Proof. +- destruct s; simpl; intros s' RT; try (monadInv RT). ++ constructor. ++ eapply sdo_sound; eauto using retype_expr_sound. ++ constructor; eauto. ++ eapply sifthenelse_sound; eauto using retype_expr_sound. ++ eapply swhile_sound; eauto using retype_expr_sound. ++ eapply sdowhile_sound; eauto using retype_expr_sound. ++ eapply sfor_sound; eauto using retype_expr_sound. ++ constructor. ++ constructor. ++ destruct o; monadInv RT. eapply sreturn_sound; eauto using retype_expr_sound. constructor. ++ eapply sswitch_sound; eauto using retype_expr_sound. ++ constructor; eauto. ++ constructor. +- destruct sl; simpl; intros sl' RT; monadInv RT. ++ constructor. ++ constructor; eauto. +Qed. + +End SOUNDNESS_CONSTRUCTORS. + +Lemma retype_function_sound: + forall ce e f f', retype_function ce e f = OK f' -> wt_function ce e f'. +Proof. + intros. monadInv H. constructor; simpl. eapply retype_stmt_sound; eauto. +Qed. + +(** * Subject reduction *) + +(** We show that reductions preserve well-typedness *) + +Lemma pres_cast_int_int: + forall sz sg n, wt_int (cast_int_int sz sg n) sz sg. +Proof. + intros. unfold cast_int_int. destruct sz; simpl. +- destruct sg. apply Int.sign_ext_idem; omega. apply Int.zero_ext_idem; omega. +- destruct sg. apply Int.sign_ext_idem; omega. apply Int.zero_ext_idem; omega. +- auto. +- destruct (Int.eq n Int.zero); auto. +Qed. + +Hint Resolve pres_cast_int_int: ty. + +Lemma pres_sem_cast: + forall v2 ty2 v1 ty1, wt_val v1 ty1 -> sem_cast v1 ty1 ty2 = Some v2 -> wt_val v2 ty2. +Proof. + unfold sem_cast, classify_cast; induction 1; simpl; intros; DestructCases; auto with ty. +- constructor. apply (pres_cast_int_int I8 s). +- constructor. apply (pres_cast_int_int I16 s). +- destruct (Int.eq n Int.zero); auto with ty. +- constructor. apply (pres_cast_int_int I8 s). +- constructor. apply (pres_cast_int_int I16 s). +- destruct (Int64.eq n Int64.zero); auto with ty. +- constructor. apply (pres_cast_int_int I8 s). +- constructor. apply (pres_cast_int_int I16 s). +- destruct (Float.cmp Ceq f Float.zero); auto with ty. +- constructor. apply (pres_cast_int_int I8 s). +- constructor. apply (pres_cast_int_int I16 s). +- destruct (Float32.cmp Ceq f Float32.zero); auto with ty. +- destruct (Int.eq n Int.zero); auto with ty. +Qed. + +Lemma pres_sem_binarith: + forall + (sem_int: signedness -> int -> int -> option val) + (sem_long: signedness -> int64 -> int64 -> option val) + (sem_float: float -> float -> option val) + (sem_single: float32 -> float32 -> option val) + v1 ty1 v2 ty2 v ty msg, + (forall sg n1 n2, + match sem_int sg n1 n2 with None | Some (Vint _) | Some Vundef => True | _ => False end) -> + (forall sg n1 n2, + match sem_long sg n1 n2 with None | Some (Vlong _) | Some Vundef => True | _ => False end) -> + (forall n1 n2, + match sem_float n1 n2 with None | Some (Vfloat _) | Some Vundef => True | _ => False end) -> + (forall n1 n2, + match sem_single n1 n2 with None | Some (Vsingle _) | Some Vundef => True | _ => False end) -> + sem_binarith sem_int sem_long sem_float sem_single v1 ty1 v2 ty2 = Some v -> + binarith_type ty1 ty2 msg = OK ty -> + wt_val v ty. +Proof with (try discriminate). + intros. unfold sem_binarith, binarith_type in *. + set (ty' := Cop.binarith_type (classify_binarith ty1 ty2)) in *. + destruct (sem_cast v1 ty1 ty') as [v1'|] eqn:CAST1... + destruct (sem_cast v2 ty2 ty') as [v2'|] eqn:CAST2... + DestructCases. +- specialize (H s i i0). rewrite H3 in H. + destruct v; auto with ty; contradiction. +- specialize (H0 s i i0). rewrite H3 in H0. + destruct v; auto with ty; contradiction. +- specialize (H1 f f0). rewrite H3 in H1. + destruct v; auto with ty; contradiction. +- specialize (H2 f f0). rewrite H3 in H2. + destruct v; auto with ty; contradiction. +Qed. + +Lemma pres_sem_binarith_int: + forall + (sem_int: signedness -> int -> int -> option val) + (sem_long: signedness -> int64 -> int64 -> option val) + v1 ty1 v2 ty2 v ty msg, + (forall sg n1 n2, + match sem_int sg n1 n2 with None | Some (Vint _) | Some Vundef => True | _ => False end) -> + (forall sg n1 n2, + match sem_long sg n1 n2 with None | Some (Vlong _) | Some Vundef => True | _ => False end) -> + sem_binarith sem_int sem_long (fun n1 n2 => None) (fun n1 n2 => None) v1 ty1 v2 ty2 = Some v -> + binarith_int_type ty1 ty2 msg = OK ty -> + wt_val v ty. +Proof. + intros. eapply pres_sem_binarith with (msg := msg); eauto. + simpl; auto. simpl; auto. + unfold binarith_int_type, binarith_type in *. + destruct (classify_binarith ty1 ty2); congruence. +Qed. + +Lemma pres_sem_shift: + forall sem_int sem_long ty1 ty2 m ty v1 v2 v, + shift_op_type ty1 ty2 m = OK ty -> + sem_shift sem_int sem_long v1 ty1 v2 ty2 = Some v -> + wt_val v ty. +Proof. + intros. unfold shift_op_type, sem_shift in *. DestructCases; auto with ty. +Qed. + +Lemma pres_sem_cmp: + forall ty1 ty2 msg ty c v1 v2 m v, + comparison_type ty1 ty2 msg = OK ty -> + sem_cmp c v1 ty1 v2 ty2 m = Some v -> + wt_val v ty. +Proof with (try discriminate). + unfold comparison_type, sem_cmp; intros. + assert (X: forall b, wt_val (Val.of_bool b) (Tint I32 Signed noattr)). + { + intros b; destruct b; constructor; exact I. + } + assert (Y: forall ob, option_map Val.of_bool ob = Some v -> wt_val v (Tint I32 Signed noattr)). + { + intros ob EQ. destruct ob as [b|]; inv EQ. eauto. + } + destruct (classify_cmp ty1 ty2). +- inv H; eauto. +- DestructCases; eauto. +- DestructCases; eauto. +- unfold sem_binarith in H0. + set (ty' := Cop.binarith_type (classify_binarith ty1 ty2)) in *. + destruct (sem_cast v1 ty1 ty') as [v1'|]... + destruct (sem_cast v2 ty2 ty') as [v2'|]... + DestructCases; auto. +Qed. + +Lemma pres_sem_binop: + forall ce op ty1 ty2 ty v1 v2 v m, + type_binop op ty1 ty2 = OK ty -> + sem_binary_operation ce op v1 ty1 v2 ty2 m = Some v -> + wt_val v1 ty1 -> wt_val v2 ty2 -> + wt_val v ty. +Proof. + intros until m; intros TY SEM WT1 WT2. + destruct op; simpl in TY; simpl in SEM. +- (* add *) + unfold sem_add in SEM; DestructCases; auto with ty. + eapply pres_sem_binarith; eauto; intros; exact I. +- (* sub *) + unfold sem_sub in SEM; DestructCases; auto with ty. + eapply pres_sem_binarith; eauto; intros; exact I. +- (* mul *) + unfold sem_mul in SEM. eapply pres_sem_binarith; eauto; intros; exact I. +- (* div *) + unfold sem_div in SEM. eapply pres_sem_binarith; eauto; intros. + simpl; DestructMatch; auto. + simpl; DestructMatch; auto. + simpl; DestructMatch; auto. + simpl; DestructMatch; auto. +- (* mod *) + unfold sem_mod in SEM. eapply pres_sem_binarith_int; eauto; intros. + simpl; DestructMatch; auto. + simpl; DestructMatch; auto. +- (* and *) + unfold sem_and in SEM. eapply pres_sem_binarith_int; eauto; intros; exact I. +- (* or *) + unfold sem_or in SEM. eapply pres_sem_binarith_int; eauto; intros; exact I. +- (* xor *) + unfold sem_xor in SEM. eapply pres_sem_binarith_int; eauto; intros; exact I. +- (* shl *) + unfold sem_shl in SEM. eapply pres_sem_shift; eauto. +- (* shr *) + unfold sem_shr in SEM. eapply pres_sem_shift; eauto. +- (* comparisons *) + eapply pres_sem_cmp; eauto. +- eapply pres_sem_cmp; eauto. +- eapply pres_sem_cmp; eauto. +- eapply pres_sem_cmp; eauto. +- eapply pres_sem_cmp; eauto. +- eapply pres_sem_cmp; eauto. +Qed. + +Lemma pres_sem_unop: + forall op ty1 ty v1 v, + type_unop op ty1 = OK ty -> + sem_unary_operation op v1 ty1 = Some v -> + wt_val v1 ty1 -> + wt_val v ty. +Proof. + intros until v; intros TY SEM WT1. + destruct op; simpl in TY; simpl in SEM. +- (* notbool *) + unfold sem_notbool in SEM; DestructCases. + destruct (Int.eq i Int.zero); constructor; auto with ty. + destruct (Float.cmp Ceq f Float.zero); constructor; auto with ty. + destruct (Float32.cmp Ceq f Float32.zero); constructor; auto with ty. + destruct (Int.eq i Int.zero); constructor; auto with ty. + constructor; auto with ty. + destruct (Int64.eq i Int64.zero); constructor; auto with ty. +- (* notint *) + unfold sem_notint in SEM; DestructCases; auto with ty. +- (* neg *) + unfold sem_neg in SEM; DestructCases; auto with ty. +- (* absfloat *) + unfold sem_absfloat in SEM; DestructCases; auto with ty. +Qed. + +Lemma wt_load_result: + forall ty chunk v, + access_mode ty = By_value chunk -> + wt_val (Val.load_result chunk v) ty. +Proof. + intros until v; intros AC. destruct ty; simpl in AC; try discriminate. + destruct i; [destruct s|destruct s|idtac|idtac]; inv AC; simpl; destruct v; auto with ty. + constructor; red. apply Int.sign_ext_idem; omega. + constructor; red. apply Int.zero_ext_idem; omega. + constructor; red. apply Int.sign_ext_idem; omega. + constructor; red. apply Int.zero_ext_idem; omega. + constructor; red. apply Int.zero_ext_idem; omega. + inv AC; simpl; destruct v; auto with ty. + destruct f; inv AC; simpl; destruct v; auto with ty. + inv AC; simpl; destruct v; auto with ty. +Qed. + +Lemma wt_decode_val: + forall ty chunk vl, + access_mode ty = By_value chunk -> + wt_val (decode_val chunk vl) ty. +Proof. + intros until vl; intros ACC. + destruct ty; simpl in ACC; try discriminate. +- destruct i; [destruct s|destruct s|idtac|idtac]; inv ACC; unfold decode_val; + destruct (proj_bytes vl); auto with ty. + constructor; red. apply Int.sign_ext_idem; omega. + constructor; red. apply Int.zero_ext_idem; omega. + constructor; red. apply Int.sign_ext_idem; omega. + constructor; red. apply Int.zero_ext_idem; omega. + apply wt_load_result. auto. + constructor; red. apply Int.zero_ext_idem; omega. +- inv ACC. unfold decode_val. destruct (proj_bytes vl); auto with ty. +- destruct f; inv ACC; unfold decode_val; destruct (proj_bytes vl); auto with ty. +- inv ACC. unfold decode_val. destruct (proj_bytes vl); auto with ty. + apply wt_load_result. auto. +Qed. + +Lemma wt_deref_loc: + forall ge ty m b ofs t v, + deref_loc ge ty m b ofs t v -> + wt_val v ty. +Proof. + induction 1. +- (* by value, non volatile *) + simpl in H1. exploit Mem.load_result; eauto. intros EQ; rewrite EQ. + apply wt_decode_val; auto. +- (* by value, volatile *) + inv H1. + + (* truly volatile *) + eapply wt_load_result; eauto. + + (* not really volatile *) + exploit Mem.load_result; eauto. intros EQ; rewrite EQ. + apply wt_decode_val; auto. +- (* by reference *) + destruct ty; simpl in H; try discriminate; auto with ty. + destruct i; destruct s; discriminate. + destruct f; discriminate. +- (* by copy *) + destruct ty; simpl in H; try discriminate; auto with ty. + destruct i; destruct s; discriminate. + destruct f; discriminate. +Qed. + +Lemma wt_bool_cast: + forall ty, wt_bool ty -> wt_cast ty type_bool. +Proof. + unfold wt_bool, wt_cast; unfold classify_bool; intros. destruct ty; simpl in *; try congruence. destruct f; congruence. +Qed. + +Lemma wt_cast_self: + forall t1 t2, wt_cast t1 t2 -> wt_cast t2 t2. +Proof. + unfold wt_cast; intros. destruct t2; simpl in *; try congruence. + destruct i; congruence. + destruct f; congruence. +Qed. + +Lemma binarith_type_int32s: + forall ty1 msg ty2, + binarith_type ty1 type_int32s msg = OK ty2 -> + ty2 = incrdecr_type ty1. +Proof. + intros. unfold incrdecr_type. + unfold binarith_type, classify_binarith in H; simpl in H. + destruct ty1; simpl; try congruence. + destruct i; destruct s; try congruence. + destruct s; congruence. + destruct f; congruence. +Qed. + +Lemma type_add_int32s: + forall ty1 ty2, + type_binop Oadd ty1 type_int32s = OK ty2 -> + ty2 = incrdecr_type ty1. +Proof. + simpl; intros. unfold classify_add in H; destruct ty1; simpl in H; + try (eapply binarith_type_int32s; eauto; fail). + destruct i; eapply binarith_type_int32s; eauto. + inv H; auto. + inv H; auto. + inv H; auto. +Qed. + +Lemma type_sub_int32s: + forall ty1 ty2, + type_binop Osub ty1 type_int32s = OK ty2 -> + ty2 = incrdecr_type ty1. +Proof. + simpl; intros. unfold classify_sub in H; destruct ty1; simpl in H; + try (eapply binarith_type_int32s; eauto; fail). + destruct i; eapply binarith_type_int32s; eauto. + inv H; auto. + inv H; auto. + inv H; auto. +Qed. + +Lemma wt_rred: + forall ge tenv a m t a' m', + rred ge a m t a' m' -> wt_rvalue ge tenv a -> wt_rvalue ge tenv a'. +Proof. + induction 1; intros WT; inversion WT. +- (* valof *) simpl in *. constructor. eapply wt_deref_loc; eauto. +- (* addrof *) constructor; auto with ty. +- (* unop *) simpl in H4. inv H2. constructor. eapply pres_sem_unop; eauto. +- (* binop *) + simpl in H6. inv H3. inv H5. constructor. eapply pres_sem_binop; eauto. +- (* cast *) inv H2. constructor. eapply pres_sem_cast; eauto. +- (* sequand true *) subst. constructor. auto. apply wt_bool_cast; auto. + red; intros. inv H0; auto with ty. +- (* sequand false *) constructor. auto with ty. +- (* seqor true *) constructor. auto with ty. +- (* seqor false *) subst. constructor. auto. apply wt_bool_cast; auto. + red; intros. inv H0; auto with ty. +- (* condition *) constructor. destruct b; auto. destruct b; auto. red; auto. +- (* sizeof *) constructor; auto with ty. +- (* alignof *) constructor; auto with ty. +- (* assign *) inversion H5. constructor. eapply pres_sem_cast; eauto. +- (* assignop *) subst tyres l r. constructor. auto. + constructor. constructor. eapply wt_deref_loc; eauto. + auto. auto. auto. +- (* postincr *) simpl in *. subst id0 l. + exploit wt_deref_loc; eauto. intros WTV1. + constructor. + constructor. auto. rewrite <- H0 in H5. constructor. + constructor; auto. constructor. constructor. auto with ty. + subst op. destruct id. + erewrite <- type_add_int32s by eauto. auto. + erewrite <- type_sub_int32s by eauto. auto. + simpl; auto. + constructor; auto. +- (* comma *) auto. +- (* paren *) inv H3. constructor. apply H5. eapply pres_sem_cast; eauto. +- (* builtin *) subst. auto with ty. +Qed. + +Lemma wt_lred: + forall tenv ge e a m a' m', + lred ge e a m a' m' -> wt_lvalue ge tenv a -> wt_lvalue ge tenv a'. +Proof. + induction 1; intros WT; constructor. +Qed. + +Lemma rred_same_type: + forall ge a m t a' m', + rred ge a m t a' m' -> typeof a' = typeof a. +Proof. + induction 1; auto. +Qed. + +Lemma lred_same_type: + forall ge e a m a' m', + lred ge e a m a' m' -> typeof a' = typeof a. +Proof. + induction 1; auto. +Qed. + +Section WT_CONTEXT. + +Variable cenv: composite_env. +Variable tenv: typenv. +Variable a a': expr. +Hypothesis SAMETY: typeof a' = typeof a. + +Lemma wt_subexpr: + forall from to C, + context from to C -> + wt_expr_kind cenv tenv to (C a) -> + wt_expr_kind cenv tenv from a +with wt_subexprlist: + forall from C, + contextlist from C -> + wt_exprlist cenv tenv (C a) -> + wt_expr_kind cenv tenv from a. +Proof. + destruct 1; intros WT; auto; inv WT; eauto. + destruct 1; intros WT; inv WT; eauto. +Qed. + +Lemma typeof_context: + forall from to C, context from to C -> typeof (C a') = typeof (C a). +Proof. + induction 1; simpl; auto. +Qed. + +Lemma wt_arguments_context: + forall k C, contextlist k C -> + forall tyl, wt_arguments (C a) tyl -> wt_arguments (C a') tyl. +Proof. + induction 1; intros. +- inv H0. constructor; auto. rewrite (typeof_context _ _ _ H); auto. + constructor; auto. +- inv H0. constructor; auto. constructor; auto. +Qed. + +Lemma wt_context: + forall from to C, + context from to C -> + wt_expr_kind cenv tenv to (C a) -> + wt_expr_kind cenv tenv from a' -> + wt_expr_kind cenv tenv to (C a') +with wt_contextlist: + forall from C, + contextlist from C -> + wt_exprlist cenv tenv (C a) -> + wt_expr_kind cenv tenv from a' -> + wt_exprlist cenv tenv (C a'). +Proof. +- induction 1; intros WT BASE; + auto; + inv WT; + try (pose (EQTY := typeof_context _ _ _ H); rewrite <- ? EQTY; econstructor; + try (apply IHcontext; assumption); rewrite ? EQTY; eauto). +* red. econstructor; eauto. eapply wt_arguments_context; eauto. +* red. econstructor; eauto. eapply wt_arguments_context; eauto. +- induction 1; intros WT BASE. +* inv WT. constructor. apply (wt_context _ _ _ H); auto. auto. +* inv WT. constructor; auto. +Qed. + +End WT_CONTEXT. + +Section WT_SWITCH. + +Lemma wt_select_switch: + forall n ce e rt sl, + wt_lblstmts ce e rt sl -> wt_lblstmts ce e rt (select_switch n sl). +Proof. + unfold select_switch; intros. + assert (A: wt_lblstmts ce e rt (select_switch_default sl)). + { + revert sl H. induction 1; simpl; intros. + constructor. + destruct case. auto. constructor; auto. + } + assert (B: forall sl', select_switch_case n sl = Some sl' -> wt_lblstmts ce e rt sl'). + { + revert H. generalize sl. induction 1; simpl; intros. + discriminate. + destruct case; eauto. destruct (zeq z n); eauto. inv H1. econstructor; eauto. + } + destruct (select_switch_case n sl); auto. +Qed. + +Lemma wt_seq_of_ls: + forall ce e rt sl, + wt_lblstmts ce e rt sl -> wt_stmt ce e rt (seq_of_labeled_statement sl). +Proof. + induction 1; simpl. + constructor. + constructor; auto. +Qed. + +End WT_SWITCH. + +Section PRESERVATION. + +Variable prog: program. +Hypothesis WTPROG: wt_program prog. +Let ge := globalenv prog. +Let gtenv := bind_globdef (PTree.empty _) prog.(prog_defs). + +Hypothesis WT_EXTERNAL: + forall id ef args res cc vargs m t vres m', + In (id, Gfun (External ef args res cc)) prog.(prog_defs) -> + external_call ef ge vargs m t vres m' -> + wt_val vres res. + +Inductive wt_expr_cont: typenv -> function -> cont -> Prop := + | wt_Kdo: forall te f k, + wt_stmt_cont te f k -> + wt_expr_cont te f (Kdo k) + | wt_Kifthenelse: forall te f s1 s2 k, + wt_stmt_cont te f k -> + wt_stmt ge te f.(fn_return) s1 -> wt_stmt ge te f.(fn_return) s2 -> + wt_expr_cont te f (Kifthenelse s1 s2 k) + | wt_Kwhile1: forall te f r s k, + wt_stmt_cont te f k -> + wt_rvalue ge te r -> wt_stmt ge te f.(fn_return) s -> wt_bool (typeof r) -> + wt_expr_cont te f (Kwhile1 r s k) + | wt_Kdowhile2: forall te f r s k, + wt_stmt_cont te f k -> + wt_rvalue ge te r -> wt_stmt ge te f.(fn_return) s -> wt_bool (typeof r) -> + wt_expr_cont te f (Kdowhile2 r s k) + | wt_Kfor2: forall te f r s2 s3 k, + wt_stmt_cont te f k -> + wt_rvalue ge te r -> wt_stmt ge te f.(fn_return) s2 -> wt_stmt ge te f.(fn_return) s3 -> + classify_bool (typeof r) <> bool_default -> + wt_expr_cont te f (Kfor2 r s2 s3 k) + | wt_Kswitch1: forall te f ls k, + wt_stmt_cont te f k -> + wt_lblstmts ge te f.(fn_return) ls -> + wt_expr_cont te f (Kswitch1 ls k) + | wt_Kreturn: forall te f k, + wt_stmt_cont te f k -> + wt_expr_cont te f (Kreturn k) + +with wt_stmt_cont: typenv -> function -> cont -> Prop := + | wt_Kseq: forall te f s k, + wt_stmt_cont te f k -> + wt_stmt ge te f.(fn_return) s -> + wt_stmt_cont te f (Kseq s k) + | wt_Kwhile2: forall te f r s k, + wt_stmt_cont te f k -> + wt_rvalue ge te r -> wt_stmt ge te f.(fn_return) s -> wt_bool (typeof r) -> + wt_stmt_cont te f (Kwhile2 r s k) + | wt_Kdowhile1: forall te f r s k, + wt_stmt_cont te f k -> + wt_rvalue ge te r -> wt_stmt ge te f.(fn_return) s -> wt_bool (typeof r) -> + wt_stmt_cont te f (Kdowhile1 r s k) + | wt_Kfor3: forall te f r s2 s3 k, + wt_stmt_cont te f k -> + wt_rvalue ge te r -> wt_stmt ge te f.(fn_return) s2 -> wt_stmt ge te f.(fn_return) s3 -> + wt_bool (typeof r) -> + wt_stmt_cont te f (Kfor3 r s2 s3 k) + | wt_Kfor4: forall te f r s2 s3 k, + wt_stmt_cont te f k -> + wt_rvalue ge te r -> wt_stmt ge te f.(fn_return) s2 -> wt_stmt ge te f.(fn_return) s3 -> + wt_bool (typeof r) -> + wt_stmt_cont te f (Kfor4 r s2 s3 k) + | wt_Kswitch2: forall te f k, + wt_stmt_cont te f k -> + wt_stmt_cont te f (Kswitch2 k) + | wt_Kstop': forall te f, + wt_stmt_cont te f Kstop + | wt_Kcall': forall te f f' e C ty k, + wt_call_cont (Kcall f' e C ty k) ty -> + ty = f.(fn_return) -> + wt_stmt_cont te f (Kcall f' e C ty k) + +with wt_call_cont: cont -> type -> Prop := + | wt_Kstop: forall ty, + wt_call_cont Kstop ty + | wt_Kcall: forall te f e C ty k, + wt_expr_cont te f k -> + wt_stmt ge te f.(fn_return) f.(fn_body) -> + (forall v, wt_val v ty -> wt_rvalue ge te (C (Eval v ty))) -> + wt_call_cont (Kcall f e C ty k) ty. + +Lemma is_wt_call_cont: + forall te f k, + is_call_cont k -> wt_stmt_cont te f k -> wt_call_cont k f.(fn_return). +Proof. + intros. inv H0; simpl in H; try contradiction. constructor. auto. +Qed. + +Lemma wt_call_cont_stmt_cont: + forall te f k, wt_call_cont k f.(fn_return) -> wt_stmt_cont te f k. +Proof. + intros. inversion H; subst. constructor. constructor; auto. +Qed. + +Lemma call_cont_wt: + forall e f k, wt_stmt_cont e f k -> wt_call_cont (call_cont k) f.(fn_return). +Proof. + induction 1; simpl; auto. + constructor. + congruence. +Qed. + +Lemma call_cont_wt': + forall e f k, wt_stmt_cont e f k -> wt_stmt_cont e f (call_cont k). +Proof. + induction 1; simpl; auto; econstructor; eauto. +Qed. + +Definition wt_fundef (fd: fundef) := + match fd with + | Internal f => wt_function ge gtenv f + | External ef targs tres cc => True + end. + +Definition fundef_return (fd: fundef) : type := + match fd with + | Internal f => f.(fn_return) + | External ef targs tres cc => tres + end. + +Lemma wt_find_funct: + forall v fd, Genv.find_funct ge v = Some fd -> wt_fundef fd. +Proof. + intros. apply Genv.find_funct_prop with (p := prog) (v := v); auto. + intros. inv WTPROG. destruct f; simpl; auto. apply H1 with id; auto. +Qed. + +Inductive wt_state: state -> Prop := + | wt_stmt_state: forall f s k e m te + (WTK: wt_stmt_cont te f k) + (WTB: wt_stmt ge te f.(fn_return) f.(fn_body)) + (WTS: wt_stmt ge te f.(fn_return) s), + wt_state (State f s k e m) + | wt_expr_state: forall f r k e m te + (WTK: wt_expr_cont te f k) + (WTB: wt_stmt ge te f.(fn_return) f.(fn_body)) + (WTE: wt_rvalue ge te r), + wt_state (ExprState f r k e m) + | wt_call_state: forall b fd vargs k m + (WTK: wt_call_cont k (fundef_return fd)) + (WTFD: wt_fundef fd) + (FIND: Genv.find_funct ge b = Some fd), + wt_state (Callstate fd vargs k m) + | wt_return_state: forall v k m ty + (WTK: wt_call_cont k ty) + (VAL: wt_val v ty), + wt_state (Returnstate v k m) + | wt_stuck_state: + wt_state Stuckstate. + +Hint Constructors wt_expr_cont wt_stmt_cont wt_stmt wt_state: ty. + +Section WT_FIND_LABEL. + +Scheme wt_stmt_ind2 := Minimality for wt_stmt Sort Prop + with wt_lblstmts2_ind2 := Minimality for wt_lblstmts Sort Prop. + +Lemma wt_find_label: + forall lbl e f s, wt_stmt ge e f.(fn_return) s -> + forall k s' k', + find_label lbl s k = Some (s', k') -> + wt_stmt_cont e f k -> + wt_stmt ge e f.(fn_return) s' /\ wt_stmt_cont e f k'. +Proof. + intros lbl e f s0 WTS0. pattern s0. + apply (wt_stmt_ind2 ge e f.(fn_return)) with + (P0 := fun ls => wt_lblstmts ge e f.(fn_return) ls -> + forall k s' k', + find_label_ls lbl ls k = Some (s', k') -> + wt_stmt_cont e f k -> + wt_stmt ge e f.(fn_return) s' /\ wt_stmt_cont e f k'); + simpl; intros; try discriminate. + + destruct (find_label lbl s1 (Kseq s2 k)) as [[sx kx] | ] eqn:F. + inv H3. eauto with ty. + eauto with ty. + + destruct (find_label lbl s1 k) as [[sx kx] | ] eqn:F. + inv H5. eauto with ty. + eauto with ty. + + eauto with ty. + + eauto with ty. + + destruct (find_label lbl s1 (Kseq (Sfor Sskip r s2 s3) k)) as [[sx kx] | ] eqn:F. + inv H7. eauto with ty. + destruct (find_label lbl s3 (Kfor3 r s2 s3 k)) as [[sx kx] | ] eqn:F2. + inv H7. eauto with ty. + eauto with ty. + + eauto with ty. + + destruct (ident_eq lbl lbl0). + inv H1. auto. + eauto. + + destruct (find_label lbl s (Kseq (seq_of_labeled_statement ls) k)) as [[sx kx] | ] eqn:F. + inv H4. eapply H0; eauto. constructor. auto. apply wt_seq_of_ls; auto. + eauto. + + assumption. +Qed. + +End WT_FIND_LABEL. + + +Lemma preservation_estep: + forall S t S', estep ge S t S' -> wt_state S -> wt_state S'. +Proof. + induction 1; intros WT; inv WT. +- (* lred *) + econstructor; eauto. change (wt_expr_kind ge te RV (C a')). + eapply wt_context with (a := a); eauto. + eapply lred_same_type; eauto. + eapply wt_lred; eauto. change (wt_expr_kind ge te LV a). eapply wt_subexpr; eauto. +- (* rred *) + econstructor; eauto. change (wt_expr_kind ge te RV (C a')). + eapply wt_context with (a := a); eauto. + eapply rred_same_type; eauto. + eapply wt_rred; eauto. change (wt_expr_kind ge te RV a). eapply wt_subexpr; eauto. +- (* call *) + assert (A: wt_expr_kind ge te RV a) by (eapply wt_subexpr; eauto). + simpl in A. inv H. inv A. simpl in H9; rewrite H4 in H9; inv H9. + assert (fundef_return fd = ty). + { destruct fd; simpl in *. + unfold type_of_function in H3. congruence. + congruence. } + econstructor. + rewrite H. econstructor; eauto. + intros. change (wt_expr_kind ge te RV (C (Eval v ty))). + eapply wt_context with (a := Ecall (Eval vf tyf) el ty); eauto. + red; constructor; auto. + eapply wt_find_funct; eauto. + eauto. +- (* stuck *) + constructor. +Qed. + +Lemma preservation_sstep: + forall S t S', sstep ge S t S' -> wt_state S -> wt_state S'. +Proof. + induction 1; intros WT; inv WT. +- inv WTS; eauto with ty. +- inv WTK; eauto with ty. +- inv WTS; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTS; eauto with ty. +- inv WTK; destruct b; eauto with ty. +- inv WTS; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTS; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTS; eauto with ty. +- inv WTS; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- econstructor. eapply call_cont_wt; eauto. constructor. +- inv WTS. eauto with ty. +- inv WTK. econstructor. eapply call_cont_wt; eauto. + inv WTE. eapply pres_sem_cast; eauto. +- econstructor. eapply is_wt_call_cont; eauto. constructor. +- inv WTS; eauto with ty. +- inv WTK. econstructor; eauto with ty. + apply wt_seq_of_ls. apply wt_select_switch; auto. +- inv WTK; eauto with ty. +- inv WTK; eauto with ty. +- inv WTS; eauto with ty. +- exploit wt_find_label. eexact WTB. eauto. eapply call_cont_wt'; eauto. + intros [A B]. eauto with ty. +- simpl in WTFD; inv WTFD. econstructor; eauto. apply wt_call_cont_stmt_cont; auto. +- exploit (Genv.find_funct_inversion prog); eauto. intros (id & A). + econstructor; eauto. +- inv WTK. eauto with ty. +Qed. + +Theorem preservation: + forall S t S', step ge S t S' -> wt_state S -> wt_state S'. +Proof. + intros. destruct H. eapply preservation_estep; eauto. eapply preservation_sstep; eauto. +Qed. + +Theorem wt_initial_state: + forall S, initial_state prog S -> wt_state S. +Proof. + intros. inv H. econstructor. constructor. + apply Genv.find_funct_ptr_prop with (p := prog) (b := b); auto. + intros. inv WTPROG. destruct f0; simpl; auto. apply H4 with id; auto. + instantiate (1 := (Vptr b Int.zero)). rewrite Genv.find_funct_find_funct_ptr. auto. +Qed. + +End PRESERVATION. + + + + + + + + + + + + + + + + + + + + + + + + + -- cgit From 2dd864217cc864d44e828a4d14dd45668e4ab095 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sat, 10 Jan 2015 11:41:41 +0100 Subject: Define a nonnegative integer "rank" for types to support structural induction over composite types. --- cfrontend/Ctypes.v | 134 ++++++++++++++++++++++++++++++++++++------ cfrontend/Initializersproof.v | 8 +-- 2 files changed, 120 insertions(+), 22 deletions(-) (limited to 'cfrontend') diff --git a/cfrontend/Ctypes.v b/cfrontend/Ctypes.v index 091c5276..a555f792 100644 --- a/cfrontend/Ctypes.v +++ b/cfrontend/Ctypes.v @@ -168,6 +168,7 @@ Record composite : Type := { co_attr: attr; co_sizeof: Z; co_alignof: Z; + co_rank: nat; co_sizeof_pos: co_sizeof >= 0; co_alignof_two_p: exists n, co_alignof = two_power_nat n; co_sizeof_alignof: (co_alignof | co_sizeof) @@ -662,6 +663,31 @@ Proof. destruct (env!i). apply X. apply Zdivide_0. Qed. +(** Type ranks *) + +(** The rank of a type is a nonnegative integer that measures the direct nesting + of arrays, struct and union types. It does not take into account indirect + nesting such as a struct type that appears under a pointer or function type. + Type ranks ensure that type expressions (ignoring pointer and function types) + have an inductive structure. *) + +Fixpoint rank_type (ce: composite_env) (t: type) : nat := + match t with + | Tarray t' _ _ => S (rank_type ce t') + | Tstruct id _ | Tunion id _ => + match ce!id with + | None => O + | Some co => S (co_rank co) + end + | _ => O + end. + +Fixpoint rank_members (ce: composite_env) (m: members) : nat := + match m with + | nil => 0%nat + | (id, t) :: m => Peano.max (rank_type ce t) (rank_members ce m) + end. + (** ** C types and back-end types *) (** Extracting a type list from a function parameter declaration. *) @@ -722,6 +748,15 @@ Fixpoint complete_members (env: composite_env) (m: members) : bool := | (id, t) :: m' => complete_type env t && complete_members env m' end. +Lemma complete_member: + forall env id t m, + In (id, t) m -> complete_members env m = true -> complete_type env t = true. +Proof. + induction m as [|[id1 t1] m]; simpl; intuition auto. + InvBooleans; inv H1; auto. + InvBooleans; eauto. +Qed. + (** Convert a composite definition to its internal representation. The size and alignment of the composite are determined at this time. The alignment takes into account the [__Alignas] attributes @@ -759,6 +794,7 @@ Program Definition composite_of_def co_attr := a; co_sizeof := align (sizeof_composite env su m) al; co_alignof := al; + co_rank := rank_members env m; co_sizeof_pos := _; co_alignof_two_p := _; co_sizeof_alignof := _ |} @@ -793,8 +829,8 @@ Fixpoint add_composite_definitions (env: composite_env) (defs: list composite_de Definition build_composite_env (defs: list composite_definition) := add_composite_definitions (PTree.empty _) defs. -(** Stability properties for alignments and sizes. If the type is - complete in a composite environment [env], its size and alignment +(** Stability properties for alignments, sizes, and ranks. If the type is + complete in a composite environment [env], its size, alignment, and rank are unchanged if we add more definitions to [env]. *) Section STABILITY. @@ -833,6 +869,16 @@ Proof. erewrite extends by eauto. auto. Qed. +Lemma rank_type_stable: + forall t, complete_type env t = true -> rank_type env' t = rank_type env t. +Proof. + induction t; simpl; intros; auto. + destruct (env!i) as [co|] eqn:E; try discriminate. + erewrite extends by eauto. auto. + destruct (env!i) as [co|] eqn:E; try discriminate. + erewrite extends by eauto. auto. +Qed. + Lemma alignof_composite_stable: forall m, complete_members env m = true -> alignof_composite env' m = alignof_composite env m. Proof. @@ -874,6 +920,14 @@ Proof. InvBooleans. rewrite complete_type_stable by auto. rewrite IHm by auto. auto. Qed. +Lemma rank_members_stable: + forall m, complete_members env m = true -> rank_members env' m = rank_members env m. +Proof. + induction m as [|[id t]]; simpl; intros. + auto. + InvBooleans. f_equal; auto. apply rank_type_stable; auto. +Qed. + End STABILITY. Lemma add_composite_definitions_incr: @@ -892,10 +946,16 @@ Qed. environment produced by [build_composite_env] are consistent with the sizes and alignments of the members of the composite types. *) -Definition composite_consistent (env: composite_env) (co: composite) : Prop := - complete_members env (co_members co) = true - /\ co_alignof co = align_attr (co_attr co) (alignof_composite env (co_members co)) - /\ co_sizeof co = align (sizeof_composite env (co_su co) (co_members co)) (co_alignof co). +Record composite_consistent (env: composite_env) (co: composite) : Prop := { + co_consistent_complete: + complete_members env (co_members co) = true; + co_consistent_alignof: + co_alignof co = align_attr (co_attr co) (alignof_composite env (co_members co)); + co_consistent_sizeof: + co_sizeof co = align (sizeof_composite env (co_su co) (co_members co)) (co_alignof co); + co_consistent_rank: + co_rank co = rank_members env (co_members co) +}. Definition composite_env_consistent (env: composite_env) : Prop := forall id co, env!id = Some co -> composite_consistent env co. @@ -920,20 +980,21 @@ Proof. { intros. unfold env1. rewrite PTree.gso; auto. congruence. } red; intros. unfold env1 in H2; rewrite PTree.gsspec in H2; destruct (peq id0 id). + subst id0. inversion H2; clear H2. subst co. - red; rewrite <- H1; simpl. +(* assert (A: alignof_composite env1 m = alignof_composite env0 m) by (apply alignof_composite_stable; assumption). - rewrite A. - split. eapply complete_members_stable; eauto. - split. auto. - f_equal. symmetry. apply sizeof_composite_stable; assumption. -+ exploit H0; eauto. intros (P & Q & R). - assert (A: alignof_composite env1 (co_members co) = alignof_composite env0 (co_members co)) - by (apply alignof_composite_stable; assumption). - red. rewrite A. - split. eapply complete_members_stable; eauto. - split. auto. - rewrite R. f_equal. symmetry. apply sizeof_composite_stable; assumption. +*) + rewrite <- H1; constructor; simpl. +* eapply complete_members_stable; eauto. +* f_equal. symmetry. apply alignof_composite_stable; auto. +* f_equal. symmetry. apply sizeof_composite_stable; auto. +* symmetry. apply rank_members_stable; auto. ++ exploit H0; eauto. intros [P Q R S]. + constructor; intros. +* eapply complete_members_stable; eauto. +* rewrite Q. f_equal. symmetry. apply alignof_composite_stable; auto. +* rewrite R. f_equal. symmetry. apply sizeof_composite_stable; auto. +* rewrite S. symmetry; apply rank_members_stable; auto. Qed. (** Moreover, every composite definition is reflected in the composite environment. *) @@ -956,3 +1017,40 @@ Proof. split. eapply add_composite_definitions_incr; eauto. apply PTree.gss. subst x; auto. Qed. + +(** As a corollay, in a consistent environment, the rank of a composite type + 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. +Proof. + induction m; simpl; intros; intuition auto. + subst a. xomega. + xomega. +Qed. + +Lemma rank_struct_member: + forall ce id a co id1 t1, + 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. +Proof. + intros; simpl. rewrite H0. + erewrite co_consistent_rank by eauto. + exploit (rank_type_members ce); eauto. + omega. +Qed. + +Lemma rank_union_member: + forall ce id a co id1 t1, + 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. +Proof. + intros; simpl. rewrite H0. + erewrite co_consistent_rank by eauto. + exploit (rank_type_members ce); eauto. + omega. +Qed. diff --git a/cfrontend/Initializersproof.v b/cfrontend/Initializersproof.v index 225dc23e..02a453cf 100644 --- a/cfrontend/Initializersproof.v +++ b/cfrontend/Initializersproof.v @@ -600,8 +600,8 @@ Proof. replace (idlsize data) with (idlsize data + 0) by omega. eapply (proj2 (transl_init_list_size il)). eauto. unfold lookup_composite in EQ. simpl. destruct (ge.(genv_cenv)!i) as [co|] eqn:?; inv EQ. - exploit ce_consistent; eauto. intros (A & B & C). - rewrite C. unfold sizeof_composite. rewrite Heqs. apply align_le. + erewrite co_consistent_sizeof by (eapply ce_consistent; eauto). + unfold sizeof_composite. rewrite Heqs. apply align_le. destruct (co_alignof_two_p x) as [n EQ]. rewrite EQ. apply two_power_nat_pos. + (* union *) simpl in H. destruct ty; try discriminate. @@ -612,8 +612,8 @@ Proof. rewrite padding_size. omega. apply Zle_trans with (sizeof_union ge (co_members x)). eapply union_field_size; eauto. - exploit ce_consistent; eauto. intros (A & B & C). - rewrite C. unfold sizeof_composite. rewrite Heqs. apply align_le. + erewrite co_consistent_sizeof by (eapply ce_consistent; eauto). + unfold sizeof_composite. rewrite Heqs. apply align_le. destruct (co_alignof_two_p x) as [n EQ]. rewrite EQ. apply two_power_nat_pos. - induction il. -- cgit From 9f2ca1049957004161834090a697cd4118e2fb08 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Tue, 20 Jan 2015 14:54:48 +0100 Subject: Protect against redefinition of the __i64_xxx helper library functions. This is achieved by declaring these functions in C2C.ml, then re-checking their declarations in the global environment during the Selection pass. In passing, the "hf" parameter for SelectLong functions was removed and replaced by fixed identifiers corresponding to the actual names of the helper functions. --- cfrontend/C2C.ml | 55 ++++++++++++++++++++++++++++++++++++++++++++++++++++++- 1 file changed, 54 insertions(+), 1 deletion(-) (limited to 'cfrontend') diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index 118b6d2d..584c265a 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -188,7 +188,60 @@ let builtins_generic = { "__compcert_va_float64", (TFloat(FDouble, []), [TPtr(TVoid [], [])], - false) + false); + (* Helper functions for int64 arithmetic *) + "__i64_dtos", + (TInt(ILongLong, []), + [TFloat(FDouble, [])], + false); + "__i64_dtou", + (TInt(IULongLong, []), + [TFloat(FDouble, [])], + false); + "__i64_stod", + (TFloat(FDouble, []), + [TInt(ILongLong, [])], + false); + "__i64_utod", + (TFloat(FDouble, []), + [TInt(IULongLong, [])], + false); + "__i64_stof", + (TFloat(FFloat, []), + [TInt(ILongLong, [])], + false); + "__i64_utof", + (TFloat(FFloat, []), + [TInt(IULongLong, [])], + false); + "__i64_sdiv", + (TInt(ILongLong, []), + [TInt(ILongLong, []); TInt(ILongLong, [])], + false); + "__i64_udiv", + (TInt(IULongLong, []), + [TInt(IULongLong, []); TInt(IULongLong, [])], + false); + "__i64_smod", + (TInt(ILongLong, []), + [TInt(ILongLong, []); TInt(ILongLong, [])], + false); + "__i64_umod", + (TInt(IULongLong, []), + [TInt(IULongLong, []); TInt(IULongLong, [])], + false); + "__i64_shl", + (TInt(ILongLong, []), + [TInt(ILongLong, []); TInt(IInt, [])], + false); + "__i64_shr", + (TInt(IULongLong, []), + [TInt(IULongLong, []); TInt(IInt, [])], + false); + "__i64_sar", + (TInt(ILongLong, []), + [TInt(ILongLong, []); TInt(IInt, [])], + false) ] } -- cgit From 1e97bb4f6297b6fa7949684e522a592aab754d99 Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Thu, 22 Jan 2015 19:26:05 +0100 Subject: Delay reads from !Machine.config before it is properly initialized. Several definitions in Cutil and elsewhere were accessing the default value of !Machine.config, before it is properly initialized in Driver. Delay evaluation of these definitions, and initialize Machine.config to nonsensical values so that lack of initialization is caught early (e.g. in Cutil.find_matching_*_kind). Also, following up on commit [3b8a094], don't use "wchar_t" typedef to type wide string literals, even if this typedef is in scope. The risk here is to hide an inconsistency between "wchar_t"'s definition in standard headers and CompCert's built-in definition. --- cfrontend/C2C.ml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) (limited to 'cfrontend') diff --git a/cfrontend/C2C.ml b/cfrontend/C2C.ml index 584c265a..fddbfd30 100644 --- a/cfrontend/C2C.ml +++ b/cfrontend/C2C.ml @@ -137,8 +137,8 @@ let builtins_generic = { (TVoid [], [TPtr(TVoid [], []); TPtr(TVoid [AConst], []); - TInt(Cutil.size_t_ikind, []); - TInt(Cutil.size_t_ikind, [])], + TInt(IUInt, []); + TInt(IUInt, [])], false); (* Annotations *) "__builtin_annot", -- cgit From f2b1c25aa56a27836652aef3feeee0856c04235c Mon Sep 17 00:00:00 2001 From: Xavier Leroy Date: Sun, 8 Feb 2015 16:41:45 +0100 Subject: Interpreter produces more detailed trace, including name of semantic rules used. Cexec: record names of rules used in every reduction. Interp: print these rule names in -trace mode. Also: simplified the exploration in "-all" mode; give unique names to states. Csem: fix name of reduction rule "red_call". --- cfrontend/Cexec.v | 299 ++++++++++++++++++++++++++++++------------------------ cfrontend/Csem.v | 2 +- 2 files changed, 165 insertions(+), 136 deletions(-) (limited to 'cfrontend') diff --git a/cfrontend/Cexec.v b/cfrontend/Cexec.v index 952d148d..ed67286f 100644 --- a/cfrontend/Cexec.v +++ b/cfrontend/Cexec.v @@ -12,6 +12,7 @@ (** Animating the CompCert C semantics *) +Require Import String. Require Import Axioms. Require Import Classical. Require Import Coqlib. @@ -31,6 +32,9 @@ Require Import Csyntax. Require Import Csem. Require Cstrategy. +Local Open Scope string_scope. +Local Open Scope list_scope. + (** Error monad with options or lists *) Notation "'do' X <- A ; B" := (match A with Some X => B | None => None end) @@ -661,9 +665,9 @@ Qed. (** * Reduction of expressions *) Inductive reduction: Type := - | Lred (l': expr) (m': mem) - | Rred (r': expr) (m': mem) (t: trace) - | Callred (fd: fundef) (args: list val) (tyres: type) (m': mem) + | Lred (rule: string) (l': expr) (m': mem) + | Rred (rule: string) (r': expr) (m': mem) (t: trace) + | Callred (rule: string) (fd: fundef) (args: list val) (tyres: type) (m': mem) | Stuckred. Section EXPRS. @@ -728,15 +732,15 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match e!x with | Some(b, ty') => check type_eq ty ty'; - topred (Lred (Eloc b Int.zero ty) m) + topred (Lred "red_var_local" (Eloc b Int.zero ty) m) | None => do b <- Genv.find_symbol ge x; - topred (Lred (Eloc b Int.zero ty) m) + topred (Lred "red_var_global" (Eloc b Int.zero ty) m) end | LV, Ederef r ty => match is_val r with | Some(Vptr b ofs, ty') => - topred (Lred (Eloc b ofs ty) m) + topred (Lred "red_deref" (Eloc b ofs ty) m) | Some _ => stuck | None => @@ -750,11 +754,11 @@ 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 (Eloc b (Int.add ofs (Int.repr delta)) ty) m) + | OK delta => topred (Lred "red_field_struct" (Eloc b (Int.add ofs (Int.repr delta)) ty) m) end | Tunion id _ => do co <- ge.(genv_cenv)!id; - topred (Lred (Eloc b ofs ty) m) + topred (Lred "red_field_union" (Eloc b ofs ty) m) | _ => stuck end | Some _ => @@ -769,20 +773,20 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := | Some(b, ofs, ty') => check type_eq ty ty'; do w',t,v <- do_deref_loc w ty m b ofs; - topred (Rred (Eval v ty) m t) + 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 (Eval (Vptr b ofs) ty) m E0) + | Some(b, ofs, ty') => topred (Rred "red_addrof" (Eval (Vptr b ofs) ty) m E0) | None => incontext (fun x => Eaddrof x ty) (step_expr LV l m) end | RV, Eunop op r1 ty => match is_val r1 with | Some(v1, ty1) => do v <- sem_unary_operation op v1 ty1; - topred (Rred (Eval v ty) m E0) + topred (Rred "red_unop" (Eval v ty) m E0) | None => incontext (fun x => Eunop op x ty) (step_expr RV r1 m) end @@ -790,7 +794,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match is_val r1, is_val r2 with | Some(v1, ty1), Some(v2, ty2) => do v <- sem_binary_operation ge op v1 ty1 v2 ty2 m; - topred (Rred (Eval v ty) m E0) + topred (Rred "red_binop" (Eval v ty) m E0) | _, _ => incontext2 (fun x => Ebinop op x r2 ty) (step_expr RV r1 m) (fun x => Ebinop op r1 x ty) (step_expr RV r2 m) @@ -799,7 +803,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match is_val r1 with | Some(v1, ty1) => do v <- sem_cast v1 ty1 ty; - topred (Rred (Eval v ty) m E0) + topred (Rred "red_cast" (Eval v ty) m E0) | None => incontext (fun x => Ecast x ty) (step_expr RV r1 m) end @@ -807,8 +811,8 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match is_val r1 with | Some(v1, ty1) => do b <- bool_val v1 ty1; - if b then topred (Rred (Eparen r2 type_bool ty) m E0) - else topred (Rred (Eval (Vint Int.zero) ty) m E0) + if b then topred (Rred "red_seqand_true" (Eparen r2 type_bool ty) m E0) + else topred (Rred "red_seqand_false" (Eval (Vint Int.zero) ty) m E0) | None => incontext (fun x => Eseqand x r2 ty) (step_expr RV r1 m) end @@ -816,8 +820,8 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match is_val r1 with | Some(v1, ty1) => do b <- bool_val v1 ty1; - if b then topred (Rred (Eval (Vint Int.one) ty) m E0) - else topred (Rred (Eparen r2 type_bool ty) m E0) + if b then topred (Rred "red_seqor_true" (Eval (Vint Int.one) ty) m E0) + else topred (Rred "red_seqor_false" (Eparen r2 type_bool ty) m E0) | None => incontext (fun x => Eseqor x r2 ty) (step_expr RV r1 m) end @@ -825,21 +829,21 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match is_val r1 with | Some(v1, ty1) => do b <- bool_val v1 ty1; - topred (Rred (Eparen (if b then r2 else r3) ty ty) m E0) + topred (Rred "red_condition" (Eparen (if b then r2 else r3) ty ty) m E0) | None => incontext (fun x => Econdition x r2 r3 ty) (step_expr RV r1 m) end | RV, Esizeof ty' ty => - topred (Rred (Eval (Vint (Int.repr (sizeof ge ty'))) ty) m E0) + topred (Rred "red_sizeof" (Eval (Vint (Int.repr (sizeof ge ty'))) ty) m E0) | RV, Ealignof ty' ty => - topred (Rred (Eval (Vint (Int.repr (alignof ge ty'))) ty) m E0) + topred (Rred "red_alignof" (Eval (Vint (Int.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) => check type_eq ty1 ty; do v <- sem_cast v2 ty2 ty1; do w',t,m' <- do_assign_loc w ty1 m b ofs v; - topred (Rred (Eval v ty) m' t) + 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) @@ -851,7 +855,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := do w',t,v1 <- do_deref_loc w ty1 m b ofs; let r' := Eassign (Eloc b ofs ty1) (Ebinop op (Eval v1 ty1) (Eval v2 ty2) tyres) ty1 in - topred (Rred r' m t) + topred (Rred "red_assignop" r' m t) | _, _ => incontext2 (fun x => Eassignop op x r2 tyres ty) (step_expr LV l1 m) (fun x => Eassignop op l1 x tyres ty) (step_expr RV r2 m) @@ -867,7 +871,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := (Ebinop op (Eval v1 ty) (Eval (Vint Int.one) type_int32s) (incrdecr_type ty)) ty) (Eval v1 ty) ty in - topred (Rred r' m t) + topred (Rred "red_postincr" r' m t) | None => incontext (fun x => Epostincr id x ty) (step_expr LV l m) end @@ -875,7 +879,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match is_val r1 with | Some _ => check type_eq (typeof r2) ty; - topred (Rred r2 m E0) + topred (Rred "red_comma" r2 m E0) | None => incontext (fun x => Ecomma x r2 ty) (step_expr RV r1 m) end @@ -883,7 +887,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := match is_val r1 with | Some (v1, ty1) => do v <- sem_cast v1 ty1 tycast; - topred (Rred (Eval v ty) m E0) + topred (Rred "red_paren" (Eval v ty) m E0) | None => incontext (fun x => Eparen x tycast ty) (step_expr RV r1 m) end @@ -895,7 +899,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := do fd <- Genv.find_funct ge vf; do vargs <- sem_cast_arguments vtl tyargs; check type_eq (type_of_fundef fd) (Tfunction tyargs tyres cconv); - topred (Callred fd vargs ty m) + topred (Callred "red_call" fd vargs ty m) | _ => stuck end | _, _ => @@ -908,7 +912,7 @@ Fixpoint step_expr (k: kind) (a: expr) (m: mem): reducts expr := do vargs <- sem_cast_arguments vtl tyargs; match do_external ef w vargs m with | None => stuck - | Some(w',t,v,m') => topred (Rred (Eval v ty) m' t) + | Some(w',t,v,m') => topred (Rred "red_builtin" (Eval v ty) m' t) end | _ => incontext (fun x => Ebuiltin ef tyargs x ty) (step_exprlist rargs m) @@ -956,21 +960,6 @@ Proof. eapply imm_safe_callred; eauto. Qed. -(* -Definition not_stuck (a: expr) (m: mem) := - forall a' k C, context k RV C -> a = C a' -> imm_safe_t k a' m. - -Lemma safe_expr_kind: - forall k C a m, - context k RV C -> - not_stuck (C a) m -> - k = Cstrategy.expr_kind a. -Proof. - intros. - symmetry. eapply Cstrategy.imm_safe_kind. eapply imm_safe_t_imm_safe. eauto. -Qed. -*) - Fixpoint exprlist_all_values (rl: exprlist) : Prop := match rl with | Enil => True @@ -1163,9 +1152,9 @@ Hint Resolve context_compose contextlist_compose. Definition reduction_ok (k: kind) (a: expr) (m: mem) (rd: reduction) : Prop := match k, rd with - | LV, Lred l' m' => lred ge e a m l' m' - | RV, Rred r' m' t => rred ge a m t r' m' /\ exists w', possible_trace w t w' - | RV, Callred fd args tyres m' => callred ge a fd args tyres /\ m' = m + | LV, Lred _ l' m' => lred ge e a m l' m' + | RV, Rred _ r' m' t => rred ge a m t r' m' /\ exists w', possible_trace w t w' + | RV, Callred _ fd args tyres m' => callred ge a fd args tyres /\ m' = m | LV, Stuckred => ~imm_safe_t k a m | RV, Stuckred => ~imm_safe_t k a m | _, _ => False @@ -1521,7 +1510,7 @@ Proof with (try (apply not_invert_ok; simpl; intro; myinv; intuition congruence; destruct (Genv.find_funct ge vf) as [fd|] eqn:?... destruct (sem_cast_arguments vtl tyargs) as [vargs|] eqn:?... destruct (type_eq (type_of_fundef fd) (Tfunction tyargs tyres cconv))... - apply topred_ok; auto. red. split; auto. eapply red_Ecall; eauto. + apply topred_ok; auto. red. split; auto. eapply red_call; eauto. eapply sem_cast_arguments_sound; eauto. apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv. congruence. apply not_invert_ok; simpl; intros; myinv. specialize (H ALLVAL). myinv. @@ -1583,73 +1572,77 @@ Qed. Lemma lred_topred: forall l1 m1 l2 m2, lred ge e l1 m1 l2 m2 -> - step_expr LV l1 m1 = topred (Lred l2 m2). + exists rule, step_expr LV l1 m1 = topred (Lred rule l2 m2). Proof. induction 1; simpl. (* var local *) - rewrite H. rewrite dec_eq_true; auto. + rewrite H. rewrite dec_eq_true. econstructor; eauto. (* var global *) - rewrite H; rewrite H0. auto. + rewrite H; rewrite H0. econstructor; eauto. (* deref *) - auto. + econstructor; eauto. (* field struct *) - rewrite H, H0; auto. + rewrite H, H0; econstructor; eauto. (* field union *) - rewrite H; auto. + rewrite H; econstructor; eauto. Qed. Lemma rred_topred: forall w' r1 m1 t r2 m2, rred ge r1 m1 t r2 m2 -> possible_trace w t w' -> - step_expr RV r1 m1 = topred (Rred r2 m2 t). + exists rule, step_expr RV r1 m1 = topred (Rred rule r2 m2 t). Proof. induction 1; simpl; intros. (* valof *) - rewrite dec_eq_true; auto. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0). auto. + rewrite dec_eq_true. + rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0). econstructor; eauto. (* addrof *) - inv H. auto. + inv H. econstructor; eauto. (* unop *) - inv H0. rewrite H; auto. + inv H0. rewrite H; econstructor; eauto. (* binop *) - inv H0. rewrite H; auto. + inv H0. rewrite H; econstructor; eauto. (* cast *) - inv H0. rewrite H; auto. + inv H0. rewrite H; econstructor; eauto. (* seqand *) - inv H0. rewrite H; auto. - inv H0. rewrite H; auto. + inv H0. rewrite H; econstructor; eauto. + inv H0. rewrite H; econstructor; eauto. (* seqor *) - inv H0. rewrite H; auto. - inv H0. rewrite H; auto. + inv H0. rewrite H; econstructor; eauto. + inv H0. rewrite H; econstructor; eauto. (* condition *) - inv H0. rewrite H; auto. + inv H0. rewrite H; econstructor; eauto. (* sizeof *) - inv H. auto. + inv H. econstructor; eauto. (* alignof *) - inv H. auto. + inv H. econstructor; eauto. (* assign *) - rewrite dec_eq_true; auto. rewrite H. rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ H0 H1). auto. + rewrite dec_eq_true. rewrite H. rewrite (do_assign_loc_complete _ _ _ _ _ _ _ _ _ H0 H1). + econstructor; eauto. (* assignop *) - rewrite dec_eq_true; auto. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0). auto. + rewrite dec_eq_true. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H0). + econstructor; eauto. (* postincr *) - rewrite dec_eq_true; auto. subst. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H1). auto. + rewrite dec_eq_true. subst. rewrite (do_deref_loc_complete _ _ _ _ _ _ _ _ H H1). + econstructor; eauto. (* comma *) - inv H0. rewrite dec_eq_true; auto. + inv H0. rewrite dec_eq_true. econstructor; eauto. (* paren *) - inv H0. rewrite H; auto. + inv H0. rewrite H; econstructor; eauto. (* builtin *) exploit sem_cast_arguments_complete; eauto. intros [vtl [A B]]. exploit do_ef_external_complete; eauto. intros C. - rewrite A. rewrite B. rewrite C. auto. + rewrite A. rewrite B. rewrite C. econstructor; eauto. Qed. Lemma callred_topred: forall a fd args ty m, callred ge a fd args ty -> - step_expr RV a m = topred (Callred fd args ty m). + exists rule, step_expr RV a m = topred (Callred rule fd args ty m). Proof. induction 1; simpl. rewrite H2. exploit sem_cast_arguments_complete; eauto. intros [vtl [A B]]. - rewrite A; rewrite H; rewrite B; rewrite H1; rewrite dec_eq_true. auto. + rewrite A; rewrite H; rewrite B; rewrite H1; rewrite dec_eq_true. econstructor; eauto. Qed. Definition reducts_incl {A B: Type} (C: A -> B) (res1: reducts A) (res2: reducts B) : Prop := @@ -1956,44 +1949,54 @@ Proof. simpl. auto. Qed. +Inductive transition : Type := TR (rule: string) (t: trace) (S': state). + Definition expr_final_state (f: function) (k: cont) (e: env) (C_rd: (expr -> expr) * reduction) := match snd C_rd with - | Lred a m => (E0, ExprState f (fst C_rd a) k e m) - | Rred a m t => (t, ExprState f (fst C_rd a) k e m) - | Callred fd vargs ty m => (E0, Callstate fd vargs (Kcall f e (fst C_rd) ty k) m) - | Stuck => (E0, Stuckstate) + | Lred rule a m => TR rule E0 (ExprState f (fst C_rd a) k e m) + | Rred rule a m t => TR rule t (ExprState f (fst C_rd a) k e m) + | Callred rule fd vargs ty m => TR rule E0 (Callstate fd vargs (Kcall f e (fst C_rd) ty k) m) + | Stuckred => TR "step_stuck" E0 Stuckstate end. Local Open Scope list_monad_scope. -Definition ret (S: state) : list (trace * state) := (E0, S) :: nil. +Definition ret (rule: string) (S: state) : list transition := + TR rule E0 S :: nil. -Definition do_step (w: world) (s: state) : list (trace * state) := +Definition do_step (w: world) (s: state) : list transition := match s with | ExprState f a k e m => match is_val a with | Some(v, ty) => match k with - | Kdo k => ret (State f Sskip k e m ) + | Kdo k => ret "step_do_2" (State f Sskip k e m ) | Kifthenelse s1 s2 k => - do b <- bool_val v ty; ret (State f (if b then s1 else s2) k e m) + do b <- bool_val v ty; + ret "step_ifthenelse_2" (State f (if b then s1 else s2) k e m) | Kwhile1 x s k => do b <- bool_val v ty; - if b then ret (State f s (Kwhile2 x s k) e m) else ret (State f Sskip k e m) + if b + then ret "step_while_true" (State f s (Kwhile2 x s k) e m) + else ret "step_while_false" (State f Sskip k e m) | Kdowhile2 x s k => do b <- bool_val v ty; - if b then ret (State f (Sdowhile x s) k e m) else ret (State f Sskip k e m) + if b + then ret "step_dowhile_true" (State f (Sdowhile x s) k e m) + else ret "step_dowhile_false" (State f Sskip k e m) | Kfor2 a2 a3 s k => do b <- bool_val v ty; - if b then ret (State f s (Kfor3 a2 a3 s k) e m) else ret (State f Sskip k e m) + if b + then ret "step_for_true" (State f s (Kfor3 a2 a3 s k) e m) + else ret "step_for_false" (State f Sskip k e m) | Kreturn k => do v' <- sem_cast v ty f.(fn_return); do m' <- Mem.free_list m (blocks_of_env ge e); - ret (Returnstate v' (call_cont k) m') + ret "step_return_2" (Returnstate v' (call_cont k) m') | Kswitch1 sl k => do n <- sem_switch_arg v ty; - ret (State f (seq_of_labeled_statement (select_switch n sl)) (Kswitch2 k) e m) + ret "step_expr_switch" (State f (seq_of_labeled_statement (select_switch n sl)) (Kswitch2 k) e m) | _ => nil end @@ -2001,48 +2004,66 @@ Definition do_step (w: world) (s: state) : list (trace * state) := map (expr_final_state f k e) (step_expr e w RV a m) end - | State f (Sdo x) k e m => ret(ExprState f x (Kdo k) e m) - - | State f (Ssequence s1 s2) k e m => ret(State f s1 (Kseq s2 k) e m) - | State f Sskip (Kseq s k) e m => ret (State f s k e m) - | State f Scontinue (Kseq s k) e m => ret (State f Scontinue k e m) - | State f Sbreak (Kseq s k) e m => ret (State f Sbreak k e m) - - | State f (Sifthenelse a s1 s2) k e m => ret (ExprState f a (Kifthenelse s1 s2 k) e m) - - | State f (Swhile x s) k e m => ret (ExprState f x (Kwhile1 x s k) e m) - | State f (Sskip|Scontinue) (Kwhile2 x s k) e m => ret (State f (Swhile x s) k e m) - | State f Sbreak (Kwhile2 x s k) e m => ret (State f Sskip k e m) - - | State f (Sdowhile a s) k e m => ret (State f s (Kdowhile1 a s k) e m) - | State f (Sskip|Scontinue) (Kdowhile1 x s k) e m => ret (ExprState f x (Kdowhile2 x s k) e m) - | State f Sbreak (Kdowhile1 x s k) e m => ret (State f Sskip k e m) + | State f (Sdo x) k e m => + ret "step_do_1" (ExprState f x (Kdo k) e m) + | State f (Ssequence s1 s2) k e m => + ret "step_seq" (State f s1 (Kseq s2 k) e m) + | State f Sskip (Kseq s k) e m => + ret "step_skip_seq" (State f s k e m) + | State f Scontinue (Kseq s k) e m => + ret "step_continue_seq" (State f Scontinue k e m) + | State f Sbreak (Kseq s k) e m => + ret "step_break_seq" (State f Sbreak k e m) + + | State f (Sifthenelse a s1 s2) k e m => + ret "step_ifthenelse_1" (ExprState f a (Kifthenelse s1 s2 k) e m) + + | State f (Swhile x s) k e m => + ret "step_while" (ExprState f x (Kwhile1 x s k) e m) + | State f (Sskip|Scontinue) (Kwhile2 x s k) e m => + ret "step_skip_or_continue_while" (State f (Swhile x s) k e m) + | State f Sbreak (Kwhile2 x s k) e m => + ret "step_break_while" (State f Sskip k e m) + + | State f (Sdowhile a s) k e m => + ret "step_dowhile" (State f s (Kdowhile1 a s k) e m) + | State f (Sskip|Scontinue) (Kdowhile1 x s k) e m => + ret "step_skip_or_continue_dowhile" (ExprState f x (Kdowhile2 x s k) e m) + | State f Sbreak (Kdowhile1 x s k) e m => + ret "step_break_dowhile" (State f Sskip k e m) | State f (Sfor a1 a2 a3 s) k e m => if is_skip a1 - then ret (ExprState f a2 (Kfor2 a2 a3 s k) e m) - else ret (State f a1 (Kseq (Sfor Sskip a2 a3 s) k) e m) - | State f Sskip (Kfor3 a2 a3 s k) e m => ret (State f a3 (Kfor4 a2 a3 s k) e m) - | State f Scontinue (Kfor3 a2 a3 s k) e m => ret (State f a3 (Kfor4 a2 a3 s k) e m) - | State f Sbreak (Kfor3 a2 a3 s k) e m => ret (State f Sskip k e m) - | State f Sskip (Kfor4 a2 a3 s k) e m => ret (State f (Sfor Sskip a2 a3 s) k e m) + then ret "step_for" (ExprState f a2 (Kfor2 a2 a3 s k) e m) + else ret "step_for_start" (State f a1 (Kseq (Sfor Sskip a2 a3 s) k) e m) + | State f (Sskip|Scontinue) (Kfor3 a2 a3 s k) e m => + ret "step_skip_or_continue_for3" (State f a3 (Kfor4 a2 a3 s k) e m) + | State f Sbreak (Kfor3 a2 a3 s k) e m => + ret "step_break_for3" (State f Sskip k e m) + | State f Sskip (Kfor4 a2 a3 s k) e m => + ret "step_skip_for4" (State f (Sfor Sskip a2 a3 s) k e m) | State f (Sreturn None) k e m => do m' <- Mem.free_list m (blocks_of_env ge e); - ret (Returnstate Vundef (call_cont k) m') - | State f (Sreturn (Some x)) k e m => ret (ExprState f x (Kreturn k) e m) + ret "step_return_0" (Returnstate Vundef (call_cont k) m') + | State f (Sreturn (Some x)) k e m => + ret "step_return_1" (ExprState f x (Kreturn k) e m) | State f Sskip ((Kstop | Kcall _ _ _ _ _) as k) e m => do m' <- Mem.free_list m (blocks_of_env ge e); - ret (Returnstate Vundef k m') + ret "step_skip_call" (Returnstate Vundef k m') - | State f (Sswitch x sl) k e m => ret (ExprState f x (Kswitch1 sl k) e m) - | State f (Sskip|Sbreak) (Kswitch2 k) e m => ret (State f Sskip k e m) - | State f Scontinue (Kswitch2 k) e m => ret (State f Scontinue k e m) + | State f (Sswitch x sl) k e m => + ret "step_switch" (ExprState f x (Kswitch1 sl k) e m) + | State f (Sskip|Sbreak) (Kswitch2 k) e m => + ret "step_skip_break_switch" (State f Sskip k e m) + | State f Scontinue (Kswitch2 k) e m => + ret "step_continue_switch" (State f Scontinue k e m) - | State f (Slabel lbl s) k e m => ret (State f s k e m) + | State f (Slabel lbl s) k e m => + ret "step_label" (State f s k e m) | State f (Sgoto lbl) k e m => match find_label lbl f.(fn_body) (call_cont k) with - | Some(s', k') => ret (State f s' k' e m) + | Some(s', k') => ret "step_goto" (State f s' k' e m) | None => nil end @@ -2050,14 +2071,15 @@ Definition do_step (w: world) (s: state) : list (trace * state) := check (list_norepet_dec ident_eq (var_names (fn_params f) ++ var_names (fn_vars f))); let (e,m1) := do_alloc_variables empty_env m (f.(fn_params) ++ f.(fn_vars)) in do m2 <- sem_bind_parameters w e m1 f.(fn_params) vargs; - ret (State f f.(fn_body) k e m2) + ret "step_internal_function" (State f f.(fn_body) k e m2) | Callstate (External ef _ _ _) vargs k m => match do_external ef w vargs m with | None => nil - | Some(w',t,v,m') => (t, Returnstate v k m') :: nil + | Some(w',t,v,m') => TR "step_external_function" t (Returnstate v k m') :: nil end - | Returnstate v (Kcall f e C ty k) m => ret (ExprState f (C (Eval v ty)) k e m) + | Returnstate v (Kcall f e C ty k) m => + ret "step_returnstate" (ExprState f (C (Eval v ty)) k e m) | _ => nil end. @@ -2065,7 +2087,7 @@ Definition do_step (w: world) (s: state) : list (trace * state) := Ltac myinv := match goal with | [ |- In _ nil -> _ ] => intro X; elim X - | [ |- In _ (ret _) -> _ ] => + | [ |- In _ (ret _ _) -> _ ] => intro X; elim X; clear X; [intro EQ; unfold ret in EQ; inv EQ; myinv | myinv] | [ |- In _ (_ :: nil) -> _ ] => @@ -2079,8 +2101,8 @@ Ltac myinv := Hint Extern 3 => exact I. Theorem do_step_sound: - forall w S t S', - In (t, S') (do_step w S) -> + forall w S rule t S', + In (TR rule t S') (do_step w S) -> Csem.step ge S t S' \/ (t = E0 /\ S' = Stuckstate /\ can_crash_world w S). Proof with try (left; right; econstructor; eauto; fail). intros until S'. destruct S; simpl. @@ -2145,45 +2167,52 @@ Proof. Qed. Theorem do_step_complete: - forall w S t S' w', possible_trace w t w' -> Csem.step ge S t S' -> In (t, S') (do_step w S). -Proof with (unfold ret; auto with coqlib). + forall w S t S' w', + possible_trace w t w' -> Csem.step ge S t S' -> exists rule, In (TR rule t S') (do_step w S). +Proof with (unfold ret; eauto with coqlib). intros until w'; intros PT H. destruct H. (* Expression step *) inversion H; subst; exploit estep_not_val; eauto; intro NOTVAL. (* lred *) unfold do_step; rewrite NOTVAL. - change (E0, ExprState f (C a') k e m') with (expr_final_state f k e (C, Lred a' m')). + exploit lred_topred; eauto. instantiate (1 := w). intros (rule & STEP). + exists rule. change (TR rule E0 (ExprState f (C a') k e m')) with (expr_final_state f k e (C, Lred rule a' m')). apply in_map. generalize (step_expr_context e w _ _ _ H1 a m). unfold reducts_incl. - intro. replace C with (fun x => C x). apply H2. - rewrite (lred_topred _ _ _ _ _ _ H0). unfold topred; auto with coqlib. + intro. replace C with (fun x => C x). apply H2. + rewrite STEP. unfold topred; auto with coqlib. apply extensionality; auto. (* rred *) unfold do_step; rewrite NOTVAL. - change (t, ExprState f (C a') k e m') with (expr_final_state f k e (C, Rred a' m' t)). + exploit rred_topred; eauto. instantiate (1 := e). intros (rule & STEP). + exists rule. + change (TR rule t (ExprState f (C a') k e m')) with (expr_final_state f k e (C, Rred rule a' m' t)). apply in_map. generalize (step_expr_context e w _ _ _ H1 a m). unfold reducts_incl. - intro. replace C with (fun x => C x). apply H2. - rewrite (rred_topred _ _ _ _ _ _ _ _ H0 PT). unfold topred; auto with coqlib. + intro. replace C with (fun x => C x). apply H2. + rewrite STEP; unfold topred; auto with coqlib. apply extensionality; auto. (* callred *) unfold do_step; rewrite NOTVAL. - change (E0, Callstate fd vargs (Kcall f e C ty k) m) with (expr_final_state f k e (C, Callred fd vargs ty m)). + exploit callred_topred; eauto. instantiate (1 := m). instantiate (1 := w). instantiate (1 := e). + intros (rule & STEP). exists rule. + change (TR rule E0 (Callstate fd vargs (Kcall f e C ty k) m)) with (expr_final_state f k e (C, Callred rule fd vargs ty m)). apply in_map. generalize (step_expr_context e w _ _ _ H1 a m). unfold reducts_incl. intro. replace C with (fun x => C x). apply H2. - rewrite (callred_topred _ _ _ _ _ _ _ H0). unfold topred; auto with coqlib. + rewrite STEP; unfold topred; auto with coqlib. apply extensionality; auto. (* stuck *) exploit not_imm_safe_stuck_red. eauto. red; intros; elim H1. eapply imm_safe_t_imm_safe. eauto. instantiate (1 := w). intros [C' IN]. - simpl do_step. rewrite NOTVAL. - change (E0, Stuckstate) with (expr_final_state f k e (C', Stuckred)). + simpl do_step. rewrite NOTVAL. + exists "step_stuck". + change (TR "step_stuck" E0 Stuckstate) with (expr_final_state f k e (C', Stuckred)). apply in_map. auto. (* Statement step *) - inv H; simpl... + inv H; simpl; econstructor... rewrite H0... rewrite H0... rewrite H0... diff --git a/cfrontend/Csem.v b/cfrontend/Csem.v index e6e3a321..fafbf29f 100644 --- a/cfrontend/Csem.v +++ b/cfrontend/Csem.v @@ -317,7 +317,7 @@ Inductive rred: expr -> mem -> trace -> expr -> mem -> Prop := (More exactly, identification of function calls that can reduce.) *) Inductive callred: expr -> fundef -> list val -> type -> Prop := - | red_Ecall: forall vf tyf tyargs tyres cconv el ty fd vargs, + | red_call: forall vf tyf tyargs tyres cconv el ty fd vargs, Genv.find_funct ge vf = Some fd -> cast_arguments el tyargs vargs -> type_of_fundef fd = Tfunction tyargs tyres cconv -> -- cgit