include "Csem.ma". include "extralib.ma". include "IOMonad.ma". include "Plogic/russell_support.ma". ndefinition P_to_P_option_res : ∀A:Type[0].∀P:A → CProp[0].option (res A) → CProp[0] ≝ λA,P,a.match a with [ None ⇒ False | Some y ⇒ match y return λ_.CProp[0] with [ Error ⇒ True | OK z ⇒ P z ]]. ndefinition err_inject : ∀A.∀P:A → Prop.∀a:option (res A).∀p:P_to_P_option_res A P a.res (sigma A P) ≝ λA.λP:A → Prop.λa:option (res A).λp:P_to_P_option_res A P a. (match a return λa'.a=a' → res (sigma A P) with [ None ⇒ λe1.? | Some b ⇒ λe1.(match b return λb'.b=b' → ? with [ Error ⇒ λ_. Error ? | OK c ⇒ λe2. OK ? (sig_intro A P c ?) ]) (refl ? b) ]) (refl ? a). ##[ nrewrite > e1 in p; nnormalize; *; ##| nrewrite > e1 in p; nrewrite > e2; nnormalize; // ##] nqed. ndefinition err_eject : ∀A.∀P: A → Prop. res (sigma A P) → res A ≝ λA,P,a.match a with [ Error ⇒ Error ? | OK b ⇒ match b with [ sig_intro w p ⇒ OK ? w] ]. ndefinition sig_eject : ∀A.∀P: A → Prop. sigma A P → A ≝ λA,P,a.match a with [ sig_intro w p ⇒ w]. ncoercion err_inject : ∀A.∀P:A → Prop.∀a.∀p:P_to_P_option_res ? P a.res (sigma A P) ≝ err_inject on a:option (res ?) to res (sigma ? ?). ncoercion err_eject : ∀A.∀P:A → Prop.∀c:res (sigma A P).res A ≝ err_eject on _c:res (sigma ? ?) to res ?. ncoercion sig_eject : ∀A.∀P:A → Prop.∀c:sigma A P.A ≝ sig_eject on _c:sigma ? ? to ?. ndefinition P_res: ∀A.∀P:A → Prop.res A → Prop ≝ λA,P,a. match a with [ Error ⇒ True | OK v ⇒ P v ]. ndefinition exec_bool_of_val : ∀v:val. ∀ty:type. res bool ≝ λv,ty. match v in val with [ Vint i ⇒ match ty with [ Tint _ _ ⇒ OK ? (¬eq i zero) | Tpointer _ _ ⇒ OK ? (¬eq i zero) | _ ⇒ Error ? ] | Vfloat f ⇒ match ty with [ Tfloat _ ⇒ OK ? (¬Fcmp Ceq f Fzero) | _ ⇒ Error ? ] | Vptr _ _ _ ⇒ match ty with [ Tint _ _ ⇒ OK ? true | Tpointer _ _ ⇒ OK ? true | _ ⇒ Error ? ] | _ ⇒ Error ? ]. nlemma 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; ncases v; ##[ ##2: #i ##| ##3: #f ##| ##4: #sp b of ##] nwhd; ##[ ##4: #H; nwhd in H:(??%?); ndestruct ##] ncases ty; ##[ ##2,11,20: #sz sg ##| ##3,12,21: #sz ##| ##4,13,22: #sp ty ##| ##5,14,23: #sp ty n ##| ##6,15,24: #args rty ##| ##7,8,16,17,25,26: #id fs ##| ##9,18,27: #id ##] #H; nwhd in H:(??%?); ndestruct; ##[ ##1,4: nlapply (eq_spec i zero); nelim (eq i zero); ##[ ##1,3: #e; nrewrite > e; napply bool_of_val_false; //; ##| ##2,4: #ne; napply bool_of_val_true; /2/; ##] ##| ##3: nelim (eq_dec f Fzero); ##[ #e; nrewrite > e; nrewrite > (Feq_zero_true …); napply bool_of_val_false; //; ##| #ne; nrewrite > (Feq_zero_false …); //; napply bool_of_val_true; /2/; ##] ##| ##2,5: napply bool_of_val_true; // ##] nqed. nlemma 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; nelim H; #v t H'; nelim H'; ##[ #i is s ne; @ true; @; //; nwhd in ⊢ (??%?); nrewrite > (eq_false … ne); //; ##| #p b i i0 s; @ true; @; // ##| #i p t ne; @ true; @; //; nwhd in ⊢ (??%?); nrewrite > (eq_false … ne); //; ##| #p b i p0 t0; @ true; @; // ##| #f s ne; @ true; @; //; nwhd in ⊢ (??%?); nrewrite > (Feq_zero_false … ne); //; ##| #i s; @ false; @; //; ##| #p t; @ false; @; //; ##| #s; @ false; @; //; nwhd in ⊢ (??%?); nrewrite > (Feq_zero_true …); //; ##] nqed. (* Prove a few minor results to make proof obligations easy. *) nlemma bind_assoc_r: ∀A,B,C,e,f,g. bind B C (bind A B e f) g = bind A C e (λx.bind B C (f x) g). #A B C e f g; ncases e; nnormalize; //; nqed. nlemma 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; nelim e; /2/; nqed. nlemma sig_bind_OK: ∀A,B. ∀P:A → Prop. ∀P':B → Prop. ∀e:res (sigma A P). ∀f:sigma A P → res B. (∀v:A. ∀p:P v. match f (sig_intro A P v p) with [ Error ⇒ True | OK v' ⇒ P' v'] ) → match bind (sigma A P) B e f with [ Error ⇒ True | OK v' ⇒ P' v' ]. #A B P P' e f; nelim e; ##[ #v0; nelim v0; #v Hv IH; napply IH; ##| #_; napply I; ##] nqed. nlemma 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; nelim e; //; #v; ncases v; /2/; nqed. nlemma sig_bind2_OK: ∀A,B,C. ∀P:A×B → Prop. ∀P':C → Prop. ∀e:res (sigma (A×B) P). ∀f:A → B → res C. (∀v1:A.∀v2:B. P 〈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 P' e f; nelim e; //; #v0; nelim v0; #v; nelim v; #v1 v2 Hv IH; napply IH; //; nqed. nlemma bool_val_distinct: Vtrue ≠ Vfalse. @; #H; nwhd in H:(??%%); ndestruct; napply (absurd ? e0 one_not_zero); nqed. nlemma 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; ncases b; #H; ninversion H; #v' ty' H' ev et ev; //; napply False_ind; napply (absurd ? ev ?); ##[ ##2: napply sym_neq ##] napply bool_val_distinct; nqed. ndefinition opt_to_res ≝ λA.λv:option A. match v with [ None ⇒ Error A | Some v ⇒ OK A v ]. nlemma opt_OK: ∀A,P,e. (∀v. e = Some ? v → P v) → match opt_to_res A e with [ Error ⇒ True | OK v ⇒ P v ]. #A P e; nelim e; /2/; nqed. nlemma opt_bind_OK: ∀A,B,P,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 e) f with [ Error ⇒ True | OK v ⇒ P v ]. #A B P e f; nelim e; nnormalize; /2/; nqed. nlemma extract_subset_pair: ∀A,B,C,P. ∀e:{e:A×B | P e}. ∀Q:A→B→res C. ∀R:C→Prop. (∀a,b. eject ?? e = 〈a,b〉 → P 〈a,b〉 → match Q a b with [ OK v ⇒ R v | Error ⇒ True]) → match match eject ?? e with [ mk_pair a b ⇒ Q a b ] with [ OK v ⇒ R v | Error ⇒ True ]. #A B C P e Q R; ncases e; #e'; ncases e'; nnormalize; ##[ #H; napply (False_ind … H); ##| #e''; ncases e''; #a b Pab H; nnormalize; /2/; ##] nqed. (* nremark err_later: ∀A,B. ∀e:res A. match e with [ Error ⇒ Error B | OK v ⇒ Error B ] = Error B. #A B e; ncases e; //; nqed. *) ndefinition try_cast_null : ∀m:mem. ∀i:int. ∀ty:type. ∀ty':type. res val ≝ λm:mem. λi:int. λty:type. λty':type. match eq i zero with [ true ⇒ match ty with [ Tint _ _ ⇒ match ty' with [ Tpointer _ _ ⇒ OK ? (Vint i) | Tarray _ _ _ ⇒ OK ? (Vint i) | Tfunction _ _ ⇒ OK ? (Vint i) | _ ⇒ Error ? ] | Tpointer _ _ ⇒ match ty' with [ Tpointer _ _ ⇒ OK ? (Vint i) | Tarray _ _ _ ⇒ OK ? (Vint i) | Tfunction _ _ ⇒ OK ? (Vint i) | _ ⇒ Error ? ] | Tarray _ _ _ ⇒ match ty' with [ Tpointer _ _ ⇒ OK ? (Vint i) | Tarray _ _ _ ⇒ OK ? (Vint i) | Tfunction _ _ ⇒ OK ? (Vint i) | _ ⇒ Error ? ] | Tfunction _ _ ⇒ match ty' with [ Tpointer _ _ ⇒ OK ? (Vint i) | Tarray _ _ _ ⇒ OK ? (Vint i) | Tfunction _ _ ⇒ OK ? (Vint i) | _ ⇒ Error ? ] | _ ⇒ Error ? ] | false ⇒ Error ? ]. nlemma 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'; nwhd in ⊢ (??%? → ?); nlapply (eq_spec i zero); ncases (eq i zero); ##[ #e; nrewrite > e; ncases ty; ##[ ##| #sz sg ##| #fs ##| #sp ty ##| #sp ty n ##| #args rty ##| #id fs ##| #id fs ##| #id ##] nwhd in ⊢ (??%? → ?); ##[ ##1,3,7,8,9: #H; ndestruct ##] ncases ty'; ##[ ##2,11,20,29: #sz' sg' ##| ##3,12,21,30: #sz' ##| ##4,13,22,31: #sp' ty' ##| ##5,14,23,32: #sp' ty' n' ##| ##6,15,24,33: #args' res' ##| ##7,8,16,17,25,26,34,35: #id' fs' ##| ##9,18,27,36: #id' ##] nwhd in ⊢ (??%? → ?); #H; ndestruct (H); ##[ ##1,5,9: napply cast_ip_z ##| ##*: napply cast_pp_z ##] //; ##| #_; nwhd in ⊢ (??%? → ?); #H; ndestruct ##] nqed. ndefinition ms_eq_dec : ∀s1,s2:memory_space. (s1 = s2) + (s1 ≠ s2). #s1; ncases s1; #s2; ncases s2; /2/; @2; @; #H; ndestruct; nqed. ndefinition exec_cast : ∀m:mem. ∀v:val. ∀ty:type. ∀ty':type. res val ≝ λm:mem. λv:val. λty:type. λty':type. match v with [ Vint i ⇒ match ty with [ Tint sz1 si1 ⇒ match ty' with [ Tint sz2 si2 ⇒ OK ? (Vint (cast_int_int sz2 si2 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 ? ] | 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 ? ] | Vfloat f ⇒ match ty with [ Tfloat sz ⇒ match ty' with [ Tint sz' si' ⇒ OK ? (Vint (cast_int_int sz' si' (cast_float_int si' f))) | Tfloat sz' ⇒ OK ? (Vfloat (cast_float_float sz' f)) | _ ⇒ Error ? ] | _ ⇒ Error ? ] | Vptr p b ofs ⇒ do s ← match ty with [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code | _ ⇒ Error ? ]; do u ← match ms_eq_dec p s with [ inl _ ⇒ OK ? something | inr _ ⇒ Error ? ]; do s' ← match ty' with [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code | _ ⇒ Error ? ]; if is_pointer_compat (block_space m b) s' then OK ? (Vptr s' b ofs) else Error ? | _ ⇒ Error ? ]. ndefinition 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'; ncases v; ##[ #H; nwhd in H:(??%?); ndestruct; ##| #i; ncases ty; ##[ #H; nwhd in H:(??%?); ndestruct; ##| ##3,9: #a; #H; nwhd in H:(??%?); ndestruct; ##| ##7,8: #a b; #H; nwhd in H:(??%?); ndestruct; ##| #sz1 si1; ncases ty'; ##[ #H; nwhd in H:(??%?); ndestruct; ##| ##3,9: #a; #H; nwhd in H:(??%?); ndestruct; // ##| ##2,7,8: #a b; #H; nwhd in H:(??%?); ndestruct; // ##| ##4,5,6: ##[ #sp ty''; nletin t ≝ (Tpointer sp ty'') ##| #sp ty'' n; nletin t ≝ (Tarray sp ty'' n) ##| #args rty; nletin t ≝ (Tfunction args rty) ##] nwhd in ⊢ (??%? → ?); nlapply (try_cast_null_sound m i (Tint sz1 si1) t v'); ncases (try_cast_null m i (Tint sz1 si1) t); ##[ ##1,3,5: #v''; #H' e; napply H'; napply e; ##| ##*: #_; nwhd in ⊢ (??%? → ?); #H; ndestruct (H); ##] ##] ##| ##*: ##[ #sp ty''; nletin t ≝ (Tpointer sp ty'') ##| #sp ty'' n; nletin t ≝ (Tarray sp ty'' n) ##| #args rty; nletin t ≝ (Tfunction args rty) ##] nwhd in ⊢ (??%? → ?); nlapply (try_cast_null_sound m i t ty' v'); ncases (try_cast_null m i t ty'); ##[ ##1,3,5: #v''; #H' e; napply H'; napply e; ##| ##*: #_; nwhd in ⊢ (??%? → ?); #H; ndestruct (H); ##] ##] ##| #f; ncases ty; ##[ ##3,9: #x; ##| ##2,4,6,7,8: #x y; ##| ##5: #x y z; ##] ##[ ncases ty'; ##[ #e; ##| ##3,9: #a e; ##| ##2,4,6,7,8: #a b e; ##| #a b c e; ##] nwhd in e:(??%?); ndestruct; //; ##| ##*: #e; nwhd in e:(??%?); ndestruct ##] ##| #sp b of; nwhd in ⊢ (??%? → ?); #e; nelim (bind_inversion ????? e); #s; *; #es e'; nelim (bind_inversion ????? e'); #u; *; #eu e''; nelim (bind_inversion ????? e''); #s'; *; #es' e'''; ncut (type_space ty s); ##[ ncases ty in es:(??%?) ⊢ %; ##[ #e; ##| ##3,9: #a e; ##| ##2,4,6,7,8: #a b e; ##| #a b c e; ##] nwhd in e:(??%?); ndestruct; //; ##| #Hty; ncut (type_space ty' s'); ##[ ncases ty' in es' ⊢ %; ##[ #e; ##| ##3,9: #a e; ##| ##2,4,6,7,8: #a b e; ##| #a b c e; ##] nwhd in e:(??%?); ndestruct; //; ##| #Hty'; ncut (s = sp). nelim (ms_eq_dec sp s) in eu; //; nnormalize; #_; #e; ndestruct. #e; nrewrite < e; nwhd in match (is_pointer_compat ??) in e'''; ncases (pointer_compat_dec (block_space m b) s') in e'''; #Hcompat e'''; nwhd in e''':(??%?); ndestruct (e'''); /2/ ##] ##] ##] nqed. ndefinition load_value_of_type' ≝ λty,m,l. match l with [ mk_pair pl ofs ⇒ match pl with [ mk_pair psp loc ⇒ load_value_of_type ty m psp loc ofs ] ]. (* 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'. *) nlet 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 i ⇒ OK ? 〈Vint i, E0〉 | Econst_float f ⇒ OK ? 〈Vfloat f, E0〉 | Evar _ ⇒ do 〈l,tr〉 ← exec_lvalue' ge en m e' ty; do v ← opt_to_res ? (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 ? (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 ? (load_value_of_type' ty m l); OK ? 〈v,tr〉 | Eaddrof a ⇒ do 〈plo,tr〉 ← exec_lvalue ge en m a; OK ? 〈match plo with [ mk_pair pl ofs ⇒ match pl with [ mk_pair pcl loc ⇒ Vptr pcl loc ofs ] ], tr〉 | Esizeof ty' ⇒ OK ? 〈Vint (repr (sizeof ty')), E0〉 | Eunop op a ⇒ do 〈v1,tr〉 ← exec_expr ge en m a; do v ← opt_to_res ? (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 ? (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 (memory_space × block × int × trace) ≝ match e' with [ Evar id ⇒ match (get … id en) with [ None ⇒ do 〈sp,l〉 ← opt_to_res ? (find_symbol ? ? ge id); OK ? 〈〈〈sp,l〉,zero〉,E0〉 (* global *) | Some loc ⇒ OK ? 〈〈〈Any,loc〉,zero〉,E0〉 (* local *) ] | Ederef a ⇒ do 〈v,tr〉 ← exec_expr ge en m a; match v with [ Vptr sp l ofs ⇒ OK ? 〈〈〈sp,l〉,ofs〉,tr〉 | _ ⇒ Error ? ] | Efield a i ⇒ match (typeof a) with [ Tstruct id fList ⇒ do 〈plofs,tr〉 ← exec_lvalue ge en m a; do delta ← field_offset i fList; OK ? 〈〈\fst plofs,add (\snd plofs) (repr delta)〉,tr〉 | Tunion id fList ⇒ do 〈plofs,tr〉 ← exec_lvalue ge en m a; OK ? 〈plofs,tr〉 | _ ⇒ Error ? ] | _ ⇒ Error ? ] and exec_lvalue (ge:genv) (en:env) (m:mem) (e:expr) on e : res (memory_space × block × int × trace) ≝ match e with [ Expr e' ty ⇒ exec_lvalue' ge en m e' ty ]. nlemma 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; nrewrite > H2 in H1; nwhd in ⊢ (% → ?); //; nqed. (* We define a special induction principle tailored to the recursive definition above. *) ndefinition is_not_lvalue: expr_descr → Prop ≝ λe. match e with [ Evar _ ⇒ False | Ederef _ ⇒ False | Efield _ _ ⇒ False | _ ⇒ True ]. ndefinition 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 ]. nlet 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 ? ]. nwhd; napply I; nqed. ntheorem 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; napply (expr_lvalue_ind ? (λe',ty.P_res ? (λr.eval_lvalue ge en m (Expr e' ty) (\fst (\fst (\fst r))) (\snd (\fst (\fst r))) (\snd (\fst r)) (\snd r)) (exec_lvalue' ge en m e' ty)) … e); ##[ ##1,2: #ty c; nwhd; //; (* expressions that are lvalues *) ##| #e' ty; ncases e'; //; ##[ #i He' ##| #e He' ##| #e i He' ##] nwhd in He' ⊢ %; napply bind2_OK; #x; ncases x; #y; ncases y; #sp loc ofs tr H; napply opt_bind_OK; #vl evl; nwhd in evl:(??%?); napply (eval_Elvalue … evl); nrewrite > H in He'; #He'; napply He'; ##| #v ty; nwhd in ⊢ (???%); nlapply (refl ? (get ident PTree block v en)); ncases (get ident PTree block v en) in ⊢ (???% → %); ##[ #eget; napply opt_bind_OK; #sploc; ncases sploc; #sp loc efind; nwhd; napply (eval_Evar_global … eget efind); ##| #loc eget; napply (eval_Evar_local … eget); ##] ##| #ty e He; nwhd in ⊢ (???%); napply bind2_OK; #v; ncases v; //; #sp l ofs tr Hv; nwhd; napply eval_Ederef; nrewrite > Hv in He; #He; napply He; ##| #ty e'' ty'' He''; napply bind2_OK; #x; ncases x; #y; ncases y; #sp loc ofs tr H; nwhd; napply eval_Eaddrof; nrewrite > H in He''; #He''; napply He''; ##| #ty op e1 He1; napply bind2_OK; #v1 tr Hv1; napply opt_bind_OK; #v ev; napply (eval_Eunop … ev); nrewrite > Hv1 in He1; #He1; napply He1; ##| #ty op e1 e2 He1 He2; napply bind2_OK; #v1 tr1 ev1; nrewrite > ev1 in He1; #He1; napply bind2_OK; #v2 tr2 ev2; nrewrite > ev2 in He2; #He2; napply opt_bind_OK; #v ev; nwhd in He1 He2; nwhd; napply (eval_Ebinop … He1 He2 ev); ##| #ty ty' e' He'; napply bind2_OK; #v tr Hv; nrewrite > Hv in He'; #He'; napply bind_OK; #v' ev'; napply (eval_Ecast … He' ?); /2/; ##| #ty e1 e2 e3 He1 He2 He3; napply bind2_OK; #vb tr1 Hvb; nrewrite > Hvb in He1; #He1; napply bind_OK; #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; napply bind2_OK; #v tr Hv; ##[ nrewrite > Hv in He2; #He2; nwhd in He2 Hv:(??%?) ⊢%; napply (eval_Econdition_true … He1 ? He2); napply (bool_of ??? Hb); ##| nrewrite > Hv in He3; #He3; nwhd in He3 Hv:(??%?) ⊢%; napply (eval_Econdition_false … He1 ? He3); napply (bool_of ??? Hb); ##] ##| #ty e1 e2 He1 He2; napply bind2_OK; #v1 tr1 Hv1; nrewrite > Hv1 in He1; #He1; napply bind_OK; #b1; ncases b1; #eb1; nlapply (exec_bool_of_val_sound … eb1); #Hb1; ##[ napply bind2_OK; #v2 tr2 Hv2; nrewrite > Hv2 in He2; #He2; napply bind_OK; #b2 eb2; napply (eval_Eandbool_2 … He1 … He2); ##[ napply (bool_of … Hb1); ##| napply (exec_bool_of_val_sound … eb2); ##] ##| napply (eval_Eandbool_1 … He1); napply (bool_of … Hb1); ##] ##| #ty e1 e2 He1 He2; napply bind2_OK; #v1 tr1 Hv1; nrewrite > Hv1 in He1; #He1; napply bind_OK; #b1; ncases b1; #eb1; nlapply (exec_bool_of_val_sound … eb1); #Hb1; ##[ napply (eval_Eorbool_1 … He1); napply (bool_of … Hb1); ##| napply bind2_OK; #v2 tr2 Hv2; nrewrite > Hv2 in He2; #He2; napply bind_OK; #b2 eb2; napply (eval_Eorbool_2 … He1 … He2); ##[ napply (bool_of … Hb1); ##| napply (exec_bool_of_val_sound … eb2); ##] ##] ##| #ty ty'; nwhd; // ##| #ty e' ty' i; ncases ty'; //; ##[ #id fs He'; napply bind2_OK; #x; ncases x; #sp l ofs H; napply bind_OK; #delta Hdelta; nrewrite > H in He'; #He'; napply (eval_Efield_struct … He' (refl ??) Hdelta); ##| #id fs He'; napply bind2_OK; #x; ncases x; #sp l ofs H; nrewrite > H in He'; #He'; napply (eval_Efield_union … He' (refl ??)); ##] ##| #ty l e' He'; napply bind2_OK; #v tr1 H; nrewrite > H in He'; #He'; napply (eval_Ecost … He'); (* exec_lvalue fails on non-lvalues. *) ##| #e' ty; ncases 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 ##] napply I; ##] nqed. nlemma addrof_eval_lvalue: ∀ge,en,m,e,sp,loc,off,tr,ty. eval_expr ge en m (Expr (Eaddrof e) ty) (Vptr sp loc off) tr → eval_lvalue ge en m e sp loc off tr. #ge en m e sp loc off tr ty H; ninversion H; ##[ ##1,2,5: #a b H; napply False_ind; ndestruct (H); ##| #a b c d e f g H1 i H2; nrewrite < H2 in H1; #H1; napply False_ind; napply (eval_lvalue_inv_ind … H1); ##[ #a b c d bad; ndestruct (bad); ##| #a b c d e f bad; ndestruct (bad); ##| #a b c d e f g bad; ndestruct (bad); ##| #a b c d e f g h i j k l m n bad; napply False_ind; ndestruct (bad); ##| #a b c d e f g h i j k l bad; ndestruct (bad); ##] ##| #e' ty' sp' loc' ofs' tr' H e1 e2 e3; ndestruct (e1 e2 e3); napply H; ##| #a b c d e f g h i bad; ndestruct (bad); ##| #a b c d e f g h i j k l k l bad; ndestruct (bad); ##| #a b c d e f g h i j k l m bad; ndestruct (bad); ##| #a b c d e f g h i j k l m bad; ndestruct (bad); ##| #a b c d e f g h bad; ndestruct (bad); ##| #a b c d e f g h i j k l m n bad; ndestruct (bad); ##| #a b c d e f g h bad; ndestruct (bad); ##| #a b c d e f g h i j k l m n bad; ndestruct (bad); ##| #a b c d e f g h i bad; ndestruct (bad); ##| #a b c d e f g bad; ndestruct (bad); ##] nqed. nlemma addrof_exec_lvalue: ∀ge,en,m,e,sp,loc,off,tr,ty. exec_lvalue ge en m e = OK ? 〈〈〈sp,loc〉,off〉,tr〉 → exec_expr ge en m (Expr (Eaddrof e) ty) = OK ? 〈Vptr sp loc off, tr〉. #ge en m e sp loc off tr ty H; nwhd in ⊢ (??%?); nrewrite > H; //; nqed. ntheorem exec_lvalue_sound: ∀ge,en,m,e. P_res ? (λr.eval_lvalue ge en m e (\fst (\fst (\fst r))) (\snd (\fst (\fst r))) (\snd (\fst r)) (\snd r)) (exec_lvalue ge en m e). #ge en m e; nlapply (refl ? (exec_lvalue ge en m e)); ncases (exec_lvalue ge en m e) in ⊢ (???% → %); ##[ #x; ncases x; #y; ncases y; #z; ncases z; #sp loc off tr H; nwhd; napply (addrof_eval_lvalue … Tvoid); nlapply (addrof_exec_lvalue … Tvoid H); #H'; nlapply (exec_expr_sound ge en m (Expr (Eaddrof e) Tvoid)); nrewrite > H'; #H''; napply H''; ##| #_; nwhd; napply I; ##] nqed. (* Plain equality versions of the above *) ndefinition 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. ndefinition exec_lvalue_sound' ≝ λge,en,m,e,sp,loc,off,tr. λH:exec_lvalue ge en m e = OK ? 〈〈〈sp,loc〉,off〉,tr〉. P_res_to_P ???? (exec_lvalue_sound ge en m e) H. (* TODO: Can we do this sensibly with a map combinator? *) nlet 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〉 ]. nlemma 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; nelim l; nwhd; //; #e1 es; #IH; napply bind2_OK; #v tr1 Hv; napply bind2_OK; #vs tr2 Hvs; nwhd; napply eval_Econs; ##[ napply (P_res_to_P … (exec_expr_sound ge e m e1) Hv); ##| napply (P_res_to_P … IH Hvs); ##] nqed. (* Don't really want to use subset rather than sigma here, but can't be bothered with *another* set of coercions. XXX: why do I have to get the recursive call's property manually? *) nlet rec exec_alloc_variables (en:env) (m:mem) (l:list (ident × type)) on l : { r:env × mem | alloc_variables en m l (\fst r) (\snd r) } ≝ match l with [ nil ⇒ Some ? 〈en, m〉 | cons h vars ⇒ match h with [ mk_pair id ty ⇒ match alloc m 0 (sizeof ty) Any with [ mk_pair m1 b1 ⇒ match exec_alloc_variables (set … id b1 en) m1 vars with [ sig_intro r p ⇒ r ] ]]]. nwhd; ##[ //; ##| nelim (exec_alloc_variables (set ident ? ? c3 c7 en) c6 c1); #H; nelim H; //; #H0; nelim H0; nnormalize; #en' m' IH; napply (alloc_variables_cons … IH); /2/; nqed. (* TODO: can we establish that length params = length vs in advance? *) nlet rec exec_bind_parameters (e:env) (m:mem) (params:list (ident × type)) (vs:list val) on params : res (Σm2:mem. bind_parameters e m params vs m2) ≝ match params with [ nil ⇒ match vs with [ nil ⇒ Some ? (OK ? m) | cons _ _ ⇒ Some ? (Error ?) ] | cons idty params' ⇒ match idty with [ mk_pair id ty ⇒ match vs with [ nil ⇒ Some ? (Error ?) | cons v1 vl ⇒ Some ? ( do b ← opt_to_res ? (get … id e); do m1 ← opt_to_res ? (store_value_of_type ty m Any b zero v1); err_eject ?? (exec_bind_parameters e m1 params' vl)) (* FIXME: don't want to have to eject here *) ] ] ]. nwhd; //; napply opt_bind_OK; #b eb; napply opt_bind_OK; #m1 em1; napply sig_bind_OK; #m2 Hm2; napply (bind_parameters_cons … eb em1 Hm2); nqed. ndefinition is_not_void : ∀t:type. res (Σu:unit. t ≠ Tvoid) ≝ λt. match t with [ Tvoid ⇒ Some ? (Error ?) | _ ⇒ Some ? (OK ??) ]. nwhd; //; @; #H; ndestruct; nqed. ninductive decide : Type ≝ | dy : decide | dn : decide. ndefinition dodecide : ∀P:Prop.∀d.∀p:(match d with [ dy ⇒ P | dn ⇒ ¬P ]).P + ¬P. #P d;ncases d;/2/; nqed. ncoercion decide_inject : ∀P:Prop.∀d.∀p:(match d with [ dy ⇒ P | dn ⇒ ¬P ]).P + ¬P ≝ dodecide on d:decide to ? + (¬?). ndefinition dodecide2 : ∀P:Prop.∀d.∀p:(match d with [ dy ⇒ P | dn ⇒ True ]).res P. #P d; ncases d; nnormalize; #p; ##[ napply (OK ? p); ##| napply Error ##] nqed. ncoercion decide_inject2 : ∀P:Prop.∀d.∀p:(match d with [ dy ⇒ P | dn ⇒ True ]).res P ≝ dodecide2 on d:decide to res ?. alias id "Tint" = "cic:/matita/c-semantics/Csyntax/type.con(0,2,0)". alias id "Tfloat" = "cic:/matita/c-semantics/Csyntax/type.con(0,3,0)". ndefinition sz_eq_dec : ∀s1,s2:intsize. (s1 = s2) + (s1 ≠ s2). #s1; ncases s1; #s2; ncases s2; /2/; @2; @; #H; ndestruct; nqed. ndefinition sg_eq_dec : ∀s1,s2:signedness. (s1 = s2) + (s1 ≠ s2). #s1; ncases s1; #s2; ncases s2; /2/; @2; @; #H; ndestruct; nqed. ndefinition fs_eq_dec : ∀s1,s2:floatsize. (s1 = s2) + (s1 ≠ s2). #s1; ncases s1; #s2; ncases s2; /2/; @2; @; #H; ndestruct; nqed. nlet rec assert_type_eq (t1,t2:type) : res (t1 = t2) ≝ match t1 with [ Tvoid ⇒ match t2 with [ Tvoid ⇒ dy | _ ⇒ dn ] | Tint sz sg ⇒ match t2 with [ Tint sz' sg' ⇒ match sz_eq_dec sz sz' with [ inl _ ⇒ match sg_eq_dec sg sg' with [ inl _ ⇒ dy | _ ⇒ dn ] | _ ⇒ dn ] | _ ⇒ dn ] | Tfloat f ⇒ match t2 with [ Tfloat f' ⇒ match fs_eq_dec f f' with [ inl _ ⇒ dy | _ ⇒ dn ] | _ ⇒ dn ] | Tpointer s t ⇒ match t2 with [ Tpointer s' t' ⇒ match ms_eq_dec s s' with [ inl _ ⇒ match assert_type_eq t t' with [ OK _ ⇒ dy | _ ⇒ dn ] | _ ⇒ dn ] | _ ⇒ dn ] | Tarray s t n ⇒ match t2 with [ Tarray s' t' n' ⇒ match ms_eq_dec s s' with [ inl _ ⇒ match assert_type_eq t t' with [ OK _ ⇒ match decidable_eq_Z_Type n n' with [ inl _ ⇒ dy | inr _ ⇒ dn ] | _ ⇒ dn ] | _ ⇒ dn ] | _ ⇒ dn ] | Tfunction tl t ⇒ match t2 with [ Tfunction tl' t' ⇒ match assert_typelist_eq tl tl' with [ OK _ ⇒ match assert_type_eq t t' with [ OK _ ⇒ dy | _ ⇒ dn ] | _ ⇒ dn ] | _ ⇒ dn ] | Tstruct i fl ⇒ match t2 with [ Tstruct i' fl' ⇒ match ident_eq i i' with [ inl _ ⇒ match assert_fieldlist_eq fl fl' with [ OK _ ⇒ dy | _ ⇒ dn ] | inr _ ⇒ dn ] | _ ⇒ dn ] | Tunion i fl ⇒ match t2 with [ Tunion i' fl' ⇒ match ident_eq i i' with [ inl _ ⇒ match assert_fieldlist_eq fl fl' with [ OK _ ⇒ dy | _ ⇒ dn ] | _ ⇒ dn ] | _ ⇒ dn ] | Tcomp_ptr i ⇒ match t2 with [ Tcomp_ptr i' ⇒ match ident_eq i i' with [ inl _ ⇒ dy | inr _ ⇒ dn ] | _ ⇒ dn ] ] and assert_typelist_eq (tl1,tl2:typelist) : res (tl1 = tl2) ≝ match tl1 with [ Tnil ⇒ match tl2 with [ Tnil ⇒ dy | _ ⇒ dn ] | Tcons t1 ts1 ⇒ match tl2 with [ Tnil ⇒ dn | Tcons t2 ts2 ⇒ match assert_type_eq t1 t2 with [ OK _ ⇒ match assert_typelist_eq ts1 ts2 with [ OK _ ⇒ dy | _ ⇒ dn ] | _ ⇒ dn ] ] ] and assert_fieldlist_eq (fl1,fl2:fieldlist) : res (fl1 = fl2) ≝ match fl1 with [ Fnil ⇒ match fl2 with [ Fnil ⇒ dy | _ ⇒ dn ] | Fcons i1 t1 fs1 ⇒ match fl2 with [ Fnil ⇒ dn | Fcons i2 t2 fs2 ⇒ match ident_eq i1 i2 with [ inl _ ⇒ match assert_type_eq t1 t2 with [ OK _ ⇒ match assert_fieldlist_eq fs1 fs2 with [ OK _ ⇒ dy | _ ⇒ dn ] | _ ⇒ dn ] | _ ⇒ dn ] ] ]. (* A poor man's clear, otherwise automation picks up recursive calls without checking that the argument is smaller. *) ngeneralize in assert_type_eq; ngeneralize in assert_typelist_eq; ngeneralize in assert_fieldlist_eq; #avoid1; #_; #avoid2; #_; #avoid3; #_; nwhd; //; (* XXX: I have no idea why the first // didn't catch these. *) //; //; //; //; //; //; //; //; //; nqed. nlet rec is_is_call_cont (k:cont) : (is_call_cont k) + (¬is_call_cont k) ≝ match k with [ Kstop ⇒ dy | Kcall _ _ _ _ ⇒ dy | _ ⇒ dn ]. nwhd; //; @; #H; nelim H; nqed. nlet rec is_Sskip (s:statement) : (s = Sskip) + (s ≠ Sskip) ≝ match s with [ Sskip ⇒ dy | _ ⇒ dn ]. ##[ //; ##| ##*: @; #H; ndestruct; ##] nqed. (* IO monad *) (* Interactions are function calls that return a value and do not change the rest of the Clight program's state. *) ndefinition io_out ≝ (ident × (list eventval)). ndefinition do_io : ident → list eventval → IO eventval io_out eventval ≝ λfn,args. Interact ?? eventval 〈fn,args〉 (λres. Value ?? eventval res). ndefinition ret: ∀T. T → (IO eventval io_out T) ≝ λT,x.(Value ?? T x). (* Checking types of values given to / received from an external function call. *) ndefinition check_eventval : ∀ev:eventval. ∀ty:typ. res (Σv:val. eventval_match ev ty v) ≝ λev,ty. match ty with [ Tint ⇒ match ev with [ EVint i ⇒ Some ? (OK ? (Vint i)) | _ ⇒ Some ? (Error ?) ] | Tfloat ⇒ match ev with [ EVfloat f ⇒ Some ? (OK ? (Vfloat f)) | _ ⇒ Some ? (Error ?) ] | _ ⇒ Some ? (Error ?) ]. nwhd; //; nqed. ndefinition check_eventval' : ∀v:val. ∀ty:typ. res (Σev:eventval. eventval_match ev ty v) ≝ λv,ty. match ty with [ Tint ⇒ match v with [ Vint i ⇒ Some ? (OK ? (EVint i)) | _ ⇒ Some ? (Error ?) ] | Tfloat ⇒ match v with [ Vfloat f ⇒ Some ? (OK ? (EVfloat f)) | _ ⇒ Some ? (Error ?) ] | _ ⇒ Some ? (Error ?) ]. nwhd; //; nqed. nlet rec check_eventval_list (vs: list val) (tys: list typ) : res (Σevs:list eventval. eventval_list_match evs tys vs) ≝ match vs with [ nil ⇒ match tys with [ nil ⇒ Some ? (OK ? (nil ?)) | _ ⇒ Some ? (Error ?) ] | cons v vt ⇒ match tys with [ nil ⇒ Some ? (Error ?) | cons ty tyt ⇒ Some ? ( do ev ← check_eventval' v ty; do evt ← check_eventval_list vt tyt; OK ? ((sig_eject ?? ev)::evt)) ] ]. nwhd; //; napply sig_bind_OK; #ev Hev; napply sig_bind_OK; #evt Hevt; nnormalize; /2/; nqed. (* execution *) ndefinition store_value_of_type' ≝ λty,m,l,v. match l with [ mk_pair pl ofs ⇒ match pl with [ mk_pair pcl loc ⇒ store_value_of_type ty m pcl loc ofs v ] ]. nlet rec exec_step (ge:genv) (st:state) on st : (IO eventval io_out (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' ← 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 ← find_funct ? ? ge vf; ! p ← err_to_io … (assert_type_eq (type_of_fundef fd) (typeof a)); (* ! k' ← match lhs with [ None ⇒ ret ? (Kcall (None ?) f e k) | Some lhs' ⇒ ! locofs ← exec_lvalue ge e m lhs'; ret ? (Kcall (Some ? 〈sig_eject ?? 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 ??? ] | Kcall _ _ _ _ ⇒ match fn_return f with [ Tvoid ⇒ ret ? 〈E0, Returnstate Vundef k (free_list m (blocks_of_env e))〉 | _ ⇒ Wrong ??? ] | 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 ??? ] | 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 ??? ] | 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 ??? ] | 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 ??? ] | Some a ⇒ ! u ← err_to_io_sig … (is_not_void (fn_return f)); ! 〈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 n ⇒ ret ? 〈tr, State f (seq_of_labeled_statement (select_switch n sl)) (Kswitch k) e m〉 | _ ⇒ Wrong ??? ] | 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' ⇒ match sk' with [ mk_pair s' k' ⇒ ret ? 〈E0, State f s' k' e m〉 ] | None ⇒ Wrong ??? ] | Scost lbl s' ⇒ ret ? 〈Echarge lbl, State f s' k e m〉 ] | Callstate f0 vargs k m ⇒ match f0 with [ Internal f ⇒ match exec_alloc_variables empty_env m ((fn_params f) @ (fn_vars f)) with [ mk_pair e m1 ⇒ ! m2 ← err_to_io_sig … (exec_bind_parameters e m1 (fn_params f) vargs); ret ? 〈E0, State f (fn_body f) k e m2〉 ] | External f argtys retty ⇒ ! evargs ← err_to_io_sig … (check_eventval_list vargs (typlist_of_typelist argtys)); ! evres ← do_io f evargs; ! vres ← err_to_io_sig … (check_eventval evres (proj_sig_res (signature_of_type argtys retty))); ret ? 〈(Eextcall f evargs evres), Returnstate vres k m〉 ] | Returnstate res k m ⇒ match k with [ Kcall r f e k' ⇒ match r with [ None ⇒ match res with [ Vundef ⇒ ret ? 〈E0, (State f Sskip k' e m)〉 | _ ⇒ Wrong ??? ] | Some r' ⇒ match r' with [ mk_pair l ty ⇒ ! m' ← store_value_of_type' ty m l res; ret ? 〈E0, (State f Sskip k' e m')〉 ] ] | _ ⇒ Wrong ??? ] ]. ntheorem exec_step_sound: ∀ge,st. P_io ??? (λr. step ge st (\fst r) (\snd r)) (exec_step ge st). #ge st; ncases st; ##[ #f s k e m; ncases s; ##[ ncases k; ##[ nwhd in ⊢ (?????%); nlapply (refl ? (fn_return f)); ncases (fn_return f) in ⊢ (???% → %); //; #H; nwhd; napply step_skip_call; //; ##| #s' k'; nwhd; //; ##| #ex s' k'; napply step_skip_or_continue_while; @; //; ##| #ex s' k'; napply res_bindIO2_OK; #v tr Hv; nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); ##[ #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; nwhd in ⊢ (?????%); ##[ napply (step_skip_or_continue_dowhile_true … (exec_expr_sound' … Hv)); ##[ @; // ##| napply (bool_of … Hb); ##] ##| napply (step_skip_or_continue_dowhile_false … (exec_expr_sound' … Hv)); ##[ @; // ##| napply (bool_of … Hb); ##] ##] ##| #_; //; ##] ##| #ex s1 s2 k'; napply step_skip_or_continue_for2; @; //; ##| #ex s1 s2 k'; napply step_skip_for3; ##| #k'; napply step_skip_break_switch; @; //; ##| #r f' e' k'; nwhd in ⊢ (?????%); nlapply (refl ? (fn_return f)); ncases (fn_return f) in ⊢ (???% → %); //; #H; nwhd; napply step_skip_call; //; ##] ##| #ex1 ex2; napply res_bindIO2_OK; #x; ncases x; #y; ncases y; #pcl loc ofs tr1 Hlval; napply res_bindIO2_OK; #v2 tr2 Hv2; napply opt_bindIO_OK; #m' em'; nwhd; napply (step_assign … (exec_lvalue_sound' … Hlval) (exec_expr_sound' … Hv2) em'); ##| #lex fex args; napply res_bindIO2_OK; #vf tr1 Hvf0; nlapply (exec_expr_sound' … Hvf0); #Hvf; napply res_bindIO2_OK; #vargs tr2 Hvargs0; nlapply (P_res_to_P ???? (exec_exprlist_sound …) Hvargs0); #Hvargs; napply opt_bindIO_OK; #fd efd; napply bindIO_OK; #ety; ncases lex; nwhd; ##[ napply (step_call_none … Hvf Hvargs efd ety); ##| #lhs'; napply res_bindIO2_OK; #x; ncases x; #y; ncases y; #pcl loc ofs tr3 Hlocofs; nwhd; napply (step_call_some … (exec_lvalue_sound' … Hlocofs) Hvf Hvargs efd ety); ##] ##| #s1 s2; nwhd; //; ##| #ex s1 s2; napply res_bindIO2_OK; #v tr Hv; nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); //; #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; ##[ napply (step_ifthenelse_true … (exec_expr_sound' … Hv)); napply (bool_of … Hb); ##| napply (step_ifthenelse_false … (exec_expr_sound' … Hv)); napply (bool_of … Hb) ##] ##| #ex s'; napply res_bindIO2_OK; #v tr Hv; nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); //; #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; ##[ napply (step_while_true … (exec_expr_sound' … Hv)); napply (bool_of … Hb); ##| napply (step_while_false … (exec_expr_sound' … Hv)); napply (bool_of … Hb); ##] ##| #ex s'; nwhd; //; ##| #s1 ex s2 s3; nwhd in ⊢ (?????%); nelim (is_Sskip s1); #Hs1; nwhd in ⊢ (?????%); ##[ nrewrite > Hs1; napply res_bindIO2_OK; #v tr Hv; nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); //; #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; ##[ napply (step_for_true … (exec_expr_sound' … Hv)); napply (bool_of … Hb); ##| napply (step_for_false … (exec_expr_sound' … Hv)); napply (bool_of … Hb); ##] ##| napply step_for_start; //; ##] ##| nwhd in ⊢ (?????%); ncases k; //; ##[ #s' k'; nwhd; //; ##| #ex s' k'; nwhd; //; ##| #ex s' k'; nwhd; //; ##| #ex s1 s2 k'; nwhd; //; ##| #k'; napply step_skip_break_switch; @2; //; ##] ##| nwhd in ⊢ (?????%); ncases k; //; ##[ #s' k'; nwhd; //; ##| #ex s' k'; nwhd; napply step_skip_or_continue_while; @2; //; ##| #ex s' k'; nwhd; napply res_bindIO2_OK; #v tr Hv; nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); //; #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; ##[ napply (step_skip_or_continue_dowhile_true … (exec_expr_sound' … Hv)); ##[ @2; // ##| napply (bool_of … Hb); ##] ##| napply (step_skip_or_continue_dowhile_false … (exec_expr_sound' … Hv)); ##[ @2; // ##| napply (bool_of … Hb); ##] ##] ##| #ex s1 s2 k'; nwhd; napply step_skip_or_continue_for2; @2; // ##| #k'; nwhd; //; ##] ##| #r; nwhd in ⊢ (?????%); ncases r; ##[ nwhd; nlapply (refl ? (fn_return f)); ncases (fn_return f) in ⊢ (???% → %); //; #H; napply step_return_0; napply H; ##| #ex; napply sig_bindIO_OK; #u Hnotvoid; napply res_bindIO2_OK; #v tr Hv; nwhd; napply (step_return_1 … Hnotvoid (exec_expr_sound' … Hv)); ##] ##| #ex ls; napply res_bindIO2_OK; #v; ncases v; //; #n tr Hv; napply step_switch; napply (exec_expr_sound' … Hv); ##| #l s'; nwhd; //; ##| #l; nwhd in ⊢ (?????%); nlapply (refl ? (find_label l (fn_body f) (call_cont k))); ncases (find_label l (fn_body f) (call_cont k)) in ⊢ (???% → %); //; #sk; ncases sk; #s' k' H; napply (step_goto … H); ##| #l s'; nwhd; //; ##] ##| #f0 vargs k m; nwhd in ⊢ (?????%); ncases f0; ##[ #fn; napply extract_subset_pair_io; #e m1 ealloc Halloc; napply sig_bindIO_OK; #m2 Hbind; nwhd; napply (step_internal_function … Halloc Hbind); ##| #id argtys rty; napply sig_bindIO_OK; #evs Hevs; napply bindIO_OK; #eres; napply sig_bindIO_OK; #res Hres; nwhd; napply step_external_function; @; ##[ napply Hevs; ##| napply Hres; ##] ##] ##| #v k' m'; nwhd in ⊢ (?????%); ncases k'; //; #r f e k; nwhd in ⊢ (?????%); ncases r; ##[ ncases v; nwhd; //; ##| #x; ncases x; #y; ncases y; #z; ncases z; #pcl b ofs ty; napply opt_bindIO_OK; #m' em'; napply step_returnstate_1; nwhd in em':(??%?); //; ##] ##] nqed. nlet rec make_initial_state (p:program) : IO eventval io_out state ≝ let ge ≝ globalenv Genv ?? p in let m0 ≝ init_mem Genv ?? p in ! 〈sp,b〉 ← find_symbol ? ? ge (prog_main ?? p); ! u ← opt_to_io … (match ms_eq_dec sp Code with [ inl _ ⇒ Some ? something | inr _ ⇒ None ? ]); ! f ← find_funct_ptr ? ? ge b; ret ? (Callstate f (nil ?) Kstop m0). nlemma make_initial_state_sound : ∀p. P_io ??? (λs.initial_state p s) (make_initial_state p). #p; ncases p; #fns main vars; nwhd in ⊢ (?????%); napply opt_bindIO2_OK; #sp b esb; napply opt_bindIO_OK; #u ecode; napply opt_bindIO_OK; #f ef; ncases sp in esb ecode; #esb ecode; nwhd in ecode:(??%%); ##[ ##1,2,3,4,5: ndestruct (ecode); ##] nwhd; napply (initial_state_intro … esb ef); nqed. ndefinition is_final_state : ∀st:state. (Σr. final_state st r) + (¬∃r. final_state st r). #st; nelim st; ##[ #f s k e m; @2; @;*; #r H; ninversion H; #i m e; ndestruct; ##| #f l k m; @2; @;*; #r H; ninversion H; #i m e; ndestruct; ##| #v k m; ncases k; ##[ ncases v; ##[ ##2: #i; @1; @ i; //; ##| ##1: @2; @; *; #r H; ninversion H; #i m e; ndestruct; ##| #f; @2; @; *; #r H; ninversion H; #i m e; ndestruct; ##| #pcl b of; @2; @; *; #r H; ninversion H; #i m e; ndestruct; ##] ##| #a b; @2; @; *; #r H; ninversion H; #i m e; ndestruct; ##| ##3,4: #a b c; @2; @; *; #r H; ninversion H; #i m e; ndestruct; ##| ##5,6,8: #a b c d; @2; @; *; #r H; ninversion H; #i m e; ndestruct; ##| #a; @2; @; *; #r H; ninversion H; #i m e; ndestruct; ##] ##] nqed. nlet rec exec_steps (n:nat) (ge:genv) (s:state) : IO eventval io_out (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'〉 ← match s with [ State f s k e m ⇒ match m with [ mk_mem c n p ⇒ exec_step ge (State f s k e (mk_mem c n p)) ] | Callstate fd args k m ⇒ match m with [ mk_mem c n p ⇒ exec_step ge (Callstate fd args k (mk_mem c n p)) ] | Returnstate r k m ⇒ match m with [ mk_mem c n p ⇒ exec_step ge (Returnstate r k (mk_mem c n p)) ] ] ;*) ! 〈t',s''〉 ← match s' with [ State f s k e m ⇒ match m with [ mk_mem c n p ⇒ exec_steps n' ge (State f s k e (mk_mem c n p)) ] | Callstate fd args k m ⇒ match m with [ mk_mem c n p ⇒ exec_steps n' ge (Callstate fd args k (mk_mem c n p)) ] | Returnstate r k m ⇒ match m with [ mk_mem c n p ⇒ exec_steps n' ge (Returnstate r k (mk_mem c n p)) ] ] ; (* ! 〈t',s''〉 ← exec_steps n' ge s';*) ret ? 〈t ⧺ t',s''〉 ] ]. ntheorem 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; nelim n; ##[ #st; nwhd in ⊢ (?????%); nelim (is_final_state st); #H; nwhd; //; ##| #n' IH st; nwhd in ⊢ (?????%); nelim (is_final_state st); #H; ##[ nwhd; //; ##| napply (P_bindIO2_OK ????????? (exec_step_sound …)); #t s'; ncases s'; ##[ #f s k e m; ##| #fd args k m; ##| #r k m; ##] nwhd in ⊢ (? → ?????(??????%?)); ncases m; #mc mn mp; #Hstep; nwhd in ⊢ (?????(??????%?)); napply (P_bindIO2_OK ????????? (IH …)); #t' s'' Hsteps; nwhd; napply (star_step ????????? Hsteps); ##[ ##2,5,8: napply Hstep; ##| ##3,6,9: // ##] ##] nqed. (* nlet rec exec_steps_without_proof (n:nat) (ge:genv) (s:state) : res (trace×state) ≝ match is_final_state s with [ inl _ ⇒ OK ? 〈E0, s〉 | inr _ ⇒ match n with [ O ⇒ OK ? 〈E0, s〉 | S n' ⇒ 〈t,s'〉 ← exec_step ge s; 〈t',s''〉 ← exec_steps_without_proof n' ge s'; OK ? 〈t ⧺ t',s''〉 ] ]. *) (* A (possibly non-terminating) execution. *) ncoinductive execution : Type ≝ | e_stop : trace → int → mem → execution | e_step : trace → state → execution → execution | e_wrong : execution | e_interact : io_out → (eventval → execution) → execution. ndefinition 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 ]. (* This definition is slightly awkward because of the need to provide resumptions. It records the last trace/state passed in, then recursively processes the next state. *) nlet corec exec_inf_aux (ge:genv) (s:IO eventval io_out (trace×state)) : execution ≝ match s with [ Wrong ⇒ e_wrong | Value v ⇒ match v with [ mk_pair t s' ⇒ match is_final_state s' with [ inl r ⇒ e_stop t (sig_eject … r) (mem_of_state s') | inr _ ⇒ e_step t s' (exec_inf_aux ge (exec_step ge s')) ] ] | Interact out k' ⇒ e_interact out (λv. exec_inf_aux ge (k' v)) ]. ndefinition exec_inf : program → execution ≝ λp. exec_inf_aux (globalenv Genv ?? p) (! s ← make_initial_state p; ret ? 〈E0,s〉). nremark execution_cases: ∀e. e = match e with [ e_stop tr r m ⇒ e_stop tr r m | e_step tr s e ⇒ e_step tr s e | e_wrong ⇒ e_wrong | e_interact o k ⇒ e_interact o k ]. #e; ncases e; //; nqed. nlemma exec_inf_aux_unfold: ∀ge,s. exec_inf_aux ge s = match s with [ Wrong ⇒ e_wrong | Value v ⇒ match v with [ mk_pair t s' ⇒ match is_final_state s' with [ inl r ⇒ e_stop t (sig_eject … r) (mem_of_state s') | inr _ ⇒ e_step t s' (exec_inf_aux ge (exec_step ge s')) ] ] | Interact out k' ⇒ e_interact out (λv. exec_inf_aux ge (k' v)) ]. #ge s; nrewrite > (execution_cases (exec_inf_aux …)); ncases s; ##[ #o k ##| #x; ncases x; #tr s'; ncases s'; ##[ #fn st k env m ##| #fd args k mem ##| #v k mem; (* for final state check *) ncases k; ##[ ncases v; ##[ ##2,3: #v' ##| ##4: #sp loc off ##] ##| #s' k' ##| ##3,4: #e s' k' ##| ##5,6: #e s' s'' k' ##| #k' ##| #a b c d ##] ##] ##| ##] nwhd in ⊢ (??%%); //; nqed. (* Finite number of steps without interaction. *) ninductive execution_steps : trace → execution → execution → Prop ≝ | steps_none : ∀e. execution_steps E0 e e | steps_one : ∀e,e',tr,tr',s. execution_steps tr' e e' → execution_steps (tr⧺tr') (e_step tr s e) e'. nlemma one_step: ∀ge,tr,s,tr',s',e. exec_inf_aux ge (Value ??? 〈tr,s〉) = e_step tr s (e_step tr' s' e) → step ge s tr' s'. #ge tr s tr' s' e; nrewrite > (exec_inf_aux_unfold ge ?); nwhd in ⊢ (??%? → ?); ncases (is_final_state s); ##[ #H E; nwhd in E:(??%?); ndestruct (E); ##| #H E; nwhd in E:(??%?); ndestruct; nrewrite > (exec_inf_aux_unfold ge ?) in e0; nlapply (refl ? (exec_step ge s)); ncases (exec_step ge s) in ⊢ (???% → %); ##[ #i o E1 E2; nwhd in E2:(??%?); ndestruct (E2); ##| #x; ncases x; #tr'' s'' E1 E2; nwhd in E2:(??%?); ncases (is_final_state s'') in E2; #H' E2; nwhd in E2:(??%?); ndestruct (E2); nlapply (exec_step_sound ge s); nrewrite > E1; nwhd in ⊢ (% → ?); #H1; napply H1 ##| #E1 E2; nwhd in E2:(??%?); ndestruct; ##] ##] nqed. (* starting after state s, zero or more steps of execution e reach state s' after which comes e'. *) ninductive execution_isteps : trace → state → execution → state → execution → Prop ≝ | isteps_none : ∀s,e. execution_isteps E0 s e s e | isteps_one : ∀e,e',tr,tr',s,s',s0. execution_isteps tr' s e s' e' → execution_isteps (tr⧺tr') s0 (e_step tr s e) s' e' | isteps_interact : ∀e,e',k,o,i,s,s',s0,tr,tr'. execution_isteps tr' s e s' e' → k i = e_step tr s e → (¬ ∃r.final_state s r) → (* used to avoid showing that k i doesn't end prog *) execution_isteps (tr⧺tr') s0 (e_interact o k) s' e'. nlemma exec_e_step: ∀ge,x,tr,s,e. exec_inf_aux ge x = e_step tr s e → exec_inf_aux ge (exec_step ge s) = e. #ge x tr s e; nrewrite > (exec_inf_aux_unfold …); ncases x; ##[ #o k EXEC; nwhd in EXEC:(??%?); ndestruct ##| #y; ncases y; #tr' s'; nwhd in ⊢ (??%? → ?); ncases (is_final_state s'); #FINAL EXEC; nwhd in EXEC:(??%?); ndestruct; napply refl; ##| #EXEC; nwhd in EXEC:(??%?); ndestruct ##] nqed. nlemma exec_e_step_inv: ∀ge,x,tr,s,e. exec_inf_aux ge x = e_step tr s e → x = Value ??? 〈tr,s〉. #ge x tr s e; nrewrite > (exec_inf_aux_unfold …); ncases x; ##[ #o k EXEC; nwhd in EXEC:(??%?); ndestruct ##| #y; ncases y; #tr' s'; nwhd in ⊢ (??%? → ?); ncases (is_final_state s'); #FINAL EXEC; nwhd in EXEC:(??%?); ndestruct; napply refl; ##| #EXEC; nwhd in EXEC:(??%?); ndestruct ##] nqed. (* NB: the E0 in the execs are irrelevant. *) nlemma several_steps: ∀ge,tr,e,e',s,s'. execution_isteps tr s e s' e' → exec_inf_aux ge (Value ??? 〈E0,s〉) = e_step E0 s e → star (mk_transrel … step) ge s tr s' ∧ exec_inf_aux ge (Value ??? 〈E0,s'〉) = e_step E0 s' e'. #ge tr0 e0 e0' s0 s0' H; nelim H; ##[ #s e EXEC; @; //; napply EXEC; ##| #e1 e2 tr1 tr2 s1 s2 s3 STEPS IH EXEC; nrewrite > (exec_inf_aux_unfold ge ?) in EXEC; nwhd in ⊢ (??%? → ?); ncases (is_final_state s3); ##[ nwhd in ⊢ (? → ??%? → ?); #FINAL BAD; ndestruct (BAD) ##] #NOTFINAL EXEC; nwhd in EXEC:(??%?); ndestruct; nrewrite > (exec_inf_aux_unfold ge ?) in e3; nlapply (refl ? (exec_step ge s3)); ncases (exec_step ge s3) in ⊢ (???% → %); ##[ #o k E EXEC'; nwhd in EXEC':(??%?); ndestruct (EXEC'); ##| #x; ncases x; #tr' s' EXEC3; nwhd in ⊢ (??%? → ?); ncases (is_final_state s'); #FINAL' EXEC'; nwhd in EXEC':(??%?); ##[ ndestruct ##] ncut (exec_inf_aux ge (exec_step ge s1) = e1); ##[ ndestruct (EXEC'); napply refl ##] #EXEC''; nlapply (IH ?); ##[ nrewrite > (exec_inf_aux_unfold ge ?); nwhd in ⊢ (??%?); ncases (is_final_state s1); #FINAL1; ##[ napply False_ind; napply (absurd ?? FINAL'); ncases FINAL1; #r' Hr; @ r'; ndestruct (EXEC'); napply Hr; ##| nwhd in ⊢ (??%?); nrewrite < EXEC''; napply refl ##] ##| *; #STARs1s2 EXEC2; @; ##[ nlapply (exec_step_sound ge s3); nrewrite > EXEC3; nwhd in ⊢ (% → ?); #STEP3; ndestruct (EXEC'); napply (star_step (mk_transrel ?? step) … STEP3 STARs1s2); //; ##| napply EXEC2; ##] ##] ##| #E EXEC'; nwhd in EXEC':(??%?); ndestruct (EXEC'); ##] ##| #e e' k o i s s' s0 tr tr' H1 H2 H3 IH; nrewrite > (exec_inf_aux_unfold ge ?); nwhd in ⊢ (??%? → ?); ncases (is_final_state s0); #FINAL EXEC; nwhd in EXEC:(??%?); ndestruct; nrewrite > (exec_inf_aux_unfold ge ?) in e1; nlapply (exec_step_sound ge s0); ncases (exec_step ge s0); ##[ #i' k' SOUND E1; nwhd in SOUND E1:(??%?); ndestruct (E1); ##| #x; ncases x; #tr' s'' SOUND; nwhd in ⊢ (??%? → ?); ncases (is_final_state s''); #FINAL'; nwhd in ⊢ (??%? → ?); #EXEC'; ndestruct (EXEC'); ##| #_; nwhd in ⊢ (??%? → ?); #EXEC'; ndestruct (EXEC'); ##] nlapply (IH ?); ##[ nrewrite > (exec_inf_aux_unfold ge ?); nwhd in ⊢ (??%?); ncases (is_final_state s); ##[ #X; ncases X; #r FINAL'; napply False_ind; napply (absurd ?? H3); @ r; napply FINAL' ##] #FINAL'; nwhd in ⊢ (??%?); nrewrite > (exec_e_step … H2); napply refl; ##| *; #STARsTOs' EXECs'; @; ##[ napply (star_step … STARsTOs'); ##[ ##2: nlapply (SOUND i); nrewrite > (exec_e_step_inv … H2); nwhd in ⊢ (% → ?); #STEP; napply STEP; ##| ##skip; ##| napply refl ##] ##| napply EXECs'; ##] ##] ##] nqed. ninductive execution_terminates : trace → state → execution → state → Prop ≝ | terminates : ∀s,s',tr,tr',r,e,m. execution_isteps tr s e s' (e_stop tr' r m) → execution_terminates (tr⧺tr') s e (Returnstate (Vint r) Kstop m). ncoinductive execution_diverging : execution → Prop ≝ | diverging_step : ∀s,e. execution_diverging e → execution_diverging (e_step E0 s e). (* Makes a finite number of interactions (including cost labels) before diverging. *) ninductive execution_diverges : trace → execution → Prop ≝ | diverges_diverging: ∀tr,s,s',e,e'. execution_isteps tr s e s' e' → execution_diverging e' → execution_diverges tr e. (* NB: "reacting" includes hitting a cost label. *) ncoinductive execution_reacts : traceinf → state → execution → Prop ≝ | reacting: ∀tr,tr',s,s',e,e'. execution_reacts tr' s' e' → execution_isteps tr s e s' e' → tr ≠ E0 → execution_reacts (tr⧻tr') s e. ninductive execution_goes_wrong: trace → state → execution → state → Prop ≝ | go_wrong: ∀tr,s,s',e. execution_isteps tr s e s' e_wrong → execution_goes_wrong tr s e s'. ninductive execution_matches_behavior: execution → program_behavior → Prop ≝ | emb_terminates: ∀s,s',e,tr,r. execution_terminates tr s e s' → execution_matches_behavior e (Terminates tr r) | emb_diverges: ∀e,tr. execution_diverges tr e → execution_matches_behavior e (Diverges tr) | emb_reacts: ∀s,e,tr. execution_reacts tr s e → execution_matches_behavior e (Reacts tr) | emb_wrong: ∀e,s,s',tr. execution_goes_wrong tr s e s' → execution_matches_behavior e (Goes_wrong tr). (* We don't morally need the cut, but the proof I tried without it failed the guarded-by-constructors check and it wasn't apparent why. *) nlet corec silent_sound ge s e (H0:execution_diverging e) (EXEC:exec_inf_aux ge (Value ??? 〈E0,s〉) = e) : forever_silent (mk_transrel ?? step) … ge s ≝ ?. ncut (∃s2.∃e2.execution_diverging e2 ∧ exec_inf_aux ge (exec_step ge s) = e_step E0 s2 e2); ##[ ncases H0 in EXEC ⊢ %; #s1 e1 H1; nrewrite > (exec_inf_aux_unfold …); nwhd in ⊢ (??%? → ?); ncases (is_final_state s); #p EXEC; nwhd in EXEC:(??%?); ndestruct; ninversion H1; #s2 e2 H2 EXEC2; @ s2; @ e2; @; //; ##| *; #s2; *; #e2; *; #H2 EXEC2; napply (forever_silent_intro (mk_transrel ?? step) … ge s s2 ? (silent_sound ge s2 (e_step E0 s2 e2) ??)); ncut (exec_step ge s = Value ??? 〈E0,s2〉); ##[ ##1,3,5: nrewrite > (exec_inf_aux_unfold …) in EXEC2; nelim (exec_step ge s); ##[ ##1,4,7: #o k IH EXEC2; nwhd in EXEC2:(??%?); ndestruct; ##| ##2,5,8: #x; ncases x; #tr s'; nwhd in ⊢ (??%? → ?); ncases (is_final_state s'); #p' EXEC2; nwhd in EXEC2:(??%?); ndestruct; napply refl; ##| ##3,6,9: #EXEC2; nwhd in EXEC2:(??%?); ndestruct ##] ##| #EXEC1; nlapply (exec_step_sound ge s); nrewrite > EXEC1; nwhd in ⊢ (% → %); // ##| #EXEC1; nrewrite > EXEC1 in EXEC2; #EXEC2; napply EXEC2; ##| #EXEC1; @; napply H2; ##] ##] nqed. nlemma final_step: ∀ge,tr,r,m,s. exec_inf_aux ge (exec_step ge s) = e_stop tr r m → step ge s tr (Returnstate (Vint r) Kstop m). #ge tr r m s; nrewrite > (exec_inf_aux_unfold …); nlapply (exec_step_sound ge s); ncases (exec_step ge s); ##[ #o k H EXEC; nwhd in EXEC:(??%?); ndestruct ##| #x; ncases x; #tr' s' SOUND; nwhd in ⊢ (??%? → ?); ncases (is_final_state s'); #FINAL EXEC; nwhd in SOUND EXEC:(??%?); ndestruct; ncases FINAL; #r' FINAL'; ninversion FINAL'; #r'' m' E1 E2; ndestruct; napply SOUND; ##| #_; #EXEC; nwhd in EXEC:(??%?); ndestruct ##] nqed. nlemma terminates_sound: ∀ge,tr,s,s',e. execution_terminates tr s e s' → exec_inf_aux ge (Value ??? 〈E0, s〉) = (e_step E0 s e) → ∃r. star (mk_transrel … step) ge s tr s' ∧ final_state s' r. #ge tr0 s0 s0' e0 H; ncases H; #s s' tr tr' r e m ESTEPS EXEC; @r; @; //; ncases (several_steps … ESTEPS EXEC); #STARs' EXECs'; napply (star_right … STARs'); ##[ ##2: napply (final_step ge tr' r m s'); napply (exec_e_step … EXECs'); ##| ##skip ##| napply refl ##] nqed. nlet corec reacts_sound ge tr s e (REACTS:execution_reacts tr s e) (EXEC:exec_inf_aux ge (Value ??? 〈E0, s〉) = e_step E0 s e) : forever_reactive (mk_transrel … step) ge s tr ≝ ?. ncut (∃s'.∃e'.∃tr'.∃tr''.execution_reacts tr'' s' e' ∧ execution_isteps tr' s e s' e' ∧ tr' ≠ E0 ∧ tr = tr'⧻tr''); ##[ ninversion REACTS; #tr0 tr' s0 s' e0 e' EREACTS ESTEPS REACTED E1 E2 E3; ndestruct (E2 E3); @ s'; @ e'; @ tr0; @ tr'; @; ##[ @; ##[ @; //; ##| napply REACTED ##] ##| napply refl ##] ##| *; #s'; *; #e'; *; #tr'; *; #tr''; *; *; *; #REACTS' ESTEPS REACTED APPTR; (* nrewrite > APPTR;*) napply (match sym_eq ??? APPTR return λx.λ_.forever_reactive (mk_transrel genv state step) ge s x with [ refl ⇒ ? ]); @; ncases (several_steps … ESTEPS EXEC); #STEPS EXEC'; ##[ ##2: napply STEPS; ##| ##skip ##| napply REACTED ##| napply reacts_sound; ##[ ##2: napply REACTS'; ##| ##skip ##| napply EXEC' ##] ##] nqed. (* is/needs completeness... nlemma wrong_sound: ∀ge,tr,s,s',e. execution_goes_wrong tr s e s' → exec_inf_aux ge (Value ??? 〈E0, s〉) = e_step E0 s e → star (mk_transrel … step) ge s tr s' ∧ nostep (mk_transrel … step) ge s' ∧ (¬∃r. final_state s' r). #ge tr0 s0 s0' e0 WRONG; ncases WRONG; #tr s s' e ESTEPS EXEC; ncases (several_steps … ESTEPS EXEC); #STAR EXEC'; *) (* ntheorem exec_inf_sound: ∀p. ∃b.execution_matches_behavior (exec_inf p) b ∧ exec_program p b. *)