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