include "Clight/Cexec.ma". (*include "Plogic/connectives.ma".*) (* Is rather careful about using destruct because it currently uses excessive normalization. *) lemma exec_bool_of_val_sound: ∀v,ty,r. exec_bool_of_val v ty = OK ? r → bool_of_val v ty (of_bool r). #v #ty #r cases v; [ | #i | #f | #r1 | #r' #b #pc #of ] cases ty; [ 2,11,20,29,38: #sz #sg | 3,12,21,30,39: #sz | 4,13,22,31,40: #rg #ty | 5,14,23,32,41: #r #ty #n | 6,15,24,33,42: #args #rty | 7,8,16,17,25,26,34,35,43,44: #id #fs | 9,18,27,36,45: #r #id ] #H whd in H:(??%?); [ 2: lapply (eq_spec i zero) cases (eq i zero) in H ⊢ % [ #E1 #E2 destruct @bool_of_val_false @is_false_int | #E1 #E2 >(?:r=¬false) [ @bool_of_val_true @is_true_int_int @E2 | destruct @refl ] ] | 8: cases (eq_dec f Fzero) [ #e >e in H ⊢ % >Feq_zero_true #E destruct @bool_of_val_false @is_false_float | #ne >Feq_zero_false in H // #E >(?:r=¬false) [ @bool_of_val_true @is_true_float @ne | destruct @refl ] ] | 14: >(?:r=false) [ @bool_of_val_false @is_false_pointer | destruct @refl ] | 15: >(?:r=true) [ @bool_of_val_true @is_true_pointer_pointer | destruct @refl ] | *: destruct ] qed. lemma bool_val_distinct: Vtrue ≠ Vfalse. % #H whd in H:(??%%); destruct; @(absurd ? e0 one_not_zero) qed. lemma bool_of: ∀v,ty,b. bool_of_val v ty (of_bool b) → if b then is_true v ty else is_false v ty. #v #ty #b cases b; #H inversion H; #v' #ty' #H' #ev #et #ev //; @False_ind @(absurd ? ev ?) [ 2: @sym_neq ] @bool_val_distinct qed. lemma try_cast_null_sound: ∀m,i,ty,ty',v'. try_cast_null m i ty ty' = OK ? v' → cast m (Vint i) ty ty' v'. #m #i #ty #ty' #v' whd in ⊢ (??%? → ?); lapply (eq_spec i zero); cases (eq i zero); [ #e >e cases ty; [ | #sz #sg | #fs | #sp #ty | #sp #ty #n | #args #rty | #id #fs | #id #fs | #r #id ] whd in ⊢ (??%? → ?); #H destruct; cases ty' in H ⊢ %; [ | #sz #sg | #fs | #sp #ty | #sp #ty #n | #args #rty | #id #fs | #id #fs | #r #id ] whd in ⊢ (??%? → ?); #H destruct (H); @cast_ip_z //; | #_ whd in ⊢ (??%? → ?); #H destruct ] qed. definition exec_cast_sound : ∀m:mem. ∀v:val. ∀ty:type. ∀ty':type. ∀v':val. exec_cast m v ty ty' = OK ? v' → cast m v ty ty' v'. #m #v #ty #ty' #v' cases v; [ #H whd in H:(??%?); destruct; | #i cases ty; [ #H whd in H:(??%?); destruct; | 3: #a #H whd in H:(??%?); destruct; | 7,8,9: #a #b #H whd in H:(??%?); destruct; | #sz1 #si1 cases ty'; [ #H whd in H:(??%?); destruct; | 3: #a #H whd in H:(??%?); destruct; // | 2,7,8,9: #a #b #H whd in H:(??%?); destruct; // | 4,5,6: [ #sp #ty'' letin t ≝ (Tpointer sp ty'') | #sp #ty'' #n letin t ≝ (Tarray sp ty'' n) | #args #rty letin t ≝ (Tfunction args rty) ] whd in ⊢ (??%? → ?); lapply (try_cast_null_sound m i (Tint sz1 si1) t v'); cases (try_cast_null m i (Tint sz1 si1) t); [ 1,3,5: #v'' #H' #e @H' @e | *: #m #_ whd in ⊢ (??%? → ?); #H destruct (H); ] ] | *: [ #sp #ty'' letin t ≝ (Tpointer sp ty'') | #sp #ty'' #n letin t ≝ (Tarray sp ty'' n) | #args #rty letin t ≝ (Tfunction args rty) ] whd in ⊢ (??%? → ?); lapply (try_cast_null_sound m i t ty' v'); cases (try_cast_null m i t ty'); [ 1,3,5: #v'' #H' #e @H' @e | *: #m #_ whd in ⊢ (??%? → ?); #H destruct (H); ] ] | #f cases ty; [ 3: #x | 2,4,6,7,8,9: #x #y | 5: #x #y #z ] [ cases ty'; [ #e | 3: #a #e | 2,4,6,7,8,9: #a #b #e | #a #b #c #e ] whd in e:(??%?); destruct; //; | *: #e whd in e:(??%?); destruct ] | #r cases ty; [ 3: #x | 2,4,6,7,8,9: #x #y | 5: #x #y #z ] whd in ⊢ (??%? → ?); #H destruct; cases (eq_region_dec r ?) in H; whd in ⊢ (? → ??%? → ?); #H1 #H2 destruct; cases ty' in H2; normalize; try #a try #b try #c try #d destruct; @cast_pp_z //; | #r #b #pc #of whd in ⊢ (??%? → ?) #e elim (bind_inversion ????? e) #s * #es #e' elim (bind_inversion ????? e') #u * #eu #e'' -e'; elim (bind_inversion ????? e'') #s' * #es' #e'''; -e''; cut (type_region ty s); [ cases ty in es:(??%?) ⊢ %; [ #e | 3: #a #e | 2,4,6,7,8,9: #a #b #e | #a #b #c #e ] whd in e:(??%?); destruct; //; | #Hty cut (type_region ty' s'); [ cases ty' in es' ⊢ %; [ #e | 3: #a #e | 2,4,6,7,8,9: #a #b #e | #a #b #c #e ] whd in e:(??%?); destruct; //; | #Hty' cut (s = r). elim (eq_region_dec r s) in eu; //; normalize; #_ #e destruct. #e >e in Hty #Hty cases (pointer_compat_dec b s') in e''' #Hcompat #e''' whd in e''':(??%?); destruct (e'''); /2/ ] ] ] qed. let rec expr_lvalue_ind (P:expr → Prop) (Q:expr_descr → type → Prop) (ci:∀ty,i.P (Expr (Econst_int i) ty)) (cf:∀ty,f.P (Expr (Econst_float f) ty)) (lv:∀e,ty. Q e ty → Plvalue P e ty) (vr:∀v,ty.Q (Evar v) ty) (dr:∀e,ty.P e → Q (Ederef e) ty) (ao:∀ty,e,ty'.Q e ty' → P (Expr (Eaddrof (Expr e ty')) ty)) (uo:∀ty,op,e.P e → P (Expr (Eunop op e) ty)) (bo:∀ty,op,e1,e2.P e1 → P e2 → P (Expr (Ebinop op e1 e2) ty)) (ca:∀ty,ty',e.P e → P (Expr (Ecast ty' e) ty)) (cd:∀ty,e1,e2,e3.P e1 → P e2 → P e3 → P (Expr (Econdition e1 e2 e3) ty)) (ab:∀ty,e1,e2.P e1 → P e2 → P (Expr (Eandbool e1 e2) ty)) (ob:∀ty,e1,e2.P e1 → P e2 → P (Expr (Eorbool e1 e2) ty)) (sz:∀ty,ty'. P (Expr (Esizeof ty') ty)) (fl:∀ty,e,ty',i. Q e ty' → Q (Efield (Expr e ty') i) ty) (co:∀ty,l,e. P e → P (Expr (Ecost l e) ty)) (xx:∀e,ty. is_not_lvalue e → Q e ty) (e:expr) on e : P e ≝ match e with [ Expr e' ty ⇒ match e' with [ Econst_int i ⇒ ci ty i | Econst_float f ⇒ cf ty f | Evar v ⇒ lv (Evar v) ty (vr v ty) | Ederef e'' ⇒ lv (Ederef e'') ty (dr e'' ty (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e'')) | Eaddrof e'' ⇒ match e'' with [ Expr e0 ty0 ⇒ ao ty e0 ty0 (lvalue_expr_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e0 ty0) ] | Eunop op e'' ⇒ uo ty op e'' (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e'') | Ebinop op e1 e2 ⇒ bo ty op e1 e2 (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e1) (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e2) | Ecast ty' e'' ⇒ ca ty ty' e'' (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e'') | Econdition e1 e2 e3 ⇒ cd ty e1 e2 e3 (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e1) (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e2) (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e3) | Eandbool e1 e2 ⇒ ab ty e1 e2 (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e1) (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e2) | Eorbool e1 e2 ⇒ ob ty e1 e2 (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e1) (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e2) | Esizeof ty' ⇒ sz ty ty' | Efield e'' i ⇒ match e'' with [ Expr ef tyf ⇒ lv (Efield (Expr ef tyf) i) ty (fl ty ef tyf i (lvalue_expr_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx ef tyf)) ] | Ecost l e'' ⇒ co ty l e'' (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e'') ] ] and lvalue_expr_ind (P:expr → Prop) (Q:expr_descr → type → Prop) (ci:∀ty,i.P (Expr (Econst_int i) ty)) (cf:∀ty,f.P (Expr (Econst_float f) ty)) (lv:∀e,ty. Q e ty → Plvalue P e ty) (vr:∀v,ty.Q (Evar v) ty) (dr:∀e,ty.P e → Q (Ederef e) ty) (ao:∀ty,e,ty'.Q e ty' → P (Expr (Eaddrof (Expr e ty')) ty)) (uo:∀ty,op,e.P e → P (Expr (Eunop op e) ty)) (bo:∀ty,op,e1,e2.P e1 → P e2 → P (Expr (Ebinop op e1 e2) ty)) (ca:∀ty,ty',e.P e → P (Expr (Ecast ty' e) ty)) (cd:∀ty,e1,e2,e3.P e1 → P e2 → P e3 → P (Expr (Econdition e1 e2 e3) ty)) (ab:∀ty,e1,e2.P e1 → P e2 → P (Expr (Eandbool e1 e2) ty)) (ob:∀ty,e1,e2.P e1 → P e2 → P (Expr (Eorbool e1 e2) ty)) (sz:∀ty,ty'. P (Expr (Esizeof ty') ty)) (fl:∀ty,e,ty',i. Q e ty' → Q (Efield (Expr e ty') i) ty) (co:∀ty,l,e. P e → P (Expr (Ecost l e) ty)) (xx:∀e,ty. is_not_lvalue e → Q e ty) (e:expr_descr) (ty:type) on e : Q e ty ≝ match e return λe0. Q e0 ty with [ Evar v ⇒ vr v ty | Ederef e'' ⇒ dr e'' ty (expr_lvalue_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e'') | Efield e' i ⇒ match e' return λe1.Q (Efield e1 i) ty with [ Expr e'' ty'' ⇒ fl ty e'' ty'' i (lvalue_expr_ind P Q ci cf lv vr dr ao uo bo ca cd ab ob sz fl co xx e'' ty'') ] | _ ⇒ xx ? ty ? ]. whd; @I qed. theorem exec_expr_sound: ∀ge:genv. ∀en:env. ∀m:mem. ∀e:expr. (P_res ? (λx.eval_expr ge en m e (\fst x) (\snd x)) (exec_expr ge en m e)). #ge #en #m #e @(expr_lvalue_ind ? (λe',ty.P_res ? (λr.eval_lvalue ge en m (Expr e' ty) (\fst (\fst r)) (\snd (\fst r)) (\snd r)) (exec_lvalue' ge en m e' ty)) … e) (* XXX // fails [ 1,2: #ty #c whd // *) [ #ty #c whd % | #ty #c whd %2 (* expressions that are lvalues *) | #e' #ty cases e'; //; [ #i #He' | #e #He' | #e #i #He' ] whd in He' ⊢ %; @bind2_OK #x cases x; #y cases y; #sp #loc #ofs #tr #H @opt_bind_OK #vl #evl whd in evl:(??%?); @(eval_Elvalue … evl) >H in He' #He' @He' | #v #ty whd in ⊢ (???%); lapply (refl ? (get ident PTree block v en)); cases (get ident PTree block v en) in ⊢ (???% → %); [ #eget @opt_bind_OK #sploc cases sploc; #sp #loc #efind whd; @(eval_Evar_global … eget efind) | #loc #eget @(eval_Evar_local … eget) ] | #ty #e #He whd in ⊢ (???%) @bind2_OK #v cases v // #r #l #pc #ofs #tr #Hv whd >Hv in He #He @eval_Ederef [ 3: @He | *: skip ] | #ty #e'' #ty'' #He'' @bind2_OK * #loc #ofs #tr #H cases ty // * #pty cases loc in H ⊢ % * #loc' #H whd try @I @eval_Eaddrof whd in H:(??%?) >H in He'' #He'' @He'' | #ty #op #e1 #He1 @bind2_OK #v1 #tr #Hv1 @opt_bind_OK #v #ev @(eval_Eunop … ev) >Hv1 in He1 #He1 @He1 | #ty #op #e1 #e2 #He1 #He2 @bind2_OK #v1 #tr1 #ev1 >ev1 in He1 #He1 @bind2_OK #v2 #tr2 #ev2 >ev2 in He2 #He2 @opt_bind_OK #v #ev whd in He1 He2; whd; @(eval_Ebinop … He1 He2 ev) | #ty #ty' #e' #He' @bind2_OK #v #tr #Hv >Hv in He' #He' @bind_OK #v' #ev' @(eval_Ecast … He' ?) (* XXX /2/; *) @(exec_cast_sound … ev') | #ty #e1 #e2 #e3 #He1 #He2 #He3 @bind2_OK #vb #tr1 #Hvb >Hvb in He1 #He1 @bind_OK #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb @bind2_OK #v #tr #Hv whd in Hv:(??%?) [ >Hv in He2 #He2 whd in He2 Hv:(??%?) ⊢%; @(eval_Econdition_true … He1 ? He2) @(bool_of ??? Hb) | >Hv in He3 #He3 whd in He3 Hv:(??%?) ⊢%; @(eval_Econdition_false … He1 ? He3) @(bool_of ??? Hb) ] | #ty #e1 #e2 #He1 #He2 @bind2_OK #v1 #tr1 #Hv1 >Hv1 in He1 #He1 @bind_OK #b1 cases b1; #eb1 lapply (exec_bool_of_val_sound … eb1); #Hb1 [ @bind2_OK #v2 #tr2 #Hv2 >Hv2 in He2 #He2 @bind_OK #b2 #eb2 @(eval_Eandbool_2 … He1 … He2) [ @(bool_of … Hb1) | @(exec_bool_of_val_sound … eb2) ] | @(eval_Eandbool_1 … He1) @(bool_of … Hb1) ] | #ty #e1 #e2 #He1 #He2 @bind2_OK #v1 #tr1 #Hv1 >Hv1 in He1 #He1 @bind_OK #b1 cases b1; #eb1 lapply (exec_bool_of_val_sound … eb1); #Hb1 [ @(eval_Eorbool_1 … He1) @(bool_of … Hb1) | @bind2_OK #v2 #tr2 #Hv2 >Hv2 in He2 #He2 @bind_OK #b2 #eb2 @(eval_Eorbool_2 … He1 … He2) [ @(bool_of … Hb1) | @(exec_bool_of_val_sound … eb2) ] ] | #ty #ty' whd; (* XXX //*) @eval_Esizeof | #ty #e' #ty' #i cases ty'; //; [ #id #fs #He' @bind2_OK #x cases x; #sp #l #ofs #H @bind_OK #delta #Hdelta whd in H:(??%?) >H in He' #He' @(eval_Efield_struct … He' (refl ??) Hdelta) | #id #fs #He' @bind2_OK #x cases x; #sp #l #ofs #H whd in H:(??%?) >H in He' #He' @(eval_Efield_union … He' (refl ??)) ] | #ty #l #e' #He' @bind2_OK #v #tr1 #H >H in He' #He' @(eval_Ecost … He') (* exec_lvalue fails on non-lvalues. *) | #e' #ty cases e'; [ 1,2,5,12: #a #H | 3,4: #a * | 13,14: #a #b * | 6,8,10,11: #a #b #H | 7,9: #a #b #c #H ] @I ] qed. lemma addrof_eval_lvalue: ∀ge,en,m,e,r,loc,pc,off,tr,ty. eval_expr ge en m (Expr (Eaddrof e) ty) (Vptr r loc pc off) tr → eval_lvalue ge en m e loc off tr. #ge #en #m #e #r #loc #pc #off #tr #ty #H inversion H; [ 1,2,5: #a #b #H @False_ind destruct (H); | #a #b #c #d #e #f #H1 #g #H2

