include "utilities/extralib.ma". include "common/IO.ma". include "common/SmallstepExec.ma". include "Clight/Csem.ma". include "Clight/TypeComparison.ma". definition P_res: ∀A.∀P:A → Prop.res A → Prop ≝ λA,P,a. match a with [ Error _ ⇒ True | OK v ⇒ P v ]. axiom ValueIsNotABoolean : String. definition exec_bool_of_val : ∀v:val. ∀ty:type. res bool ≝ λv,ty. match v in val with [ Vint sz i ⇒ match ty with [ Tint sz' _ ⇒ intsize_eq_elim ? sz sz' ? i (λi. OK ? (¬eq_bv ? i (zero ?))) (Error ? (msg TypeMismatch)) | _ ⇒ Error ? (msg TypeMismatch) ] | Vfloat f ⇒ match ty with [ Tfloat _ ⇒ OK ? (¬Fcmp Ceq f Fzero) | _ ⇒ Error ? (msg TypeMismatch) ] | Vptr _ ⇒ match ty with [ Tpointer _ _ ⇒ OK ? true | _ ⇒ Error ? (msg TypeMismatch) ] | Vnull _ ⇒ match ty with [ Tpointer _ _ ⇒ OK ? false | _ ⇒ Error ? (msg TypeMismatch) ] | _ ⇒ Error ? (msg ValueIsNotABoolean) ]. lemma bool_of_val_complete : ∀v,ty,r. bool_of_val v ty r → ∃b. r = of_bool b ∧ exec_bool_of_val v ty = OK ? b. #v #ty #r #H elim H; #v #t #H' elim H'; [ * #sg #i #ne %{ true} % // whd in ⊢ (??%?); >(eq_bv_false … ne) // | #ptr #r #ty %{ true} % // | #f #s #ne %{ true} % //; whd in ⊢ (??%?); >(Feq_zero_false … ne) //; | * #sg %{ false} % // | #r #r' #t %{ false} % //; | #s %{ false} % //; whd in ⊢ (??%?); >(Feq_zero_true …) //; ] qed. (* Prove a few minor results to make proof obligations easy. *) lemma bind_OK: ∀A,B,P,e,f. (∀v. e = OK A v → match f v with [ Error _ ⇒ True | OK v' ⇒ P v' ]) → match bind A B e f with [ Error _ ⇒ True | OK v ⇒ P v ]. #A #B #P #e #f elim e; /2/; qed. lemma sig_bind_OK: ∀A,B. ∀P:A → Prop. ∀P':B → Prop. ∀e:res (Sig A P). ∀f:Sig A P → res B. (∀v:A. ∀p:P v. match f (mk_Sig A P v p) with [ Error _ ⇒ True | OK v' ⇒ P' v'] ) → match bind (Sig A P) B e f with [ Error _ ⇒ True | OK v' ⇒ P' v' ]. #A #B #P #P' #e #f elim e; [ #v0 elim v0; #v #Hv #IH @IH | #m #_ @I ] qed. lemma bind2_OK: ∀A,B,C,P,e,f. (∀v1,v2. e = OK ? 〈v1,v2〉 → match f v1 v2 with [ Error _ ⇒ True | OK v' ⇒ P v' ]) → match bind2 A B C e f with [ Error _ ⇒ True | OK v ⇒ P v ]. #A #B #C #P #e #f elim e; //; #v cases v; /2/; qed. lemma opt_bind_OK: ∀A,B,P,m,e,f. (∀v. e = Some A v → match f v with [ Error _ ⇒ True | OK v' ⇒ P v' ]) → match bind A B (opt_to_res A m e) f with [ Error _ ⇒ True | OK v ⇒ P v ]. #A #B #P #m #e #f elim e; normalize; /2/; qed. axiom BadCast : String. (* TODO: refine into more precise errors? *) definition try_cast_null : ∀m:mem. ∀sz. ∀i:BitVector (bitsize_of_intsize sz). ∀ty:type. ∀ty':type. res val ≝ λm:mem. λsz. λi. λty:type. λty':type. match eq_bv ? i (zero ?) with [ true ⇒ match ty with [ Tint sz' _ ⇒ if eq_intsize sz sz' then match ty' with [ Tpointer r _ ⇒ OK ? (Vnull r) | Tarray r _ _ ⇒ OK ? (Vnull r) | Tfunction _ _ ⇒ OK ? (Vnull Code) | _ ⇒ Error ? (msg BadCast) ] else Error ? (msg TypeMismatch) | _ ⇒ Error ? (msg BadCast) ] | false ⇒ Error ? (msg BadCast) ]. definition exec_cast : ∀m:mem. ∀v:val. ∀ty:type. ∀ty':type. res val ≝ λm:mem. λv:val. λty:type. λty':type. match v with [ Vint sz i ⇒ match ty with [ Tint sz1 si1 ⇒ intsize_eq_elim ? sz sz1 ? i (λi. match ty' with [ Tint sz2 si2 ⇒ OK ? (Vint ? (cast_int_int sz1 si1 sz2 i)) | Tfloat sz2 ⇒ OK ? (Vfloat (cast_float_float sz2 (cast_int_float si1 ? i))) | Tpointer _ _ ⇒ do r ← try_cast_null m ? i ty ty'; OK val r | Tarray _ _ _ ⇒ do r ← try_cast_null m ? i ty ty'; OK val r | Tfunction _ _ ⇒ do r ← try_cast_null m ? i ty ty'; OK val r | _ ⇒ Error ? (msg BadCast) ]) (Error ? (msg BadCast)) | Tpointer _ _ ⇒ do r ← try_cast_null m ? i ty ty'; OK val r | Tarray _ _ _ ⇒ do r ← try_cast_null m ? i ty ty'; OK val r | Tfunction _ _ ⇒ do r ← try_cast_null m ? i ty ty'; OK val r | _ ⇒ Error ? (msg TypeMismatch) ] | Vfloat f ⇒ match ty with [ Tfloat sz ⇒ match ty' with [ Tint sz' si' ⇒ OK ? (Vint sz' (cast_float_int sz' si' f)) | Tfloat sz' ⇒ OK ? (Vfloat (cast_float_float sz' f)) | _ ⇒ Error ? (msg BadCast) ] | _ ⇒ Error ? (msg TypeMismatch) ] | Vptr ptr ⇒ do s ← match ty with [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code | _ ⇒ Error ? (msg TypeMismatch) ]; do u ← match eq_region_dec (ptype ptr) s with [ inl _ ⇒ OK ? it | inr _ ⇒ Error ? (msg BadCast) ]; do s' ← match ty' with [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code | _ ⇒ Error ? (msg BadCast)]; match pointer_compat_dec (pblock ptr) s' with [ inl p' ⇒ OK ? (Vptr (mk_pointer s' (pblock ptr) p' (poff ptr))) | inr _ ⇒ Error ? (msg BadCast) ] | Vnull r ⇒ do s ← match ty with [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code | _ ⇒ Error ? (msg TypeMismatch) ]; do u ← match eq_region_dec r s with [ inl _ ⇒ OK ? it | inr _ ⇒ Error ? (msg BadCast) ]; do s' ← match ty' with [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code | _ ⇒ Error ? (msg BadCast) ]; OK ? (Vnull s') | _ ⇒ Error ? (msg BadCast) ]. definition load_value_of_type' ≝ λty,m,l. let 〈loc,ofs〉 ≝ l in load_value_of_type ty m loc ofs. axiom BadlyTypedTerm : String. axiom UnknownIdentifier : String. axiom BadLvalueTerm : String. axiom FailedLoad : String. axiom FailedOp : String. (* To make the evaluation of bare lvalue expressions invoke exec_lvalue with a structurally smaller value, we break out the surrounding Expr constructor and use exec_lvalue'. *) let rec exec_expr (ge:genv) (en:env) (m:mem) (e:expr) on e : res (val×trace) ≝ match e with [ Expr e' ty ⇒ match e' with [ Econst_int sz i ⇒ match ty with [ Tint sz' _ ⇒ if eq_intsize sz sz' then OK ? 〈Vint sz i, E0〉 else Error ? (msg BadlyTypedTerm) | _ ⇒ Error ? (msg BadlyTypedTerm) ] | Econst_float f ⇒ OK ? 〈Vfloat f, E0〉 | Evar _ ⇒ do 〈l,tr〉 ← exec_lvalue' ge en m e' ty; do v ← opt_to_res ? (msg FailedLoad) (load_value_of_type' ty m l); OK ? 〈v,tr〉 | Ederef _ ⇒ do 〈l,tr〉 ← exec_lvalue' ge en m e' ty; do v ← opt_to_res ? (msg FailedLoad) (load_value_of_type' ty m l); OK ? 〈v,tr〉 | Efield _ _ ⇒ do 〈l,tr〉 ← exec_lvalue' ge en m e' ty; do v ← opt_to_res ? (msg FailedLoad) (load_value_of_type' ty m l); OK ? 〈v,tr〉 | Eaddrof a ⇒ do 〈lo,tr〉 ← exec_lvalue ge en m a; match ty with [ Tpointer r _ ⇒ let 〈loc,ofs〉 ≝ lo in match pointer_compat_dec loc r with [ inl pc ⇒ OK ? 〈Vptr (mk_pointer r loc pc ofs), tr〉 | inr _ ⇒ Error ? (msg TypeMismatch) ] | _ ⇒ Error ? (msg BadlyTypedTerm) ] | Esizeof ty' ⇒ match ty with [ Tint sz _ ⇒ OK ? 〈Vint sz (repr ? (sizeof ty')), E0〉 | _ ⇒ Error ? (msg BadlyTypedTerm) ] | Eunop op a ⇒ do 〈v1,tr〉 ← exec_expr ge en m a; do v ← opt_to_res ? (msg FailedOp) (sem_unary_operation op v1 (typeof a)); OK ? 〈v,tr〉 | Ebinop op a1 a2 ⇒ do 〈v1,tr1〉 ← exec_expr ge en m a1; do 〈v2,tr2〉 ← exec_expr ge en m a2; do v ← opt_to_res ? (msg FailedOp) (sem_binary_operation op v1 (typeof a1) v2 (typeof a2) m); OK ? 〈v,tr1⧺tr2〉 | Econdition a1 a2 a3 ⇒ do 〈v,tr1〉 ← exec_expr ge en m a1; do b ← exec_bool_of_val v (typeof a1); do 〈v',tr2〉 ← match b return λ_.res (val×trace) with [ true ⇒ (exec_expr ge en m a2) | false ⇒ (exec_expr ge en m a3) ]; OK ? 〈v',tr1⧺tr2〉 (* if b then exec_expr ge en m a2 else exec_expr ge en m a3)*) | Eorbool a1 a2 ⇒ do 〈v1,tr1〉 ← exec_expr ge en m a1; do b1 ← exec_bool_of_val v1 (typeof a1); match b1 return λ_.res (val×trace) with [ true ⇒ OK ? 〈Vtrue,tr1〉 | false ⇒ do 〈v2,tr2〉 ← exec_expr ge en m a2; do b2 ← exec_bool_of_val v2 (typeof a2); OK ? 〈of_bool b2, tr1⧺tr2〉 ] | Eandbool a1 a2 ⇒ do 〈v1,tr1〉 ← exec_expr ge en m a1; do b1 ← exec_bool_of_val v1 (typeof a1); match b1 return λ_.res (val×trace) with [ true ⇒ do 〈v2,tr2〉 ← exec_expr ge en m a2; do b2 ← exec_bool_of_val v2 (typeof a2); OK ? 〈of_bool b2, tr1⧺tr2〉 | false ⇒ OK ? 〈Vfalse,tr1〉 ] | Ecast ty' a ⇒ do 〈v,tr〉 ← exec_expr ge en m a; do v' ← exec_cast m v (typeof a) ty'; OK ? 〈v',tr〉 | Ecost l a ⇒ do 〈v,tr〉 ← exec_expr ge en m a; OK ? 〈v,tr⧺(Echarge l)〉 ] ] and exec_lvalue' (ge:genv) (en:env) (m:mem) (e':expr_descr) (ty:type) on e' : res (block × offset × trace) ≝ match e' with [ Evar id ⇒ match (lookup ?? en id) with [ None ⇒ do l ← opt_to_res ? [MSG UnknownIdentifier; CTX ? id] (find_symbol ? ? ge id); OK ? 〈〈l,zero_offset〉,E0〉 (* global *) | Some loc ⇒ OK ? 〈〈loc,zero_offset〉,E0〉 (* local *) ] | Ederef a ⇒ do 〈v,tr〉 ← exec_expr ge en m a; match v with [ Vptr ptr ⇒ OK ? 〈〈pblock ptr, poff ptr〉,tr〉 | _ ⇒ Error ? (msg TypeMismatch) ] | Efield a i ⇒ match (typeof a) with [ Tstruct id fList ⇒ do 〈lofs,tr〉 ← exec_lvalue ge en m a; do delta ← field_offset i fList; OK ? 〈〈\fst lofs,shift_offset ? (\snd lofs) (repr I32 delta)〉,tr〉 | Tunion id fList ⇒ do 〈lofs,tr〉 ← exec_lvalue ge en m a; OK ? 〈lofs,tr〉 | _ ⇒ Error ? (msg BadlyTypedTerm) ] | _ ⇒ Error ? (msg BadLvalueTerm) ] and exec_lvalue (ge:genv) (en:env) (m:mem) (e:expr) on e : res (block × offset × trace) ≝ match e with [ Expr e' ty ⇒ exec_lvalue' ge en m e' ty ]. lemma P_res_to_P: ∀A,P,e,v. P_res A P e → e = OK A v → P v. #A #P #e #v #H1 #H2 >H2 in H1; whd in ⊢ (% → ?); //; qed. (* We define a special induction principle tailored to the recursive definition above. *) definition is_not_lvalue: expr_descr → Prop ≝ λe. match e with [ Evar _ ⇒ False | Ederef _ ⇒ False | Efield _ _ ⇒ False | _ ⇒ True ]. definition Plvalue ≝ λP:(expr → Prop).λe,ty. match e return λ_.Prop with [ Evar _ ⇒ P (Expr e ty) | Ederef _ ⇒ P (Expr e ty) | Efield _ _ ⇒ P (Expr e ty) | _ ⇒ True ]. (* TODO: Can we do this sensibly with a map combinator? *) let rec exec_exprlist (ge:genv) (e:env) (m:mem) (l:list expr) on l : res (list val×trace) ≝ match l with [ nil ⇒ OK ? 〈nil val, E0〉 | cons e1 es ⇒ do 〈v,tr1〉 ← exec_expr ge e m e1; do 〈vs,tr2〉 ← exec_exprlist ge e m es; OK ? 〈cons val v vs, tr1⧺tr2〉 ]. let rec exec_alloc_variables (en:env) (m:mem) (l:list (ident × type)) on l : env × mem ≝ match l with [ nil ⇒ 〈en, m〉 | cons h vars ⇒ let 〈id,ty〉 ≝ h in let 〈m1,b1〉 ≝ alloc m 0 (sizeof ty) Any in exec_alloc_variables (add ?? en id b1) m1 vars ]. axiom WrongNumberOfParameters : String. axiom FailedStore : String. (* TODO: can we establish that length params = length vs in advance? *) let rec exec_bind_parameters (e:env) (m:mem) (params:list (ident × type)) (vs:list val) on params : res mem ≝ match params with [ nil ⇒ match vs with [ nil ⇒ OK ? m | cons _ _ ⇒ Error ? (msg WrongNumberOfParameters) ] | cons idty params' ⇒ let 〈id,ty〉 ≝ idty in match vs with [ nil ⇒ Error ? (msg WrongNumberOfParameters) | cons v1 vl ⇒ do b ← opt_to_res ? [MSG UnknownIdentifier; CTX ? id] (lookup ?? e id); do m1 ← opt_to_res ? [MSG FailedStore; CTX ? id] (store_value_of_type ty m b zero_offset v1); exec_bind_parameters e m1 params' vl ] ]. let rec is_is_call_cont (k:cont) : Sum (is_call_cont k) (Not (is_call_cont k)) ≝ match k return λk'.(is_call_cont k') + (¬is_call_cont k') with [ Kstop ⇒ ? | Kcall _ _ _ _ ⇒ ? | _ ⇒ ? ]. [ 1,8: %1 ; // | *: %2 ; /2/ ] qed. definition is_Sskip : ∀s:statement. (s = Sskip) ⊎ (s ≠ Sskip) ≝ λs.match s return λs'.(s' = Sskip) ⊎ (s' ≠ Sskip) with [ Sskip ⇒ inl ?? (refl ??) | _ ⇒ inr ?? (nmk ? (λH.?)) ]. destruct qed. (* execution *) definition store_value_of_type' ≝ λty,m,l,v. let 〈loc,ofs〉 ≝ l in store_value_of_type ty m loc ofs v. axiom NonsenseState : String. axiom ReturnMismatch : String. axiom UnknownLabel : String. axiom BadFunctionValue : String. let rec exec_step (ge:genv) (st:state) on st : (IO io_out io_in (trace × state)) ≝ match st with [ State f s k e m ⇒ match s with [ Sassign a1 a2 ⇒ ! 〈l,tr1〉 ← exec_lvalue ge e m a1; ! 〈v2,tr2〉 ← exec_expr ge e m a2; ! m' ← opt_to_io … (msg FailedStore) (store_value_of_type' (typeof a1) m l v2); ret ? 〈tr1⧺tr2, State f Sskip k e m'〉 | Scall lhs a al ⇒ ! 〈vf,tr2〉 ← exec_expr ge e m a; ! 〈vargs,tr3〉 ← exec_exprlist ge e m al; ! fd ← opt_to_io … (msg BadFunctionValue) (find_funct ? ? ge vf); ! p ← err_to_io … (assert_type_eq (type_of_fundef fd) (fun_typeof a)); (* requires associativity of IOMonad, so rearrange it below ! k' ← match lhs with [ None ⇒ ret ? (Kcall (None ?) f e k) | Some lhs' ⇒ ! locofs ← exec_lvalue ge e m lhs'; ret ? (Kcall (Some ? 〈locofs, typeof lhs'〉) f e k) ]; ret ? 〈E0, Callstate fd vargs k' m〉) *) match lhs with [ None ⇒ ret ? 〈tr2⧺tr3, Callstate fd vargs (Kcall (None ?) f e k) m〉 | Some lhs' ⇒ ! 〈locofs,tr1〉 ← exec_lvalue ge e m lhs'; ret ? 〈tr1⧺tr2⧺tr3, Callstate fd vargs (Kcall (Some ? 〈locofs, typeof lhs'〉) f e k) m〉 ] | Ssequence s1 s2 ⇒ ret ? 〈E0, State f s1 (Kseq s2 k) e m〉 | Sskip ⇒ match k with [ Kseq s k' ⇒ ret ? 〈E0, State f s k' e m〉 | Kstop ⇒ match fn_return f with [ Tvoid ⇒ ret ? 〈E0, Returnstate Vundef k (free_list m (blocks_of_env e))〉 | _ ⇒ Wrong ??? (msg NonsenseState) ] | Kcall _ _ _ _ ⇒ match fn_return f with [ Tvoid ⇒ ret ? 〈E0, Returnstate Vundef k (free_list m (blocks_of_env e))〉 | _ ⇒ Wrong ??? (msg NonsenseState) ] | Kwhile a s' k' ⇒ ret ? 〈E0, State f (Swhile a s') k' e m〉 | Kdowhile a s' k' ⇒ ! 〈v,tr〉 ← exec_expr ge e m a; ! b ← err_to_io … (exec_bool_of_val v (typeof a)); match b with [ true ⇒ ret ? 〈tr, State f (Sdowhile a s') k' e m〉 | false ⇒ ret ? 〈tr, State f Sskip k' e m〉 ] | Kfor2 a2 a3 s' k' ⇒ ret ? 〈E0, State f a3 (Kfor3 a2 a3 s' k') e m〉 | Kfor3 a2 a3 s' k' ⇒ ret ? 〈E0, State f (Sfor Sskip a2 a3 s') k' e m〉 | Kswitch k' ⇒ ret ? 〈E0, State f Sskip k' e m〉 | _ ⇒ Wrong ??? (msg NonsenseState) ] | Scontinue ⇒ match k with [ Kseq s' k' ⇒ ret ? 〈E0, State f Scontinue k' e m〉 | Kwhile a s' k' ⇒ ret ? 〈E0, State f (Swhile a s') k' e m〉 | Kdowhile a s' k' ⇒ ! 〈v,tr〉 ← exec_expr ge e m a; ! b ← err_to_io … (exec_bool_of_val v (typeof a)); match b with [ true ⇒ ret ? 〈tr, State f (Sdowhile a s') k' e m〉 | false ⇒ ret ? 〈tr, State f Sskip k' e m〉 ] | Kfor2 a2 a3 s' k' ⇒ ret ? 〈E0, State f a3 (Kfor3 a2 a3 s' k') e m〉 | Kswitch k' ⇒ ret ? 〈E0, State f Scontinue k' e m〉 | _ ⇒ Wrong ??? (msg NonsenseState) ] | Sbreak ⇒ match k with [ Kseq s' k' ⇒ ret ? 〈E0, State f Sbreak k' e m〉 | Kwhile a s' k' ⇒ ret ? 〈E0, State f Sskip k' e m〉 | Kdowhile a s' k' ⇒ ret ? 〈E0, State f Sskip k' e m〉 | Kfor2 a2 a3 s' k' ⇒ ret ? 〈E0, State f Sskip k' e m〉 | Kswitch k' ⇒ ret ? 〈E0, State f Sskip k' e m〉 | _ ⇒ Wrong ??? (msg NonsenseState) ] | Sifthenelse a s1 s2 ⇒ ! 〈v,tr〉 ← exec_expr ge e m a; ! b ← err_to_io … (exec_bool_of_val v (typeof a)); ret ? 〈tr, State f (if b then s1 else s2) k e m〉 | Swhile a s' ⇒ ! 〈v,tr〉 ← exec_expr ge e m a; ! b ← err_to_io … (exec_bool_of_val v (typeof a)); ret ? 〈tr, if b then State f s' (Kwhile a s' k) e m else State f Sskip k e m〉 | Sdowhile a s' ⇒ ret ? 〈E0, State f s' (Kdowhile a s' k) e m〉 | Sfor a1 a2 a3 s' ⇒ match is_Sskip a1 with [ inl _ ⇒ ! 〈v,tr〉 ← exec_expr ge e m a2; ! b ← err_to_io … (exec_bool_of_val v (typeof a2)); ret ? 〈tr, State f (if b then s' else Sskip) (if b then (Kfor2 a2 a3 s' k) else k) e m〉 | inr _ ⇒ ret ? 〈E0, State f a1 (Kseq (Sfor Sskip a2 a3 s') k) e m〉 ] | Sreturn a_opt ⇒ match a_opt with [ None ⇒ match fn_return f with [ Tvoid ⇒ ret ? 〈E0, Returnstate Vundef (call_cont k) (free_list m (blocks_of_env e))〉 | _ ⇒ Wrong ??? (msg ReturnMismatch) ] | Some a ⇒ match type_eq_dec (fn_return f) Tvoid with [ inl _ ⇒ Wrong ??? (msg ReturnMismatch) | inr _ ⇒ ! 〈v,tr〉 ← exec_expr ge e m a; ret ? 〈tr, Returnstate v (call_cont k) (free_list m (blocks_of_env e))〉 ] ] | Sswitch a sl ⇒ ! 〈v,tr〉 ← exec_expr ge e m a; match v with [ Vint sz n ⇒ ret ? 〈tr, State f (seq_of_labeled_statement (select_switch sz n sl)) (Kswitch k) e m〉 | _ ⇒ Wrong ??? (msg TypeMismatch)] | Slabel lbl s' ⇒ ret ? 〈E0, State f s' k e m〉 | Sgoto lbl ⇒ match find_label lbl (fn_body f) (call_cont k) with [ Some sk' ⇒ let 〈s',k'〉 ≝ sk' in ret ? 〈E0, State f s' k' e m〉 | None ⇒ Wrong ??? [MSG UnknownLabel; CTX ? lbl] ] | Scost lbl s' ⇒ ret ? 〈Echarge lbl, State f s' k e m〉 ] | Callstate f0 vargs k m ⇒ match f0 with [ CL_Internal f ⇒ let 〈e,m1〉 ≝ exec_alloc_variables empty_env m ((fn_params f) @ (fn_vars f)) in ! m2 ← exec_bind_parameters e m1 (fn_params f) vargs; ret ? 〈E0, State f (fn_body f) k e m2〉 | CL_External f argtys retty ⇒ ! evargs ← check_eventval_list vargs (typlist_of_typelist argtys); ! evres ← do_io f evargs (proj_sig_res (signature_of_type argtys retty)); ret ? 〈Eextcall f evargs (mk_eventval ? evres), Returnstate (mk_val ? evres) k m〉 ] | Returnstate res k m ⇒ match k with [ Kcall r f e k' ⇒ match r with [ None ⇒ ret ? 〈E0, (State f Sskip k' e m)〉 | Some r' ⇒ let 〈l,ty〉 ≝ r' in ! m' ← opt_to_io … (msg FailedStore) (store_value_of_type' ty m l res); ret ? 〈E0, (State f Sskip k' e m')〉 ] | _ ⇒ Wrong ??? (msg NonsenseState) ] ]. axiom MainMissing : String. definition make_global : clight_program → genv ≝ λp. globalenv Genv ?? (fst ??) p. let rec make_initial_state (p:clight_program) : res state ≝ let ge ≝ make_global p in do m0 ← init_mem Genv ?? (fst ??) p; do b ← opt_to_res ? (msg MainMissing) (find_symbol ? ? ge (prog_main ?? p)); do f ← opt_to_res ? (msg MainMissing) (find_funct_ptr ? ? ge b); OK ? (Callstate f (nil ?) Kstop m0). let rec is_final (s:state) : option int ≝ match s with [ Returnstate res k m ⇒ match k with [ Kstop ⇒ match res with [ Vint sz r ⇒ intsize_eq_elim ? sz I32 ? r (λr. Some ? r) (None ?) | _ ⇒ None ? ] | _ ⇒ None ? ] | _ ⇒ None ? ]. definition is_final_state : ∀st:state. (Sig ? (final_state st)) + (¬∃r. final_state st r). #st elim st; [ #f #s #k #e #m %2 ; % *; #r #H inversion H; #i #m #e destruct; | #f #l #k #m %2 ; % *; #r #H inversion H; #i #m #e destruct; | #v #k #m cases k; [ cases v; [ 2: * #i [ 1,2: %2 % * #r #H inversion H #i' #m #e destruct | %1 ; %{ i} //; ] | 1: %2 ; % *; #r #H inversion H; #i #m #e destruct; | #f %2 ; % *; #r #H inversion H; #i #m #e destruct; | #r %2 ; % *; #r #H inversion H; #i #m #e destruct; | #ptr %2 ; % *; #r #H inversion H; #i #m #e destruct; ] | #a #b %2 ; % *; #r #H inversion H; #i #m #e destruct; | 3,4: #a #b #c %2 ; % *; #r #H inversion H; #i #m #e destruct; | 5,6,8: #a #b #c #d %2 ; % *; #r #H inversion H; #i #m #e destruct; | #a %2 ; % *; #r #H inversion H; #i #m #e destruct; ] ] qed. let rec exec_steps (n:nat) (ge:genv) (s:state) : IO io_out io_in (trace×state) ≝ match is_final_state s with [ inl _ ⇒ ret ? 〈E0, s〉 | inr _ ⇒ match n with [ O ⇒ ret ? 〈E0, s〉 | S n' ⇒ ! 〈t,s'〉 ← exec_step ge s; ! 〈t',s''〉 ← exec_steps n' ge s'; ret ? 〈t ⧺ t',s''〉 ] ]. definition mem_of_state : state → mem ≝ λs.match s with [ State f s k e m ⇒ m | Callstate f a k m ⇒ m | Returnstate r k m ⇒ m ]. definition clight_exec : trans_system io_out io_in ≝ mk_trans_system ?? genv (λ_.state) (λ_.is_final) exec_step. definition clight_fullexec : fullexec io_out io_in ≝ mk_fullexec ??? clight_exec make_global make_initial_state.