H whd in ⊢ (??%?) cases r @refl qed. theorem exec_lvalue_sound: ∀ge,en,m,e. P_res ? (λr.eval_lvalue ge en m e (\fst (\fst r)) (\snd (\fst r)) (\snd r)) (exec_lvalue ge en m e). #ge #en #m #e lapply (refl ? (exec_lvalue ge en m e)); cases (exec_lvalue ge en m e) in ⊢ (???% → %); [ #x cases x #y cases y #loc #off #tr #H whd cases loc in H ⊢ % #locr #loci #H @(addrof_eval_lvalue … locr … (Tpointer locr Tvoid)) [ @same_compat ] lapply (addrof_exec_lvalue … H) #H' lapply (exec_expr_sound ge en m (Expr (Eaddrof e) (Tpointer locr Tvoid))) >H' #H'' @H'' | #msg #_ whd @I ] qed. (* Plain equality versions of the above *) definition exec_expr_sound' ≝ λge,en,m,e,v. λH:exec_expr ge en m e = OK ? v. P_res_to_P ???? (exec_expr_sound ge en m e) H. definition exec_lvalue_sound' ≝ λge,en,m,e,loc,off,tr. λH:exec_lvalue ge en m e = OK ? 〈〈loc,off〉,tr〉. P_res_to_P ???? (exec_lvalue_sound ge en m e) H. lemma exec_exprlist_sound: ∀ge,e,m,l. P_res ? (λvltr:list val×trace. eval_exprlist ge e m l (\fst vltr) (\snd vltr)) (exec_exprlist ge e m l). #ge #e #m #l elim l; whd; (* XXX //; *) [ % | #e1 #es #IH @bind2_OK #v #tr1 #Hv @bind2_OK #vs #tr2 #Hvs whd; @eval_Econs [ @(P_res_to_P … (exec_expr_sound ge e m e1) Hv) | @(P_res_to_P … IH Hvs) ] ] qed. lemma exec_alloc_variables_sound : ∀l,en,m,en',m'. exec_alloc_variables en m l = 〈en',m'〉 → alloc_variables en m l en' m'. #l elim l [ #en #m #en' #m' #EXEC whd in EXEC:(??%?); destruct % | * #id #ty #t #IH #en #m #en' #m' lapply (refl ? (alloc m O (sizeof ty) Any)) #ALLOC #EXEC whd in EXEC:(??%?) ALLOC:(???%) @(alloc_variables_cons … ALLOC) @IH @EXEC qed. lemma exec_bind_parameters_sound : ∀ps,vs,en,m. P_res ? (λm'. bind_parameters en m ps vs m') (exec_bind_parameters en m ps vs). #ps elim ps [ * // | * #id #ty #ps' #IH * [ // | #v #vs #en #m @opt_bind_OK #b #GET @opt_bind_OK #m' #STORE lapply (refl ? (exec_bind_parameters en m' ps' vs)) cases (exec_bind_parameters en m' ps' vs) in ⊢ (???% → %) [2: #msg #_ %] #m'' #BIND @(bind_parameters_cons … GET STORE) lapply (IH vs en m') whd in ⊢ (% → ?) >BIND // ] ] qed. lemma check_eventval_list_sound : ∀vs,tys. P_res ? (λevs. eventval_list_match evs tys vs) (check_eventval_list vs tys). #vs0 elim vs0 [ * // | #v #vs #IH * [ // | #ty #tys whd in ⊢ (???%) cases ty cases v // #v' #sz try #sg @bind_OK #evs #CHECK @(evl_match_cons ??????? (P_res_to_P ???? (IH ?) CHECK)) // ] ] qed. theorem exec_step_sound: ∀ge,st. P_io ??? (λr. step ge st (\fst r) (\snd r)) (exec_step ge st). #ge #st cases st; [ #f #s #k #e #m cases s; [ cases k; [ whd in ⊢ (?????%); lapply (refl ? (fn_return f)); cases (fn_return f) in ⊢ (???% → %); //; #H whd; @step_skip_call //; | #s' #k' whd; (* XXX //; *) @step_skip_seq | #ex #s' #k' @step_skip_or_continue_while % //; | #ex #s' #k' @res_bindIO2_OK #v #tr #Hv letin bexpr ≝ (exec_bool_of_val v (typeof ex)); lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); [ #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb whd in ⊢ (?????%); [ @(step_skip_or_continue_dowhile_true … (exec_expr_sound' … Hv)) [ % // | @(bool_of … Hb) ] | @(step_skip_or_continue_dowhile_false … (exec_expr_sound' … Hv)) [ % // | @(bool_of … Hb) ] ] | #msg #_ //; ] | #ex #s1 #s2 #k' @step_skip_or_continue_for2 % //; | #ex #s1 #s2 #k' @step_skip_for3 | #k' @step_skip_break_switch % //; | #r #f' #e' #k' whd in ⊢ (?????%); lapply (refl ? (fn_return f)); cases (fn_return f) in ⊢ (???% → %); //; #H whd; @step_skip_call //; ] | #ex1 #ex2 @res_bindIO2_OK #x cases x; #y cases y; #pcl #loc #ofs #tr1 #Hlval @res_bindIO2_OK #v2 #tr2 #Hv2 @opt_bindIO_OK #m' #em' whd; @(step_assign … (exec_lvalue_sound' … Hlval) (exec_expr_sound' … Hv2) em') | #lex #fex #args @res_bindIO2_OK #vf #tr1 #Hvf0 lapply (exec_expr_sound' … Hvf0); #Hvf @res_bindIO2_OK #vargs #tr2 #Hvargs0 lapply (P_res_to_P ???? (exec_exprlist_sound …) Hvargs0); #Hvargs @opt_bindIO_OK #fd #efd @bindIO_OK #ety cases lex; whd; [ @(step_call_none … Hvf Hvargs efd ety) | #lhs' @res_bindIO2_OK #x cases x; #y cases y; #pcl #loc #ofs #tr3 #Hlocofs whd; @(step_call_some … (exec_lvalue_sound' … Hlocofs) Hvf Hvargs efd ety) ] | #s1 #s2 whd; (* XXX //; *) @step_seq | #ex #s1 #s2 @res_bindIO2_OK #v #tr #Hv letin bexpr ≝ (exec_bool_of_val v (typeof ex)); lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); //; #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb [ @(step_ifthenelse_true … (exec_expr_sound' … Hv)) @(bool_of … Hb) | @(step_ifthenelse_false … (exec_expr_sound' … Hv)) @(bool_of … Hb) ] | #ex #s' @res_bindIO2_OK #v #tr #Hv letin bexpr ≝ (exec_bool_of_val v (typeof ex)); lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); //; #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb [ @(step_while_true … (exec_expr_sound' … Hv)) @(bool_of … Hb) | @(step_while_false … (exec_expr_sound' … Hv)) @(bool_of … Hb) ] | #ex #s' whd; (* XXX //; *) @step_dowhile | #s1 #ex #s2 #s3 whd in ⊢ (?????%); elim (is_Sskip s1); #Hs1 whd in ⊢ (?????%); [ >Hs1 @res_bindIO2_OK #v #tr #Hv letin bexpr ≝ (exec_bool_of_val v (typeof ex)); lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); //; #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb [ @(step_for_true … (exec_expr_sound' … Hv)) @(bool_of … Hb) | @(step_for_false … (exec_expr_sound' … Hv)) @(bool_of … Hb) ] | @step_for_start //; ] | whd in ⊢ (?????%); cases k; //; [ #s' #k' whd (* XXX // *) @step_break_seq | #ex #s' #k' whd (* //; *) @step_break_while | #ex #s' #k' whd (* //; *) @step_break_dowhile | #ex #s1 #s2 #k' whd (* //; *) @step_break_for2 | #k' @step_skip_break_switch %2 // ] | whd in ⊢ (?????%); cases k; //; [ #s' #k' whd; (* XXX //;*) @step_continue_seq | #ex #s' #k' whd; @step_skip_or_continue_while %2 ; //; | #ex #s' #k' whd; @res_bindIO2_OK #v #tr #Hv letin bexpr ≝ (exec_bool_of_val v (typeof ex)); lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); //; #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb [ @(step_skip_or_continue_dowhile_true … (exec_expr_sound' … Hv)) [ %2 ; // | @(bool_of … Hb) ] | @(step_skip_or_continue_dowhile_false … (exec_expr_sound' … Hv)) [ %2 ; // | @(bool_of … Hb) ] ] | #ex #s1 #s2 #k' whd; @step_skip_or_continue_for2 %2 ; // | #k' whd; (* XXX //;*) @step_continue_switch ] | #r whd in ⊢ (?????%); cases r; [ whd; lapply (refl ? (fn_return f)); cases (fn_return f) in ⊢ (???% → %); //; #H @step_return_0 @H | #ex cases (type_eq_dec (fn_return f) Tvoid); //; whd in ⊢ (% → ?????%); #Hnotvoid @res_bindIO2_OK #v #tr #Hv whd; @(step_return_1 … Hnotvoid (exec_expr_sound' … Hv)) ] | #ex #ls @res_bindIO2_OK #v cases v; //; #n #tr #Hv @step_switch @(exec_expr_sound' … Hv) | #l #s' whd; @step_label (* XXX //; *) | #l whd in ⊢ (?????%); lapply (refl ? (find_label l (fn_body f) (call_cont k))); cases (find_label l (fn_body f) (call_cont k)) in ⊢ (???% → %); //; #sk cases sk; #s' #k' #H @(step_goto … H) | #l #s' whd; (* XXX //; *) @step_cost ] | #f0 #vargs #k #m whd in ⊢ (?????%); cases f0; [ #fn whd in ⊢ (?????%) lapply (refl ? (exec_alloc_variables empty_env m (fn_params fn@fn_vars fn))) cases (exec_alloc_variables empty_env m (fn_params fn@fn_vars fn)) in ⊢ (???% → %) #en' #m' #ALLOC whd in ⊢ (?????%) @res_bindIO_OK #m2 #BIND whd; @(step_internal_function … (exec_alloc_variables_sound … ALLOC)) @(P_res_to_P … (exec_bind_parameters_sound …) BIND) | #id #argtys #rty @res_bindIO_OK #evs #Hevs @bindIO_OK #eres whd; @step_external_function % [ @(P_res_to_P … (check_eventval_list_sound …) Hevs) | @mk_val_correct ] ] | #v #k' #m' whd in ⊢ (?????%); cases k'; //; #r #f #e #k whd in ⊢ (?????%); cases r; [ whd; @step_returnstate_0 | #x cases x; #y cases y; #z cases z; #pcl #b #ofs #ty @opt_bindIO_OK #m' #em' @step_returnstate_1 whd in em':(??%?); //; ] ] qed. lemma make_initial_state_sound : ∀p. P_res ? (λgs.globalenv Genv ?? p = OK ? (\fst gs) ∧ initial_state p (\snd gs)) (make_initial_state p). #p cases p; #fns #main #vars whd in ⊢ (???%); @bind_OK #ge #Ege @bind_OK #m #Em @opt_bind_OK #x cases x; #sp #b #esb @opt_bind_OK #f #ef whd; % [ whd in ⊢ (???(??%)) // | @(initial_state_intro … Ege Em esb ef) ] qed. theorem exec_steps_sound: ∀ge,n,st. P_io ??? (λts:trace×state. star (mk_transrel ?? step) ge st (\fst ts) (\snd ts)) (exec_steps n ge st). #ge #n elim n; [ #st whd in ⊢ (?????%); elim (is_final_state st); #H whd; % | #n' #IH #st whd in ⊢ (?????%); elim (is_final_state st); #H [ whd; % | @(P_bindIO2_OK ????????? (exec_step_sound …)) #t #s' cases s'; [ #f #s #k #e #m | #fd #args #k #m | #r #k #m ] whd in ⊢ (? → ?????(??????%?)); cases m; #mc #mn #mp #Hstep whd in ⊢ (?????(??????%?)); @(P_bindIO2_OK ????????? (IH …)) #t' #s'' #Hsteps whd; @(star_step ????????? Hsteps) [ 2,5,8: @Hstep | 3,6,9: // ] ] qed. lemma is_final_sound: ∀s,r. is_final s = Some ? r → final_state s r. * [ 3: #v * [ #m #r cases v [ 2: #r' #E normalize in E; destruct % | *: normalize #x1 try #x2 try #x3 try #x4 try #x5 destruct ] | *: normalize #x1 try #x2 try #x3 try #x4 try #x5 try #x6 try #x7 destruct ] | *: normalize #x1 #x2 #x3 #x4 #x5 #x6 try #x7 destruct ] qed.