Changeset 487 for Deliverables/D3.1/C-semantics
- Timestamp:
- Feb 9, 2011, 11:49:17 AM (9 years ago)
- Location:
- Deliverables/D3.1/C-semantics
- Files:
-
- 23 edited
Legend:
- Unmodified
- Added
- Removed
-
Deliverables/D3.1/C-semantics/AST.ma
r485 r487 17 17 the abstract syntax trees of many of the intermediate languages. *) 18 18 19 include " datatypes/sums.ma".19 include "basics/types.ma". 20 20 include "extralib.ma". 21 21 include "Integers.ma". … … 28 28 etc) are represented by the type [positive] of positive integers. *) 29 29 30 ndefinition ident ≝ Pos.31 32 ndefinition ident_eq : ∀x,y:ident. (x=y) + (x≠y).33 #x y; nlapply (pos_compare_to_Prop x y); ncases (pos_compare x y);34 ##[ #H; @2; /2/; ##| #H; @1; //; ##| #H; @2; /2/ ##] nqed.30 definition ident ≝ Pos. 31 32 definition ident_eq : ∀x,y:ident. (x=y) + (x≠y). 33 #x #y lapply (pos_compare_to_Prop x y); cases (pos_compare x y); 34 [ #H %2 /2/; | #H %1 //; | #H %2 /2/ ] qed. 35 35 36 36 (* 37 37 (* XXX: we use nats for now, but if in future we use binary like compcert 38 38 then the maps will be easier to define. *) 39 ndefinition ident ≝ nat.40 ndefinition ident_eq : ∀x,y:ident. (x=y) + (x≠y).41 #x ; nelim x;42 ##[ #y; ncases y; /3/;43 ##| #x'; #IH; #y; ncases y;44 ##[ @2; @; #H; ndestruct45 ##| #y'; nelim (IH y');46 ##[ #e; ndestruct; /2/47 ##| #ne; @2;/2/;48 ##]49 ##]50 ##] nqed.39 definition ident ≝ nat. 40 definition ident_eq : ∀x,y:ident. (x=y) + (x≠y). 41 #x elim x; 42 [ #y cases y; /3/; 43 | #x' #IH #y cases y; 44 [ %{2} % #H destruct 45 | #y' elim (IH y'); 46 [ #e destruct; /2/ 47 | #ne %{2} /2/; 48 ] 49 ] 50 ] qed. 51 51 *) 52 52 (* * The intermediate languages are weakly typed, using only two types: … … 54 54 numbers. *) 55 55 56 ninductive typ : Type≝56 inductive typ : Type[0] ≝ 57 57 | ASTint : typ 58 58 | ASTfloat : typ. 59 59 60 ndefinition typesize : typ → Z ≝ λty.60 definition typesize : typ → Z ≝ λty. 61 61 match ty return λ_.Z with [ ASTint ⇒ 4 | ASTfloat ⇒ 8 ]. 62 62 63 nlemma typesize_pos: ∀ty. typesize ty > 0.64 #ty ; ncases ty; //; nqed.65 66 nlemma typ_eq: ∀t1,t2: typ. (t1=t2) + (t1≠t2).67 #t1 ;#t2;ncases t1;ncases t2;/2/; @2; napply nmk; #H; ndestruct; nqed.68 69 nlemma opt_typ_eq: ∀t1,t2: option typ. (t1=t2) + (t1≠t2).70 #t1 ;#t2;ncases t1;ncases t2;/2/;71 ##[ ##1,2: #ty; @2; napply nmk; #H; ndestruct;72 ##| #ty1;#ty2; nelim (typ_eq ty1 ty2); /2/; #neq; @2; napply nmk; 73 #H ; ndestruct; /2/;74 ##] nqed.63 lemma typesize_pos: ∀ty. typesize ty > 0. 64 #ty cases ty; //; qed. 65 66 lemma typ_eq: ∀t1,t2: typ. (t1=t2) + (t1≠t2). 67 #t1 #t2 cases t1;cases t2;/2/; %2 @nmk #H destruct; qed. 68 69 lemma opt_typ_eq: ∀t1,t2: option typ. (t1=t2) + (t1≠t2). 70 #t1 #t2 cases t1;cases t2;/2/; 71 [ 1,2: #ty %2 @nmk #H destruct; 72 | #ty1 #ty2 elim (typ_eq ty1 ty2); /2/; #neq %2 @nmk 73 #H destruct; /2/; 74 ] qed. 75 75 76 76 (* * Additionally, function definitions and function calls are annotated … … 80 80 for the function. *) 81 81 82 nrecord signature : Type≝ {82 record signature : Type[0] ≝ { 83 83 sig_args: list typ; 84 84 sig_res: option typ 85 85 }. 86 86 87 ndefinition proj_sig_res : signature → typ ≝ λs.87 definition proj_sig_res : signature → typ ≝ λs. 88 88 match sig_res s with 89 89 [ None ⇒ ASTint … … 93 93 (* Memory spaces *) 94 94 95 ninductive region : Type≝95 inductive region : Type[0] ≝ 96 96 | Any : region 97 97 | Data : region … … 101 101 | Code : region. 102 102 103 ndefinition eq_region : region → region → bool ≝103 definition eq_region : region → region → bool ≝ 104 104 λr1,r2. 105 105 match r1 with … … 112 112 ]. 113 113 114 nlemma eq_region_elim : ∀P:bool → Type. ∀r1,r2.114 lemma eq_region_elim : ∀P:bool → Type[0]. ∀r1,r2. 115 115 (r1 = r2 → P true) → (r1 ≠ r2 → P false) → 116 116 P (eq_region r1 r2). 117 #P r1 r2; ncases r1; ncases r2; #Ptrue Pfalse;118 ntry ( napply Ptrue // ); 119 napply Pfalse; @; #E; ndestruct;120 nqed.121 122 ndefinition eq_region_dec : ∀r1,r2:region. (r1=r2)+(r1≠r2).123 #r1 r2; napply (eq_region_elim ? r1 r2); /2/; nqed.117 #P #r1 #r2 cases r1; cases r2; #Ptrue #Pfalse 118 try ( @Ptrue // ) 119 @Pfalse % #E destruct; 120 qed. 121 122 definition eq_region_dec : ∀r1,r2:region. (r1=r2)+(r1≠r2). 123 #r1 #r2 @(eq_region_elim ? r1 r2) /2/; qed. 124 124 125 125 (* * Memory accesses (load and store instructions) are annotated by … … 127 127 chunk of memory being accessed. *) 128 128 129 ninductive memory_chunk : Type[0] ≝129 inductive memory_chunk : Type[0] ≝ 130 130 | Mint8signed : memory_chunk (*r 8-bit signed integer *) 131 131 | Mint8unsigned : memory_chunk (*r 8-bit unsigned integer *) … … 139 139 (* * Initialization data for global variables. *) 140 140 141 ninductive init_data: Type[0] ≝141 inductive init_data: Type[0] ≝ 142 142 | Init_int8: int → init_data 143 143 | Init_int16: int → init_data … … 160 160 programs are common to all languages. *) 161 161 162 nrecord program (F,V: Type) : Type:= {162 record program (F,V: Type[0]) : Type[0] := { 163 163 prog_funct: list (ident × F); 164 164 prog_main: ident; … … 166 166 }. 167 167 168 ndefinition prog_funct_names ≝ λF,V: Type. λp: program F V.168 definition prog_funct_names ≝ λF,V: Type[0]. λp: program F V. 169 169 map ?? (fst ident F) (prog_funct ?? p). 170 170 171 ndefinition prog_var_names ≝ λF,V: Type. λp: program F V.171 definition prog_var_names ≝ λF,V: Type[0]. λp: program F V. 172 172 map ?? (λx: ident × (list init_data) × region × V. fst ?? (fst ?? (fst ?? x))) (prog_vars ?? p). 173 173 (* … … 184 184 *) 185 185 186 ndefinition transf_program : ∀A,B. (A → B) → list (ident × A) → list (ident × B) ≝186 definition transf_program : ∀A,B. (A → B) → list (ident × A) → list (ident × B) ≝ 187 187 λA,B,transf,l. 188 188 map ?? (λid_fn. 〈fst ?? id_fn, transf (snd ?? id_fn)〉) l. 189 189 190 ndefinition transform_program : ∀A,B,V. (A → B) → program A V → program B V ≝190 definition transform_program : ∀A,B,V. (A → B) → program A V → program B V ≝ 191 191 λA,B,V,transf,p. 192 192 mk_program B V … … 194 194 (prog_main A V p) 195 195 (prog_vars A V p). 196 197 nlemma transform_program_function:196 (* 197 lemma transform_program_function: 198 198 ∀A,B,V,transf,p,i,tf. 199 199 in_list ? 〈i, tf〉 (prog_funct ?? (transform_program A B V transf p)) → 200 200 ∃f. in_list ? 〈i, f〉 (prog_funct ?? p) ∧ transf f = tf. 201 n normalize; #A B V transf p i tf H; nelim (list_in_map_inv ????? H);202 #x ; nelim x; #i' tf'; *; #e H; ndestruct; @tf';/2/;203 nqed.204 201 normalize; #A #B #V #transf #p #i #tf #H elim (list_in_map_inv ????? H); 202 #x elim x; #i' #tf' *; #e #H destruct; %{tf'} /2/; 203 qed. 204 *) 205 205 (* 206 206 End TRANSF_PROGRAM. … … 432 432 a type for such functions and some generic transformation functions. *) 433 433 434 nrecord external_function : Type≝ {434 record external_function : Type[0] ≝ { 435 435 ef_id: ident; 436 436 ef_sig: signature 437 437 }. 438 438 (* 439 ninductive fundef (F: Type): Type ≝439 inductive fundef (F: Type): Type ≝ 440 440 | Internal: F → fundef F 441 441 | External: external_function → fundef F. … … 448 448 Variable transf: A -> B. 449 449 *) 450 ndefinition transf_fundef : ∀A,B. (A→B) → fundef A → fundef B ≝450 definition transf_fundef : ∀A,B. (A→B) → fundef A → fundef B ≝ 451 451 λA,B,transf,fd. 452 452 match fd with -
Deliverables/D3.1/C-semantics/Animation.ma
r478 r487 6 6 we stop at a continuation - it can be too large to work with. *) 7 7 8 ndefinition get_input : ∀o:io_out. eventval → res (io_in o) ≝8 definition get_input : ∀o:io_out. eventval → res (io_in o) ≝ 9 9 λo,ev. 10 10 match io_in_typ o return λt. res (eventval_type t) with … … 13 13 ]. 14 14 15 nlet rec up_to_nth_step (n:nat) (input:list eventval) (e:execution) (t:trace) : res (trace × state) ≝15 let rec up_to_nth_step (n:nat) (input:list eventval) (e:execution) (t:trace) : res (trace × state) ≝ 16 16 match n with 17 17 [ O ⇒ match e with [ e_step tr s _ ⇒ OK ? 〈t⧺tr, s〉 … … 30 30 ]. 31 31 32 ndefinition exec_up_to : clight_program → nat → list eventval → res (trace × state) ≝32 definition exec_up_to : clight_program → nat → list eventval → res (trace × state) ≝ 33 33 λp,n,i. up_to_nth_step n i (exec_inf p) E0. 34 34 … … 36 36 For example, 37 37 38 nremarkexec: result ? (exec_up_to myprog 20 [EVint (repr 12)]).39 n normalize;(* you can examine the result here *)40 @; nqed.38 example exec: result ? (exec_up_to myprog 20 [EVint (repr 12)]). 39 normalize (* you can examine the result here *) 40 % qed. 41 41 42 42 *) 43 43 44 ninductive result (A:Type) : A → Type≝44 inductive result (A:Type[0]) : A → Type[0] ≝ 45 45 | witness : ∀a:A. result A a. 46 46 -
Deliverables/D3.1/C-semantics/Cexec.ma
r485 r487 5 5 include "IOMonad.ma". 6 6 7 (* 7 8 include "Plogic/russell_support.ma". 8 9 ndefinition P_to_P_option_res : ∀A:Type[0].∀P:A → CProp[0].option (res A) → CProp[0] ≝9 *) 10 definition P_to_P_option_res : ∀A:Type[0].∀P:A → CProp[0].option (res A) → CProp[0] ≝ 10 11 λA,P,a.match a with [ None ⇒ False | Some y ⇒ match y return λ_.CProp[0] with [ Error ⇒ True | OK z ⇒ P z ]]. 11 12 12 ndefinition err_inject : ∀A.∀P:A → Prop.∀a:option (res A).∀p:P_to_P_option_res A P a.res (sigmaA P) ≝13 definition err_inject : ∀A.∀P:A → Prop.∀a:option (res A).∀p:P_to_P_option_res A P a.res (Sig A P) ≝ 13 14 λA.λP:A → Prop.λa:option (res A).λp:P_to_P_option_res A P a. 14 (match a return λa'.a=a' → res ( sigmaA P) with15 (match a return λa'.a=a' → res (Sig A P) with 15 16 [ None ⇒ λe1.? 16 17 | Some b ⇒ λe1.(match b return λb'.b=b' → ? with 17 18 [ Error ⇒ λ_. Error ? 18 | OK c ⇒ λe2. OK ? ( sig_introA P c ?)19 | OK c ⇒ λe2. OK ? (dp A P c ?) 19 20 ]) (refl ? b) 20 21 ]) (refl ? a). 21 ##[ nrewrite > e1 in p; nnormalize; *;22 ##| nrewrite > e1 in p; nrewrite > e2; nnormalize; //23 ##] nqed.24 25 ndefinition err_eject : ∀A.∀P: A → Prop. res (sigmaA P) → res A ≝22 [ >e1 in p; normalize; *; 23 | >e1 in p >e2 normalize; // 24 ] qed. 25 26 definition err_eject : ∀A.∀P: A → Prop. res (Sig A P) → res A ≝ 26 27 λA,P,a.match a with [ Error ⇒ Error ? | OK b ⇒ 27 match b with [ sig_introw p ⇒ OK ? w] ].28 29 ndefinition sig_eject : ∀A.∀P: A → Prop. sigmaA P → A ≝30 λA,P,a.match a with [ sig_introw p ⇒ w].31 32 ncoercion err_inject :33 ∀A.∀P:A → Prop.∀a.∀p:P_to_P_option_res ? P a.res ( sigmaA P) ≝ err_inject34 on a:option (res ?) to res ( sigma? ?).35 ncoercion err_eject : ∀A.∀P:A → Prop.∀c:res (sigmaA P).res A ≝ err_eject36 on _c:res ( sigma? ?) to res ?.37 ncoercion sig_eject : ∀A.∀P:A → Prop.∀c:sigmaA P.A ≝ sig_eject38 on _c: sigma? ? to ?.39 40 ndefinition P_res: ∀A.∀P:A → Prop.res A → Prop ≝28 match b with [ dp w p ⇒ OK ? w] ]. 29 30 definition sig_eject : ∀A.∀P: A → Prop. Sig A P → A ≝ 31 λA,P,a.match a with [ dp w p ⇒ w]. 32 33 coercion err_inject : 34 ∀A.∀P:A → Prop.∀a.∀p:P_to_P_option_res ? P a.res (Sig A P) ≝ err_inject 35 on a:option (res ?) to res (Sig ? ?). 36 coercion err_eject : ∀A.∀P:A → Prop.∀c:res (Sig A P).res A ≝ err_eject 37 on _c:res (Sig ? ?) to res ?. 38 coercion sig_eject : ∀A.∀P:A → Prop.∀c:Sig A P.A ≝ sig_eject 39 on _c:Sig ? ? to ?. 40 41 definition P_res: ∀A.∀P:A → Prop.res A → Prop ≝ 41 42 λA,P,a. match a with [ Error ⇒ True | OK v ⇒ P v ]. 42 43 43 ndefinition exec_bool_of_val : ∀v:val. ∀ty:type. res bool ≝44 definition exec_bool_of_val : ∀v:val. ∀ty:type. res bool ≝ 44 45 λv,ty. match v in val with 45 46 [ Vint i ⇒ match ty with … … 62 63 ]. 63 64 64 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.65 #v ty r H; nelim H; #v t H'; nelim H';66 ##[ #i is s ne; @ true; @; //; nwhd in ⊢ (??%?); nrewrite > (eq_false … ne);//;67 ##| #p b i i0 s; @ true; @;//68 ##| #f s ne; @ true; @; //; nwhd in ⊢ (??%?); nrewrite > (Feq_zero_false … ne);//;69 ##| #i s; @ false; @;//;70 ##| #r r' t; @ false; @;//;71 ##| #s; @ false; @; //; nwhd in ⊢ (??%?); nrewrite > (Feq_zero_true …);//;72 ##]73 nqed.65 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. 66 #v #ty #r #H elim H; #v #t #H' elim H'; 67 [ #i #is #s #ne %{ true} % //; whd in ⊢ (??%?); >(eq_false … ne) //; 68 | #p #b #i #i0 #s %{ true} % // 69 | #f #s #ne %{ true} % //; whd in ⊢ (??%?); >(Feq_zero_false … ne) //; 70 | #i #s %{ false} % //; 71 | #r #r' #t %{ false} % //; 72 | #s %{ false} % //; whd in ⊢ (??%?); >(Feq_zero_true …) //; 73 ] 74 qed. 74 75 75 76 (* Prove a few minor results to make proof obligations easy. *) 76 77 77 nlemma bind_OK: ∀A,B,P,e,f.78 lemma bind_OK: ∀A,B,P,e,f. 78 79 (∀v. e = OK A v → match f v with [ Error ⇒ True | OK v' ⇒ P v' ]) → 79 80 match bind A B e f with [ Error ⇒ True | OK v ⇒ P v ]. 80 #A B P e f; nelim e; /2/; nqed.81 82 nlemma sig_bind_OK: ∀A,B. ∀P:A → Prop. ∀P':B → Prop. ∀e:res (sigma A P). ∀f:sigmaA P → res B.83 (∀v:A. ∀p:P v. match f ( sig_introA P v p) with [ Error ⇒ True | OK v' ⇒ P' v'] ) →84 match bind ( sigmaA P) B e f with [ Error ⇒ True | OK v' ⇒ P' v' ].85 #A B P P' e f; nelim e;86 ##[ #v0; nelim v0; #v Hv IH; napply IH; 87 ##| #_; napply I; 88 ##] nqed.89 90 nlemma bind2_OK: ∀A,B,C,P,e,f.81 #A #B #P #e #f elim e; /2/; qed. 82 83 lemma sig_bind_OK: ∀A,B. ∀P:A → Prop. ∀P':B → Prop. ∀e:res (Sig A P). ∀f:Sig A P → res B. 84 (∀v:A. ∀p:P v. match f (dp A P v p) with [ Error ⇒ True | OK v' ⇒ P' v'] ) → 85 match bind (Sig A P) B e f with [ Error ⇒ True | OK v' ⇒ P' v' ]. 86 #A #B #P #P' #e #f elim e; 87 [ #v0 elim v0; #v #Hv #IH @IH 88 | #_ @I 89 ] qed. 90 91 lemma bind2_OK: ∀A,B,C,P,e,f. 91 92 (∀v1,v2. e = OK ? 〈v1,v2〉 → match f v1 v2 with [ Error ⇒ True | OK v' ⇒ P v' ]) → 92 93 match bind2 A B C e f with [ Error ⇒ True | OK v ⇒ P v ]. 93 #A B C P e f; nelim e; //; #v; ncases v; /2/; nqed.94 95 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.94 #A #B #C #P #e #f elim e; //; #v cases v; /2/; qed. 95 96 lemma sig_bind2_OK: ∀A,B,C. ∀P:A×B → Prop. ∀P':C → Prop. ∀e:res (Sig (A×B) P). ∀f:A → B → res C. 96 97 (∀v1:A.∀v2:B. P 〈v1,v2〉 → match f v1 v2 with [ Error ⇒ True | OK v' ⇒ P' v'] ) → 97 98 match bind2 A B C e f with [ Error ⇒ True | OK v' ⇒ P' v' ]. 98 #A B C P P' e f; nelim e; //;99 #v0 ; nelim v0; #v; nelim v; #v1 v2 Hv IH; napply IH; //; nqed.100 101 nlemma opt_bind_OK: ∀A,B,P,e,f.99 #A #B #C #P #P' #e #f elim e; //; 100 #v0 elim v0; #v elim v; #v1 #v2 #Hv #IH @IH //; qed. 101 102 lemma opt_bind_OK: ∀A,B,P,e,f. 102 103 (∀v. e = Some A v → match f v with [ Error ⇒ True | OK v' ⇒ P v' ]) → 103 104 match bind A B (opt_to_res A e) f with [ Error ⇒ True | OK v ⇒ P v ]. 104 #A B P e f; nelim e; nnormalize; /2/; nqed.105 106 nlemma extract_subset_pair: ∀A,B,C,P. ∀e:{e:A×B | P e}. ∀Q:A→B→res C. ∀R:C→Prop.105 #A #B #P #e #f elim e; normalize; /2/; qed. 106 (* 107 lemma extract_subset_pair: ∀A,B,C,P. ∀e:{e:A×B | P e}. ∀Q:A→B→res C. ∀R:C→Prop. 107 108 (∀a,b. eject ?? e = 〈a,b〉 → P 〈a,b〉 → match Q a b with [ OK v ⇒ R v | Error ⇒ True]) → 108 match match eject ?? e with [ mk_pair a b ⇒ Q a b ] with [ OK v ⇒ R v | Error ⇒ True ].109 #A B C P e Q R; ncases e; #e'; ncases e'; nnormalize;110 ##[ #H; napply (False_ind … H); 111 ##| #e''; ncases e''; #a b Pab H; nnormalize; /2/;112 ##] nqed.113 109 match match eject ?? e with [ pair a b ⇒ Q a b ] with [ OK v ⇒ R v | Error ⇒ True ]. 110 #A #B #C #P #e #Q #R cases e; #e' cases e'; normalize; 111 [ #H @(False_ind … H) 112 | #e'' cases e''; #a #b #Pab #H normalize; /2/; 113 ] qed. 114 *) 114 115 (* 115 116 nremark err_later: ∀A,B. ∀e:res A. match e with [ Error ⇒ Error B | OK v ⇒ Error B ] = Error B. 116 #A B e; ncases e; //; nqed.117 #A #B #e cases e; //; qed. 117 118 *) 118 119 119 ndefinition try_cast_null : ∀m:mem. ∀i:int. ∀ty:type. ∀ty':type. res val ≝120 definition try_cast_null : ∀m:mem. ∀i:int. ∀ty:type. ∀ty':type. res val ≝ 120 121 λm:mem. λi:int. λty:type. λty':type. 121 122 match eq i zero with … … 134 135 ]. 135 136 136 ndefinition exec_cast : ∀m:mem. ∀v:val. ∀ty:type. ∀ty':type. res val ≝137 definition exec_cast : ∀m:mem. ∀v:val. ∀ty:type. ∀ty':type. res val ≝ 137 138 λm:mem. λv:val. λty:type. λty':type. 138 139 match v with … … 165 166 | Vptr p b ofs ⇒ 166 167 do s ← match ty with [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code | _ ⇒ Error ? ]; 167 do u ← match eq_region_dec p s with [ inl _ ⇒ OK ? something| inr _ ⇒ Error ? ];168 do u ← match eq_region_dec p s with [ inl _ ⇒ OK ? it | inr _ ⇒ Error ? ]; 168 169 do s' ← match ty' with 169 170 [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code … … 174 175 | Vnull r ⇒ 175 176 do s ← match ty with [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code | _ ⇒ Error ? ]; 176 do u ← match eq_region_dec r s with [ inl _ ⇒ OK ? something| inr _ ⇒ Error ? ];177 do u ← match eq_region_dec r s with [ inl _ ⇒ OK ? it | inr _ ⇒ Error ? ]; 177 178 do s' ← match ty' with 178 179 [ Tpointer s _ ⇒ OK ? s | Tarray s _ _ ⇒ OK ? s | Tfunction _ _ ⇒ OK ? Code … … 182 183 ]. 183 184 184 ndefinition load_value_of_type' ≝185 λty,m,l. match l with [ mk_pair pl ofs ⇒ match pl with [ mk_pair psp loc ⇒185 definition load_value_of_type' ≝ 186 λty,m,l. match l with [ pair pl ofs ⇒ match pl with [ pair psp loc ⇒ 186 187 load_value_of_type ty m psp loc ofs ] ]. 187 188 … … 190 191 and use exec_lvalue'. *) 191 192 192 nlet rec exec_expr (ge:genv) (en:env) (m:mem) (e:expr) on e : res (val×trace) ≝193 let rec exec_expr (ge:genv) (en:env) (m:mem) (e:expr) on e : res (val×trace) ≝ 193 194 match e with 194 195 [ Expr e' ty ⇒ … … 210 211 | Eaddrof a ⇒ 211 212 do 〈plo,tr〉 ← exec_lvalue ge en m a; 212 OK ? 〈match plo with [ mk_pair pl ofs ⇒ match pl with [ mk_pair pcl loc ⇒ Vptr pcl loc ofs ] ], tr〉213 OK ? 〈match plo with [ pair pl ofs ⇒ match pl with [ pair pcl loc ⇒ Vptr pcl loc ofs ] ], tr〉 213 214 | Esizeof ty' ⇒ OK ? 〈Vint (repr (sizeof ty')), E0〉 214 215 | Eunop op a ⇒ … … 282 283 match e with [ Expr e' ty ⇒ exec_lvalue' ge en m e' ty ]. 283 284 284 nlemma P_res_to_P: ∀A,P,e,v. P_res A P e → e = OK A v → P v.285 #A P e v H1 H2; nrewrite > H2 in H1; nwhd in ⊢ (% → ?); //; nqed.285 lemma P_res_to_P: ∀A,P,e,v. P_res A P e → e = OK A v → P v. 286 #A #P #e #v #H1 #H2 >H2 in H1; whd in ⊢ (% → ?); //; qed. 286 287 287 288 (* We define a special induction principle tailored to the recursive definition 288 289 above. *) 289 290 290 ndefinition is_not_lvalue: expr_descr → Prop ≝291 definition is_not_lvalue: expr_descr → Prop ≝ 291 292 λe. match e with [ Evar _ ⇒ False | Ederef _ ⇒ False | Efield _ _ ⇒ False | _ ⇒ True ]. 292 293 293 ndefinition Plvalue ≝ λP:(expr → Prop).λe,ty.294 definition Plvalue ≝ λP:(expr → Prop).λe,ty. 294 295 match e return λ_.Prop with [ Evar _ ⇒ P (Expr e ty) | Ederef _ ⇒ P (Expr e ty) | Efield _ _ ⇒ P (Expr e ty) | _ ⇒ True ]. 295 296 296 297 (* TODO: Can we do this sensibly with a map combinator? *) 297 nlet rec exec_exprlist (ge:genv) (e:env) (m:mem) (l:list expr) on l : res (list val×trace) ≝298 let rec exec_exprlist (ge:genv) (e:env) (m:mem) (l:list expr) on l : res (list val×trace) ≝ 298 299 match l with 299 300 [ nil ⇒ OK ? 〈nil val, E0〉 … … 304 305 ]. 305 306 306 (* Don't really want to use subset rather than sigma here, but can't be bothered 307 with *another* set of coercions. XXX: why do I have to get the recursive 308 call's property manually? *) 309 310 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) } ≝ 307 let rec exec_alloc_variables (en:env) (m:mem) (l:list (ident × type)) on l : env × mem ≝ 311 308 match l with 312 [ nil ⇒ Some ?〈en, m〉309 [ nil ⇒ 〈en, m〉 313 310 | cons h vars ⇒ 314 match h with [ mk_pair id ty ⇒ 315 match alloc m 0 (sizeof ty) Any with [ mk_pair m1 b1 ⇒ 316 match exec_alloc_variables (set … id b1 en) m1 vars with 317 [ sig_intro r p ⇒ r ] 318 ]]]. nwhd; 319 ##[ //; 320 ##| nelim (exec_alloc_variables (set ident ? ? c3 c7 en) c6 c1); 321 #H; nelim H; //; #H0; nelim H0; nnormalize; #en' m' IH; 322 napply (alloc_variables_cons … IH); /2/; 323 nqed. 311 match h with [ pair id ty ⇒ 312 match alloc m 0 (sizeof ty) Any with [ pair m1 b1 ⇒ 313 exec_alloc_variables (set … id b1 en) m1 vars 314 ]]]. 324 315 325 316 (* TODO: can we establish that length params = length vs in advance? *) 326 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)≝317 let rec exec_bind_parameters (e:env) (m:mem) (params:list (ident × type)) (vs:list val) on params : res mem ≝ 327 318 match params with 328 [ nil ⇒ match vs with [ nil ⇒ Some ? (OK ? m) | cons _ _ ⇒ Some ? (Error ?)]329 | cons idty params' ⇒ match idty with [ mk_pair id ty ⇒319 [ nil ⇒ match vs with [ nil ⇒ OK ? m | cons _ _ ⇒ Error ? ] 320 | cons idty params' ⇒ match idty with [ pair id ty ⇒ 330 321 match vs with 331 [ nil ⇒ Some ? (Error ?)332 | cons v1 vl ⇒ Some ? (322 [ nil ⇒ Error ? 323 | cons v1 vl ⇒ 333 324 do b ← opt_to_res ? (get … id e); 334 325 do m1 ← opt_to_res ? (store_value_of_type ty m Any b zero v1); 335 e rr_eject ?? (exec_bind_parameters e m1 params' vl)) (* FIXME: don't want to have to eject here *)326 exec_bind_parameters e m1 params' vl 336 327 ] 337 328 ] ]. 338 nwhd; //; 339 napply opt_bind_OK; #b eb; 340 napply opt_bind_OK; #m1 em1; 341 napply sig_bind_OK; #m2 Hm2; 342 napply (bind_parameters_cons … eb em1 Hm2); 343 nqed. 344 345 alias id "Tint" = "cic:/matita/c-semantics/Csyntax/type.con(0,2,0)". 346 alias id "Tfloat" = "cic:/matita/c-semantics/Csyntax/type.con(0,3,0)". 347 ndefinition sz_eq_dec : ∀s1,s2:intsize. (s1 = s2) + (s1 ≠ s2). 348 #s1; ncases s1; #s2; ncases s2; /2/; @2; @; #H; ndestruct; nqed. 349 ndefinition sg_eq_dec : ∀s1,s2:signedness. (s1 = s2) + (s1 ≠ s2). 350 #s1; ncases s1; #s2; ncases s2; /2/; @2; @; #H; ndestruct; nqed. 351 ndefinition fs_eq_dec : ∀s1,s2:floatsize. (s1 = s2) + (s1 ≠ s2). 352 #s1; ncases s1; #s2; ncases s2; /2/; @2; @; #H; ndestruct; nqed. 353 354 nlet rec type_eq_dec (t1,t2:type) : (t1 = t2) + (t1 ≠ t2) ≝ 329 330 definition sz_eq_dec : ∀s1,s2:intsize. (s1 = s2) + (s1 ≠ s2). 331 #s1 cases s1; #s2 cases s2; /2/; %2 ; % #H destruct; qed. 332 definition sg_eq_dec : ∀s1,s2:signedness. (s1 = s2) + (s1 ≠ s2). 333 #s1 cases s1; #s2 cases s2; /2/; %2 ; % #H destruct; qed. 334 definition fs_eq_dec : ∀s1,s2:floatsize. (s1 = s2) + (s1 ≠ s2). 335 #s1 cases s1; #s2 cases s2; /2/; %2 ; % #H destruct; qed. 336 337 let rec type_eq_dec (t1,t2:type) : (t1 = t2) + (t1 ≠ t2) ≝ 355 338 match t1 return λt'. (t' = t2) + (t' ≠ t2) with 356 339 [ Tvoid ⇒ match t2 return λt'. (Tvoid = t') + (Tvoid ≠ t') with [ Tvoid ⇒ inl ?? (refl ??) | _ ⇒ inr ?? (nmk ? (λH.?)) ] … … 424 407 | inr e ⇒ inr ?? (nmk ? (λH.match e with [ nmk e' ⇒ e' ? ])) ] 425 408 | inr e ⇒ inr ?? (nmk ? (λH.match e with [ nmk e' ⇒ e' ? ])) ] ] 426 ]. ndestruct; //; nqed.427 428 ndefinition assert_type_eq : ∀t1,t2:type. res (t1 = t2) ≝409 ]. destruct; //; qed. 410 411 definition assert_type_eq : ∀t1,t2:type. res (t1 = t2) ≝ 429 412 λt1,t2. match type_eq_dec t1 t2 with [ inl p ⇒ OK ? p | inr _ ⇒ Error ? ]. 430 413 431 nlet rec is_is_call_cont (k:cont) : (is_call_cont k) + (¬is_call_cont k) ≝414 let rec is_is_call_cont (k:cont) : (is_call_cont k) + (¬is_call_cont k) ≝ 432 415 match k return λk'.(is_call_cont k') + (¬is_call_cont k') with 433 416 [ Kstop ⇒ ? … … 435 418 | _ ⇒ ? 436 419 ]. 437 ##[ ##1,8: @1; //438 ##| ##*: @2; /2/439 ##] nqed.440 441 ndefinition is_Sskip : ∀s:statement. (s = Sskip) + (s ≠ Sskip) ≝420 [ 1,8: %1 ; // 421 | *: %2 ; /2/ 422 ] qed. 423 424 definition is_Sskip : ∀s:statement. (s = Sskip) + (s ≠ Sskip) ≝ 442 425 λs.match s return λs'.(s' = Sskip) + (s' ≠ Sskip) with 443 426 [ Sskip ⇒ inl ?? (refl ??) 444 427 | _ ⇒ inr ?? (nmk ? (λH.?)) 445 ]. ndestruct;446 nqed.428 ]. destruct 429 qed. 447 430 448 431 (* Checking types of values given to / received from an external function call. *) 449 432 450 ndefinition eventval_type : ∀ty:typ. Type≝433 definition eventval_type : ∀ty:typ. Type[0] ≝ 451 434 λty. match ty with [ ASTint ⇒ int | ASTfloat ⇒ float ]. 452 435 453 ndefinition mk_eventval: ∀ty:typ. eventval_type ty → eventval ≝436 definition mk_eventval: ∀ty:typ. eventval_type ty → eventval ≝ 454 437 λty:typ. match ty return λty'.eventval_type ty' → eventval with [ ASTint ⇒ λv.EVint v | ASTfloat ⇒ λv.EVfloat v ]. 455 438 456 ndefinition mk_val: ∀ty:typ. eventval_type ty → val ≝439 definition mk_val: ∀ty:typ. eventval_type ty → val ≝ 457 440 λty:typ. match ty return λty'.eventval_type ty' → val with [ ASTint ⇒ λv.Vint v | ASTfloat ⇒ λv.Vfloat v ]. 458 441 459 nlemma mk_val_correct: ∀ty:typ. ∀v:eventval_type ty.442 lemma mk_val_correct: ∀ty:typ. ∀v:eventval_type ty. 460 443 eventval_match (mk_eventval ty v) ty (mk_val ty v). 461 #ty ; ncases ty; nnormalize; //; nqed.462 463 ndefinition convert_eventval : ∀ev:eventval. ∀ty:typ. res (Σv:val. eventval_match ev ty v)≝444 #ty cases ty; normalize; //; qed. 445 446 definition convert_eventval : ∀ev:eventval. ∀ty:typ. res val ≝ 464 447 λev,ty. 465 448 match ty with 466 [ ASTint ⇒ match ev with [ EVint i ⇒ Some ? (OK ? (Vint i)) | _ ⇒ Some ? (Error ?)]467 | ASTfloat ⇒ match ev with [ EVfloat f ⇒ Some ? (OK ? (Vfloat f)) | _ ⇒ Some ? (Error ?)]468 | _ ⇒ Some ? (Error ?)469 ]. nwhd; //; nqed.470 471 ndefinition check_eventval' : ∀v:val. ∀ty:typ. res (Σev:eventval. eventval_match ev ty v)≝449 [ ASTint ⇒ match ev with [ EVint i ⇒ OK ? (Vint i) | _ ⇒ Error ? ] 450 | ASTfloat ⇒ match ev with [ EVfloat f ⇒ OK ? (Vfloat f) | _ ⇒ Error ? ] 451 | _ ⇒ Error ? 452 ]. 453 454 definition check_eventval' : ∀v:val. ∀ty:typ. res eventval ≝ 472 455 λv,ty. 473 456 match ty with 474 [ ASTint ⇒ match v with [ Vint i ⇒ Some ? (OK ? (EVint i)) | _ ⇒ Some ? (Error ?)]475 | ASTfloat ⇒ match v with [ Vfloat f ⇒ Some ? (OK ? (EVfloat f)) | _ ⇒ Some ? (Error ?)]457 [ ASTint ⇒ match v with [ Vint i ⇒ OK ? (EVint i) | _ ⇒ Error ? ] 458 | ASTfloat ⇒ match v with [ Vfloat f ⇒ OK ? (EVfloat f) | _ ⇒ Error ? ] 476 459 | _ ⇒ Some ? (Error ?) 477 ]. nwhd; //; nqed.478 479 nlet rec check_eventval_list (vs: list val) (tys: list typ) : res (Σevs:list eventval. eventval_list_match evs tys vs) ≝460 ]. 461 462 let rec check_eventval_list (vs: list val) (tys: list typ) : res (list eventval) ≝ 480 463 match vs with 481 [ nil ⇒ match tys with [ nil ⇒ Some ? (OK ? (nil ?)) | _ ⇒ Some ? (Error ?)]464 [ nil ⇒ match tys with [ nil ⇒ OK ? (nil ?) | _ ⇒ Error ? ] 482 465 | cons v vt ⇒ 483 466 match tys with 484 [ nil ⇒ Some ? (Error ?)485 | cons ty tyt ⇒ Some ? (467 [ nil ⇒ Error ? 468 | cons ty tyt ⇒ 486 469 do ev ← check_eventval' v ty; 487 470 do evt ← check_eventval_list vt tyt; 488 OK ? ((sig_eject ?? ev)::evt)) 489 ] 490 ]. nwhd; //; 491 napply sig_bind_OK; #ev Hev; 492 napply sig_bind_OK; #evt Hevt; 493 nnormalize; /2/; 494 nqed. 471 OK ? (ev::evt) 472 ] 473 ]. 495 474 496 475 (* IO monad *) … … 498 477 (* Interactions are function calls that return a value and do not change 499 478 the rest of the Clight program's state. *) 500 nrecord io_out : Type≝479 record io_out : Type[0] ≝ 501 480 { io_function: ident 502 481 ; io_args: list eventval … … 504 483 }. 505 484 506 ndefinition io_in ≝ λo. eventval_type (io_in_typ o).507 508 ndefinition do_io : ident → list eventval → ∀t:typ. IO io_out io_in (eventval_type t) ≝485 definition io_in ≝ λo. eventval_type (io_in_typ o). 486 487 definition do_io : ident → list eventval → ∀t:typ. IO io_out io_in (eventval_type t) ≝ 509 488 λfn,args,t. Interact ??? (mk_io_out fn args t) (λres. Value ??? res). 510 489 511 ndefinition ret: ∀T. T → (IO io_out io_in T) ≝490 definition ret: ∀T. T → (IO io_out io_in T) ≝ 512 491 λT,x.(Value ?? T x). 513 492 514 493 (* execution *) 515 494 516 ndefinition store_value_of_type' ≝495 definition store_value_of_type' ≝ 517 496 λty,m,l,v. 518 match l with [ mk_pair pl ofs ⇒519 match pl with [ mk_pair pcl loc ⇒497 match l with [ pair pl ofs ⇒ 498 match pl with [ pair pcl loc ⇒ 520 499 store_value_of_type ty m pcl loc ofs v ] ]. 521 500 522 nlet rec exec_step (ge:genv) (st:state) on st : (IO io_out io_in (trace × state)) ≝501 let rec exec_step (ge:genv) (st:state) on st : (IO io_out io_in (trace × state)) ≝ 523 502 match st with 524 503 [ State f s k e m ⇒ … … 639 618 | Sgoto lbl ⇒ 640 619 match find_label lbl (fn_body f) (call_cont k) with 641 [ Some sk' ⇒ match sk' with [ mk_pair s' k' ⇒ ret ? 〈E0, State f s' k' e m〉 ]620 [ Some sk' ⇒ match sk' with [ pair s' k' ⇒ ret ? 〈E0, State f s' k' e m〉 ] 642 621 | None ⇒ Wrong ??? 643 622 ] … … 647 626 match f0 with 648 627 [ Internal f ⇒ 649 match exec_alloc_variables empty_env m ((fn_params f) @ (fn_vars f)) with [ mk_pair e m1 ⇒650 ! m2 ← e rr_to_io_sig … (exec_bind_parameters e m1 (fn_params f) vargs);628 match exec_alloc_variables empty_env m ((fn_params f) @ (fn_vars f)) with [ pair e m1 ⇒ 629 ! m2 ← exec_bind_parameters e m1 (fn_params f) vargs; 651 630 ret ? 〈E0, State f (fn_body f) k e m2〉 652 631 ] 653 632 | External f argtys retty ⇒ 654 ! evargs ← err_to_io_sig … (check_eventval_list vargs (typlist_of_typelist argtys));633 ! evargs ← check_eventval_list vargs (typlist_of_typelist argtys); 655 634 ! evres ← do_io f evargs (proj_sig_res (signature_of_type argtys retty)); 656 635 ret ? 〈Eextcall f evargs (mk_eventval ? evres), Returnstate (mk_val ? evres) k m〉 … … 662 641 [ None ⇒ ret ? 〈E0, (State f Sskip k' e m)〉 663 642 | Some r' ⇒ 664 match r' with [ mk_pair l ty ⇒643 match r' with [ pair l ty ⇒ 665 644 666 645 ! m' ← store_value_of_type' ty m l res; … … 672 651 ]. 673 652 674 nlet rec make_initial_state (p:clight_program) : res (genv × state) ≝653 let rec make_initial_state (p:clight_program) : res (genv × state) ≝ 675 654 do ge ← globalenv Genv ?? p; 676 655 do m0 ← init_mem Genv ?? p; 677 656 do 〈sp,b〉 ← opt_to_res ? (find_symbol ? ? ge (prog_main ?? p)); 678 do u ← opt_to_res ? (match eq_region_dec sp Code with [ inl _ ⇒ Some ? something| inr _ ⇒ None ? ]);657 do u ← opt_to_res ? (match eq_region_dec sp Code with [ inl _ ⇒ Some ? it | inr _ ⇒ None ? ]); 679 658 do f ← opt_to_res ? (find_funct_ptr ? ? ge b); 680 659 OK ? 〈ge,Callstate f (nil ?) Kstop m0〉. 681 660 682 ndefinition is_final_state : ∀st:state. (Σr. final_state st r) + (¬∃r. final_state st r).683 #st ; nelim st;684 ##[ #f s k e m; @2; @;*; #r H; ninversion H; #i m e; ndestruct;685 ##| #f l k m; @2; @;*; #r H; ninversion H; #i m e; ndestruct;686 ##| #v k m; ncases k;687 ##[ ncases v;688 ##[ ##2: #i; @1; @ i;//;689 ##| ##1: @2; @; *; #r H; ninversion H; #i m e; ndestruct;690 ##| #f; @2; @; *; #r H; ninversion H; #i m e; ndestruct;691 ##| #r; @2; @; *; #r H; ninversion H; #i m e; ndestruct;692 ##| #r b of; @2; @; *; #r H; ninversion H; #i m e; ndestruct;693 ##]694 ##| #a b; @2; @; *; #r H; ninversion H; #i m e; ndestruct;695 ##| ##3,4: #a b c; @2; @; *; #r H; ninversion H; #i m e; ndestruct;696 ##| ##5,6,8: #a b c d; @2; @; *; #r H; ninversion H; #i m e; ndestruct;697 ##| #a; @2; @; *; #r H; ninversion H; #i m e; ndestruct;698 ##]699 ##] nqed.700 701 nlet rec exec_steps (n:nat) (ge:genv) (s:state) :661 definition is_final_state : ∀st:state. (Sig ? (final_state st)) + (¬∃r. final_state st r). 662 #st elim st; 663 [ #f #s #k #e #m %2 ; % *; #r #H inversion H; #i #m #e destruct; 664 | #f #l #k #m %2 ; % *; #r #H inversion H; #i #m #e destruct; 665 | #v #k #m cases k; 666 [ cases v; 667 [ 2: #i %1 ; %{ i} //; 668 | 1: %2 ; % *; #r #H inversion H; #i #m #e destruct; 669 | #f %2 ; % *; #r #H inversion H; #i #m #e destruct; 670 | #r %2 ; % *; #r #H inversion H; #i #m #e destruct; 671 | #r #b #of %2 ; % *; #r #H inversion H; #i #m #e destruct; 672 ] 673 | #a #b %2 ; % *; #r #H inversion H; #i #m #e destruct; 674 | 3,4: #a #b #c %2 ; % *; #r #H inversion H; #i #m #e destruct; 675 | 5,6,8: #a #b #c #d %2 ; % *; #r #H inversion H; #i #m #e destruct; 676 | #a %2 ; % *; #r #H inversion H; #i #m #e destruct; 677 ] 678 ] qed. 679 680 let rec exec_steps (n:nat) (ge:genv) (s:state) : 702 681 IO io_out io_in (trace×state) ≝ 703 682 match is_final_state s with … … 724 703 725 704 (* A (possibly non-terminating) execution. *) 726 ncoinductive execution : Type≝705 coinductive execution : Type[0] ≝ 727 706 | e_stop : trace → int → mem → execution 728 707 | e_step : trace → state → execution → execution … … 730 709 | e_interact : ∀o:io_out. (io_in o → execution) → execution. 731 710 732 ndefinition mem_of_state : state → mem ≝711 definition mem_of_state : state → mem ≝ 733 712 λs.match s with [ State f s k e m ⇒ m | Callstate f a k m ⇒ m | Returnstate r k m ⇒ m ]. 734 713 … … 737 716 state. *) 738 717 739 nlet corec exec_inf_aux (ge:genv) (s:IO io_out io_in (trace×state)) : execution ≝718 let corec exec_inf_aux (ge:genv) (s:IO io_out io_in (trace×state)) : execution ≝ 740 719 match s with 741 720 [ Wrong ⇒ e_wrong 742 | Value v ⇒ match v with [ mk_pair t s' ⇒721 | Value v ⇒ match v with [ pair t s' ⇒ 743 722 match is_final_state s' with 744 723 [ inl r ⇒ e_stop t (sig_eject … r) (mem_of_state s') … … 748 727 749 728 750 ndefinition exec_inf : clight_program → execution ≝729 definition exec_inf : clight_program → execution ≝ 751 730 λp. 752 731 match make_initial_state p with … … 755 734 ]. 756 735 757 nremarkexecution_cases: ∀e.736 lemma execution_cases: ∀e. 758 737 e = match e with [ e_stop tr r m ⇒ e_stop tr r m | e_step tr s e ⇒ e_step tr s e 759 738 | e_wrong ⇒ e_wrong | e_interact o k ⇒ e_interact o k ]. 760 #e ; ncases e; //; nqed.761 762 nlemma exec_inf_aux_unfold: ∀ge,s. exec_inf_aux ge s =739 #e cases e; //; qed. 740 741 lemma exec_inf_aux_unfold: ∀ge,s. exec_inf_aux ge s = 763 742 match s with 764 743 [ Wrong ⇒ e_wrong 765 | Value v ⇒ match v with [ mk_pair t s' ⇒744 | Value v ⇒ match v with [ pair t s' ⇒ 766 745 match is_final_state s' with 767 746 [ inl r ⇒ e_stop t (sig_eject … r) (mem_of_state s') … … 769 748 | Interact out k' ⇒ e_interact out (λv. exec_inf_aux ge (k' v)) 770 749 ]. 771 #ge s; nrewrite > (execution_cases (exec_inf_aux …)); ncases s;772 ##[ #ok773 ##| #x; ncases x; #tr s'; ncases s';774 ##[ #fn st k envm775 ##| #fd args kmem776 ##| #v k mem; (* for final state check *) ncases k;777 ##[ ncases v; ##[ ##2,3: #v' ##| ##4: #r ##| ##5: #sp loc off ##]778 ##| #s' k' ##| ##3,4: #e s' k' ##| ##5,6: #e s' s'' k' ##| #k' ##| #a b c d ##]779 ##]780 ##| ##]781 nwhd in ⊢ (??%%); //;782 nqed.783 750 #ge #s >(execution_cases (exec_inf_aux …)) cases s; 751 [ #o #k 752 | #x cases x; #tr #s' cases s'; 753 [ #fn #st #k #env #m 754 | #fd #args #k #mem 755 | #v #k #mem (* for final state check *) cases k; 756 [ cases v; [ 2,3: #v' | 4: #r | 5: #sp #loc #off ] 757 | #s' #k' | 3,4: #e #s' #k' | 5,6: #e #s' #s'' #k' | #k' | #a #b #c #d ] 758 ] 759 | ] 760 whd in ⊢ (??%%); //; 761 qed. 762 -
Deliverables/D3.1/C-semantics/CexecComplete.ma
r485 r487 1 1 include "Cexec.ma". 2 include "Plogic/connectives.ma". 3 4 ndefinition yields ≝ λA.λa:res A.λv':A. 2 3 definition yields ≝ λA.λa:res A.λv':A. 5 4 match a with [ OK v ⇒ v' = v | _ ⇒ False ]. 6 5 … … 9 8 that one particular one exists corresponding to a derivation in the inductive 10 9 semantics.) *) 11 nlet rec yieldsIO (A:Type) (a:IO io_out io_in A) (v':A) on a : Prop ≝10 let rec yieldsIO (A:Type[0]) (a:IO io_out io_in A) (v':A) on a : Prop ≝ 12 11 match a with [ Value v ⇒ v' = v | Interact _ k ⇒ ∃r.yieldsIO A (k r) v' | _ ⇒ False ]. 13 12 14 ndefinition yields_sig : ∀A,P. res (Σx:A. P x) → A → Prop ≝15 λA,P,e,v. match e with [ OK v' ⇒ match v' with [ sig_introv'' _ ⇒ v = v'' ] | _ ⇒ False].16 17 nlet rec yieldsIO_sig (A:Type) (P:A → Prop) (e:IO io_out io_in (Σx:A. P x)) (v:A) on e : Prop ≝13 definition yields_sig : ∀A,P. res (Sig A P) → A → Prop ≝ 14 λA,P,e,v. match e with [ OK v' ⇒ match v' with [ dp v'' _ ⇒ v = v'' ] | _ ⇒ False]. 15 16 let rec yieldsIO_sig (A:Type[0]) (P:A → Prop) (e:IO io_out io_in (Sig A P)) (v:A) on e : Prop ≝ 18 17 match e with 19 [ Value v' ⇒ match v' with [ sig_introv'' _ ⇒ v = v'' ]18 [ Value v' ⇒ match v' with [ dp v'' _ ⇒ v = v'' ] 20 19 | Interact _ k ⇒ ∃r.yieldsIO_sig A P (k r) v 21 20 | _ ⇒ False]. 22 21 23 nlemma remove_io_sig: ∀A. ∀P:A → Prop. ∀a,v',p.22 lemma remove_io_sig: ∀A. ∀P:A → Prop. ∀a,v',p. 24 23 yieldsIO A a v' → 25 yieldsIO_sig A P (io_inject io_out io_in A (λx.P x)(Some ? a) p) v'.26 #A P a; nelim a;27 ##[ #a k IH v' p H; nwhd in H ⊢ %; nelim H; #r H'; @ r; napply IH; napply H'; 28 ##| #v v' p H; napply H; 29 ##| #a b;*;30 ##] nqed.31 32 nlemma yields_eq: ∀A,a,v'. yields A a v' → a = OK ? v'.33 #A a v'; ncases a; //; nwhd in ⊢ (% → ?); *;34 nqed.35 36 nlemma yields_sig_eq: ∀A,P,e,v. yields_sig A P e v → ∃p. e = OK ? (sig_intro… v p).37 #A P e v; ncases e;38 ##[ #vp; ncases vp; #v' p H; nwhd in H; nrewrite > H; @ p; napply refl; 39 ##| *;40 ##] nqed.41 42 nlemma is_pointer_compat_true: ∀m,b,sp.24 yieldsIO_sig A P (io_inject io_out io_in A P (Some ? a) p) v'. 25 #A #P #a elim a; 26 [ #a #k #IH #v' #p #H whd in H ⊢ %; elim H; #r #H' %{ r} @IH @H' 27 | #v #v' #p #H @H 28 | #a #b *; 29 ] qed. 30 31 lemma yields_eq: ∀A,a,v'. yields A a v' → a = OK ? v'. 32 #A #a #v' cases a; //; whd in ⊢ (% → ?); *; 33 qed. 34 35 lemma yields_sig_eq: ∀A,P,e,v. yields_sig A P e v → ∃p. e = OK ? (dp … v p). 36 #A #P #e #v cases e; 37 [ #vp cases vp; #v' #p #H whd in H; >H %{ p} @refl 38 | *; 39 ] qed. 40 41 lemma is_pointer_compat_true: ∀m,b,sp. 43 42 pointer_compat (block_space m b) sp → 44 43 is_pointer_compat (block_space m b) sp = true. 45 #m b sp H; nwhd in ⊢ (??%?);46 nelim (pointer_compat_dec (block_space m b) sp);47 ##[ //48 ##| #H'; napply False_ind; napply (absurd … H H'); 49 ##] nqed.50 51 nlemma eq_region_dec_true: ∀s. eq_region_dec s s = inl ???.52 ##[ #s; ncases s; napply refl; 53 ##| ##skip54 ##] nqed.55 56 ntheorem is_det: ∀p,s,s'.44 #m #b #sp #H whd in ⊢ (??%?); 45 elim (pointer_compat_dec (block_space m b) sp); 46 [ // 47 | #H' @False_ind @(absurd … H H') 48 ] qed. 49 50 lemma eq_region_dec_true: ∀s. eq_region_dec s s = inl ???. 51 [ #s cases s; @refl 52 | skip 53 ] qed. 54 55 theorem is_det: ∀p,s,s'. 57 56 initial_state p s → initial_state p s' → s = s'. 58 #p s s' H1 H2;59 ninversion H1; #b1 f1 ge1 m1 e11 e12 e13 e14 e15; 60 ninversion H2; #b2 f2 ge2 m2 e21 e22 e23 e24 e25; 61 nrewrite > e11 in e21; #e1; nrewrite > (?:ge1 = ge2) in e13 e14; 62 ##[ ##2: ndestruct (e1) skip (e11); napply refl; ##]63 #e13 e14;64 65 nrewrite > e12 in e22; #e2; ndestruct (e2) skip (e12);66 67 nrewrite > e13 in e23; #e3; nrewrite > (?:b1 = b2) in e14; 68 ##[ nrewrite > e24; #e4; nrewrite > (?:f2 = f1); 69 ##[ //;70 ##| ndestruct (e4) skip (e24 e25); //;71 ##]72 ##| ndestruct (e3) skip (e13); //73 ##] nqed.74 75 nlemma remove_res_sig: ∀A. ∀P:A → Prop. ∀a,v',p.57 #p #s #s' #H1 #H2 58 inversion H1; #b1 #f1 #ge1 #m1 #e11 #e12 #e13 #e14 #e15 59 inversion H2; #b2 #f2 #ge2 #m2 #e21 #e22 #e23 #e24 #e25 60 >e11 in e21 #e1 >(?:ge1 = ge2) in e13 e14 61 [ 2: destruct (e1) skip (e11); @refl ] 62 #e13 #e14 63 64 >e12 in e22 #e2 destruct (e2) skip (e12); 65 66 >e13 in e23 #e3 >(?:b1 = b2) in e14 67 [ >e24 #e4 >(?:f2 = f1) 68 [ //; 69 | destruct (e4) skip (e24 e25); //; 70 ] 71 | destruct (e3) skip (e13); // 72 ] qed. 73 74 lemma remove_res_sig: ∀A. ∀P:A → Prop. ∀a,v',p. 76 75 yields A a v' → 77 yields_sig A P (err_inject A (λx.P x)(Some ? a) p) v'.78 #A P a; ncases a;79 ##[ #v v' p H; napply H; 80 ##| #a b;*;81 ##] nqed.82 83 84 ntheorem the_initial_state:76 yields_sig A P (err_inject A P (Some ? a) p) v'. 77 #A #P #a cases a; 78 [ #v #v' #p #H @H 79 | #a #b *; 80 ] qed. 81 82 83 theorem the_initial_state: 85 84 ∀p,s. initial_state p s → ∃ge. yields ? (make_initial_state p) 〈ge,s〉. 86 #p s; ncases p; #fns main globs H;87 ninversion H;88 #b f ge m e1 e2 e3 e4 e5; @ge;89 nwhd in ⊢ (??%?);90 nrewrite > e1; 91 nwhd in ⊢ (??%?);92 nrewrite > e2; 93 nwhd in ⊢ (??%?);94 nrewrite > e3; 95 nwhd in ⊢ (??%?);96 nrewrite > e4; 97 nwhd; napply refl; 98 nqed.99 100 nlemma cast_complete: ∀m,v,ty,ty',v'.85 #p #s cases p; #fns #main #globs #H 86 inversion H; 87 #b #f #ge #m #e1 #e2 #e3 #e4 #e5 %{ge} 88 whd in ⊢ (??%?); 89 >e1 90 whd in ⊢ (??%?); 91 >e2 92 whd in ⊢ (??%?); 93 >e3 94 whd in ⊢ (??%?); 95 >e4 96 whd; @refl 97 qed. 98 99 lemma cast_complete: ∀m,v,ty,ty',v'. 101 100 cast m v ty ty' v' → yields ? (exec_cast m v ty ty') v'. 102 #m v ty ty' v' H;103 nelim H;104 ##[ #m i sz1 sz2 sg1 sg2; napply refl; 105 ##| #m f sz szi sg; napply refl; 106 ##| #m i sz sz' sg; napply refl; 107 ##| #m f sz sz'; napply refl; 108 ##| #m sp sp' ty ty' b ofs H1 H2 H3; 109 nelim H1; ##[ #sp1 ty1 ##| #sp1 ty1 n1 ##| #tys1 ty1; nletin sp1 ≝ Code ##]110 nwhd in ⊢ (??%?);111 ##[ ##1,2: nrewrite > (eq_region_dec_true …); nwhd in ⊢ (??%?); ##]112 nelim H2 in H3 ⊢ %; ##[ ##1,4,7: #sp2 ty2 ##| ##2,5,8: #sp2 ty2 n2 ##| ##3,6,9: #tys2 ty2; nletin sp2 ≝ Code ##]113 #H3 ; nwhd in ⊢ (??%?);114 nrewrite > (is_pointer_compat_true …);//;115 ##| #m sz si ty'' r H; ncases H; //;116 ##| #m t t' r r' H H'; ncases H; ntry #a; ntry #b; ntry #c; ncases H'; ntry #d; ntry #e; ntry #f; 117 nwhd in ⊢ (??%?); ntry napply refl; nrewrite > eq_region_dec_true; napply refl;118 ##] nqed.101 #m #v #ty #ty' #v' #H 102 elim H; 103 [ #m #i #sz1 #sz2 #sg1 #sg2 @refl 104 | #m #f #sz #szi #sg @refl 105 | #m #i #sz #sz' #sg @refl 106 | #m #f #sz #sz' @refl 107 | #m #sp #sp' #ty #ty' #b #ofs #H1 #H2 #H3 108 elim H1; [ #sp1 #ty1 | #sp1 #ty1 #n1 | #tys1 #ty1 letin sp1 ≝ Code ] 109 whd in ⊢ (??%?); 110 [ 1,2: >(eq_region_dec_true …) whd in ⊢ (??%?); ] 111 elim H2 in H3 ⊢ %; [ 1,4,7: #sp2 #ty2 | 2,5,8: #sp2 #ty2 #n2 | 3,6,9: #tys2 #ty2 letin sp2 ≝ Code ] 112 #H3 whd in ⊢ (??%?); 113 >(is_pointer_compat_true …) //; 114 | #m #sz #si #ty'' #r #H cases H; //; 115 | #m #t #t' #r #r' #H #H' cases H; try #a try #b try #c cases H'; try #d try #e try #f 116 whd in ⊢ (??%?); try @refl >eq_region_dec_true @refl 117 ] qed. 119 118 120 119 (* Use to narrow down the choice of expression to just the lvalues. *) 121 nlemma lvalue_expr: ∀ge,env,m,e,ty,sp,l,ofs,tr. ∀P:expr_descr → Prop.120 lemma lvalue_expr: ∀ge,env,m,e,ty,sp,l,ofs,tr. ∀P:expr_descr → Prop. 122 121 eval_lvalue ge env m (Expr e ty) sp l ofs tr → 123 122 (∀id. P (Evar id)) → (∀e'. P (Ederef e')) → (∀e',id. P (Efield e' id)) → 124 123 P e. 125 #ge env m e ty sp l ofs tr P H; napply (eval_lvalue_inv_ind … H);126 ##[ #id l ty e1 e2 e3 e4 e5 e6; ndestruct; //127 ##| #id sp l ty e1 e2 e3 e4 e5 e6 e7; ndestruct; //128 ##| #e ty sp l ofs tr H e1 e2 e3 e4 e5; ndestruct; //129 ##| #e id ty sp l ofs id' fs d tr H e1 e2;(* bogus? *) #_; #e3 e4 e5 e6 e7; ndestruct; //130 ##| #e id ty sp l ofs id' fs tr H e1;(* bogus? *) #_; #e2 e3 e4 e5 e6; ndestruct; //131 ##] nqed.132 133 nlemma bool_of_val_3_complete : ∀v,ty,r. bool_of_val v ty r → ∃b. r = of_bool b ∧ yields ? (exec_bool_of_val v ty) b.134 #v ty r H; nelim H; #v t H'; nelim H';135 ##[ #i is s ne; @ true; @; //; nwhd; nrewrite > (eq_false … ne);//;136 ##| #p b i i0 s; @ true; @;//137 ##| #f s ne; @ true; @; //; nwhd; nrewrite > (Feq_zero_false … ne);//;138 ##| #i s; @ false; @;//;139 ##| #r r' t; @ false; @;//;140 ##| #s; @ false; @; //; nwhd; nrewrite > (Feq_zero_true …);//;141 ##]142 nqed.143 144 nlemma bool_of_true: ∀v,ty. is_true v ty → yields ? (exec_bool_of_val v ty) true.145 #v ty H; nelim H;146 ##[ #i is s ne; nwhd; nrewrite > (eq_false … ne);//;147 ##| #p b i i0 s;//148 ##| #f s ne; nwhd; nrewrite > (Feq_zero_false … ne);//;149 ##]150 nqed.151 152 nlemma bool_of_false: ∀v,ty. is_false v ty → yields ? (exec_bool_of_val v ty) false.153 #v ty H; nelim H;154 ##[ #i s;//;155 ##| #r r' t;//;156 ##| #s; nwhd; nrewrite > (Feq_zero_true …);//;157 ##]158 nqed.159 160 nlemma expr_lvalue_complete: ∀ge,env,m.124 #ge #env #m #e #ty #sp #l #ofs #tr #P #H @(eval_lvalue_inv_ind … H) 125 [ #id #l #ty #e1 #e2 #e3 #e4 #e5 #e6 destruct; // 126 | #id #sp #l #ty #e1 #e2 #e3 #e4 #e5 #e6 #e7 destruct; // 127 | #e #ty #sp #l #ofs #tr #H #e1 #e2 #e3 #e4 #e5 destruct; // 128 | #e #id #ty #sp #l #ofs #id' #fs #d #tr #H #e1 #e2 (* bogus? *) #_ #e3 #e4 #e5 #e6 #e7 destruct; // 129 | #e #id #ty #sp #l #ofs #id' #fs #tr #H #e1 (* bogus? *) #_ #e2 #e3 #e4 #e5 #e6 destruct; // 130 ] qed. 131 132 lemma bool_of_val_3_complete : ∀v,ty,r. bool_of_val v ty r → ∃b. r = of_bool b ∧ yields ? (exec_bool_of_val v ty) b. 133 #v #ty #r #H elim H; #v #t #H' elim H'; 134 [ #i #is #s #ne %{ true} % //; whd; >(eq_false … ne) //; 135 | #p #b #i #i0 #s %{ true} % // 136 | #f #s #ne %{ true} % //; whd; >(Feq_zero_false … ne) //; 137 | #i #s %{ false} % //; 138 | #r #r' #t %{ false} % //; 139 | #s %{ false} % //; whd; >(Feq_zero_true …) //; 140 ] 141 qed. 142 143 lemma bool_of_true: ∀v,ty. is_true v ty → yields ? (exec_bool_of_val v ty) true. 144 #v #ty #H elim H; 145 [ #i #is #s #ne whd; >(eq_false … ne) //; 146 | #p #b #i #i0 #s // 147 | #f #s #ne whd; >(Feq_zero_false … ne) //; 148 ] 149 qed. 150 151 lemma bool_of_false: ∀v,ty. is_false v ty → yields ? (exec_bool_of_val v ty) false. 152 #v #ty #H elim H; 153 [ #i #s //; 154 | #r #r' #t //; 155 | #s whd; >(Feq_zero_true …) //; 156 ] 157 qed. 158 159 lemma expr_lvalue_complete: ∀ge,env,m. 161 160 (∀e,v,tr. eval_expr ge env m e v tr → yields ? (exec_expr ge env m e) (〈v,tr〉)) ∧ 162 161 (∀e,sp,l,off,tr. eval_lvalue ge env m e sp l off tr → yields ? (exec_lvalue ge env m e) (〈〈〈sp,l〉,off〉,tr〉)). 163 #ge env m;164 napply(combined_expr_lvalue_ind ge env m162 #ge #env #m 163 @(combined_expr_lvalue_ind ge env m 165 164 (λe,v,tr,H. yields ? (exec_expr ge env m e) (〈v,tr〉)) 166 165 (λe,sp,l,off,tr,H. yields ? (exec_lvalue ge env m e) (〈〈〈sp,l〉,off〉,tr〉))); 167 ##[ #i ty; napply refl; 168 ##| #f ty; napply refl; 169 ##| #e ty sp l off v tr H1 H2; napply (lvalue_expr … H1); 170 ##[ #id ##| #e' ##| #e' id ##] #H3;171 nwhd in ⊢ (??%?);172 nrewrite > (yields_eq ??? H3);173 nwhd in ⊢ (??%?); nrewrite > H2; napply refl;174 ##| #e ty sp l off tr H1 H2; nwhd in ⊢ (??%?);175 nrewrite > (yields_eq ??? H2);176 napply refl;177 ##| #ty' ty; napply refl; 178 ##| #op e ty v1 v tr H1 H2 H3; nwhd in ⊢ (??%?);179 nrewrite > (yields_eq ??? H3);180 nwhd in ⊢ (??%?); nrewrite > H2; napply refl;181 ##| #op e1 e2 ty v1 v2 v tr1 tr2 H1 H2 e3 H4 H5; nwhd in ⊢ (??%?);182 nrewrite > (yields_eq ??? H4); nwhd in ⊢ (??%?);183 nrewrite > (yields_eq ??? H5); nwhd in ⊢ (??%?);184 nrewrite > e3; napply refl;185 ##| #e1 e2 e3 ty v1 v2 tr1 tr2 H1 H2 H3 H4 H5; nwhd in ⊢ (??%?);186 nrewrite > (yields_eq ??? H4); nwhd in ⊢ (??%?);187 nrewrite > (yields_eq ??? (bool_of_true ?? H2));188 nrewrite > (yields_eq ??? H5);189 napply refl;190 ##| #e1 e2 e3 ty v1 v2 tr1 tr2 H1 H2 H3 H4 H5; nwhd in ⊢ (??%?);191 nrewrite > (yields_eq ??? H4); nwhd in ⊢ (??%?);192 nrewrite > (yields_eq ??? (bool_of_false ?? H2));193 nrewrite > (yields_eq ??? H5);194 napply refl;195 ##| #e1 e2 ty v1 tr H1 H2 H3; nwhd in ⊢ (??%?);196 nrewrite > (yields_eq ??? H3); nwhd in ⊢ (??%?);197 nrewrite > (yields_eq ??? (bool_of_true ?? H2));198 napply refl;199 ##| #e1 e2 ty v1 v2 v tr1 tr2 H1 H2 H3 H4 H5 H6; nwhd in ⊢ (??%?);200 nrewrite > (yields_eq ??? H5); nwhd in ⊢ (??%?);201 nrewrite > (yields_eq ??? (bool_of_false ?? H2));202 nrewrite > (yields_eq ??? H6); nwhd in ⊢ (??%?);203 nelim (bool_of_val_3_complete … H4); #b; *; #evb Hb;204 nrewrite > (yields_eq ??? Hb); nwhd in ⊢ (??%?); nrewrite < evb;205 napply refl;206 ##| #e1 e2 ty v1 tr H1 H2 H3; nwhd in ⊢ (??%?);207 nrewrite > (yields_eq ??? H3); nwhd in ⊢ (??%?);208 nrewrite > (yields_eq ??? (bool_of_false ?? H2));209 napply refl;210 ##| #e1 e2 ty v1 v2 v tr1 tr2 H1 H2 H3 H4 H5 H6; nwhd in ⊢ (??%?);211 nrewrite > (yields_eq ??? H5); nwhd in ⊢ (??%?);212 nrewrite > (yields_eq ??? (bool_of_true ?? H2));213 nrewrite > (yields_eq ??? H6); nwhd in ⊢ (??%?);214 nelim (bool_of_val_3_complete … H4); #b; *; #evb Hb;215 nrewrite > (yields_eq ??? Hb); nwhd in ⊢ (??%?); nrewrite < evb;216 napply refl;217 ##| #e ty ty' v1 v tr H1 H2 H3; nwhd in ⊢ (??%?);218 nrewrite > (yields_eq ??? H3); nwhd in ⊢ (??%?);219 nrewrite > (yields_eq ??? (cast_complete … H2));220 napply refl;221 ##| #e ty v l tr H1 H2; nwhd in ⊢ (??%?);222 nrewrite > (yields_eq ??? H2); nwhd in ⊢ (??%?);223 napply refl;166 [ #i #ty @refl 167 | #f #ty @refl 168 | #e #ty #sp #l #off #v #tr #H1 #H2 @(lvalue_expr … H1) 169 [ #id | #e' | #e' #id ] #H3 170 whd in ⊢ (??%?); 171 >(yields_eq ??? H3) 172 whd in ⊢ (??%?); >H2 @refl 173 | #e #ty #sp #l #off #tr #H1 #H2 whd in ⊢ (??%?); 174 >(yields_eq ??? H2) 175 @refl 176 | #ty' #ty @refl 177 | #op #e #ty #v1 #v #tr #H1 #H2 #H3 whd in ⊢ (??%?); 178 >(yields_eq ??? H3) 179 whd in ⊢ (??%?); >H2 @refl 180 | #op #e1 #e2 #ty #v1 #v2 #v #tr1 #tr2 #H1 #H2 #e3 #H4 #H5 whd in ⊢ (??%?); 181 >(yields_eq ??? H4) whd in ⊢ (??%?); 182 >(yields_eq ??? H5) whd in ⊢ (??%?); 183 >e3 @refl 184 | #e1 #e2 #e3 #ty #v1 #v2 #tr1 #tr2 #H1 #H2 #H3 #H4 #H5 whd in ⊢ (??%?); 185 >(yields_eq ??? H4) whd in ⊢ (??%?); 186 >(yields_eq ??? (bool_of_true ?? H2)) 187 >(yields_eq ??? H5) 188 @refl 189 | #e1 #e2 #e3 #ty #v1 #v2 #tr1 #tr2 #H1 #H2 #H3 #H4 #H5 whd in ⊢ (??%?); 190 >(yields_eq ??? H4) whd in ⊢ (??%?); 191 >(yields_eq ??? (bool_of_false ?? H2)) 192 >(yields_eq ??? H5) 193 @refl 194 | #e1 #e2 #ty #v1 #tr #H1 #H2 #H3 whd in ⊢ (??%?); 195 >(yields_eq ??? H3) whd in ⊢ (??%?); 196 >(yields_eq ??? (bool_of_true ?? H2)) 197 @refl 198 | #e1 #e2 #ty #v1 #v2 #v #tr1 #tr2 #H1 #H2 #H3 #H4 #H5 #H6 whd in ⊢ (??%?); 199 >(yields_eq ??? H5) whd in ⊢ (??%?); 200 >(yields_eq ??? (bool_of_false ?? H2)) 201 >(yields_eq ??? H6) whd in ⊢ (??%?); 202 elim (bool_of_val_3_complete … H4); #b *; #evb #Hb 203 >(yields_eq ??? Hb) whd in ⊢ (??%?); <evb 204 @refl 205 | #e1 #e2 #ty #v1 #tr #H1 #H2 #H3 whd in ⊢ (??%?); 206 >(yields_eq ??? H3) whd in ⊢ (??%?); 207 >(yields_eq ??? (bool_of_false ?? H2)) 208 @refl 209 | #e1 #e2 #ty #v1 #v2 #v #tr1 #tr2 #H1 #H2 #H3 #H4 #H5 #H6 whd in ⊢ (??%?); 210 >(yields_eq ??? H5) whd in ⊢ (??%?); 211 >(yields_eq ??? (bool_of_true ?? H2)) 212 >(yields_eq ??? H6) whd in ⊢ (??%?); 213 elim (bool_of_val_3_complete … H4); #b *; #evb #Hb 214 >(yields_eq ??? Hb) whd in ⊢ (??%?); <evb 215 @refl 216 | #e #ty #ty' #v1 #v #tr #H1 #H2 #H3 whd in ⊢ (??%?); 217 >(yields_eq ??? H3) whd in ⊢ (??%?); 218 >(yields_eq ??? (cast_complete … H2)) 219 @refl 220 | #e #ty #v #l #tr #H1 #H2 whd in ⊢ (??%?); 221 >(yields_eq ??? H2) whd in ⊢ (??%?); 222 @refl 224 223 225 224 (* lvalues *) 226 ##| #id l ty e1; nwhd in ⊢ (??%?); nrewrite > e1; napply refl; 227 ##| #id sp l ty e1 e2; nwhd in ⊢ (??%?); nrewrite > e1; 228 nrewrite > e2; napply refl;229 ##| #e ty sp l ofs tr H1 H2; nwhd in ⊢ (??%?);230 nrewrite > (yields_eq ??? H2);231 napply refl;232 ##| #e i ty sp l ofs id fList delta tr H1 H2 H3 H4; ncases e in H2 H4 ⊢ %;233 #e' ty' H2; nwhd in H2:(??%?); nrewrite > H2; #H4; nwhd in ⊢ (??%?);234 nrewrite > (yields_eq ??? H4); nwhd in ⊢ (??%?);235 nrewrite > H3; napply refl;236 ##| #e i ty sp l ofs id fList tr; ncases e; #e' ty' H1 H2; 237 nwhd in H2:(??%?); nrewrite > H2; #H3; nwhd in ⊢ (??%?);238 nrewrite > (yields_eq ??? H3); napply refl;239 ##] nqed.240 241 ntheorem expr_complete: ∀ge,env,m.225 | #id #l #ty #e1 whd in ⊢ (??%?); >e1 @refl 226 | #id #sp #l #ty #e1 #e2 whd in ⊢ (??%?); >e1 227 >e2 @refl 228 | #e #ty #sp #l #ofs #tr #H1 #H2 whd in ⊢ (??%?); 229 >(yields_eq ??? H2) 230 @refl 231 | #e #i #ty #sp #l #ofs #id #fList #delta #tr #H1 #H2 #H3 #H4 cases e in H2 H4 ⊢ %; 232 #e' #ty' #H2 whd in H2:(??%?); >H2 #H4 whd in ⊢ (??%?); 233 >(yields_eq ??? H4) whd in ⊢ (??%?); 234 >H3 @refl 235 | #e #i #ty #sp #l #ofs #id #fList #tr cases e; #e' #ty' #H1 #H2 236 whd in H2:(??%?); >H2 #H3 whd in ⊢ (??%?); 237 >(yields_eq ??? H3) @refl 238 ] qed. 239 240 theorem expr_complete: ∀ge,env,m. 242 241 ∀e,v,tr. eval_expr ge env m e v tr → yields ? (exec_expr ge env m e) (〈v,tr〉). 243 #ge env m; nelim (expr_lvalue_complete ge env m); /2/; nqed.244 245 ntheorem exprlist_complete: ∀ge,env,m,es,vs,tr.242 #ge #env #m elim (expr_lvalue_complete ge env m); /2/; qed. 243 244 theorem exprlist_complete: ∀ge,env,m,es,vs,tr. 246 245 eval_exprlist ge env m es vs tr → yields ? (exec_exprlist ge env m es) (〈vs,tr〉). 247 #ge env m es vs tr H; nelim H;248 ##[ napply refl; 249 ##| #e et v vt tr trt H1 H2 H3; nwhd in ⊢ (??%?);250 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);251 nrewrite > (yields_eq ??? H3);252 napply refl;253 ##] nqed.254 255 ntheorem lvalue_complete: ∀ge,env,m.246 #ge #env #m #es #vs #tr #H elim H; 247 [ @refl 248 | #e #et #v #vt #tr #trt #H1 #H2 #H3 whd in ⊢ (??%?); 249 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 250 >(yields_eq ??? H3) 251 @refl 252 ] qed. 253 254 theorem lvalue_complete: ∀ge,env,m. 256 255 ∀e,sp,l,off,tr. eval_lvalue ge env m e sp l off tr → yields ? (exec_lvalue ge env m e) (〈〈〈sp,l〉,off〉,tr〉). 257 #ge env m; nelim (expr_lvalue_complete ge env m); /2/; nqed.258 259 nlet rec P_typelist (P:type → Prop) (l:typelist) on l : Prop ≝256 #ge #env #m elim (expr_lvalue_complete ge env m); /2/; qed. 257 258 let rec P_typelist (P:type → Prop) (l:typelist) on l : Prop ≝ 260 259 match l with 261 260 [ Tnil ⇒ True … … 263 262 ]. 264 263 265 nlet rec type_ind2l264 let rec type_ind2l 266 265 (P:type → Prop) (Q:typelist → Prop) 267 266 (vo:P Tvoid) … … 308 307 ]. 309 308 310 nlemma assert_type_eq_true: ∀t. ∃p.assert_type_eq t t = OK ? p.311 #t ; nwhd in ⊢ (??(λ_.??%?)); ncases (type_eq_dec t t); #E;312 ##[ @ E;//313 ##| napply False_ind; napply (absurd ?? E);//314 ##] nqed.315 316 nlemma alloc_vars_complete: ∀env,m,l,env',m'.309 lemma assert_type_eq_true: ∀t. ∃p.assert_type_eq t t = OK ? p. 310 #t whd in ⊢ (??(λ_.??%?)); cases (type_eq_dec t t); #E 311 [ %{ E} // 312 | @False_ind @(absurd ?? E) // 313 ] qed. 314 315 lemma alloc_vars_complete: ∀env,m,l,env',m'. 317 316 alloc_variables env m l env' m' → 318 ∃p.exec_alloc_variables env m l = sig_intro ?? (Some ? 〈env', m'〉) p. 319 #env m l env' m' H; nelim H; 320 ##[ #env'' m''; @ ?; nwhd; //; 321 ##| #env1 m1 id ty l1 m2 loc m3 env2 H1 H2 H3; 322 nwhd in H1:(??%?) ⊢ (??(λ_.??%?)); 323 ndestruct (H1); 324 nelim H3; #p3 e3; nrewrite > e3; nwhd in ⊢ (??(λ_.??%?)); @ ?; //; 325 ##] nqed. 326 327 nlemma bind_params_complete: ∀e,m,params,vs,m2. 317 exec_alloc_variables env m l = 〈env', m'〉. 318 #env #m #l #env' #m' #H elim H; 319 [ #env'' #m'' % 320 | #env1 #m1 #id #ty #l1 #m2 #loc #m3 #env2 #H1 #H2 #H3 321 < H3 whd in H1:(??%?) ⊢ (??%?) 322 destruct (H1) // 323 ] qed. 324 325 lemma bind_params_complete: ∀e,m,params,vs,m2. 328 326 bind_parameters e m params vs m2 → 329 yields_sig ?? (exec_bind_parameters e m params vs) m2. 330 #e m params vs m2 H; nelim H; 331 ##[ //; 332 ##| #env1 m1 id ty l v tl loc m2 m3 H1 H2 H3 H4; 333 napply remove_res_sig; 334 nrewrite > H1; nwhd in ⊢ (??%?); 335 nrewrite > H2; nwhd in ⊢ (??%?); 336 nelim (yields_sig_eq ???? H4); #p4 e4; nrewrite > e4; 337 napply refl; 338 ##] nqed. 339 340 nlemma eventval_match_complete': ∀ev,ty,v. 341 eventval_match ev ty v → yields_sig ?? (check_eventval' v ty) ev. 342 #ev ty v H; nelim H; //; nqed. 343 344 nlemma eventval_list_match_complete: ∀vs,tys,evs. 345 eventval_list_match evs tys vs → yields_sig ?? (check_eventval_list vs tys) evs. 346 #vs tys evs H; nelim H; 347 ##[ // 348 ##| #e etl ty tytl v vtl H1 H2 H3; napply remove_res_sig; 349 nelim (yields_sig_eq ???? (eventval_match_complete' … H1)); #p1 e1; nrewrite > e1; nwhd in ⊢ (??%?); 350 nelim (yields_sig_eq ???? H3); #p3 e3; nrewrite > e3; nwhd in ⊢ (??%?); 351 napply refl; 352 ##] nqed. 353 354 nlemma dec_true: ∀P:Prop.∀f:P + ¬P.∀p:P.∀Q:(P + ¬P) → Prop. (∀p'.Q (inl ?? p')) → Q f. 355 #P f p Q H; ncases f; 356 ##[ napply H 357 ##| #np; napply False_ind; napply (absurd ? p np); 358 ##] nqed. 359 360 nlemma dec_false: ∀P:Prop.∀f:P + ¬P.∀p:¬P.∀Q:(P + ¬P) → Prop. (∀p'.Q (inr ?? p')) → Q f. 361 #P f p Q H; ncases f; 362 ##[ #np; napply False_ind; napply (absurd ? np p); 363 ##| napply H 364 ##] nqed. 365 366 ntheorem step_complete: ∀ge,s,tr,s'. 327 yields ? (exec_bind_parameters e m params vs) m2. 328 #e #m #params #vs #m2 #H elim H; 329 [ //; 330 | #env1 #m1 #id #ty #l #v #tl #loc #m2 #m3 #H1 #H2 #H3 #H4 331 whd in ⊢ (??%?) 332 >H1 whd in ⊢ (??%?); 333 >H2 whd in ⊢ (??%?); 334 @H4 335 ] qed. 336 337 lemma eventval_match_complete': ∀ev,ty,v. 338 eventval_match ev ty v → yields ? (check_eventval' v ty) ev. 339 #ev #ty #v #H elim H; //; qed. 340 341 lemma eventval_list_match_complete: ∀vs,tys,evs. 342 eventval_list_match evs tys vs → yields ? (check_eventval_list vs tys) evs. 343 #vs #tys #evs #H elim H; 344 [ // 345 | #e #etl #ty #tytl #v #vtl #H1 #H2 #H3 whd in ⊢ (??%?) 346 >(yields_eq ??? (eventval_match_complete' … H1)) whd in ⊢ (??%?) 347 >(yields_eq ??? H3) whd in ⊢ (??%?) // 348 ] qed. 349 350 lemma dec_true: ∀P:Prop.∀f:P + ¬P.∀p:P.∀Q:(P + ¬P) → Prop. (∀p'.Q (inl ?? p')) → Q f. 351 #P #f #p #Q #H cases f; 352 [ @H 353 | #np @False_ind @(absurd ? p np) 354 ] qed. 355 356 lemma dec_false: ∀P:Prop.∀f:P + ¬P.∀p:¬P.∀Q:(P + ¬P) → Prop. (∀p'.Q (inr ?? p')) → Q f. 357 #P #f #p #Q #H cases f; 358 [ #np @False_ind @(absurd ? np p) 359 | @H 360 ] qed. 361 362 theorem step_complete: ∀ge,s,tr,s'. 367 363 step ge s tr s' → yieldsIO ? (exec_step ge s) 〈tr,s'〉. 368 #ge s tr s' H; nelim H;369 ##[ #f e e1 k e2 m sp loc ofs v m' tr1 tr2 H1 H2 H3; nwhd in ⊢ (??%?);370 nrewrite > (yields_eq ??? (lvalue_complete … H1)); nwhd in ⊢ (??%?);371 nrewrite > (yields_eq ??? (expr_complete … H2)); nwhd in ⊢ (??%?);372 nrewrite > H3; napply refl;373 ##| #f e eargs k ef m vf vargs f' tr1 tr2 H1 H2 H3 H4; nwhd in ⊢ (??%?);374 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);375 nrewrite > (yields_eq ??? (exprlist_complete … H2)); nwhd in ⊢ (??%?);376 nrewrite > H3; nwhd in ⊢ (??%?);377 nrewrite > H4; nelim (assert_type_eq_true (fun_typeof e)); #pty ety; nrewrite > ety;378 napply refl;379 ##| #f el ef eargs k env m sp loc ofs vf vargs f' tr1 tr2 tr3 H1 H2 H3 H4 H5; nwhd in ⊢ (??%?);380 nrewrite > (yields_eq ??? (expr_complete … H2)); nwhd in ⊢ (??%?);381 nrewrite > (yields_eq ??? (exprlist_complete … H3)); nwhd in ⊢ (??%?);382 nrewrite > H4; nwhd in ⊢ (??%?);383 nrewrite > H5; nelim (assert_type_eq_true (fun_typeof ef)); #pty ety; nrewrite > ety;384 nwhd in ⊢ (??%?);385 nrewrite > (yields_eq ??? (lvalue_complete … H1)); nwhd in ⊢ (??%?);386 napply refl;387 ##| #f s1 s2 k env m; napplyrefl388 ##| ##5,6,7: #f s k env m; napplyrefl389 ##| #f e s1 s2 k env m v tr H1 H2; nwhd in ⊢ (??%?);390 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);391 nrewrite > (yields_eq ??? (bool_of_true ?? H2));392 napplyrefl393 ##| #f e s1 s2 k env m v tr H1 H2; nwhd in ⊢ (??%?);394 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);395 nrewrite > (yields_eq ??? (bool_of_false ?? H2));396 napplyrefl397 ##| #f e s k env m v tr H1 H2; nwhd in ⊢ (??%?);398 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);399 nrewrite > (yields_eq ??? (bool_of_false ?? H2));400 napplyrefl401 ##| #f e s k env m v tr H1 H2; nwhd in ⊢ (??%?);402 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);403 nrewrite > (yields_eq ??? (bool_of_true ?? H2));404 napplyrefl405 ##| #f s1 e s2 k env m H; ncases H; #es1; nrewrite > es1; napply refl; 406 ##| ##13,14: #f e s k env m; napplyrefl407 ##| #f s1 e s2 k env m v tr; *; #es1; nrewrite > es1; #H1 H2; nwhd in ⊢ (??%?);408 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);409 nrewrite > (yields_eq ??? (bool_of_false ?? H2));410 napplyrefl411 ##| #f s1 e s2 k env m v tr; *; #es1; nrewrite > es1; #H1 H2; nwhd in ⊢ (??%?);412 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);413 nrewrite > (yields_eq ??? (bool_of_true ?? H2));414 napplyrefl415 ##| #f e s k env m; napply refl; 416 ##| #f s1 e s2 s3 k env m nskip; nwhd in ⊢ (??%?); ncases (is_Sskip s1);417 ##[ #H; napply False_ind; napply(absurd ? H nskip)418 ##| #H; nwhd in ⊢ (??%?); napply refl ##]419 ##| #f e s1 s2 k env m v tr H1 H2; nwhd in ⊢ (??%?);420 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);421 nrewrite > (yields_eq ??? (bool_of_false ?? H2));422 napply refl;423 ##| #f e s1 s2 k env m v tr H1 H2; nwhd in ⊢ (??%?);424 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);425 nrewrite > (yields_eq ??? (bool_of_true ?? H2));426 napply refl;427 ##| #f s1 e s2 s3 k env m; *; #es1; nrewrite > es1; napply refl; 428 ##| ##22,23: #f e s1 s2 k env m; napply refl; 429 ##| #f k env m H; nwhd in ⊢ (??%?); nrewrite > H; napply refl; 430 ##| #f e k env m v tr H1 H2; nwhd in ⊢ (??%?);431 napply (dec_false ? (type_eq_dec (fn_return f) Tvoid) H1); #pf';432 nwhd in ⊢ (??%?);433 nrewrite > (yields_eq ??? (expr_complete … H2)); nwhd in ⊢ (??%?);434 napply refl;435 ##| #f k env m; ncases k;436 ##[ #H1 H2; nwhd in ⊢ (??%?); nrewrite > H2; napply refl;437 ##| #s' k'; nwhd in ⊢ (% → ?); *;438 ##| ##3,4: #e' s' k'; nwhd in ⊢ (% → ?); *;439 ##| ##5,6: #e' s1' s2' k'; nwhd in ⊢ (% → ?); *;440 ##| #k'; nwhd in ⊢ (% → ?); *;441 ##| #r f' env' k' H1 H2; nwhd in ⊢ (??%?); nrewrite > H2; napplyrefl442 ##]443 ##| #f e s k env m i tr H1; nwhd in ⊢ (??%?);444 nrewrite > (yields_eq ??? (expr_complete … H1)); nwhd in ⊢ (??%?);445 napplyrefl446 ##| #f s k env m; *; #es; nrewrite > es; napply refl; 447 ##| #f k env m; napplyrefl448 ##| #f l s k env m; napplyrefl449 ##| #f l k env m s k' H1; nwhd in ⊢ (??%?); nrewrite > H1; napply refl; 450 ##| #f args k m1 env m2 m3 H1 H2; nwhd in ⊢ (??%?);451 nelim (alloc_vars_complete … H1); #p1 e1; nrewrite > e1; nwhd in ⊢ (??%?);452 nelim (yields_sig_eq ???? (bind_params_complete … H2)); #p2 e2; nrewrite > e2;453 napply refl;454 ##| #id tys rty args k m rv tr H; nwhd in ⊢ (??%?);455 ninversion H; #f' args' rv' eargs erv H1 H2 e1 e2 e3 e4; nrewrite < e1 in H1 H2;456 #H1 H2;457 nelim (yields_sig_eq ???? (eventval_list_match_complete … H1)); #p1 e1; nrewrite > e1; nwhd in ⊢ (??%?);458 nwhd; ninversion H2; #x e5 e6 e7; @ x; nwhd in ⊢ (??%?);459 napplyrefl460 ##| #v f env k m; napplyrefl461 ##| #v f env k m1 m2 sp loc ofs ty H; nwhd in ⊢ (??%?);462 nrewrite > H; napplyrefl463 ##| #f l s k env m; napplyrefl464 ##] nqed.364 #ge #s #tr #s' #H elim H; 365 [ #f #e #e1 #k #e2 #m #sp #loc #ofs #v #m' #tr1 #tr2 #H1 #H2 #H3 whd in ⊢ (??%?); 366 >(yields_eq ??? (lvalue_complete … H1)) whd in ⊢ (??%?); 367 >(yields_eq ??? (expr_complete … H2)) whd in ⊢ (??%?); 368 >H3 @refl 369 | #f #e #eargs #k #ef #m #vf #vargs #f' #tr1 #tr2 #H1 #H2 #H3 #H4 whd in ⊢ (??%?); 370 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 371 >(yields_eq ??? (exprlist_complete … H2)) whd in ⊢ (??%?); 372 >H3 whd in ⊢ (??%?); 373 >H4 elim (assert_type_eq_true (fun_typeof e)); #pty #ety >ety 374 @refl 375 | #f #el #ef #eargs #k #env #m #sp #loc #ofs #vf #vargs #f' #tr1 #tr2 #tr3 #H1 #H2 #H3 #H4 #H5 whd in ⊢ (??%?); 376 >(yields_eq ??? (expr_complete … H2)) whd in ⊢ (??%?); 377 >(yields_eq ??? (exprlist_complete … H3)) whd in ⊢ (??%?); 378 >H4 whd in ⊢ (??%?); 379 >H5 elim (assert_type_eq_true (fun_typeof ef)); #pty #ety >ety 380 whd in ⊢ (??%?); 381 >(yields_eq ??? (lvalue_complete … H1)) whd in ⊢ (??%?); 382 @refl 383 | #f #s1 #s2 #k #env #m @refl 384 | 5,6,7: #f #s #k #env #m @refl 385 | #f #e #s1 #s2 #k #env #m #v #tr #H1 #H2 whd in ⊢ (??%?); 386 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 387 >(yields_eq ??? (bool_of_true ?? H2)) 388 @refl 389 | #f #e #s1 #s2 #k #env #m #v #tr #H1 #H2 whd in ⊢ (??%?); 390 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 391 >(yields_eq ??? (bool_of_false ?? H2)) 392 @refl 393 | #f #e #s #k #env #m #v #tr #H1 #H2 whd in ⊢ (??%?); 394 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 395 >(yields_eq ??? (bool_of_false ?? H2)) 396 @refl 397 | #f #e #s #k #env #m #v #tr #H1 #H2 whd in ⊢ (??%?); 398 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 399 >(yields_eq ??? (bool_of_true ?? H2)) 400 @refl 401 | #f #s1 #e #s2 #k #env #m #H cases H; #es1 >es1 @refl 402 | 13,14: #f #e #s #k #env #m @refl 403 | #f #s1 #e #s2 #k #env #m #v #tr *; #es1 >es1 #H1 #H2 whd in ⊢ (??%?); 404 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 405 >(yields_eq ??? (bool_of_false ?? H2)) 406 @refl 407 | #f #s1 #e #s2 #k #env #m #v #tr *; #es1 >es1 #H1 #H2 whd in ⊢ (??%?); 408 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 409 >(yields_eq ??? (bool_of_true ?? H2)) 410 @refl 411 | #f #e #s #k #env #m @refl 412 | #f #s1 #e #s2 #s3 #k #env #m #nskip whd in ⊢ (??%?); cases (is_Sskip s1); 413 [ #H @False_ind @(absurd ? H nskip) 414 | #H whd in ⊢ (??%?); @refl ] 415 | #f #e #s1 #s2 #k #env #m #v #tr #H1 #H2 whd in ⊢ (??%?); 416 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 417 >(yields_eq ??? (bool_of_false ?? H2)) 418 @refl 419 | #f #e #s1 #s2 #k #env #m #v #tr #H1 #H2 whd in ⊢ (??%?); 420 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 421 >(yields_eq ??? (bool_of_true ?? H2)) 422 @refl 423 | #f #s1 #e #s2 #s3 #k #env #m *; #es1 >es1 @refl 424 | 22,23: #f #e #s1 #s2 #k #env #m @refl 425 | #f #k #env #m #H whd in ⊢ (??%?); >H @refl 426 | #f #e #k #env #m #v #tr #H1 #H2 whd in ⊢ (??%?); 427 @(dec_false ? (type_eq_dec (fn_return f) Tvoid) H1) #pf' 428 whd in ⊢ (??%?); 429 >(yields_eq ??? (expr_complete … H2)) whd in ⊢ (??%?); 430 @refl 431 | #f #k #env #m cases k; 432 [ #H1 #H2 whd in ⊢ (??%?); >H2 @refl 433 | #s' #k' whd in ⊢ (% → ?); *; 434 | 3,4: #e' #s' #k' whd in ⊢ (% → ?); *; 435 | 5,6: #e' #s1' #s2' #k' whd in ⊢ (% → ?); *; 436 | #k' whd in ⊢ (% → ?); *; 437 | #r #f' #env' #k' #H1 #H2 whd in ⊢ (??%?); >H2 @refl 438 ] 439 | #f #e #s #k #env #m #i #tr #H1 whd in ⊢ (??%?); 440 >(yields_eq ??? (expr_complete … H1)) whd in ⊢ (??%?); 441 @refl 442 | #f #s #k #env #m *; #es >es @refl 443 | #f #k #env #m @refl 444 | #f #l #s #k #env #m @refl 445 | #f #l #k #env #m #s #k' #H1 whd in ⊢ (??%?); >H1 @refl 446 | #f #args #k #m1 #env #m2 #m3 #H1 #H2 whd in ⊢ (??%?); 447 >(alloc_vars_complete … H1) whd in ⊢ (??%?); 448 >(yields_eq ??? (bind_params_complete … H2)) 449 // 450 | #id #tys #rty #args #k #m #rv #tr #H whd in ⊢ (??%?); 451 inversion H; #f' #args' #rv' #eargs #erv #H1 #H2 #e1 #e2 #e3 #e4 <e1 in H1 H2 452 #H1 #H2 453 >(yields_eq ??? (eventval_list_match_complete … H1)) whd in ⊢ (??%?); 454 whd; inversion H2; #x #e5 #e6 #e7 %{ x} whd in ⊢ (??%?); 455 @refl 456 | #v #f #env #k #m @refl 457 | #v #f #env #k #m1 #m2 #sp #loc #ofs #ty #H whd in ⊢ (??%?); 458 >H @refl 459 | #f #l #s #k #env #m @refl 460 ] qed. -
Deliverables/D3.1/C-semantics/CexecEquiv.ma
r485 r487 3 3 include "extralib.ma". 4 4 5 include "Plogic/jmeq.ma". 6 include "Plogic/connectives.ma". 5 include "basics/jmeq.ma". 7 6 8 7 (* A "single execution" - where all of the input values are made explicit. *) 9 ncoinductive s_execution : Type≝8 coinductive s_execution : Type[0] ≝ 10 9 | se_stop : trace → int → mem → s_execution 11 10 | se_step : trace → state → s_execution → s_execution … … 13 12 | se_interact : ∀o:io_out. (io_in o → execution) → io_in o → s_execution → s_execution. 14 13 15 ncoinductive single_exec_of : execution → s_execution → Prop ≝14 coinductive single_exec_of : execution → s_execution → Prop ≝ 16 15 | seo_stop : ∀tr,r,m. single_exec_of (e_stop tr r m) (se_stop tr r m) 17 16 | seo_step : ∀tr,s,e,se. … … 25 24 (* starting after state s, zero or more steps of execution e reach state s' 26 25 after which comes e'. *) 27 ninductive execution_isteps : trace → state → s_execution → state → s_execution → Prop ≝26 inductive execution_isteps : trace → state → s_execution → state → s_execution → Prop ≝ 28 27 | isteps_none : ∀s,e. execution_isteps E0 s e s e 29 28 | isteps_one : ∀e,e',tr,tr',s,s',s0. … … 34 33 execution_isteps (tr⧺tr') s0 (se_interact o k i (se_step tr s e)) s' e'. 35 34 36 nlemma isteps_trans: ∀tr1,tr2,s1,s2,s3,e1,e2,e3.35 lemma isteps_trans: ∀tr1,tr2,s1,s2,s3,e1,e2,e3. 37 36 execution_isteps tr1 s1 e1 s2 e2 → 38 37 execution_isteps tr2 s2 e2 s3 e3 → 39 38 execution_isteps (tr1⧺tr2) s1 e1 s3 e3. 40 #tr1 tr2 s1 s2 s3 e1 e2 e3 H1; nelim H1;41 ##[ #s e;//;42 ##| #e e' tr tr' s1' s2' s3' H1 H2 H3; 43 nrewrite > (Eapp_assoc …);44 napply isteps_one;45 napply H2; napply H3;46 ##| #e e' o k i s1' s2' s3' tr tr' H1 H2 H3; 47 nrewrite > (Eapp_assoc …);48 napply isteps_interact;39 #tr1 #tr2 #s1 #s2 #s3 #e1 #e2 #e3 #H1 elim H1; 40 [ #s #e //; 41 | #e #e' #tr #tr' #s1' #s2' #s3' #H1 #H2 #H3 42 >(Eapp_assoc …) 43 @isteps_one 44 @H2 @H3 45 | #e #e' #o #k #i #s1' #s2' #s3' #tr #tr' #H1 #H2 #H3 46 >(Eapp_assoc …) 47 @isteps_interact 49 48 /2/ 50 ##] nqed.51 52 nlemma exec_e_step: ∀ge,x,tr,s,e.49 ] qed. 50 51 lemma exec_e_step: ∀ge,x,tr,s,e. 53 52 exec_inf_aux ge x = e_step tr s e → 54 53 exec_inf_aux ge (exec_step ge s) = e. 55 #ge x tr s e;56 nrewrite > (exec_inf_aux_unfold …); ncases x;57 ##[ #o k EXEC; nwhd in EXEC:(??%?); ndestruct58 ##| #y; ncases y; #tr' s'; nwhd in ⊢ (??%? → ?);59 ncases (is_final_state s'); #FINAL EXEC; nwhd in EXEC:(??%?); ndestruct;60 napply refl;61 ##| #EXEC; nwhd in EXEC:(??%?); ndestruct62 ##] nqed.63 64 nlemma exec_e_step_inv: ∀ge,x,tr,s,e.54 #ge #x #tr #s #e 55 >(exec_inf_aux_unfold …) cases x; 56 [ #o #k #EXEC whd in EXEC:(??%?); destruct 57 | #y cases y; #tr' #s' whd in ⊢ (??%? → ?); 58 cases (is_final_state s'); #FINAL #EXEC whd in EXEC:(??%?); destruct; 59 @refl 60 | #EXEC whd in EXEC:(??%?); destruct 61 ] qed. 62 63 lemma exec_e_step_inv: ∀ge,x,tr,s,e. 65 64 exec_inf_aux ge x = e_step tr s e → 66 65 x = Value ??? 〈tr,s〉. 67 #ge x tr s e;68 nrewrite > (exec_inf_aux_unfold …); ncases x;69 ##[ #o k EXEC; nwhd in EXEC:(??%?); ndestruct70 ##| #y; ncases y; #tr' s'; nwhd in ⊢ (??%? → ?);71 ncases (is_final_state s'); #FINAL EXEC; nwhd in EXEC:(??%?); ndestruct;72 napply refl;73 ##| #EXEC; nwhd in EXEC:(??%?); ndestruct74 ##] nqed.75 76 nlemma exec_e_step_inv2: ∀ge,x,tr,s,e.66 #ge #x #tr #s #e 67 >(exec_inf_aux_unfold …) cases x; 68 [ #o #k #EXEC whd in EXEC:(??%?); destruct 69 | #y cases y; #tr' #s' whd in ⊢ (??%? → ?); 70 cases (is_final_state s'); #FINAL #EXEC whd in EXEC:(??%?); destruct; 71 @refl 72 | #EXEC whd in EXEC:(??%?); destruct 73 ] qed. 74 75 lemma exec_e_step_inv2: ∀ge,x,tr,s,e. 77 76 exec_inf_aux ge x = e_step tr s e → 78 77 ¬∃r.final_state s r. 79 #ge x tr s e;80 nrewrite > (exec_inf_aux_unfold …); ncases x;81 ##[ #o k EXEC; nwhd in EXEC:(??%?); ndestruct82 ##| #y; ncases y; #tr' s'; nwhd in ⊢ (??%? → ?);83 ncases (is_final_state s'); #FINAL EXEC; nwhd in EXEC:(??%?); ndestruct;84 napply FINAL;85 ##| #EXEC; nwhd in EXEC:(??%?); ndestruct86 ##] nqed.87 88 ndefinition exec_from : genv → state → s_execution → Prop ≝78 #ge #x #tr #s #e 79 >(exec_inf_aux_unfold …) cases x; 80 [ #o #k #EXEC whd in EXEC:(??%?); destruct 81 | #y cases y; #tr' #s' whd in ⊢ (??%? → ?); 82 cases (is_final_state s'); #FINAL #EXEC whd in EXEC:(??%?); destruct; 83 @FINAL 84 | #EXEC whd in EXEC:(??%?); destruct 85 ] qed. 86 87 definition exec_from : genv → state → s_execution → Prop ≝ 89 88 λge,s,se. single_exec_of (exec_inf_aux ge (exec_step ge s)) se. 90 89 91 nlemma exec_from_step : ∀ge,s,tr,s',e.90 lemma exec_from_step : ∀ge,s,tr,s',e. 92 91 exec_from ge s (se_step tr s' e) → 93 92 exec_step ge s = Value ??? 〈tr,s'〉 ∧ exec_from ge s' e. 94 #ge s0 tr0 s0' e0 H; ninversion H;95 ##[ #tr r m E1 E2; ndestruct96 ##| #tr s e se H1 H2 E; ndestruct (E);97 nrewrite > (exec_e_step_inv … H2);98 nrewrite < (exec_e_step … H2) in H1; #H1; @;//99 ##| #_; #E; ndestruct100 ##| #o k i se H1 H2 E; ndestruct101 ##] nqed.102 103 nlemma exec_from_interact : ∀ge,s,o,k,i,tr,s',e.93 #ge #s0 #tr0 #s0' #e0 #H inversion H; 94 [ #tr #r #m #E1 #E2 destruct 95 | #tr #s #e #se #H1 #H2 #E destruct (E); 96 >(exec_e_step_inv … H2) 97 <(exec_e_step … H2) in H1 #H1 % // 98 | #_ #E destruct 99 | #o #k #i #se #H1 #H2 #E destruct 100 ] qed. 101 102 lemma exec_from_interact : ∀ge,s,o,k,i,tr,s',e. 104 103 exec_from ge s (se_interact o k i (se_step tr s' e)) → 105 104 step ge s tr s' ∧ 106 105 (*exec_step ge s = Value ??? 〈tr,s'〉 ∧*) exec_from ge s' e. 107 #ge s0 o0 k0 i0 tr0 s0' e0 H; ninversion H;108 ##[ #tr r m E1 E2; ndestruct109 ##| #tr s e se H1 H2 E; ndestruct (E)110 ##| #_; #E; ndestruct111 ##| #o k i se H1 H2 E; ndestruct (E);112 nlapply (exec_step_sound ge s0);113 ncases (exec_step ge s0) in H2 ⊢ %;114 ##[ #o' k'; nrewrite > (exec_inf_aux_unfold …);115 #E' ; nwhd in E':(??%?); ndestruct (E');116 #STEP ;117 ninversion H1;118 ##[ #tr r m E1 E2; ndestruct119 ##| #tr' s' e' se' H2 H3 E2; ndestruct (E2);120 nrewrite < (exec_e_step … H3) in H2; #H2; @;//;121 nlapply (STEP i);122 nrewrite > (exec_e_step_inv … H3);123 #S ; napply S;124 ##| #_; #E; ndestruct125 ##| #o k i se H1 H2 E; ndestruct126 ##]127 ##| #x; ncases x; #tr' s'; nrewrite > (exec_inf_aux_unfold …);128 nwhd in ⊢ (??%? → ?); ncases (is_final_state s');129 #F E; nwhd in E:(??%?); ndestruct130 ##| nrewrite > (exec_inf_aux_unfold …);131 #E' ; nwhd in E':(??%?); ndestruct (E');132 ##]133 ##] nqed.134 135 nlemma exec_from_interact_stop : ∀ge,s,o,k,i,tr,r,m.106 #ge #s0 #o0 #k0 #i0 #tr0 #s0' #e0 #H inversion H; 107 [ #tr #r #m #E1 #E2 destruct 108 | #tr #s #e #se #H1 #H2 #E destruct (E) 109 | #_ #E destruct 110 | #o #k #i #se #H1 #H2 #E destruct (E); 111 lapply (exec_step_sound ge s0); 112 cases (exec_step ge s0) in H2 ⊢ %; 113 [ #o' #k' >(exec_inf_aux_unfold …) 114 #E' whd in E':(??%?); destruct (E'); 115 #STEP 116 inversion H1; 117 [ #tr #r #m #E1 #E2 destruct 118 | #tr' #s' #e' #se' #H2 #H3 #E2 destruct (E2); 119 <(exec_e_step … H3) in H2 #H2 % //; 120 lapply (STEP i); 121 >(exec_e_step_inv … H3) 122 #S @S 123 | #_ #E destruct 124 | #o #k #i #se #H1 #H2 #E destruct 125 ] 126 | #x cases x; #tr' #s' >(exec_inf_aux_unfold …) 127 whd in ⊢ (??%? → ?); cases (is_final_state s'); 128 #F #E whd in E:(??%?); destruct 129 | >(exec_inf_aux_unfold …) 130 #E' whd in E':(??%?); destruct (E'); 131 ] 132 ] qed. 133 134 lemma exec_from_interact_stop : ∀ge,s,o,k,i,tr,r,m. 136 135 exec_from ge s (se_interact o k i (se_stop tr r m)) → 137 136 step ge s tr (Returnstate (Vint r) Kstop m). 138 #ge s0 o0 k0 i0 tr0 s0' e0 H; ninversion H;139 ##[ #tr r m E1 E2; ndestruct140 ##| #tr s e se H1 H2 E; ndestruct (E)141 ##| #_; #E; ndestruct142 ##| #o k i se H1 H2 E; ndestruct (E);143 nlapply (exec_step_sound ge s0);144 nrewrite >(exec_inf_aux_unfold …) in H2;145 ncases (exec_step ge s0);146 ##[ #o' k';147 #E' ; nwhd in E':(??%?); ndestruct (E');148 #STEP ;149 ninversion H1;150 ##[ #tr r m E1 E2; nlapply (STEP i); ndestruct;151 nrewrite >(exec_inf_aux_unfold …) in E1;152 ncases (k' i);153 ##[ #o2 k2 E; nwhd in E:(??%?); ndestruct (E)154 ##| #x; ncases x; #tr2 s2; nwhd in ⊢ (??%? → ?);155 ncases (is_final_state s2);156 ##[ #F; ncases F; #r' FINAL E; nwhd in E:(??%?);157 ndestruct (E);158 ninversion FINAL;159 #r'' m'' E1 E2; ndestruct (E1 E2); //;160 ##| #NF E; nwhd in E:(??%?); ndestruct (E)161 ##]162 ##| #E; nwhd in E:(??%?); ndestruct (E)163 ##]164 ##| #tr s e e' H EXEC E; ndestruct (E)165 ##| #EXEC E; ndestruct (E)166 ##| #o2 k2 i2 e2 H EXEC E; ndestruct (E)167 ##]168 ##| #x; ncases x; #tr s; nwhd in ⊢ (??%? → ?);169 ncases (is_final_state s); #F E; nwhd in E:(??%?); ndestruct (E)170 ##| #E; nwhd in E:(??%?); ndestruct (E)171 ##]172 ##] nqed.137 #ge #s0 #o0 #k0 #i0 #tr0 #s0' #e0 #H inversion H; 138 [ #tr #r #m #E1 #E2 destruct 139 | #tr #s #e #se #H1 #H2 #E destruct (E) 140 | #_ #E destruct 141 | #o #k #i #se #H1 #H2 #E destruct (E); 142 lapply (exec_step_sound ge s0); 143 >(exec_inf_aux_unfold …) in H2; 144 cases (exec_step ge s0); 145 [ #o' #k' 146 #E' whd in E':(??%?); destruct (E'); 147 #STEP 148 inversion H1; 149 [ #tr #r #m #E1 #E2 lapply (STEP i); destruct; 150 >(exec_inf_aux_unfold …) in E1; 151 cases (k' i); 152 [ #o2 #k2 #E whd in E:(??%?); destruct (E) 153 | #x cases x; #tr2 #s2 whd in ⊢ (??%? → ?); 154 cases (is_final_state s2); 155 [ #F cases F; #r' #FINAL #E whd in E:(??%?); 156 destruct (E); 157 inversion FINAL; 158 #r'' #m'' #E1 #E2 destruct (E1 E2); //; 159 | #NF #E whd in E:(??%?); destruct (E) 160 ] 161 | #E whd in E:(??%?); destruct (E) 162 ] 163 | #tr #s #e #e' #H #EXEC #E destruct (E) 164 | #EXEC #E destruct (E) 165 | #o2 #k2 #i2 #e2 #H #EXEC #E destruct (E) 166 ] 167 | #x cases x; #tr #s whd in ⊢ (??%? → ?); 168 cases (is_final_state s); #F #E whd in E:(??%?); destruct (E) 169 | #E whd in E:(??%?); destruct (E) 170 ] 171 ] qed. 173 172 174 173 (* NB: the E0 in the execs are irrelevant. *) 175 nlemma several_steps: ∀ge,tr,e,e',s,s'.174 lemma several_steps: ∀ge,tr,e,e',s,s'. 176 175 execution_isteps tr s e s' e' → 177 176 exec_from ge s e → 178 177 star (mk_transrel … step) ge s tr s' ∧ 179 178 exec_from ge s' e'. 180 #ge tr0 e0 e0' s0 s0' H;181 nelim H;182 ##[ #s e EXEC; @;//;183 ##| #e1 e2 tr1 tr2 s1 s2 s3 STEPS IH EXEC; 184 nelim (exec_from_step … EXEC); #EXEC3 EXEC1;185 nelim (IH EXEC1);186 #STAR12 EXEC2; @;//;187 nlapply (exec_step_sound ge s3);188 nrewrite > EXEC3; #STEP3;189 napply (star_step (mk_transrel ?? step) … STEP3 STAR12);190 napplyrefl191 ##| #e1 e2 o k i s1 s2 s3 tr1 tr2 STEPS IH EXEC; 192 nelim (exec_from_interact … EXEC);193 #STEP3 EXEC1;194 nelim (IH EXEC1); #STAR EXEC2;195 @;196 ##[ napply (star_step (mk_transrel ?? step) … STEP3 STAR);197 napplyrefl198 ##| //199 ##]200 ##] nqed.201 202 ninductive execution_terminates : trace → state → s_execution → int → mem → Prop ≝179 #ge #tr0 #e0 #e0' #s0 #s0' #H 180 elim H; 181 [ #s #e #EXEC % //; 182 | #e1 #e2 #tr1 #tr2 #s1 #s2 #s3 #STEPS #IH #EXEC 183 elim (exec_from_step … EXEC); #EXEC3 #EXEC1 184 elim (IH EXEC1); 185 #STAR12 #EXEC2 % //; 186 lapply (exec_step_sound ge s3); 187 >EXEC3 #STEP3 188 @(star_step (mk_transrel ?? step) … STEP3 STAR12) 189 @refl 190 | #e1 #e2 #o #k #i #s1 #s2 #s3 #tr1 #tr2 #STEPS #IH #EXEC 191 elim (exec_from_interact … EXEC); 192 #STEP3 #EXEC1 193 elim (IH EXEC1); #STAR #EXEC2 194 % 195 [ @(star_step (mk_transrel ?? step) … STEP3 STAR) 196 @refl 197 | // 198 ] 199 ] qed. 200 201 inductive execution_terminates : trace → state → s_execution → int → mem → Prop ≝ 203 202 | terminates : ∀s,s',tr,tr',r,e,m. 204 203 execution_isteps tr s e s' (se_stop tr' r m) → … … 209 208 execution_terminates (tr⧺tr') s (se_step E0 s e) r m. 210 209 211 ncoinductive execution_diverging : s_execution → Prop ≝210 coinductive execution_diverging : s_execution → Prop ≝ 212 211 | diverging_step : ∀s,e. execution_diverging e → execution_diverging (se_step E0 s e). 213 212 214 213 (* Makes a finite number of interactions (including cost labels) before diverging. *) 215 ninductive execution_diverges : trace → state → s_execution → Prop ≝214 inductive execution_diverges : trace → state → s_execution → Prop ≝ 216 215 | diverges_diverging: ∀tr,s,s',e,e'. 217 216 execution_isteps tr s e s' e' → … … 220 219 221 220 (* NB: "reacting" includes hitting a cost label. *) 222 ncoinductive execution_reacting : traceinf → state → s_execution → Prop ≝221 coinductive execution_reacting : traceinf → state → s_execution → Prop ≝ 223 222 | reacting: ∀tr,tr',s,s',e,e'. 224 223 execution_reacting tr' s' e' → … … 227 226 execution_reacting (tr⧻tr') s e. 228 227 229 ninductive execution_reacts : traceinf → state → s_execution → Prop ≝228 inductive execution_reacts : traceinf → state → s_execution → Prop ≝ 230 229 | reacts: ∀tr,s,e. 231 230 execution_reacting tr s e → 232 231 execution_reacts tr s (se_step E0 s e). 233 232 234 ninductive execution_goes_wrong: trace → state → s_execution → state → Prop ≝233 inductive execution_goes_wrong: trace → state → s_execution → state → Prop ≝ 235 234 | go_wrong: ∀tr,s,s',e. 236 235 execution_isteps tr s e s' se_wrong → 237 236 execution_goes_wrong tr s (se_step E0 s e) s'. 238 237 239 nlet corec silent_sound ge s e238 let corec silent_sound ge s e 240 239 (H0:execution_diverging e) 241 240 (EXEC:exec_from ge s e) 242 241 : forever_silent (mk_transrel ?? step) … ge s ≝ ?. 243 ncut (∃s2.∃e2.And (And (execution_diverging e2) (step ge s E0 s2)) (exec_from ge s2 e2));244 ##[ ncases H0 in EXEC ⊢ %; #s1 e1 H1 EXEC; 245 nelim (exec_from_step … EXEC);246 #EXEC0 EXEC1;247 @ s1; @ e1; @; //; @;//;248 nlapply (exec_step_sound ge s); nrewrite > EXEC0; nwhd in ⊢ (% → ?); //;249 ##| *; #s2; *; #e2; *; *; #H2 STEP2 EXEC2; 250 napply (forever_silent_intro (mk_transrel ?? step) … ge s s2 ? (silent_sound ge s2 e2 ??));242 cut (∃s2.∃e2.And (And (execution_diverging e2) (step ge s E0 s2)) (exec_from ge s2 e2)); 243 [ cases H0 in EXEC ⊢ %; #s1 #e1 #H1 #EXEC 244 elim (exec_from_step … EXEC); 245 #EXEC0 #EXEC1 246 %{ s1} %{ e1} % //; % //; 247 lapply (exec_step_sound ge s); >EXEC0 whd in ⊢ (% → ?); #H @H 248 | *; #s2 *; #e2 *; *; #H2 #STEP2 #EXEC2 249 @(forever_silent_intro (mk_transrel ?? step) … ge s s2 ? (silent_sound ge s2 e2 ??)) 251 250 //; 252 ##] nqed.253 254 nlemma final_step: ∀ge,tr,r,m,s.251 ] qed. 252 253 lemma final_step: ∀ge,tr,r,m,s. 255 254 exec_from ge s (se_stop tr r m) → 256 255 step ge s tr (Returnstate (Vint r) Kstop m). 257 #ge tr r m s EXEC;258 nwhd in EXEC;259 ninversion EXEC;260 ##[ #tr' r' m' EXEC' E; ndestruct (E);261 nlapply (exec_step_sound ge s);262 nrewrite >(exec_inf_aux_unfold …) in EXEC';263 ncases (exec_step ge s);264 ##[ #o k EXEC'; nwhd in EXEC':(??%?); ndestruct (EXEC');265 ##| #x; ncases x; #tr'' s'; nwhd in ⊢ (??%? → ?);266 ncases (is_final_state s'); #F; ##[ ##1: ncases F; #r'' FINAL; ##]267 #EXEC' ; nwhd in EXEC':(??%?); ndestruct (EXEC');268 ##| #EXEC'; nwhd in EXEC':(??%?); ndestruct (EXEC');269 ##]270 ninversion FINAL; #r''' m' E1 E2 H; ndestruct (E1 E2);271 napply H;272 ##| #tr' s' e' se' H EXEC' E; ndestruct273 ##| #EXEC' E; ndestruct274 ##| #o k i e H EXEC E; ndestruct275 ##] nqed.276 277 278 nlemma e_stop_inv: ∀ge,x,tr,r,m.256 #ge #tr #r #m #s #EXEC 257 whd in EXEC; 258 inversion EXEC; 259 [ #tr' #r' #m' #EXEC' #E destruct (E); 260 lapply (exec_step_sound ge s); 261 >(exec_inf_aux_unfold …) in EXEC'; 262 cases (exec_step ge s); 263 [ #o #k #EXEC' whd in EXEC':(??%?); destruct (EXEC'); 264 | #x cases x; #tr'' #s' whd in ⊢ (??%? → ?); 265 cases (is_final_state s'); #F [ 1: cases F; #r'' #FINAL ] 266 #EXEC' whd in EXEC':(??%?); destruct (EXEC'); 267 | #EXEC' whd in EXEC':(??%?); destruct (EXEC'); 268 ] 269 inversion FINAL; #r''' #m' #E1 #E2 #H destruct (E1 E2); 270 @H 271 | #tr' #s' #e' #se' #H #EXEC' #E destruct 272 | #EXEC' #E destruct 273 | #o #k #i #e #H #EXEC #E destruct 274 ] qed. 275 276 277 lemma e_stop_inv: ∀ge,x,tr,r,m. 279 278 exec_inf_aux ge x = e_stop tr r m → 280 279 x = Value ??? 〈tr,Returnstate (Vint r) Kstop m〉. 281 #ge x tr r m;282 nrewrite > (exec_inf_aux_unfold …); ncases x;283 ##[ #o k EXEC; nwhd in EXEC:(??%?); ndestruct;284 ##| #z; ncases z; #tr' s'; nwhd in ⊢ (??%? → ?); ncases (is_final_state s');285 ##[ #F; ncases F; #r' FINAL; ncases FINAL; #r'' m' EXEC; nwhd in EXEC:(??%?);286 ndestruct (EXEC); napply refl;287 ##| #F EXEC; nwhd in EXEC:(??%?); ndestruct (EXEC);288 ##]289 ##| #EXEC; nwhd in EXEC:(??%?); ndestruct (EXEC);290 ##] nqed.291 292 nlemma terminates_sound: ∀ge,tr,s,r,m,e.280 #ge #x #tr #r #m 281 >(exec_inf_aux_unfold …) cases x; 282 [ #o #k #EXEC whd in EXEC:(??%?); destruct; 283 | #z cases z; #tr' #s' whd in ⊢ (??%? → ?); cases (is_final_state s'); 284 [ #F cases F; #r' #FINAL cases FINAL; #r'' #m' #EXEC whd in EXEC:(??%?); 285 destruct (EXEC); @refl 286 | #F #EXEC whd in EXEC:(??%?); destruct (EXEC); 287 ] 288 | #EXEC whd in EXEC:(??%?); destruct (EXEC); 289 ] qed. 290 291 lemma terminates_sound: ∀ge,tr,s,r,m,e. 293 292 execution_terminates tr s (se_step E0 s e) r m → 294 293 exec_from ge s e → 295 294 star (mk_transrel … step) ge s tr (Returnstate (Vint r) Kstop m). 296 #ge tr0 s0 r m e0 H; ninversion H;297 ##[ #s s' tr tr' r e m ESTEPS E1 E2 E3 E4 E5 EXEC; 298 ndestruct (E1 E2 E3 E4 E5);299 ncases (several_steps … ESTEPS EXEC); #STARs' EXECs';300 napply (star_right … STARs');301 ##[ ##2: napply (final_step ge tr' r m s' … EXECs');302 ##| ##skip303 ##| napplyrefl304 ##]305 ##| #s s' tr tr' r e m o k i ESTEPS E1 E2 E3 E4 E5 EXEC; 306 ndestruct;307 ncases (several_steps … ESTEPS EXEC); #STARs' EXECs';308 napply (star_right … STARs');309 ##[ napplytr'310 ##| napply (exec_from_interact_stop … EXECs');311 ##| napplyrefl312 ##]313 ##] nqed.314 315 nlet corec reacts_sound ge tr s e295 #ge #tr0 #s0 #r #m #e0 #H inversion H; 296 [ #s #s' #tr #tr' #r #e #m #ESTEPS #E1 #E2 #E3 #E4 #E5 #EXEC 297 destruct (E1 E2 E3 E4 E5); 298 cases (several_steps … ESTEPS EXEC); #STARs' #EXECs' 299 @(star_right … STARs') 300 [ 2: @(final_step ge tr' r m s' … EXECs') 301 | skip 302 | @refl 303 ] 304 | #s #s' #tr #tr' #r #e #m #o #k #i #ESTEPS #E1 #E2 #E3 #E4 #E5 #EXEC 305 destruct; 306 cases (several_steps … ESTEPS EXEC); #STARs' #EXECs' 307 @(star_right … STARs') 308 [ @tr' 309 | @(exec_from_interact_stop … EXECs') 310 | @refl 311 ] 312 ] qed. 313 314 let corec reacts_sound ge tr s e 316 315 (REACTS:execution_reacting tr s e) 317 316 (EXEC:exec_from ge s e) : 318 317 forever_reactive (mk_transrel … step) ge s tr ≝ ?. 319 ncut (∃s'.∃e'.∃tr'.∃tr''.(And (And (And (execution_reacting tr'' s' e') (execution_isteps tr' s e s' e')) (tr' ≠ E0)) (tr = tr'⧻tr'')));320 ##[ ninversion REACTS;321 #tr0 tr' s0 s' e0 e' EREACTS ESTEPS REACTED E1 E2 E3; ndestruct (E2 E3);322 @ s'; @ e'; @ tr0; @ tr'; @; ##[ @; ##[ @; //; ##| napply REACTED ##] ##| napply refl ##]323 ##| *; #s'; *; #e'; *; #tr'; *; #tr''; 324 *; *; *; #REACTS' ESTEPS REACTED APPTR;325 (* nrewrite > APPTR;*)326 napply (match sym_eq ??? APPTR return λx.λ_.forever_reactive (mk_transrel genv state step) ge s x with [ refl ⇒ ? ]);327 @;328 ncases (several_steps … ESTEPS EXEC); #STEPS EXEC';329 ##[ ##2: napply STEPS;330 ##| ##skip331 ##| napplyREACTED332 ##| napply reacts_sound;333 ##[ ##2: napply REACTS';334 ##| ##skip335 ##| napplyEXEC'336 ##]337 ##]338 nqed.339 340 nlemma exec_from_wrong: ∀ge,s.318 cut (∃s'.∃e'.∃tr'.∃tr''.(And (And (And (execution_reacting tr'' s' e') (execution_isteps tr' s e s' e')) (tr' ≠ E0)) (tr = tr'⧻tr''))); 319 [ inversion REACTS; 320 #tr0 #tr' #s0 #s' #e0 #e' #EREACTS #ESTEPS #REACTED #E1 #E2 #E3 destruct (E2 E3); 321 %{ s'} %{ e'} %{ tr0} %{ tr'} % [ % [ % //; | @REACTED ] | @refl ] 322 | *; #s' *; #e' *; #tr' *; #tr'' 323 *; *; *; #REACTS' #ESTEPS #REACTED #APPTR 324 (* >APPTR *) 325 @(match sym_eq ??? APPTR return λx.λ_.forever_reactive (mk_transrel genv state step) ge s x with [ refl ⇒ ? ]) 326 % 327 cases (several_steps … ESTEPS EXEC); #STEPS #EXEC' 328 [ 2: @STEPS 329 | skip 330 | @REACTED 331 | @reacts_sound 332 [ 2: @REACTS' 333 | skip 334 | @EXEC' 335 ] 336 ] 337 qed. 338 339 lemma exec_from_wrong: ∀ge,s. 341 340 exec_from ge s se_wrong → 342 341 exec_step ge s = Wrong ???. 343 #ge s H; nwhd in H;344 ninversion H;345 ##[ #tr r m EXEC E; ndestruct (E)346 ##| #tr s' e e' H EXEC E; ndestruct (E)347 ##| nrewrite > (exec_inf_aux_unfold …); 348 ncases (exec_step ge s);349 ##[ #o k EXEC; nwhd in EXEC:(??%?); ndestruct350 ##| #x; ncases x; #tr s'; nwhd in ⊢ (??%? → ?); ncases (is_final_state s');351 #F EXEC; nwhd in EXEC:(??%?); ndestruct352 ##| //353 ##]354 ##| #o k i e H EXEC E; ndestruct355 ##] nqed.356 357 nlemma exec_from_step_notfinal: ∀ge,s,tr,s',e.342 #ge #s #H whd in H; 343 inversion H; 344 [ #tr #r #m #EXEC #E destruct (E) 345 | #tr #s' #e #e' #H #EXEC #E destruct (E) 346 | >(exec_inf_aux_unfold …) 347 cases (exec_step ge s); 348 [ #o #k #EXEC whd in EXEC:(??%?); destruct 349 | #x cases x; #tr #s' whd in ⊢ (??%? → ?); cases (is_final_state s'); 350 #F #EXEC whd in EXEC:(??%?); destruct 351 | // 352 ] 353 | #o #k #i #e #H #EXEC #E destruct 354 ] qed. 355 356 lemma exec_from_step_notfinal: ∀ge,s,tr,s',e. 358 357 exec_from ge s (se_step tr s' e) → 359 358 ¬(∃r. final_state s' r). 360 #ge s tr s' e H; nwhd in H; ninversion H;361 ##[ #tr' r m EXEC E; ndestruct362 ##| #tr' s'' e' e'' H EXEC E; ndestruct (E);363 nrewrite >(exec_inf_aux_unfold …) in EXEC;364 ncases (exec_step ge s);365 ##[ #o k EXEC; nwhd in EXEC:(??%?); ndestruct366 ##| #x; ncases x; #tr1 s1; nwhd in ⊢ (??%? → ?); ncases (is_final_state s1);367 #F E; nwhd in E:(??%?); ndestruct; napply F;368 ##| #E; nwhd in E:(??%?); ndestruct369 ##]370 ##| #EXEC E; ndestruct371 ##| #o k i e' H EXEC E; ndestruct372 ##] nqed.373 374 nlemma exec_from_interact_step_notfinal: ∀ge,s,o,k,i,tr,s',e.359 #ge #s #tr #s' #e #H whd in H; inversion H; 360 [ #tr' #r #m #EXEC #E destruct 361 | #tr' #s'' #e' #e'' #H #EXEC #E destruct (E); 362 >(exec_inf_aux_unfold …) in EXEC; 363 cases (exec_step ge s); 364 [ #o #k #EXEC whd in EXEC:(??%?); destruct 365 | #x cases x; #tr1 #s1 whd in ⊢ (??%? → ?); cases (is_final_state s1); 366 #F #E whd in E:(??%?); destruct; @F 367 | #E whd in E:(??%?); destruct 368 ] 369 | #EXEC #E destruct 370 | #o #k #i #e' #H #EXEC #E destruct 371 ] qed. 372 373 lemma exec_from_interact_step_notfinal: ∀ge,s,o,k,i,tr,s',e. 375 374 exec_from ge s (se_interact o k i (se_step tr s' e)) → 376 375 ¬(∃r. final_state s' r). 377 #ge s o k i tr s' e H;378 @; *; #r F; ncases F in H; #r' m H; 379 ninversion H;380 ##[ #tr' r m EXEC E; ndestruct381 ##| #tr' s'' e' e'' H EXEC E; ndestruct (E);382 ##| #EXEC E; ndestruct383 ##| #o' k' i' e' H EXEC E; ndestruct;384 nrewrite >(exec_inf_aux_unfold …) in EXEC;385 ncases (exec_step ge s);386 ##[ #o1 k1 EXEC1; nwhd in EXEC1:(??%?); ndestruct (EXEC1);387 ninversion H;388 ##[ #tr1 r1 m1 EXECK E; ndestruct (E);389 ##| #tr1 s1 e1 e2 H1 EXECK E; ndestruct (E);390 nrewrite >(exec_inf_aux_unfold …) in EXECK;391 ncases (k1 i');392 ##[ #o2 k2 EXECK; nwhd in EXECK:(??%?); ndestruct393 ##| #x; ncases x; #tr2 s2; nwhd in ⊢ (??%? → ?);394 ncases (is_final_state s2); #F EXECK;395 nwhd in EXECK:(??%?); ndestruct;396 napply (absurd ?? F);397 @ r';//;398 ##| #E; nwhd in E:(??%?); ndestruct399 ##]400 ##| #EXECK E; nwhd in E:(??%?); ndestruct401 ##| #o2 k2 i2 e2 H2 EXECK E; ndestruct402 ##]403 ##| #x; ncases x; #tr1 s1; nwhd in ⊢ (??%? → ?);404 ncases (is_final_state s1); #F E; nwhd in E:(??%?); ndestruct;405 ##| #E; nwhd in E:(??%?); ndestruct406 ##]407 ##] nqed.408 409 nlemma wrong_sound: ∀ge,tr,s,s',e.376 #ge #s #o #k #i #tr #s' #e #H 377 % *; #r #F cases F in H; #r' #m #H 378 inversion H; 379 [ #tr' #r #m #EXEC #E destruct 380 | #tr' #s'' #e' #e'' #H #EXEC #E destruct (E); 381 | #EXEC #E destruct 382 | #o' #k' #i' #e' #H #EXEC #E destruct; 383 >(exec_inf_aux_unfold …) in EXEC; 384 cases (exec_step ge s); 385 [ #o1 #k1 #EXEC1 whd in EXEC1:(??%?); destruct (EXEC1); 386 inversion H; 387 [ #tr1 #r1 #m1 #EXECK #E destruct (E); 388 | #tr1 #s1 #e1 #e2 #H1 #EXECK #E destruct (E); 389 >(exec_inf_aux_unfold …) in EXECK; 390 cases (k1 i'); 391 [ #o2 #k2 #EXECK whd in EXECK:(??%?); destruct 392 | #x cases x; #tr2 #s2 whd in ⊢ (??%? → ?); 393 cases (is_final_state s2); #F #EXECK 394 whd in EXECK:(??%?); destruct; 395 @(absurd ?? F) 396 %{ r'} //; 397 | #E whd in E:(??%?); destruct 398 ] 399 | #EXECK #E whd in E:(??%?); destruct 400 | #o2 #k2 #i2 #e2 #H2 #EXECK #E destruct 401 ] 402 | #x cases x; #tr1 #s1 whd in ⊢ (??%? → ?); 403 cases (is_final_state s1); #F #E whd in E:(??%?); destruct; 404 | #E whd in E:(??%?); destruct 405 ] 406 ] qed. 407 408 lemma wrong_sound: ∀ge,tr,s,s',e. 410 409 execution_goes_wrong tr s (se_step E0 s e) s' → 411 410 exec_from ge s e → … … 414 413 nostep (mk_transrel … step) ge s' ∧ 415 414 (¬∃r. final_state s' r). 416 #ge tr0 s0 s0' e0 WRONG; ninversion WRONG;417 #tr s s' e ESTEPS E1 E2 E3 E4 EXEC NOTFINAL; ndestruct (E1 E2 E3 E4);418 ncases (several_steps … ESTEPS EXEC);419 #STAR EXEC'; @;420 ##[ @; ##[ napply STAR; 421 ##| #badtr bads; @; #badSTEP;422 nlapply (step_complete … badSTEP);423 nrewrite > (exec_from_wrong … EXEC');415 #ge #tr0 #s0 #s0' #e0 #WRONG inversion WRONG; 416 #tr #s #s' #e #ESTEPS #E1 #E2 #E3 #E4 #EXEC #NOTFINAL destruct (E1 E2 E3 E4); 417 cases (several_steps … ESTEPS EXEC); 418 #STAR #EXEC' % 419 [ % [ @STAR 420 | #badtr #bads % #badSTEP 421 lapply (step_complete … badSTEP); 422 >(exec_from_wrong … EXEC') 424 423 //; 425 ##]426 ##| @; 427 nelim ESTEPS in NOTFINAL EXEC ⊢ %;428 ##[ #s1 e1 NF EX F; napply (absurd ? F NF);429 ##| #e1 e2 tr1 tr2 s1 s2 s3 ESTEPS1 IH NF EXEC;430 ncases (exec_from_step … EXEC); #EXEC3 EXEC1;431 napply (IH … EXEC1);432 napply (exec_from_step_notfinal … EXEC);433 ##| #e1 e2 o k i s1 s2 s3 tr1 tr2 ESTEPS1 IH NF EXEC;434 napply IH;435 ##[ napply (exec_from_interact_step_notfinal … EXEC);436 ##| ncases (exec_from_interact … EXEC); //;437 ##]438 ##]439 ##] nqed.440 441 ninductive execution_characterisation : state → s_execution → Prop ≝424 ] 425 | % 426 elim ESTEPS in NOTFINAL EXEC ⊢ %; 427 [ #s1 #e1 #NF #EX #F @(absurd ? F NF) 428 | #e1 #e2 #tr1 #tr2 #s1 #s2 #s3 #ESTEPS1 #IH #NF #EXEC 429 cases (exec_from_step … EXEC); #EXEC3 #EXEC1 430 @(IH … EXEC1) 431 @(exec_from_step_notfinal … EXEC) 432 | #e1 #e2 #o #k #i #s1 #s2 #s3 #tr1 #tr2 #ESTEPS1 #IH #NF #EXEC 433 @IH 434 [ @(exec_from_interact_step_notfinal … EXEC) 435 | cases (exec_from_interact … EXEC); //; 436 ] 437 ] 438 ] qed. 439 440 inductive execution_characterisation : state → s_execution → Prop ≝ 442 441 | ec_terminates: ∀s,r,m,e,tr. 443 442 execution_terminates tr s e r m → … … 454 453 455 454 (* bit of a hack to avoid inability to reduce term in match *) 456 ndefinition interact_prop : ∀A:Type.(∀o:io_out. (io_in o → IO io_out io_in A) → Prop) → IO io_out io_in A → Prop ≝455 definition interact_prop : ∀A:Type[0].(∀o:io_out. (io_in o → IO io_out io_in A) → Prop) → IO io_out io_in A → Prop ≝ 457 456 λA,P,e. match e return λ_.Prop with [ Interact o k ⇒ P o k | _ ⇒ True ]. 458 457 459 nlemma err_does_not_interact: ∀A,B,P,e1,e2.458 lemma err_does_not_interact: ∀A,B,P,e1,e2. 460 459 (∀x:B.interact_prop A P (e2 x)) → 461 460 interact_prop A P (bindIO ?? B A (err_to_io ??? e1) e2). 462 #A B P e1 e2 H;463 ncases e1; //; nqed.464 465 nlemma err2_does_not_interact: ∀A,B,C,P,e1,e2.461 #A #B #P #e1 #e2 #H 462 cases e1; //; qed. 463 464 lemma err2_does_not_interact: ∀A,B,C,P,e1,e2. 466 465 (∀x,y.interact_prop A P (e2 x y)) → 467 466 interact_prop A P (bindIO2 ?? B C A (err_to_io ??? e1) e2). 468 #A B C P e1 e2 H;469 ncases e1; ##[ #z; ncases z; ##] //; nqed.470 471 nlemma err_sig_does_not_interact: ∀A,B,P.∀Q:B→Prop.∀e1,e2.467 #A #B #C #P #e1 #e2 #H 468 cases e1; [ #z cases z; ] //; qed. 469 470 lemma err_sig_does_not_interact: ∀A,B,P.∀Q:B→Prop.∀e1,e2. 472 471 (∀x.interact_prop A P (e2 x)) → 473 interact_prop A P (bindIO ?? ( sigmaB Q) A (err_to_io_sig ??? Q e1) e2).474 #A B P Q e1 e2 H;475 ncases e1; //; nqed.476 477 nlemma opt_does_not_interact: ∀A,B,P,e1,e2.472 interact_prop A P (bindIO ?? (Sig B Q) A (err_to_io_sig ??? Q e1) e2). 473 #A #B #P #Q #e1 #e2 #H 474 cases e1; //; qed. 475 476 lemma opt_does_not_interact: ∀A,B,P,e1,e2. 478 477 (∀x:B.interact_prop A P (e2 x)) → 479 478 interact_prop A P (bindIO ?? B A (opt_to_io ??? e1) e2). 480 #A B P e1 e2 H;481 ncases e1; //; nqed.482 483 nlemma exec_step_interaction:479 #A #B #P #e1 #e2 #H 480 cases e1; //; qed. 481 482 lemma exec_step_interaction: 484 483 ∀ge,s. interact_prop ? (λo,k. ∀i.∃tr.∃s'. k i = Value ??? 〈tr,s'〉 ∧ tr ≠ E0) (exec_step ge s). 485 #ge s; ncases s; 486 ##[ #f st kk e m; ncases st; 487 ##[ ##11,14: #a ##| ##2,4,6,7,12,13,15: #a b ##| ##3,5: #a b c ##| ##8: #a b c d ##] 488 ##[ ##4,6,8,9: napply I ##] 489 nwhd in ⊢ (???%); 490 ##[ ncases a; ##[ ncases (fn_return f); //; ##| #e; nwhd nodelta in ⊢ (???%); 491 ncases (type_eq_dec (fn_return f) Tvoid); #x; //; napply err2_does_not_interact; // ##] 492 ##| ncases (find_label a (fn_body f) (call_cont kk)); ##[ napply I ##| #z; ncases z; #x y; napply I ##] 493 ##| napply err2_does_not_interact; #x1 x2; napply err2_does_not_interact; #x3 x4; napply opt_does_not_interact; #x5; napply I 494 ##| ##4,7: napply err2_does_not_interact; #x1 x2; napply err_does_not_interact; #x3; napply I 495 ##| napply err2_does_not_interact; #x1 x2; ncases x1; //; 496 ##| napply err2_does_not_interact; #x1 x2; napply err2_does_not_interact; #x3 x4; napply opt_does_not_interact; #x5; napply err_does_not_interact; #x6; ncases a; 497 ##[ napply I; ##| #x7; napply err2_does_not_interact; #x8 x9; napply I ##] 498 ##| ncases (is_Sskip a); #H; ##[ napply err2_does_not_interact; #x1 x2; napply err_does_not_interact; #x3; napply I 499 ##| napply I ##] 500 ##| ncases kk; ##[ ##1,8: ncases (fn_return f); //; ##| ##2,3,5,6,7: //; 501 ##| #z1 z2 z3; napply err2_does_not_interact; #x1 x2; napply err_does_not_interact; #x3; ncases x3; napply I ##] 502 ##| ncases kk; //; 503 ##| ncases kk; ##[ ##4: #z1 z2 z3; napply err2_does_not_interact; #x1 x2; napply err_does_not_interact; #x3; ncases x3; napply I 504 ##| ##*: // ##] 505 ##] 506 ##| #f args kk m; ncases f; 507 ##[ #f'; nwhd in ⊢ (???%); ncases (exec_alloc_variables empty_env m (fn_params f'@fn_vars f')); 508 #x; ncases x; ##[ *; ##| #z; ncases z; #x1 x2 H; 509 napply err_sig_does_not_interact; //; ##] 484 #ge #s cases s; 485 [ #f #st #kk #e #m cases st; 486 [ 11,14: #a | 2,4,6,7,12,13,15: #a #b | 3,5: #a #b #c | 8: #a #b #c #d ] 487 [ 4,6,8,9: @I ] 488 whd in ⊢ (???%); 489 [ cases a; [ cases (fn_return f); //; | #e whd nodelta in ⊢ (???%); 490 cases (type_eq_dec (fn_return f) Tvoid); #x //; @err2_does_not_interact // ] 491 | cases (find_label a (fn_body f) (call_cont kk)); [ @I | #z cases z #x #y @I ] 492 | @err2_does_not_interact #x1 #x2 @err2_does_not_interact #x3 #x4 @opt_does_not_interact #x5 @I 493 | 4,7: @err2_does_not_interact #x1 #x2 @err_does_not_interact #x3 @I 494 | @err2_does_not_interact #x1 #x2 cases x1; //; 495 | @err2_does_not_interact #x1 #x2 @err2_does_not_interact #x3 #x4 @opt_does_not_interact #x5 @err_does_not_interact #x6 cases a; 496 [ @I | #x7 @err2_does_not_interact #x8 #x9 @I ] 497 | cases (is_Sskip a); #H [ @err2_does_not_interact #x1 #x2 @err_does_not_interact #x3 @I 498 | @I ] 499 | cases kk; [ 1,8: cases (fn_return f); //; | 2,3,5,6,7: //; 500 | #z1 #z2 #z3 @err2_does_not_interact #x1 #x2 @err_does_not_interact #x3 cases x3; @I ] 501 | cases kk; //; 502 | cases kk; [ 4: #z1 #z2 #z3 @err2_does_not_interact #x1 #x2 @err_does_not_interact #x3 cases x3; @I 503 | *: // ] 504 ] 505 | #f #args #kk #m cases f; 506 [ #f' whd in ⊢ (???%); cases (exec_alloc_variables empty_env m (fn_params f'@fn_vars f')) 507 #x1 #x2 whd in ⊢ (???%) @err_does_not_interact // 510 508 (* This is the only case that actually matters! *) 511 ##| #fn argtys rty; nwhd in ⊢ (???%);512 napply err_sig_does_not_interact; #x1;513 nwhd; #i; @; ##[ ##2: @; ##[ ##2: @; ##[ @; nwhd in ⊢ (??%?); napply refl;514 ##| @; #E; nwhd in E:(??%%); ndestruct (E); ##] ##] ##]515 ##]516 ##| #v kk m; nwhd in ⊢ (???%); ncases kk;517 ##[ ##8: #x1 x2 x3 x4; ncases x1;518 ##[ nwhd in ⊢ (???%); ncases v; // ##| #x5; nwhd in ⊢ (???%); ncases x5;519 #x6 x7; napply opt_does_not_interact; // ##]520 ##| ##*: // ##]521 ##] nqed.509 | #fn #argtys #rty whd in ⊢ (???%); 510 @err_does_not_interact #x1 511 whd; #i % [ 2: % [ 2: % [ % whd in ⊢ (??%?); @refl 512 | % #E whd in E:(??%%); destruct (E); ] ] ] 513 ] 514 | #v #kk #m whd in ⊢ (???%); cases kk; 515 [ 8: #x1 #x2 #x3 #x4 cases x1; 516 [ whd in ⊢ (???%); cases v; // | #x5 whd in ⊢ (???%); cases x5; 517 #x6 #x7 @opt_does_not_interact // ] 518 | *: // ] 519 ] qed. 522 520 523 521 524 522 (* Some classical logic (roughly like a fragment of Coq's library) *) 525 nlemma classical_doubleneg:523 lemma classical_doubleneg: 526 524 ∀classic:(∀P:Prop.P ∨ ¬P). 527 525 ∀P:Prop. ¬ (¬ P) → P. 528 #classic P; *; #H;529 ncases (classic P);530 ##[ // ##| #H'; napply False_ind; /2/; ##]531 nqed.532 533 nlemma classical_not_all_not_ex:526 #classic #P *; #H 527 cases (classic P); 528 [ // | #H' @False_ind /2/; ] 529 qed. 530 531 lemma classical_not_all_not_ex: 534 532 ∀classic:(∀P:Prop.P ∨ ¬P). 535 ∀A:Type .∀P:A → Prop. ¬ (∀x. ¬ P x) → ∃x. P x.536 #classic A P; *; #H;537 napply (classical_doubleneg classic); @; *; #H'; 538 napply H; #x; @; #H''; napply H'; @x; napply H''; 539 nqed.540 541 nlemma classical_not_all_ex_not:533 ∀A:Type[0].∀P:A → Prop. ¬ (∀x. ¬ P x) → ∃x. P x. 534 #classic #A #P *; #H 535 @(classical_doubleneg classic) % *; #H' 536 @H #x % #H'' @H' %{x} @H'' 537 qed. 538 539 lemma classical_not_all_ex_not: 542 540 ∀classic:(∀P:Prop.P ∨ ¬P). 543 ∀A:Type .∀P:A → Prop. ¬ (∀x. P x) → ∃x. ¬ P x.544 #classic A P; *; #H;545 napply (classical_not_all_not_ex classic A (λx.¬ P x)); 546 @; #H'; napply H; #x; napply (classical_doubleneg classic); 547 napply H'; 548 nqed.549 550 nlemma not_ex_all_not:551 ∀A:Type .∀P:A → Prop. ¬ (∃x. P x) → ∀x. ¬ P x.552 #A P; *; #H x; @; #H';553 napply H; @ x; napply H'; 554 nqed.555 556 nlemma not_imply_elim:541 ∀A:Type[0].∀P:A → Prop. ¬ (∀x. P x) → ∃x. ¬ P x. 542 #classic #A #P *; #H 543 @(classical_not_all_not_ex classic A (λx.¬ P x)) 544 % #H' @H #x @(classical_doubleneg classic) 545 @H' 546 qed. 547 548 lemma not_ex_all_not: 549 ∀A:Type[0].∀P:A → Prop. ¬ (∃x. P x) → ∀x. ¬ P x. 550 #A #P *; #H #x % #H' 551 @H %{ x} @H' 552 qed. 553 554 lemma not_imply_elim: 557 555 ∀classic:(∀P:Prop.P ∨ ¬P). 558 556 ∀P,Q:Prop. ¬ (P → Q) → P. 559 #classic P Q; *; #H;560 napply (classical_doubleneg classic); @; *; #H'; 561 napply H; #H''; napply False_ind; napply H'; napply H''; 562 nqed.563 564 nlemma not_imply_elim2:557 #classic #P #Q *; #H 558 @(classical_doubleneg classic) % *; #H' 559 @H #H'' @False_ind @H' @H'' 560 qed. 561 562 lemma not_imply_elim2: 565 563 ∀P,Q:Prop. ¬ (P → Q) → ¬ Q. 566 #P Q; *; #H; @; #H';567 napply H; #_; napply H'; 568 nqed.569 570 nlemma imply_to_and:564 #P #Q *; #H % #H' 565 @H #_ @H' 566 qed. 567 568 lemma imply_to_and: 571 569 ∀classic:(∀P:Prop.P ∨ ¬P). 572 570 ∀P,Q:Prop. ¬ (P → Q) → P ∧ ¬Q. 573 #classic P Q H; @;574 ##[ napply (not_imply_elim classic P Q H); 575 ##| napply (not_imply_elim2 P Q H); 576 ##] nqed.577 578 nlemma not_and_to_imply:571 #classic #P #Q #H % 572 [ @(not_imply_elim classic P Q H) 573 | @(not_imply_elim2 P Q H) 574 ] qed. 575 576 lemma not_and_to_imply: 579 577 ∀classic:(∀P:Prop.P ∨ ¬P). 580 578 ∀P,Q:Prop. ¬ (P ∧ Q) → P → ¬Q. 581 #classic P Q; *; #H H';582 @; #H''; napply H; @;//;583 nqed.584 585 ninductive execution_not_over : s_execution → Prop ≝579 #classic #P #Q *; #H #H' 580 % #H'' @H % //; 581 qed. 582 583 inductive execution_not_over : s_execution → Prop ≝ 586 584 | eno_step: ∀tr,s,e. execution_not_over (se_step tr s e) 587 585 | eno_interact: ∀o,k,tr,s,e,i. execution_not_over (se_interact o k i (se_step tr s e)). 588 586 589 nlemma eno_stop: ∀tr,r,m. execution_not_over (se_stop tr r m) → False.590 #tr0 r0 m0 H; ninversion H;591 ##[ #tr s e E; ndestruct592 ##| #o k tr s e i E; ndestruct593 ##] nqed.594 595 nlemma eno_wrong: execution_not_over se_wrong → False.596 #H ; ninversion H;597 ##[ #tr s e E; ndestruct598 ##| #o k tr s e i E; ndestruct599 ##] nqed.600 601 nlet corec show_divergence s e587 lemma eno_stop: ∀tr,r,m. execution_not_over (se_stop tr r m) → False. 588 #tr0 #r0 #m0 #H inversion H; 589 [ #tr #s #e #E destruct 590 | #o #k #tr #s #e #i #E destruct 591 ] qed. 592 593 lemma eno_wrong: execution_not_over se_wrong → False. 594 #H inversion H; 595 [ #tr #s #e #E destruct 596 | #o #k #tr #s #e #i #E destruct 597 ] qed. 598 599 let corec show_divergence s e 602 600 (NONTERMINATING:∀tr1,s1,e1. execution_isteps tr1 s e s1 e1 → 603 601 execution_not_over e1) 604 602 (UNREACTIVE:∀tr2,s2,e2. execution_isteps tr2 s e s2 e2 → tr2 = E0) 605 (CONTINUES:∀tr2,s2,o,k,i,e'. execution_isteps tr2 s e s2 (se_interact o k i e') → ∃tr3.∃s3.∃e3. e' = se_step tr3 s3 e3 ∧ tr3 ≠ E0)603 (CONTINUES:∀tr2,s2,o,k,i,e'. execution_isteps tr2 s e s2 (se_interact o k i e') → ∃tr3.∃s3.∃e3. And (e' = se_step tr3 s3 e3) (tr3 ≠ E0)) 606 604 : execution_diverging e ≝ ?. 607 nlapply (NONTERMINATING E0 s e ?); //; 608 ncases e in UNREACTIVE NONTERMINATING CONTINUES ⊢ %; 609 ##[ #tr i m; #_; #_; #_; #ENO; nelim (eno_stop … ENO); 610 ##| #tr s' e' UNREACTIVE; nlapply (UNREACTIVE tr s' e' ?); 611 ##[ nrewrite < (E0_right tr) in ⊢ (?%????); 612 napply isteps_one; napply isteps_none; 613 ##| #TR; napply (match sym_eq ??? TR with [ refl ⇒ ? ]); (* nrewrite > TR in UNREACTIVE ⊢ %;*) 614 #NONTERMINATING CONTINUES; #_; @; 615 napply (show_divergence s'); 616 ##[ #tr1 s1 e1 S; napply (NONTERMINATING tr1 s1 e1); 617 nchange in ⊢ (?%????) with (Eapp E0 tr1); napply isteps_one; 618 napply S; 619 ##| #tr2 s2 e2 S; nrewrite > TR in UNREACTIVE; #UNREACTIVE; napply (UNREACTIVE tr2 s2 e2); 620 nchange in ⊢ (?%????) with (Eapp E0 tr2); 621 napply isteps_one; napply S; 622 ##| #tr2 s2 o k i e2 S; napply (CONTINUES tr2 s2 o k i); 623 nchange in ⊢ (?%????) with (Eapp E0 tr2); 624 napply (isteps_one … S); 625 ##] 626 ##] 627 ##| #_; #_; #_; #ENO; nelim (eno_wrong … ENO); 628 ##| #o k i e' UNREACTIVE NONTERMINATING CONTINUES; #_; 629 nlapply (CONTINUES E0 s o k i e' (isteps_none …)); 630 *; #tr'; *; #s'; *; #e'; *; #EXEC NOTSILENT; 631 napply False_ind; napply (absurd ?? NOTSILENT); 632 napply (UNREACTIVE … s' e'); 633 nrewrite < (E0_right tr') in ⊢ (?%????); 634 nrewrite > EXEC; 635 napply isteps_interact; //; 636 ##] nqed. 637 638 nlemma se_inv: ∀e1,e2. 605 lapply (NONTERMINATING E0 s e ?); //; 606 cases e in UNREACTIVE NONTERMINATING CONTINUES ⊢ %; 607 [ #tr #i #m #_ #_ #_ #ENO elim (eno_stop … ENO); 608 | #tr #s' #e' #UNREACTIVE lapply (UNREACTIVE tr s' e' ?); 609 [ <(E0_right tr) in ⊢ (?%????) 610 @isteps_one @isteps_none 611 | #TR @(match sym_eq ??? TR with [ refl ⇒ ? ]) (* >TR in UNREACTIVE ⊢ % *) 612 #NONTERMINATING #CONTINUES #_ % 613 @(show_divergence s') 614 [ #tr1 #s1 #e1 #S @(NONTERMINATING tr1 s1 e1) 615 change in ⊢ (?%????) with (Eapp E0 tr1); @isteps_one 616 @S 617 | #tr2 #s2 #e2 #S >TR in UNREACTIVE #UNREACTIVE @(UNREACTIVE tr2 s2 e2) 618 change in ⊢ (?%????) with (Eapp E0 tr2); 619 @isteps_one @S 620 | #tr2 #s2 #o #k #i #e2 #S @(CONTINUES tr2 s2 o k i) 621 change in ⊢ (?%????) with (Eapp E0 tr2); 622 @(isteps_one … S) 623 ] 624 ] 625 | #_ #_ #_ #ENO elim (eno_wrong … ENO); 626 | #o #k #i #e' #UNREACTIVE #NONTERMINATING #CONTINUES #_ 627 lapply (CONTINUES E0 s o k i e' (isteps_none …)); 628 *; #tr' *; #s' *; #e' *; #EXEC #NOTSILENT 629 @False_ind @(absurd ?? NOTSILENT) 630 @(UNREACTIVE … s' e') 631 <(E0_right tr') in ⊢ (?%????) 632 >EXEC 633 @isteps_interact //; 634 ] qed. 635 636 (* XXX == > jmeq notation and coercion *) 637 638 lemma jmeq_to_eq : ∀A:Type[0].∀a,b:A.jmeq A a A b → a = b. 639 #A #a #b #E @gral @jm_to_eq_sigma @E 640 qed. 641 642 coercion jmeq_to_eq : ∀A:Type[0].∀a,b:A.∀p:jmeq A a A b.a = b ≝ 643 jmeq_to_eq on _p: jmeq ???? to eq ???. 644 645 notation > "hvbox(a break ≃ b)" 646 non associative with precedence 45 647 for @{ 'jmeq ? $a ? $b }. 648 649 notation < "hvbox(term 46 a break maction (≃) (≃\sub(t,u)) term 46 b)" 650 non associative with precedence 45 651 for @{ 'jmeq $t $a $u $b }. 652 653 interpretation "john major's equality" 'jmeq t x u y = (jmeq t x u y). 654 655 (* XXX < == *) 656 657 lemma se_inv: ∀e1,e2. 639 658 single_exec_of e1 e2 → 640 659 match e1 with … … 644 663 | e_interact o k ⇒ match e2 with [ se_interact o' k' i e ⇒ o' = o ∧ k' ≃ k ∧ single_exec_of (k' i) e | _ ⇒ False ] 645 664 ]. 646 #e01 e02 H;647 ncases H;648 ##[ #tr r m; nwhd; @; ##[ @ ##] //649 ##| #tr s e1' e2' H'; nwhd; @; ##[ @ ##] //650 ##| nwhd; //651 ##| #o k i e H'; nwhd; @; ##[ @ ##] //652 ##] nqed.653 654 nlemma interaction_is_not_silent: ∀ge,o,k,i,tr,s,s',e.665 #e01 #e02 #H 666 cases H; 667 [ #tr #r #m whd; % [ % ] // 668 | #tr #s #e1' #e2' #H' whd; % [ % ] // 669 | whd; // 670 | #o #k #i #e #H' whd; % [ % ] // 671 ] qed. 672 673 lemma interaction_is_not_silent: ∀ge,o,k,i,tr,s,s',e. 655 674 exec_from ge s (se_interact o k i (se_step tr s' e)) → 656 675 tr ≠ E0. 657 #ge o k i tr s s' e; nwhd in ⊢ (% → ?); nrewrite > (exec_inf_aux_unfold …);658 nlapply (exec_step_interaction ge s);659 ncases (exec_step ge s);660 ##[ #o' k' ; nwhd in ⊢ (% → ?%? → ?); #H K; ncases (se_inv … K);661 *; #E1 E2 H1; ndestruct (E1);662 nlapply (H i); *; #tr'; *; #s''; *; #K' TR;663 nrewrite > E2 in H1; #H1; nwhd in H1:(?%?); nrewrite > K' in H1;664 nrewrite > (exec_inf_aux_unfold …); nwhd in ⊢ (?%? → ?);665 ncases (is_final_state s'');666 ##[ #F; nwhd in ⊢ (?%? → ?); #S;667 napply False_ind; napply (absurd ? S); ncases (se_inv … S)668 ##| #NF S; nwhd in S:(?%?); ncases (se_inv … S);669 *; #E1 E2 S'; nrewrite < E1; napplyTR670 ##]671 ##| #x; ncases x; #tr' s'' H; nwhd in ⊢ (?%? → ?);672 ncases (is_final_state s''); #F E; nwhd in E:(?%?);673 ninversion E;674 ##[ ##1,5: #tr1 e1 m1 E1 E2; ndestruct675 ##| ##2,6: #tr s1 e1 e2 H E1 E2; ndestruct676 ##| ##3,7: #E; ndestruct677 ##| ##4,8: #o1 k1 i1 e1 H1 E1 E2; ndestruct678 ##]679 ##| #_; #E; nwhd in E:(?%?);680 ninversion E;681 ##[ ##1,5: #tr1 e1 m1 E1 E2; ndestruct682 ##| ##2,6: #tr s1 e1 e2 H E1 E2; ndestruct683 ##| ##3,7: #E1 E2; ndestruct684 ##| ##4,8: #o1 k1 i1 e1 H1 E1 E2; ndestruct685 ##]686 ##] nqed.687 688 nlet corec reactive_traceinf' ge s e676 #ge #o #k #i #tr #s #s' #e whd in ⊢ (% → ?); >(exec_inf_aux_unfold …) 677 lapply (exec_step_interaction ge s); 678 cases (exec_step ge s); 679 [ #o' #k' ; whd in ⊢ (% → ?%? → ?); #H #K cases (se_inv … K); 680 *; #E1 #E2 #H1 destruct (E1); 681 lapply (H i); *; #tr' *; #s'' *; #K' #TR 682 >E2 in H1 #H1 whd in H1:(?%?); >K' in H1 683 >(exec_inf_aux_unfold …) whd in ⊢ (?%? → ?); 684 cases (is_final_state s''); 685 [ #F whd in ⊢ (?%? → ?); #S 686 @False_ind @(absurd ? S) cases (se_inv … S) 687 | #NF #S whd in S:(?%?); cases (se_inv … S); 688 *; #E1 #E2 #S' <E1 @TR 689 ] 690 | #x cases x; #tr' #s'' #H whd in ⊢ (?%? → ?); 691 cases (is_final_state s''); #F #E whd in E:(?%?); 692 inversion E; 693 [ 1,5: #tr1 #e1 #m1 #E1 #E2 destruct 694 | 2,6: #tr #s1 #e1 #e2 #H #E1 #E2 destruct 695 | 3,7: #E destruct 696 | 4,8: #o1 #k1 #i1 #e1 #H1 #E1 #E2 destruct 697 ] 698 | #_ #E whd in E:(?%?); 699 inversion E; 700 [ 1,5: #tr1 #e1 #m1 #E1 #E2 destruct 701 | 2,6: #tr #s1 #e1 #e2 #H #E1 #E2 destruct 702 | 3,7: #E1 #E2 destruct 703 | 4,8: #o1 #k1 #i1 #e1 #H1 #E1 #E2 destruct 704 ] 705 ] qed. 706 707 let corec reactive_traceinf' ge s e 689 708 (EXEC:exec_from ge s e) 690 709 (REACTIVE: ∀tr,s1,e1. 691 710 execution_isteps tr s e s1 e1 → 692 Σx.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0)711 (Sig ? (λx.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0))) 693 712 : traceinf' ≝ ?. 694 nlapply (REACTIVE E0 s e (isteps_none …));695 *; #x ; ncases x; #tr; #y; ncases y; #s' e'; *; #STEPS H;696 @ tr ? H; 697 napply (reactive_traceinf' ge s' e' ?); 698 ##[ ncases (several_steps … STEPS EXEC); #_; // 699 ##| #tr1 s1 e1 STEPS1; 700 napply REACTIVE;701 ##[ ##2:702 napply (isteps_trans … STEPS STEPS1);703 ##| ##skip704 ##]705 ##]706 nqed.713 lapply (REACTIVE E0 s e (isteps_none …)); 714 *; #x cases x; #tr #y cases y; #s' #e' *; #STEPS #H 715 %{ tr ? H} 716 @(reactive_traceinf' ge s' e' ?) 717 [ cases (several_steps … STEPS EXEC); #_ #H' @H' 718 | #tr1 #s1 #e1 #STEPS1 719 @REACTIVE 720 [ 2: 721 @(isteps_trans … STEPS STEPS1) 722 | skip 723 ] 724 ] 725 qed. 707 726 708 727 (* A slightly different version of the above, to work around a problem with the 709 728 next result. *) 710 nlet corec reactive_traceinf'' ge s e729 let corec reactive_traceinf'' ge s e 711 730 (EXEC:exec_from ge s e) 712 (REACTIVE0: Σx.execution_isteps (\fst x) s e (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0)731 (REACTIVE0: Sig ? (λx.execution_isteps (\fst x) s e (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0)) 713 732 (REACTIVE: ∀tr,s1,e1. 714 733 execution_isteps tr s e s1 e1 → 715 Σx.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0)734 (Sig ? (λx.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0))) 716 735 : traceinf' ≝ ?. 717 ncases REACTIVE0; #x; ncases x; #tr; #y; ncases y; #s' e'; *; #STEPS H; 718 @ tr ? H; 719 napply (reactive_traceinf'' ge s' e' ?); 720 ##[ ncases (several_steps … STEPS EXEC); #_; // 721 ##| napply (REACTIVE … STEPS); 722 ##| #tr1 s1 e1 STEPS1; 723 napply REACTIVE;724 ##[ ##2:725 napply (isteps_trans … STEPS STEPS1);726 ##| ##skip727 ##]728 ##] nqed.736 cases REACTIVE0; #x cases x; #tr #y cases y; #s' #e' *; #STEPS #H 737 %{ tr ? H} 738 @(reactive_traceinf'' ge s' e' ?) 739 [ cases (several_steps … STEPS EXEC); #_ #H' @H' 740 | @(REACTIVE … STEPS) 741 | #tr1 #s1 #e1 #STEPS1 742 @REACTIVE 743 [ 2: 744 @(isteps_trans … STEPS STEPS1) 745 | skip 746 ] 747 ] qed. 729 748 730 749 (* We want to prove 731 750 732 nlemma show_reactive : ∀ge,s.751 lemma show_reactive : ∀ge,s. 733 752 ∀REACTIVE:∀tr,s1,e1. 734 753 execution_isteps tr s (exec_inf_aux ge (exec_step ge s)) s1 e1 → … … 740 759 we can do case analysis on, then get it into the desired form afterwards. 741 760 *) 742 nlet corec show_reactive' ge s e761 let corec show_reactive' ge s e 743 762 (EXEC:exec_from ge s e) 744 (REACTIVE0: Σx.execution_isteps (\fst x) s e (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0)763 (REACTIVE0: Sig ? (λx.execution_isteps (\fst x) s e (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0)) 745 764 (REACTIVE: ∀tr1,s1,e1. 746 765 execution_isteps tr1 s e s1 e1 → 747 Σx.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0)766 (Sig ? (λx.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0))) 748 767 : execution_reacting (traceinf_of_traceinf' (reactive_traceinf'' ge s e EXEC REACTIVE0 REACTIVE)) s e ≝ ?. 749 (* nrewrite > (unroll_traceinf' (reactive_traceinf'' …));*)750 napply (match sym_eq ??? (unroll_traceinf' (reactive_traceinf'' …)) with [ refl ⇒ ? ]); 751 ncases REACTIVE0;752 #x ; ncases x; #tr1; #y; ncases y; #s1 e1; #z; ncases z; #STEPS NOTSILENT;753 nwhd in ⊢ (?(?%)??);754 (* nrewrite > (traceinf_traceinfp_app …);*)755 napply (match sym_eq ??? (traceinf_traceinfp_app …) with [ refl ⇒ ? ]); 756 napply (reacting … STEPS NOTSILENT); 757 napply show_reactive'; 758 nqed.759 760 nlemma show_reactive : ∀ge,s,e.768 (*>(unroll_traceinf' (reactive_traceinf'' …)) *) 769 @(match sym_eq ??? (unroll_traceinf' (reactive_traceinf'' …)) with [ refl ⇒ ? ]) 770 cases REACTIVE0; 771 #x cases x; #tr1 #y cases y; #s1 #e1 #z cases z; #STEPS #NOTSILENT 772 whd in ⊢ (?(?%)??); 773 (*>(traceinf_traceinfp_app …) *) 774 @(match sym_eq ??? (traceinf_traceinfp_app …) with [ refl ⇒ ? ]) 775 @(reacting … STEPS NOTSILENT) 776 @show_reactive' 777 qed. 778 779 lemma show_reactive : ∀ge,s,e. 761 780 ∀EXEC:exec_from ge s e. 762 781 ∀REACTIVE:∀tr,s1,e1. 763 782 execution_isteps tr s e s1 e1 → 764 Σx.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0.783 (Sig ? (λx.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0)). 765 784 execution_reacting (traceinf_of_traceinf' (reactive_traceinf'' ge s e EXEC ? REACTIVE)) s e. 766 ##[ #ge s e EXEC REACTIVE; 767 napply show_reactive';768 ##| napply (REACTIVE … (isteps_none …)); 769 ##] nqed.770 771 nlemma execution_characterisation_complete:785 [ #ge #s #e #EXEC #REACTIVE 786 @show_reactive' 787 | @(REACTIVE … (isteps_none …)) 788 ] qed. 789 790 lemma execution_characterisation_complete: 772 791 ∀classic:(∀P:Prop.P ∨ ¬P). 773 ∀constructive_indefinite_description:(∀A:Type . ∀P:A→Prop. (∃x. P x) → Σx : A. P x).792 ∀constructive_indefinite_description:(∀A:Type[0]. ∀P:A→Prop. (∃x. P x) → Sig A P). 774 793 ∀ge,s,e. 775 794 exec_from ge s e → 776 795 execution_characterisation s (se_step E0 s e). 777 #classic constructive_indefinite_description ge s e EXEC;778 ncases (classic (∀tr1,s1,e1. execution_isteps tr1 s e s1 e1 →796 #classic #constructive_indefinite_description #ge #s #e #EXEC 797 cases (classic (∀tr1,s1,e1. execution_isteps tr1 s e s1 e1 → 779 798 execution_not_over e1)); 780 ##[ #NONTERMINATING; 781 ncases (classic (∃tr,s1,e1. execution_isteps tr s e s1 e1 ∧799 [ #NONTERMINATING 800 cases (classic (∃tr,s1,e1. execution_isteps tr s e s1 e1 ∧ 782 801 ∀tr2,s2,e2. execution_isteps tr2 s1 e1 s2 e2 → tr2 = E0)); 783 ##[ *; #tr; *; #s1; *; #e1; *; #INITIAL UNREACTIVE;784 napply (ec_diverges … s ? tr);785 napply (diverges_diverging … INITIAL);786 napply (show_divergence s1 e1);787 ##[ #tr2 s2 e2 S; napply (NONTERMINATING (Eapp tr tr2) s2 e2);788 napply (isteps_trans … INITIAL S);789 ##| #tr2 s2 e2 S; napply (UNREACTIVE … S);790 ##| #tr2 s2 o k i e2 STEPS;791 nlapply (NONTERMINATING (Eapp tr tr2) s2 (se_interact o k i e2) ?);792 ##[ napply (isteps_trans … INITIAL STEPS) ##]793 #NOTOVER ; ninversion NOTOVER;794 ##[ #tr' s' e' E; ndestruct (E);795 ##| #o' k' tr' s' e' i' E; ndestruct (E);796 @ tr'; @s'; @e'; @;//;797 ncases (several_steps … INITIAL EXEC); #_; #EXEC1;798 ncases (several_steps … STEPS EXEC1); #_; #EXEC2;799 napply (interaction_is_not_silent … EXEC2);800 ##]801 ##]802 803 ##| *; #NOTUNREACTIVE;804 ncut (∀tr,s1,e1.execution_isteps tr s e s1 e1 →802 [ *; #tr *; #s1 *; #e1 *; #INITIAL #UNREACTIVE 803 @(ec_diverges … s ? tr) 804 @(diverges_diverging … INITIAL) 805 @(show_divergence s1 e1) 806 [ #tr2 #s2 #e2 #S @(NONTERMINATING (Eapp tr tr2) s2 e2) 807 @(isteps_trans … INITIAL S) 808 | #tr2 #s2 #e2 #S @(UNREACTIVE … S) 809 | #tr2 #s2 #o #k #i #e2 #STEPS 810 lapply (NONTERMINATING (Eapp tr tr2) s2 (se_interact o k i e2) ?); 811 [ @(isteps_trans … INITIAL STEPS) ] 812 #NOTOVER inversion NOTOVER; 813 [ #tr' #s' #e' #E destruct (E); 814 | #o' #k' #tr' #s' #e' #i' #E destruct (E); 815 %{ tr'} %{s'} %{e'} % //; 816 cases (several_steps … INITIAL EXEC); #_ #EXEC1 817 cases (several_steps … STEPS EXEC1); #_ #EXEC2 818 @(interaction_is_not_silent … EXEC2) 819 ] 820 ] 821 822 | *; #NOTUNREACTIVE 823 cut (∀tr,s1,e1.execution_isteps tr s e s1 e1 → 805 824 ∃x.execution_isteps (\fst x) s1 e1 (\fst (\snd x)) (\snd (\snd x)) ∧ (\fst x) ≠ E0); 806 ##[ #tr s1 e1 STEPS;807 napply (classical_doubleneg classic); @; #NOREACTION;808 napply NOTUNREACTIVE;809 @ tr; @s1; @e1; @;//;810 #tr2 s2 e2 STEPS2;811 nlapply (not_ex_all_not … NOREACTION); #NR1;812 nlapply (not_and_to_imply classic … (NR1 〈tr2,〈s2,e2〉〉)); #NR2;813 napply (classical_doubleneg classic);814 napply NR2; //;815 ##| #REACTIVE;816 napply ec_reacts;817 ##[ ##2: napply reacts;818 napply (show_reactive ge s … EXEC);819 #tr s1 e1 STEPS;820 napply constructive_indefinite_description;821 napply (REACTIVE … tr s1 e1 STEPS);822 ##| ##skip823 ##]824 ##]825 ##]825 [ #tr #s1 #e1 #STEPS 826 @(classical_doubleneg classic) % #NOREACTION 827 @NOTUNREACTIVE 828 %{ tr} %{s1} %{e1} % //; 829 #tr2 #s2 #e2 #STEPS2 830 lapply (not_ex_all_not … NOREACTION); #NR1 831 lapply (not_and_to_imply classic … (NR1 〈tr2,〈s2,e2〉〉)); #NR2 832 @(classical_doubleneg classic) 833 @NR2 normalize // 834 | #REACTIVE 835 @ec_reacts 836 [ 2: @reacts 837 @(show_reactive ge s … EXEC) 838 #tr #s1 #e1 #STEPS 839 @constructive_indefinite_description 840 @(REACTIVE … tr s1 e1 STEPS) 841 | skip 842 ] 843 ] 844 ] 826 845 827 ##| #NOTNONTERMINATING; nlapply (classical_not_all_ex_not classic … NOTNONTERMINATING);828 *; #tr NNT2; nlapply (classical_not_all_ex_not classic … NNT2);829 *; #s' NNT3; nlapply (classical_not_all_ex_not classic … NNT3);830 *; #e NNT4; nelim (imply_to_and classic … NNT4);831 ncases e;832 ##[ #tr' r m; #STEPS NOSTEP;833 napply (ec_terminates s r m ? (Eapp tr tr')); @;834 ##[ napplys'835 ##| napplySTEPS836 ##]837 ##| #tr' s'' e' STEPS; *; #NOSTEP; napply False_rect_Type0;838 napply NOSTEP;//839 ##| #STEPS NOSTEP;840 napply (ec_wrong ? s s' tr); @;//;846 | #NOTNONTERMINATING lapply (classical_not_all_ex_not classic … NOTNONTERMINATING); 847 *; #tr #NNT2 lapply (classical_not_all_ex_not classic … NNT2); 848 *; #s' #NNT3 lapply (classical_not_all_ex_not classic … NNT3); 849 *; #e #NNT4 elim (imply_to_and classic … NNT4); 850 cases e; 851 [ #tr' #r #m #STEPS #NOSTEP 852 @(ec_terminates s r m ? (Eapp tr tr')) % 853 [ @s' 854 | @STEPS 855 ] 856 | #tr' #s'' #e' #STEPS *; #NOSTEP @False_rect_Type0 857 @NOSTEP // 858 | #STEPS #NOSTEP 859 @(ec_wrong ? s s' tr) % //; 841 860 (* The following is stupidly complicated when most of the cases are impossible. 842 861 It ought to be simplified. *) 843 ##| #o k i e' STEPS NOSTEP;844 ncases e' in STEPS NOSTEP;845 ##[ #tr' r m STEPS NOSTEP;846 napply (ec_terminates s ???);847 ##[ ##4: napply (annoying_corner_case_terminates … STEPS); ##]848 ##| #tr1 s1 e1 STEPS; *; #NOSTEP;849 napply False_ind; napply NOSTEP;//850 ##| #STEPS NOSTEP;851 nlapply (exec_step_interaction ge s');852 ncases (several_steps … STEPS EXEC); #_;853 nwhd in ⊢ (% → ?);854 nrewrite > (exec_inf_aux_unfold …);855 ncases (exec_step ge s');856 ##[ #o1 k1 EXEC' H; nwhd in EXEC':(?%?) H;857 ncases (se_inv … EXEC'); *; #E1 E2 H2; ndestruct (E1 E2);858 ncases (H i); #tr1; *; #s1; *; #K E; nrewrite > K in H2;859 nrewrite > (exec_inf_aux_unfold …);860 nwhd in ⊢ (?%? → ?); ncases (is_final_state s1);861 #F S; nwhd in S:(?%?); ncases (se_inv … S);862 ##| #x; ncases x; #tr' s'; nwhd in ⊢ (?%? → ?);863 ncases (is_final_state s'); #F S; nwhd in S:(?%?);864 ncases (se_inv … S);865 ##| #S; ncases (se_inv … S);866 ##]867 ##| #o1 k1 i1 e1 STEPS NOSTEP;868 nlapply (exec_step_interaction ge s');869 ncases (several_steps … STEPS EXEC); #_;870 nwhd in ⊢ (% → ?);871 nrewrite > (exec_inf_aux_unfold …);872 ncases (exec_step ge s');873 ##[ #o1 k1 EXEC' H; nwhd in EXEC':(?%?) H;874 ncases (se_inv … EXEC'); *; #E1 E2 H2; ndestruct (E1 E2);875 ncases (H i); #tr1; *; #s1; *; #K E; nrewrite > K in H2;876 nrewrite > (exec_inf_aux_unfold …);877 nwhd in ⊢ (?%? → ?); ncases (is_final_state s1);878 #F S; nwhd in S:(?%?); ncases (se_inv … S);879 ##| #x; ncases x; #tr' s'; nwhd in ⊢ (?%? → ?);880 ncases (is_final_state s'); #F S; nwhd in S:(?%?);881 ncases (se_inv … S);882 ##| #S; ncases (se_inv … S);883 ##]884 ##]885 ##]886 ##]887 nqed.888 889 ninductive execution_matches_behavior: s_execution → program_behavior → Prop ≝862 | #o #k #i #e' #STEPS #NOSTEP 863 cases e' in STEPS NOSTEP; 864 [ #tr' #r #m #STEPS #NOSTEP 865 @(ec_terminates s ???) 866 [ 4: @(annoying_corner_case_terminates … STEPS) ] 867 | #tr1 #s1 #e1 #STEPS *; #NOSTEP 868 @False_ind @NOSTEP // 869 | #STEPS #NOSTEP 870 lapply (exec_step_interaction ge s'); 871 cases (several_steps … STEPS EXEC); #_ 872 whd in ⊢ (% → ?); 873 >(exec_inf_aux_unfold …) 874 cases (exec_step ge s'); 875 [ #o1 #k1 #EXEC' #H whd in EXEC':(?%?) H; 876 cases (se_inv … EXEC'); *; #E1 #E2 #H2 destruct (E1 E2); 877 cases (H i); #tr1 *; #s1 *; #K #E >K in H2 878 >(exec_inf_aux_unfold …) 879 whd in ⊢ (?%? → ?); cases (is_final_state s1); 880 #F #S whd in S:(?%?); cases (se_inv … S); 881 | #x cases x; #tr' #s' whd in ⊢ (?%? → ?); 882 cases (is_final_state s'); #F #S whd in S:(?%?); 883 cases (se_inv … S); 884 | #S cases (se_inv … S); 885 ] 886 | #o1 #k1 #i1 #e1 #STEPS #NOSTEP 887 lapply (exec_step_interaction ge s'); 888 cases (several_steps … STEPS EXEC); #_ 889 whd in ⊢ (% → ?); 890 >(exec_inf_aux_unfold …) 891 cases (exec_step ge s'); 892 [ #o1 #k1 #EXEC' #H whd in EXEC':(?%?) H; 893 cases (se_inv … EXEC'); *; #E1 #E2 #H2 destruct (E1 E2); 894 cases (H i); #tr1 *; #s1 *; #K #E >K in H2 895 >(exec_inf_aux_unfold …) 896 whd in ⊢ (?%? → ?); cases (is_final_state s1); 897 #F #S whd in S:(?%?); cases (se_inv … S); 898 | #x cases x; #tr' #s' whd in ⊢ (?%? → ?); 899 cases (is_final_state s'); #F #S whd in S:(?%?); 900 cases (se_inv … S); 901 | #S cases (se_inv … S); 902 ] 903 ] 904 ] 905 ] 906 qed. 907 908 inductive execution_matches_behavior: s_execution → program_behavior → Prop ≝ 890 909 | emb_terminates: ∀s,e,tr,r,m. 891 910 execution_terminates tr s e r m → … … 903 922 execution_matches_behavior se_wrong (Goes_wrong E0). 904 923 905 nlemma exec_state_terminates: ∀tr,tr',s,s',e,r,m.924 lemma exec_state_terminates: ∀tr,tr',s,s',e,r,m. 906 925 execution_terminates tr s (se_step tr' s' e) r m → s = s'. 907 #tr tr' s s' e r m H; ninversion H;908 ##[ #s1 s2 tr1 tr2 r' e' m' H' E1 E2 E3 E4 E5; ndestruct; napplyrefl909 ##| #s1 s2 tr1 tr2 r' e' m' o k i H' E1 E2 E3 E4 E5; ndestruct; napplyrefl910 ##] nqed.911 912 nlemma exec_state_diverges: ∀tr,tr',s,s',e.926 #tr #tr' #s #s' #e #r #m #H inversion H; 927 [ #s1 #s2 #tr1 #tr2 #r' #e' #m' #H' #E1 #E2 #E3 #E4 #E5 destruct; @refl 928 | #s1 #s2 #tr1 #tr2 #r' #e' #m' #o #k #i #H' #E1 #E2 #E3 #E4 #E5 destruct; @refl 929 ] qed. 930 931 lemma exec_state_diverges: ∀tr,tr',s,s',e. 913 932 execution_diverges tr s (se_step tr' s' e) → s = s'. 914 #tr tr' s s' e H; ninversion H;915 #tr1 s1 s2 e1 e2 H' E1 E2 E3 E4; ndestruct; napply refl; nqed.916 917 nlemma exec_state_reacts: ∀tr,tr',s,s',e.933 #tr #tr' #s #s' #e #H inversion H; 934 #tr1 #s1 #s2 #e1 #e2 #H' #E1 #E2 #E3 #E4 destruct; @refl qed. 935 936 lemma exec_state_reacts: ∀tr,tr',s,s',e. 918 937 execution_reacts tr s (se_step tr' s' e) → s = s'. 919 #tr tr' s s' e H; ninversion H;920 #tr1 s1 e1 H' E1 E2 E3; ndestruct; napply refl; nqed.921 922 nlemma exec_state_wrong: ∀tr,tr',s,s',s'',e.938 #tr #tr' #s #s' #e #H inversion H; 939 #tr1 #s1 #e1 #H' #E1 #E2 #E3 destruct; @refl qed. 940 941 lemma exec_state_wrong: ∀tr,tr',s,s',s'',e. 923 942 execution_goes_wrong tr s (se_step tr' s' e) s'' → s = s'. 924 #tr tr' s s' s'' e H; ninversion H;925 #tr1 s1 s2 e1 H' E1 E2 E3 E4; ndestruct; napply refl; nqed.926 927 nlemma behavior_of_execution: ∀s,e.943 #tr #tr' #s #s' #s'' #e #H inversion H; 944 #tr1 #s1 #s2 #e1 #H' #E1 #E2 #E3 #E4 destruct; @refl qed. 945 946 lemma behavior_of_execution: ∀s,e. 928 947 execution_characterisation s e → 929 948 ∃b:program_behavior. execution_matches_behavior e b. 930 #s0 e0 exec;931 ncases exec;932 ##[ #s r m e tr TERM; 933 @ (Terminates tr r);934 napply (emb_terminates … TERM);935 ##| #s e tr DIV; 936 @ (Diverges tr);937 napply (emb_diverges … DIV);938 ##| #s e tr REACTS; 939 @ (Reacts tr);940 napply (emb_reacts … REACTS);941 ##| #e s s' tr WRONG; 942 @ (Goes_wrong tr);943 napply (emb_wrong … WRONG);944 ##] nqed.945 946 nlemma initial_state_not_final: ∀ge,s.949 #s0 #e0 #exec 950 cases exec; 951 [ #s #r #m #e #tr #TERM 952 %{ (Terminates tr r)} 953 @(emb_terminates … TERM) 954 | #s #e #tr #DIV 955 %{ (Diverges tr)} 956 @(emb_diverges … DIV) 957 | #s #e #tr #REACTS 958 %{ (Reacts tr)} 959 @(emb_reacts … REACTS) 960 | #e #s #s' #tr #WRONG 961 %{ (Goes_wrong tr)} 962 @(emb_wrong … WRONG) 963 ] qed. 964 965 lemma initial_state_not_final: ∀ge,s. 947 966 initial_state ge s → 948 967 ¬ ∃r.final_state s r. 949 #ge s H; ncases H;950 #b f ge m E1 E2 E3 E4; @; *; #r H2;951 ninversion H2;952 #r' m' E5 E6; ndestruct (E5);953 nqed.954 955 nlemma initial_step: ∀ge,s,e.968 #ge #s #H cases H; 969 #b #f #ge #m #E1 #E2 #E3 #E4 % *; #r #H2 970 inversion H2; 971 #r' #m' #E5 #E6 destruct (E5); 972 qed. 973 974 lemma initial_step: ∀ge,s,e. 956 975 exec_inf_aux ge (Value ??? 〈E0,s〉) = e → 957 976 ¬(∃r.final_state s r) → 958 977 ∃e'.e = e_step E0 s e'. 959 #ge s e; nrewrite > (exec_inf_aux_unfold …);960 nwhd in ⊢ (??%? → ?); ncases (is_final_state s);961 ##[ #FINAL EXEC NOTFINAL; 962 napply False_ind; napply (absurd ?? NOTFINAL);963 ncases FINAL;964 #r F; @r; napply F;965 ##| #F1 EXEC F2; nwhd in EXEC:(??%?); @; ##[ ##2: nrewrite < EXEC; napply refl ##]966 nqed.967 968 ntheorem exec_inf_equivalence:978 #ge #s #e >(exec_inf_aux_unfold …) 979 whd in ⊢ (??%? → ?); cases (is_final_state s); 980 [ #FINAL #EXEC #NOTFINAL 981 @False_ind @(absurd ?? NOTFINAL) 982 cases FINAL; 983 #r #F %{r} @F 984 | #F1 #EXEC #F2 whd in EXEC:(??%?); % [ 2: <EXEC @refl ] 985 qed. 986 987 theorem exec_inf_equivalence: 969 988 ∀classic:(∀P:Prop.P ∨ ¬P). 970 ∀constructive_indefinite_description:(∀A:Type . ∀P:A→Prop. (∃x. P x) → Σx : A. P x).989 ∀constructive_indefinite_description:(∀A:Type[0]. ∀P:A→Prop. (∃x. P x) → Sig A P). 971 990 ∀p,e. single_exec_of (exec_inf p) e → 972 991 ∃b.execution_matches_behavior e b ∧ exec_program p b. 973 #classic constructive_indefinite_description p e;974 nwhd in ⊢ (?%? → ??(λ_.?(?%?)%));975 nlapply (make_initial_state_sound p);976 nlapply (the_initial_state p);977 ncases (make_initial_state p);978 ##[ #gs; ncases gs; #ge s INITIAL' INITIAL; nwhd in INITIAL ⊢ (?%? → ?);979 ncases INITIAL; #Ege INITIAL'';980 nrewrite > (exec_inf_aux_unfold …);981 nwhd in ⊢ (?%? → ?);982 ncases (is_final_state s);983 ##[ #F; napply False_ind;984 napply (absurd ?? (initial_state_not_final … INITIAL''));985 ncases F; #r F'; @r; napply F';986 ##| #NOTFINAL; nwhd in ⊢ (?%? → ?); ncases e;987 ##[ #tr r m EXEC0 ##| #tr s' e0 EXEC0 ##| #EXEC0 ##| #o k i e EXEC0 ##]988 ncases (se_inv … EXEC0); *; #E1 E2; nrewrite < E1; nrewrite < E2; #EXEC';989 nlapply (behavior_of_execution ??992 #classic #constructive_indefinite_description #p #e 993 whd in ⊢ (?%? → ??(λ_.?(?%?)%)); 994 lapply (make_initial_state_sound p); 995 lapply (the_initial_state p); 996 cases (make_initial_state p); 997 [ #gs cases gs; #ge #s #INITIAL' #INITIAL whd in INITIAL ⊢ (?%? → ?); 998 cases INITIAL; #Ege #INITIAL'' 999 >(exec_inf_aux_unfold …) 1000 whd in ⊢ (?%? → ?); 1001 cases (is_final_state s); 1002 [ #F @False_ind 1003 @(absurd ?? (initial_state_not_final … INITIAL'')) 1004 cases F; #r #F' %{r} @F' 1005 | #NOTFINAL whd in ⊢ (?%? → ?); cases e; 1006 [ #tr #r #m #EXEC0 | #tr #s' #e0 #EXEC0 | #EXEC0 | #o #k #i #e #EXEC0 ] 1007 cases (se_inv … EXEC0); *; #E1 #E2 <E1 <E2 #EXEC' 1008 lapply (behavior_of_execution ?? 990 1009 (execution_characterisation_complete classic constructive_indefinite_description ge s ? EXEC')); 991 *; #b MATCHES; @b; @;//;992 #ge' ; nrewrite > Ege; #Ege'; nrewrite > (?:ge' = ge); ##[ ##2: ndestruct (Ege') skip (INITIAL Ege EXEC0 EXEC'); // ##]993 ninversion MATCHES;994 ##[ #s0 e1 tr1 r m TERM EXEC BEHAVES; nrewrite < EXEC in TERM;995 #TERM ;996 nlapply (exec_state_terminates … TERM); #E1;997 nrewrite > E1 in TERM; #TERM;998 napply (program_terminates (mk_transrel … step) ?? ge s);999 ##[ ##2: napplyINITIAL''1000 ##| ##3: napply (terminates_sound … TERM EXEC');1001 ##| ##skip1002 ##| //;1003 ##]1004 ##| #s0 e tr DIVERGES EXEC E2; nrewrite < EXEC in DIVERGES; #DIVERGES;1005 nlapply (exec_state_diverges … DIVERGES);1006 #E1 ; nrewrite > E1 in DIVERGES; #DIVERGES;1007 ninversion DIVERGES; #tr' s1 s2 e1 e2 INITSTEPS DIVERGING E4 E5 E6;1008 nrewrite < E4 in INITSTEPS ⊢ %; nrewrite < E5 in E6 ⊢ %; #E6 INITSTEPS;1009 ncut (e0 = e1); ##[ ndestruct (E6) skip (MATCHES EXEC0 EXEC'); // ##]1010 #E7 ; nrewrite < E7 in INITSTEPS; #INITSTEPS;1011 ncases (several_steps … INITSTEPS EXEC'); #INITSTAR EXECDIV;1012 napply (program_diverges (mk_transrel … step) ?? ge s … INITIAL'' INITSTAR);1013 napply (silent_sound … DIVERGING EXECDIV);1014 ##| #s0 e tr REACTS EXEC E2; nrewrite < EXEC in REACTS; #REACTS;1015 nlapply (exec_state_reacts … REACTS);1016 #E1 ; nrewrite > E1 in REACTS; #REACTS;1017 ninversion REACTS; #tr' s' e'' REACTING E4 E5;1018 nrewrite < E4 in REACTING ⊢ %; nrewrite < E5; #REACTING E6;1019 ncut (e0 = e''); ##[ ndestruct (E6) skip (MATCHES EXEC0 EXEC'); // ##]1020 #E7 ; nrewrite < E7 in REACTING; #REACTING;1021 napply (program_reacts (mk_transrel … step) ?? ge s … INITIAL'');1022 napply (reacts_sound … REACTING EXEC');1023 ##| #e s1 s2 tr WRONG EXEC E2; nrewrite < EXEC in WRONG; #WRONG;1024 nlapply (exec_state_wrong … WRONG);1025 #E1 ; nrewrite > E1 in WRONG; #WRONG;1026 ninversion WRONG; #tr' s1' s2' e'' GOESWRONG E4 E5 E6 E7;1027 nrewrite < E4 in GOESWRONG ⊢ %; nrewrite < E5; nrewrite < E7; #GOESWRONG;1028 ncut (e0 = e''); ##[ ndestruct (E6) skip (INITIAL MATCHES EXEC0 EXEC'); // ##]1029 #E8 ; nrewrite < E8 in GOESWRONG; #GOESWRONG;1030 nelim (wrong_sound … WRONG EXEC' NOTFINAL); *; #STAR STOP FINAL;1031 napply (program_goes_wrong (mk_transrel … step) ?? ge s … INITIAL'' STAR STOP);1032 #r ; @; #F; napply (absurd ?? FINAL); @r; napply F;1033 ##| #E; ndestruct (E);1034 ##]1035 ##]1036 ##| nwhd in ⊢ ((∀_.? → %) → ?);1037 #NOINIT ; #_; #EXEC;1038 @ (Goes_wrong E0); @;1039 ##[ nwhd in EXEC:(?%?);1040 ncases e in EXEC;1041 ##[ #tr r m EXEC0 ##| #tr s' e0 EXEC0 ##| #EXEC0 ##| #o k i e EXEC0 ##]1042 ncases (se_inv … EXEC0);1043 napply emb_initially_wrong;1044 ##| #ge Ege;1045 napply program_goes_initially_wrong;1046 #s ; @; #INIT; ncases (NOINIT s INIT); #ge' H; napply H;1047 ##]1048 ##] nqed.1049 1010 *; #b #MATCHES %{b} % //; 1011 #ge' >Ege #Ege' >(?:ge' = ge) [ 2: destruct (Ege') skip (INITIAL Ege EXEC0 EXEC'); // ] 1012 inversion MATCHES; 1013 [ #s0 #e1 #tr1 #r #m #TERM #EXEC #BEHAVES <EXEC in TERM 1014 #TERM 1015 lapply (exec_state_terminates … TERM); #E1 1016 >E1 in TERM #TERM 1017 @(program_terminates (mk_transrel … step) ?? ge s) 1018 [ 2: @INITIAL'' 1019 | 3: @(terminates_sound … TERM EXEC') 1020 | skip 1021 | //; 1022 ] 1023 | #s0 #e #tr #DIVERGES #EXEC #E2 <EXEC in DIVERGES #DIVERGES 1024 lapply (exec_state_diverges … DIVERGES); 1025 #E1 >E1 in DIVERGES #DIVERGES 1026 inversion DIVERGES; #tr' #s1 #s2 #e1 #e2 #INITSTEPS #DIVERGING #E4 #E5 #E6 1027 <E4 in INITSTEPS ⊢ % <E5 in E6 ⊢ % #E6 #INITSTEPS 1028 cut (e0 = e1); [ destruct (E6) skip (MATCHES EXEC0 EXEC'); // ] 1029 #E7 <E7 in INITSTEPS #INITSTEPS 1030 cases (several_steps … INITSTEPS EXEC'); #INITSTAR #EXECDIV 1031 @(program_diverges (mk_transrel … step) ?? ge s … INITIAL'' INITSTAR) 1032 @(silent_sound … DIVERGING EXECDIV) 1033 | #s0 #e #tr #REACTS #EXEC #E2 <EXEC in REACTS #REACTS 1034 lapply (exec_state_reacts … REACTS); 1035 #E1 >E1 in REACTS #REACTS 1036 inversion REACTS; #tr' #s' #e'' #REACTING #E4 #E5 1037 <E4 in REACTING ⊢ % <E5 #REACTING #E6 1038 cut (e0 = e''); [ destruct (E6) skip (MATCHES EXEC0 EXEC'); // ] 1039 #E7 <E7 in REACTING #REACTING 1040 @(program_reacts (mk_transrel … step) ?? ge s … INITIAL'') 1041 @(reacts_sound … REACTING EXEC') 1042 | #e #s1 #s2 #tr #WRONG #EXEC #E2 <EXEC in WRONG #WRONG 1043 lapply (exec_state_wrong … WRONG); 1044 #E1 >E1 in WRONG #WRONG 1045 inversion WRONG; #tr' #s1' #s2' #e'' #GOESWRONG #E4 #E5 #E6 #E7 1046 <E4 in GOESWRONG ⊢ % <E5 <E7 #GOESWRONG 1047 cut (e0 = e''); [ destruct (E6) skip (INITIAL Ege MATCHES EXEC0 EXEC'); // ] 1048 #E8 <E8 in GOESWRONG #GOESWRONG 1049 elim (wrong_sound … WRONG EXEC' NOTFINAL); *; #STAR #STOP #FINAL 1050 @(program_goes_wrong (mk_transrel … step) ?? ge s … INITIAL'' STAR STOP) 1051 #r % #F @(absurd ?? FINAL) %{r} @F 1052 | #E destruct (E); 1053 ] 1054 ] 1055 | whd in ⊢ ((∀_.? → %) → ?); 1056 #NOINIT #_ #EXEC 1057 %{ (Goes_wrong E0)} % 1058 [ whd in EXEC:(?%?); 1059 cases e in EXEC; 1060 [ #tr #r #m #EXEC0 | #tr #s' #e0 #EXEC0 | #EXEC0 | #o #k #i #e #EXEC0 ] 1061 cases (se_inv … EXEC0); 1062 @emb_initially_wrong 1063 | #ge #Ege 1064 @program_goes_initially_wrong 1065 #s % #INIT cases (NOINIT s INIT); #ge' #H @H 1066 ] 1067 ] qed. 1068 -
Deliverables/D3.1/C-semantics/CexecSound.ma
r485 r487 1 1 include "Cexec.ma". 2 2 3 include "Plogic/connectives.ma". 4 5 nlemma exec_bool_of_val_sound: ∀v,ty,r.3 (*include "Plogic/connectives.ma".*) 4 5 lemma exec_bool_of_val_sound: ∀v,ty,r. 6 6 exec_bool_of_val v ty = OK ? r → bool_of_val v ty (of_bool r). 7 #v ty r;8 ncases v; ##[ ##| #i ##| #f ##| #r1 ##| #r b of ##]9 ncases ty; ##[ ##2,11,20,29,38: #sz sg ##| ##3,12,21,30,39: #sz ##| ##4,13,22,31,40: #r 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 ##]10 #H ; nwhd in H:(??%?); ndestruct;11 ##[ nlapply (eq_spec i zero); nelim (eq i zero);12 ##[ #e; nrewrite > e; napply bool_of_val_false;//;13 ##| #ne; napply bool_of_val_true;/2/;14 ##]15 ##| nelim (eq_dec f Fzero);16 ##[ #e; nrewrite > e; nrewrite > (Feq_zero_true …); napply bool_of_val_false;//;17 ##| #ne; nrewrite > (Feq_zero_false …); //; napply bool_of_val_true;/2/;18 ##]19 ##| napply bool_of_val_false;//20 ##| napply bool_of_val_true;//21 ##] nqed.22 23 nlemma bool_val_distinct: Vtrue ≠ Vfalse.24 @; #H; nwhd in H:(??%%); ndestruct; napply (absurd ? e0 one_not_zero); 25 nqed.26 27 nlemma bool_of: ∀v,ty,b. bool_of_val v ty (of_bool b) →7 #v #ty #r 8 cases v; [ | #i | #f | #r1 | #r #b #of ] 9 cases ty; [ 2,11,20,29,38: #sz #sg | 3,12,21,30,39: #sz | 4,13,22,31,40: #r #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 ] 10 #H whd in H:(??%?); destruct; 11 [ lapply (eq_spec i zero); elim (eq i zero); 12 [ #e >e @bool_of_val_false //; 13 | #ne @bool_of_val_true /2/; 14 ] 15 | elim (eq_dec f Fzero); 16 [ #e >e >(Feq_zero_true …) @bool_of_val_false //; 17 | #ne >(Feq_zero_false …) //; @bool_of_val_true /2/; 18 ] 19 | @bool_of_val_false // 20 | @bool_of_val_true // 21 ] qed. 22 23 lemma bool_val_distinct: Vtrue ≠ Vfalse. 24 % #H whd in H:(??%%); destruct; @(absurd ? e0 one_not_zero) 25 qed. 26 27 lemma bool_of: ∀v,ty,b. bool_of_val v ty (of_bool b) → 28 28 if b then is_true v ty else is_false v ty. 29 #v ty b; ncases b; #H; ninversion H; #v' ty' H' ev et ev;//;30 napply False_ind; napply (absurd ? ev ?); 31 ##[ ##2: napply sym_neq ##] napply bool_val_distinct; 32 nqed.33 34 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'.35 #m i ty ty' v';36 nwhd in ⊢ (??%? → ?);37 nlapply (eq_spec i zero); ncases (eq i zero);38 ##[ #e; nrewrite > e; 39 ncases ty; ##[ ##| #sz sg ##| #fs ##| #sp ty ##| #sp ty n ##| #args rty ##| #id fs ##| #id fs ##| #r id ##]40 nwhd in ⊢ (??%? → ?); #H; ndestruct;41 ncases ty' in H ⊢ %; ##[ ##| #sz sg ##| #fs ##| #sp ty ##| #sp ty n ##| #args rty ##| #id fs ##| #id fs ##| #r id ##]42 nwhd in ⊢ (??%? → ?); #H; ndestruct (H); napply cast_ip_z;//;43 ##| #_; nwhd in ⊢ (??%? → ?); #H; ndestruct44 ##]45 nqed.46 47 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'.48 #m v ty ty' v';49 ncases v;50 ##[ #H; nwhd in H:(??%?); ndestruct;51 ##| #i; ncases ty;52 ##[ #H; nwhd in H:(??%?); ndestruct;53 ##| ##3: #a; #H; nwhd in H:(??%?); ndestruct;54 ##| ##7,8,9: #a b; #H; nwhd in H:(??%?); ndestruct;55 ##| #sz1 si1; ncases ty';56 ##[ #H; nwhd in H:(??%?); ndestruct;57 ##| ##3: #a; #H; nwhd in H:(??%?); ndestruct; //58 ##| ##2,7,8,9: #a b; #H; nwhd in H:(??%?); ndestruct; //59 ##| ##4,5,6: ##[ #sp ty''; nletin t ≝ (Tpointer sp ty'')60 ##| #sp ty'' n; nletin t ≝ (Tarray sp ty'' n)61 ##| #args rty; nletin t ≝ (Tfunction args rty) ##]62 nwhd in ⊢ (??%? → ?);63 nlapply (try_cast_null_sound m i (Tint sz1 si1) t v');64 ncases (try_cast_null m i (Tint sz1 si1) t);65 ##[ ##1,3,5: #v''; #H' e; napply H'; napply e;66 ##| ##*: #_; nwhd in ⊢ (??%? → ?); #H; ndestruct (H);67 ##]68 ##]69 ##| ##*: ##[ #sp ty''; nletin t ≝ (Tpointer sp ty'')70 ##| #sp ty'' n; nletin t ≝ (Tarray sp ty'' n)71 ##| #args rty; nletin t ≝ (Tfunction args rty) ##]72 nwhd in ⊢ (??%? → ?);73 nlapply (try_cast_null_sound m i t ty' v');74 ncases (try_cast_null m i t ty');75 ##[ ##1,3,5: #v''; #H' e; napply H'; napply e;76 ##| ##*: #_; nwhd in ⊢ (??%? → ?); #H; ndestruct (H);77 ##]78 ##]79 ##| #f; ncases ty; ##[ ##3: #x; ##| ##2,4,6,7,8,9: #x y; ##| ##5: #x y z; ##]80 ##[ ncases ty'; ##[ #e; ##| ##3: #a e; ##| ##2,4,6,7,8,9: #a b e; ##| #a b c e; ##]81 nwhd in e:(??%?); ndestruct; //;82 ##| ##*: #e; nwhd in e:(??%?); ndestruct83 ##]84 ##| #r; ncases ty; ##[ ##3: #x; ##| ##2,4,6,7,8,9: #x y; ##| ##5: #x y z; ##]85 nwhd in ⊢ (??%? → ?); #H; ndestruct; ncases (eq_region_dec r ?) in H;86 nwhd in ⊢ (? → ??%? → ?); #H1 H2; ndestruct;87 ncases ty' in H2; nnormalize; ntry #a; ntry #b; ntry #c; ntry #d; ndestruct;88 napply cast_pp_z;//;89 ##| #sp b of; nwhd in ⊢ (??%? → ?); #e; 90 nelim (bind_inversion ????? e); #s; *; #es e';91 nelim (bind_inversion ????? e'); #u; *; #eu e'';92 nelim (bind_inversion ????? e''); #s'; *; #es' e''';93 ncut (type_region ty s);94 ##[ ncases ty in es:(??%?) ⊢ %; ##[ #e; ##| ##3: #a e; ##| ##2,4,6,7,8,9: #a b e; ##| #a b c e; ##]95 nwhd in e:(??%?); ndestruct; //;96 ##| #Hty;97 ncut (type_region ty' s');98 ##[ ncases ty' in es' ⊢ %; ##[ #e; ##| ##3: #a e; ##| ##2,4,6,7,8,9: #a b e; ##| #a b c e; ##]99 nwhd in e:(??%?); ndestruct; //;100 ##| #Hty';101 ncut (s = sp). nelim (eq_region_dec sp s) in eu; //; nnormalize; #_; #e; ndestruct.102 #e ; nrewrite < e;103 nwhd in match (is_pointer_compat ??) in e''';104 ncases (pointer_compat_dec (block_space m b) s') in e'''; #Hcompat e''';105 nwhd in e''':(??%?); ndestruct (e'''); /2/106 ##]107 ##]108 ##] nqed.109 110 nlet rec expr_lvalue_ind29 #v #ty #b cases b; #H inversion H; #v' #ty' #H' #ev #et #ev //; 30 @False_ind @(absurd ? ev ?) 31 [ 2: @sym_neq ] @bool_val_distinct 32 qed. 33 34 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'. 35 #m #i #ty #ty' #v' 36 whd in ⊢ (??%? → ?); 37 lapply (eq_spec i zero); cases (eq i zero); 38 [ #e >e 39 cases ty; [ | #sz #sg | #fs | #sp #ty | #sp #ty #n | #args #rty | #id #fs | #id #fs | #r #id ] 40 whd in ⊢ (??%? → ?); #H destruct; 41 cases ty' in H ⊢ %; [ | #sz #sg | #fs | #sp #ty | #sp #ty #n | #args #rty | #id #fs | #id #fs | #r #id ] 42 whd in ⊢ (??%? → ?); #H destruct (H); @cast_ip_z //; 43 | #_ whd in ⊢ (??%? → ?); #H destruct 44 ] 45 qed. 46 47 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'. 48 #m #v #ty #ty' #v' 49 cases v; 50 [ #H whd in H:(??%?); destruct; 51 | #i cases ty; 52 [ #H whd in H:(??%?); destruct; 53 | 3: #a #H whd in H:(??%?); destruct; 54 | 7,8,9: #a #b #H whd in H:(??%?); destruct; 55 | #sz1 #si1 cases ty'; 56 [ #H whd in H:(??%?); destruct; 57 | 3: #a #H whd in H:(??%?); destruct; // 58 | 2,7,8,9: #a #b #H whd in H:(??%?); destruct; // 59 | 4,5,6: [ #sp #ty'' letin t ≝ (Tpointer sp ty'') 60 | #sp #ty'' #n letin t ≝ (Tarray sp ty'' n) 61 | #args #rty letin t ≝ (Tfunction args rty) ] 62 whd in ⊢ (??%? → ?); 63 lapply (try_cast_null_sound m i (Tint sz1 si1) t v'); 64 cases (try_cast_null m i (Tint sz1 si1) t); 65 [ 1,3,5: #v'' #H' #e @H' @e 66 | *: #_ whd in ⊢ (??%? → ?); #H destruct (H); 67 ] 68 ] 69 | *: [ #sp #ty'' letin t ≝ (Tpointer sp ty'') 70 | #sp #ty'' #n letin t ≝ (Tarray sp ty'' n) 71 | #args #rty letin t ≝ (Tfunction args rty) ] 72 whd in ⊢ (??%? → ?); 73 lapply (try_cast_null_sound m i t ty' v'); 74 cases (try_cast_null m i t ty'); 75 [ 1,3,5: #v'' #H' #e @H' @e 76 | *: #_ whd in ⊢ (??%? → ?); #H destruct (H); 77 ] 78 ] 79 | #f cases ty; [ 3: #x | 2,4,6,7,8,9: #x #y | 5: #x #y #z ] 80 [ cases ty'; [ #e | 3: #a #e | 2,4,6,7,8,9: #a #b #e | #a #b #c #e ] 81 whd in e:(??%?); destruct; //; 82 | *: #e whd in e:(??%?); destruct 83 ] 84 | #r cases ty; [ 3: #x | 2,4,6,7,8,9: #x #y | 5: #x #y #z ] 85 whd in ⊢ (??%? → ?); #H destruct; cases (eq_region_dec r ?) in H; 86 whd in ⊢ (? → ??%? → ?); #H1 #H2 destruct; 87 cases ty' in H2; normalize; try #a try #b try #c try #d destruct; 88 @cast_pp_z //; 89 | #sp #b #of whd in ⊢ (??%? → ?); #e 90 elim (bind_inversion ????? e); #s *; #es #e' 91 elim (bind_inversion ????? e'); #u *; #eu #e'' 92 elim (bind_inversion ????? e''); #s' *; #es' #e''' 93 cut (type_region ty s); 94 [ cases ty in es:(??%?) ⊢ %; [ #e | 3: #a #e | 2,4,6,7,8,9: #a #b #e | #a #b #c #e ] 95 whd in e:(??%?); destruct; //; 96 | #Hty 97 cut (type_region ty' s'); 98 [ cases ty' in es' ⊢ %; [ #e | 3: #a #e | 2,4,6,7,8,9: #a #b #e | #a #b #c #e ] 99 whd in e:(??%?); destruct; //; 100 | #Hty' 101 cut (s = sp). elim (eq_region_dec sp s) in eu; //; normalize; #_ #e destruct. 102 #e <e 103 whd in match (is_pointer_compat ??) in e'''; 104 cases (pointer_compat_dec (block_space m b) s') in e'''; #Hcompat #e''' 105 whd in e''':(??%?); destruct (e'''); /2/ 106 ] 107 ] 108 ] qed. 109 110 let rec expr_lvalue_ind 111 111 (P:expr → Prop) 112 112 (Q:expr_descr → type → Prop) … … 172 172 | 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'') ] 173 173 | _ ⇒ xx ? ty ? 174 ]. nwhd; napply I; nqed.175 176 177 ntheorem exec_expr_sound: ∀ge:genv. ∀en:env. ∀m:mem. ∀e:expr.174 ]. whd; @I qed. 175 176 177 theorem exec_expr_sound: ∀ge:genv. ∀en:env. ∀m:mem. ∀e:expr. 178 178 (P_res ? (λx.eval_expr ge en m e (\fst x) (\snd x)) (exec_expr ge en m e)). 179 #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); 180 ##[ ##1,2: #ty c; nwhd; //; 179 #ge #en #m #e @(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) 180 (* XXX // fails [ 1,2: #ty #c whd // *) 181 [ #ty #c whd % 182 | #ty #c whd %2 181 183 (* expressions that are lvalues *) 182 ##| #e' ty; ncases e'; //; ##[ #i He' ##| #e He' ##| #e i He' ##] nwhd in He' ⊢ %; 183 napply bind2_OK; #x; ncases x; #y; ncases y; #sp loc ofs tr H; 184 napply opt_bind_OK; #vl evl; nwhd in evl:(??%?); napply (eval_Elvalue … evl); 185 nrewrite > H in He'; #He'; napply He'; 186 ##| #v ty; 187 nwhd in ⊢ (???%); 188 nlapply (refl ? (get ident PTree block v en)); 189 ncases (get ident PTree block v en) in ⊢ (???% → %); 190 ##[ #eget; napply opt_bind_OK; #sploc; ncases sploc; #sp loc efind; 191 nwhd; napply (eval_Evar_global … eget efind); 192 ##| #loc eget; napply (eval_Evar_local … eget); 193 ##] 194 ##| #ty e He; nwhd in ⊢ (???%); 195 napply bind2_OK; #v; ncases v; //; #sp l ofs tr Hv; nwhd; 196 napply eval_Ederef; nrewrite > Hv in He; #He; napply He; 197 ##| #ty e'' ty'' He''; napply bind2_OK; #x; ncases x; #y; ncases y; #sp loc ofs tr H; 198 nwhd; napply eval_Eaddrof; nrewrite > H in He''; #He''; napply He''; 199 ##| #ty op e1 He1; napply bind2_OK; #v1 tr Hv1; 200 napply opt_bind_OK; #v ev; 201 napply (eval_Eunop … ev); 202 nrewrite > Hv1 in He1; #He1; napply He1; 203 ##| #ty op e1 e2 He1 He2; 204 napply bind2_OK; #v1 tr1 ev1; nrewrite > ev1 in He1; #He1; 205 napply bind2_OK; #v2 tr2 ev2; nrewrite > ev2 in He2; #He2; 206 napply opt_bind_OK; #v ev; nwhd in He1 He2; nwhd; 207 napply (eval_Ebinop … He1 He2 ev); 208 ##| #ty ty' e' He'; 209 napply bind2_OK; #v tr Hv; nrewrite > Hv in He'; #He'; 210 napply bind_OK; #v' ev'; 211 napply (eval_Ecast … He' ?); 212 /2/; 213 ##| #ty e1 e2 e3 He1 He2 He3; 214 napply bind2_OK; #vb tr1 Hvb; nrewrite > Hvb in He1; #He1; 215 napply bind_OK; #b; 216 ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; 217 napply bind2_OK; #v tr Hv; 218 ##[ nrewrite > Hv in He2; #He2; nwhd in He2 Hv:(??%?) ⊢%; 219 napply (eval_Econdition_true … He1 ? He2); napply (bool_of ??? Hb); 220 ##| nrewrite > Hv in He3; #He3; nwhd in He3 Hv:(??%?) ⊢%; 221 napply (eval_Econdition_false … He1 ? He3); napply (bool_of ??? Hb); 222 ##] 223 ##| #ty e1 e2 He1 He2; 224 napply bind2_OK; #v1 tr1 Hv1; nrewrite > Hv1 in He1; #He1; 225 napply bind_OK; #b1; ncases b1; #eb1; nlapply (exec_bool_of_val_sound … eb1); #Hb1; 226 ##[ napply bind2_OK; #v2 tr2 Hv2; nrewrite > Hv2 in He2; #He2; 227 napply bind_OK; #b2 eb2; 228 napply (eval_Eandbool_2 … He1 … He2); 229 ##[ napply (bool_of … Hb1); ##| napply (exec_bool_of_val_sound … eb2); ##] 230 ##| napply (eval_Eandbool_1 … He1); napply (bool_of … Hb1); 231 ##] 232 ##| #ty e1 e2 He1 He2; 233 napply bind2_OK; #v1 tr1 Hv1; nrewrite > Hv1 in He1; #He1; 234 napply bind_OK; #b1; ncases b1; #eb1; nlapply (exec_bool_of_val_sound … eb1); #Hb1; 235 ##[ napply (eval_Eorbool_1 … He1); napply (bool_of … Hb1); 236 ##| napply bind2_OK; #v2 tr2 Hv2; nrewrite > Hv2 in He2; #He2; 237 napply bind_OK; #b2 eb2; 238 napply (eval_Eorbool_2 … He1 … He2); 239 ##[ napply (bool_of … Hb1); ##| napply (exec_bool_of_val_sound … eb2); ##] 240 ##] 241 ##| #ty ty'; nwhd; // 242 ##| #ty e' ty' i; ncases ty'; //; 243 ##[ #id fs He'; napply bind2_OK; #x; ncases x; #sp l ofs H; 244 napply bind_OK; #delta Hdelta; nrewrite > H in He'; #He'; 245 napply (eval_Efield_struct … He' (refl ??) Hdelta); 246 ##| #id fs He'; napply bind2_OK; #x; ncases x; #sp l ofs H; 247 nrewrite > H in He'; #He'; 248 napply (eval_Efield_union … He' (refl ??)); 249 ##] 250 ##| #ty l e' He'; napply bind2_OK; #v tr1 H; nrewrite > H in He'; #He'; 251 napply (eval_Ecost … He'); 184 | #e' #ty cases e'; //; [ #i #He' | #e #He' | #e #i #He' ] whd in He' ⊢ %; 185 @bind2_OK #x cases x; #y cases y; #sp #loc #ofs #tr #H 186 @opt_bind_OK #vl #evl whd in evl:(??%?); @(eval_Elvalue … evl) 187 >H in He' #He' @He' 188 | #v #ty 189 whd in ⊢ (???%); 190 lapply (refl ? (get ident PTree block v en)); 191 cases (get ident PTree block v en) in ⊢ (???% → %); 192 [ #eget @opt_bind_OK #sploc cases sploc; #sp #loc #efind 193 whd; @(eval_Evar_global … eget efind) 194 | #loc #eget @(eval_Evar_local … eget) 195 ] 196 | #ty #e #He whd in ⊢ (???%); 197 @bind2_OK #v cases v; //; #sp #l #ofs #tr #Hv whd; 198 @eval_Ederef >Hv in He #He @He 199 | #ty #e'' #ty'' #He'' @bind2_OK #x cases x; #y cases y; #sp #loc #ofs #tr #H 200 whd; @eval_Eaddrof >H in He'' #He'' @He'' 201 | #ty #op #e1 #He1 @bind2_OK #v1 #tr #Hv1 202 @opt_bind_OK #v #ev 203 @(eval_Eunop … ev) 204 >Hv1 in He1 #He1 @He1 205 | #ty #op #e1 #e2 #He1 #He2 206 @bind2_OK #v1 #tr1 #ev1 >ev1 in He1 #He1 207 @bind2_OK #v2 #tr2 #ev2 >ev2 in He2 #He2 208 @opt_bind_OK #v #ev whd in He1 He2; whd; 209 @(eval_Ebinop … He1 He2 ev) 210 | #ty #ty' #e' #He' 211 @bind2_OK #v #tr #Hv >Hv in He' #He' 212 @bind_OK #v' #ev' 213 @(eval_Ecast … He' ?) 214 (* XXX /2/; *) 215 @(exec_cast_sound … ev') 216 | #ty #e1 #e2 #e3 #He1 #He2 #He3 217 @bind2_OK #vb #tr1 #Hvb >Hvb in He1 #He1 218 @bind_OK #b 219 cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb 220 @bind2_OK #v #tr #Hv 221 [ >Hv in He2 #He2 whd in He2 Hv:(??%?) ⊢%; 222 @(eval_Econdition_true … He1 ? He2) @(bool_of ??? Hb) 223 | >Hv in He3 #He3 whd in He3 Hv:(??%?) ⊢%; 224 @(eval_Econdition_false … He1 ? He3) @(bool_of ??? Hb) 225 ] 226 | #ty #e1 #e2 #He1 #He2 227 @bind2_OK #v1 #tr1 #Hv1 >Hv1 in He1 #He1 228 @bind_OK #b1 cases b1; #eb1 lapply (exec_bool_of_val_sound … eb1); #Hb1 229 [ @bind2_OK #v2 #tr2 #Hv2 >Hv2 in He2 #He2 230 @bind_OK #b2 #eb2 231 @(eval_Eandbool_2 … He1 … He2) 232 [ @(bool_of … Hb1) | @(exec_bool_of_val_sound … eb2) ] 233 | @(eval_Eandbool_1 … He1) @(bool_of … Hb1) 234 ] 235 | #ty #e1 #e2 #He1 #He2 236 @bind2_OK #v1 #tr1 #Hv1 >Hv1 in He1 #He1 237 @bind_OK #b1 cases b1; #eb1 lapply (exec_bool_of_val_sound … eb1); #Hb1 238 [ @(eval_Eorbool_1 … He1) @(bool_of … Hb1) 239 | @bind2_OK #v2 #tr2 #Hv2 >Hv2 in He2 #He2 240 @bind_OK #b2 #eb2 241 @(eval_Eorbool_2 … He1 … He2) 242 [ @(bool_of … Hb1) | @(exec_bool_of_val_sound … eb2) ] 243 ] 244 | #ty #ty' whd; (* XXX //*) @eval_Esizeof 245 | #ty #e' #ty' #i cases ty'; //; 246 [ #id #fs #He' @bind2_OK #x cases x; #sp #l #ofs #H 247 @bind_OK #delta #Hdelta >H in He' #He' 248 @(eval_Efield_struct … He' (refl ??) Hdelta) 249 | #id #fs #He' @bind2_OK #x cases x; #sp #l #ofs #H 250 >H in He' #He' 251 @(eval_Efield_union … He' (refl ??)) 252 ] 253 | #ty #l #e' #He' @bind2_OK #v #tr1 #H >H in He' #He' 254 @(eval_Ecost … He') 252 255 (* exec_lvalue fails on non-lvalues. *) 253 ##| #e' ty; ncases e';254 ##[ ##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 ##]255 napply I;256 ##] nqed.257 258 nlemma addrof_eval_lvalue: ∀ge,en,m,e,sp,loc,off,tr,ty.256 | #e' #ty cases e'; 257 [ 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 ] 258 @I 259 ] qed. 260 261 lemma addrof_eval_lvalue: ∀ge,en,m,e,sp,loc,off,tr,ty. 259 262 eval_expr ge en m (Expr (Eaddrof e) ty) (Vptr sp loc off) tr → 260 263 eval_lvalue ge en m e sp loc off tr. 261 #ge en m e sp loc off tr ty H; ninversion H;262 ##[ ##1,2,5: #a b H; napply False_ind; ndestruct (H);263 ##| #a b c d e f g H1 i H2; nrewrite < H2 in H1; #H1; napply False_ind; 264 napply (eval_lvalue_inv_ind … H1);265 ##[ #a b c d bad; ndestruct (bad);266 ##| #a b c d e f bad; ndestruct (bad);267 ##| #a b c d e f g bad; ndestruct (bad);268 ##| #a b c d e f g h i j k l m n bad; napply False_ind; ndestruct (bad);269 ##| #a b c d e f g h i j k l bad; ndestruct (bad);270 ##]271 ##| #e' ty' sp' loc' ofs' tr' H e1 e2 e3; ndestruct (e1 e2 e3); napply H; 272 ##| #a b c d e f g h i bad; ndestruct (bad);273 ##| #a b c d e f g h i j k l k l bad; ndestruct (bad);274 ##| #a b c d e f g h i j k l m bad; ndestruct (bad);275 ##| #a b c d e f g h i j k l m bad; ndestruct (bad);276 ##| #a b c d e f g h bad; ndestruct (bad);277 ##| #a b c d e f g h i j k l m n bad; ndestruct (bad);278 ##| #a b c d e f g h bad; ndestruct (bad);279 ##| #a b c d e f g h i j k l m n bad; ndestruct (bad);280 ##| #a b c d e f g h i bad; ndestruct (bad);281 ##| #a b c d e f g bad; ndestruct (bad);282 ##] nqed.283 284 nlemma addrof_exec_lvalue: ∀ge,en,m,e,sp,loc,off,tr,ty.264 #ge #en #m #e #sp #loc #off #tr #ty #H inversion H; 265 [ 1,2,5: #a #b #H @False_ind destruct (H); 266 | #a #b #c #d #e #f #g #H1 #i #H2 <H2 in H1 #H1 @False_ind 267 @(eval_lvalue_inv_ind … H1) 268 [ #a #b #c #d #bad destruct (bad); 269 | #a #b #c #d #e #f #bad destruct (bad); 270 | #a #b #c #d #e #f #g #bad destruct (bad); 271 | #a #b #c #d #e #f #g #h #i #j #k #l #m #n #bad @False_ind destruct (bad); 272 | #a #b #c #d #e #f #g #h #i #j #k #l #bad destruct (bad); 273 ] 274 | #e' #ty' #sp' #loc' #ofs' #tr' #H #e1 #e2 #e3 destruct (e1 e2 e3); @H 275 | #a #b #c #d #e #f #g #h #i #bad destruct (bad); 276 | #a #b #c #d #e #f #g #h #i #j #k #l #k #l #bad destruct (bad); 277 | #a #b #c #d #e #f #g #h #i #j #k #l #m #bad destruct (bad); 278 | #a #b #c #d #e #f #g #h #i #j #k #l #m #bad destruct (bad); 279 | #a #b #c #d #e #f #g #h #bad destruct (bad); 280 | #a #b #c #d #e #f #g #h #i #j #k #l #m #n #bad destruct (bad); 281 | #a #b #c #d #e #f #g #h #bad destruct (bad); 282 | #a #b #c #d #e #f #g #h #i #j #k #l #m #n #bad destruct (bad); 283 | #a #b #c #d #e #f #g #h #i #bad destruct (bad); 284 | #a #b #c #d #e #f #g #bad destruct (bad); 285 ] qed. 286 287 lemma addrof_exec_lvalue: ∀ge,en,m,e,sp,loc,off,tr,ty. 285 288 exec_lvalue ge en m e = OK ? 〈〈〈sp,loc〉,off〉,tr〉 → 286 289 exec_expr ge en m (Expr (Eaddrof e) ty) = OK ? 〈Vptr sp loc off, tr〉. 287 #ge en m e sp loc off tr ty H; nwhd in ⊢ (??%?);288 nrewrite > H;//;289 nqed.290 291 ntheorem exec_lvalue_sound: ∀ge,en,m,e.290 #ge #en #m #e #sp #loc #off #tr #ty #H whd in ⊢ (??%?); 291 >H //; 292 qed. 293 294 theorem exec_lvalue_sound: ∀ge,en,m,e. 292 295 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). 293 #ge en m e; nlapply (refl ? (exec_lvalue ge en m e));294 ncases (exec_lvalue ge en m e) in ⊢ (???% → %);295 ##[ #x; ncases x; #y; ncases y; #z; ncases z; #sp loc off tr H; nwhd;296 napply (addrof_eval_lvalue … Tvoid);297 nlapply (addrof_exec_lvalue … Tvoid H); #H';298 nlapply (exec_expr_sound ge en m (Expr (Eaddrof e) Tvoid));299 nrewrite > H'; #H''; napply H'';300 ##| #_; nwhd; napply I; 301 ##] nqed.296 #ge #en #m #e lapply (refl ? (exec_lvalue ge en m e)); 297 cases (exec_lvalue ge en m e) in ⊢ (???% → %); 298 [ #x cases x; #y cases y; #z cases z; #sp #loc #off #tr #H whd; 299 @(addrof_eval_lvalue … Tvoid) 300 lapply (addrof_exec_lvalue … Tvoid H); #H' 301 lapply (exec_expr_sound ge en m (Expr (Eaddrof e) Tvoid)); 302 >H' #H'' @H'' 303 | #_ whd; @I 304 ] qed. 302 305 303 306 (* Plain equality versions of the above *) 304 307 305 ndefinition exec_expr_sound' ≝ λge,en,m,e,v.308 definition exec_expr_sound' ≝ λge,en,m,e,v. 306 309 λH:exec_expr ge en m e = OK ? v. 307 310 P_res_to_P ???? (exec_expr_sound ge en m e) H. 308 311 309 ndefinition exec_lvalue_sound' ≝ λge,en,m,e,sp,loc,off,tr.312 definition exec_lvalue_sound' ≝ λge,en,m,e,sp,loc,off,tr. 310 313 λH:exec_lvalue ge en m e = OK ? 〈〈〈sp,loc〉,off〉,tr〉. 311 314 P_res_to_P ???? (exec_lvalue_sound ge en m e) H. 312 315 313 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). 314 #ge e m l; nelim l; 315 nwhd; //; 316 #e1 es; #IH; 317 napply bind2_OK; #v tr1 Hv; 318 napply bind2_OK; #vs tr2 Hvs; 319 nwhd; napply eval_Econs; 320 ##[ napply (P_res_to_P … (exec_expr_sound ge e m e1) Hv); 321 ##| napply (P_res_to_P … IH Hvs); 322 ##] nqed. 323 324 ntheorem exec_step_sound: ∀ge,st. 316 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). 317 #ge #e #m #l elim l; 318 whd; (* XXX //; *) 319 [ % 320 | #e1 #es #IH 321 @bind2_OK #v #tr1 #Hv 322 @bind2_OK #vs #tr2 #Hvs 323 whd; @eval_Econs 324 [ @(P_res_to_P … (exec_expr_sound ge e m e1) Hv) 325 | @(P_res_to_P … IH Hvs) 326 ] 327 ] qed. 328 329 lemma exec_alloc_variables_sound : ∀l,en,m,en',m'. 330 exec_alloc_variables en m l = 〈en',m'〉 → 331 alloc_variables en m l en' m'. 332 #l elim l 333 [ #en #m #en' #m' #EXEC whd in EXEC:(??%?); destruct % 334 | * #id #ty #t #IH #en #m #en' #m' 335 lapply (refl ? (alloc m O (sizeof ty) Any)) #ALLOC #EXEC 336 whd in EXEC:(??%?) ALLOC:(???%) 337 @(alloc_variables_cons … ALLOC) 338 @IH @EXEC 339 qed. 340 341 lemma exec_bind_parameters_sound : ∀ps,vs,en,m. 342 P_res ? (λm'. bind_parameters en m ps vs m') (exec_bind_parameters en m ps vs). 343 #ps elim ps 344 [ * // 345 | * #id #ty #ps' #IH * 346 [ // 347 | #v #vs #en #m 348 @opt_bind_OK #b #GET 349 @opt_bind_OK #m' #STORE 350 lapply (refl ? (exec_bind_parameters en m' ps' vs)) 351 cases (exec_bind_parameters en m' ps' vs) in ⊢ (???% → %) [2: #_ %] 352 #m'' #BIND 353 @(bind_parameters_cons … GET STORE) 354 lapply (IH vs en m') whd in ⊢ (% → ?) >BIND // 355 ] 356 ] qed. 357 358 lemma check_eventval_list_sound : ∀vs,tys. 359 P_res ? (λevs. eventval_list_match evs tys vs) (check_eventval_list vs tys). 360 #vs0 elim vs0 361 [ * // 362 | #v #vs #IH * 363 [ // 364 | #ty #tys whd in ⊢ (???%) 365 cases ty cases v // #v' 366 @bind_OK #evs #CHECK 367 @(evl_match_cons ??????? (P_res_to_P … CHECK)) // 368 ] 369 ] qed. 370 371 theorem exec_step_sound: ∀ge,st. 325 372 P_io ??? (λr. step ge st (\fst r) (\snd r)) (exec_step ge st). 326 #ge st; ncases st; 327 ##[ #f s k e m; ncases s; 328 ##[ ncases k; 329 ##[ nwhd in ⊢ (?????%); 330 nlapply (refl ? (fn_return f)); ncases (fn_return f) in ⊢ (???% → %); 331 //; #H; nwhd; 332 napply step_skip_call; //; 333 ##| #s' k'; nwhd; //; 334 ##| #ex s' k'; napply step_skip_or_continue_while; @; //; 335 ##| #ex s' k'; 336 napply res_bindIO2_OK; #v tr Hv; 337 nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); 338 nlapply (refl ? bexpr); 339 ncases bexpr in ⊢ (???% → %); 340 ##[ #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; 341 nwhd in ⊢ (?????%); 342 ##[ napply (step_skip_or_continue_dowhile_true … (exec_expr_sound' … Hv)); 343 ##[ @; // ##| napply (bool_of … Hb); ##] 344 ##| napply (step_skip_or_continue_dowhile_false … (exec_expr_sound' … Hv)); 345 ##[ @; // ##| napply (bool_of … Hb); ##] 346 ##] 347 ##| #_; //; 348 ##] 349 ##| #ex s1 s2 k'; napply step_skip_or_continue_for2; @; //; 350 ##| #ex s1 s2 k'; napply step_skip_for3; 351 ##| #k'; napply step_skip_break_switch; @; //; 352 ##| #r f' e' k'; nwhd in ⊢ (?????%); 353 nlapply (refl ? (fn_return f)); ncases (fn_return f) in ⊢ (???% → %); 354 //; #H; nwhd; 355 napply step_skip_call; //; 356 ##] 357 ##| #ex1 ex2; 358 napply res_bindIO2_OK; #x; ncases x; #y; ncases y; #pcl loc ofs tr1 Hlval; 359 napply res_bindIO2_OK; #v2 tr2 Hv2; 360 napply opt_bindIO_OK; #m' em'; 361 nwhd; napply (step_assign … (exec_lvalue_sound' … Hlval) (exec_expr_sound' … Hv2) em'); 362 ##| #lex fex args; 363 napply res_bindIO2_OK; #vf tr1 Hvf0; nlapply (exec_expr_sound' … Hvf0); #Hvf; 364 napply res_bindIO2_OK; #vargs tr2 Hvargs0; nlapply (P_res_to_P ???? (exec_exprlist_sound …) Hvargs0); #Hvargs; 365 napply opt_bindIO_OK; #fd efd; 366 napply bindIO_OK; #ety; 367 ncases lex; nwhd; 368 ##[ napply (step_call_none … Hvf Hvargs efd ety); 369 ##| #lhs'; 370 napply res_bindIO2_OK; #x; ncases x; #y; ncases y; #pcl loc ofs tr3 Hlocofs; 371 nwhd; napply (step_call_some … (exec_lvalue_sound' … Hlocofs) Hvf Hvargs efd ety); 372 ##] 373 ##| #s1 s2; nwhd; //; 374 ##| #ex s1 s2; 375 napply res_bindIO2_OK; #v tr Hv; 376 nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); 377 nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); //; 378 #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; 379 ##[ napply (step_ifthenelse_true … (exec_expr_sound' … Hv)); napply (bool_of … Hb); 380 ##| napply (step_ifthenelse_false … (exec_expr_sound' … Hv)); napply (bool_of … Hb) 381 ##] 382 ##| #ex s'; 383 napply res_bindIO2_OK; #v tr Hv; 384 nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); 385 nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); //; 386 #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; 387 ##[ napply (step_while_true … (exec_expr_sound' … Hv)); napply (bool_of … Hb); 388 ##| napply (step_while_false … (exec_expr_sound' … Hv)); napply (bool_of … Hb); 389 ##] 390 ##| #ex s'; nwhd; //; 391 ##| #s1 ex s2 s3; nwhd in ⊢ (?????%); nelim (is_Sskip s1); #Hs1; nwhd in ⊢ (?????%); 392 ##[ nrewrite > Hs1; 393 napply res_bindIO2_OK; #v tr Hv; 394 nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); 395 nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); //; 396 #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; 397 ##[ napply (step_for_true … (exec_expr_sound' … Hv)); napply (bool_of … Hb); 398 ##| napply (step_for_false … (exec_expr_sound' … Hv)); napply (bool_of … Hb); 399 ##] 400 ##| napply step_for_start; //; 401 ##] 402 ##| nwhd in ⊢ (?????%); ncases k; //; 403 ##[ #s' k'; nwhd; //; 404 ##| #ex s' k'; nwhd; //; 405 ##| #ex s' k'; nwhd; //; 406 ##| #ex s1 s2 k'; nwhd; //; 407 ##| #k'; napply step_skip_break_switch; @2; //; 408 ##] 409 ##| nwhd in ⊢ (?????%); ncases k; //; 410 ##[ #s' k'; nwhd; //; 411 ##| #ex s' k'; nwhd; napply step_skip_or_continue_while; @2; //; 412 ##| #ex s' k'; nwhd; 413 napply res_bindIO2_OK; #v tr Hv; 414 nletin bexpr ≝ (exec_bool_of_val v (typeof ex)); 415 nlapply (refl ? bexpr); ncases bexpr in ⊢ (???% → %); //; 416 #b; ncases b; #eb; nlapply (exec_bool_of_val_sound … eb); #Hb; 417 ##[ napply (step_skip_or_continue_dowhile_true … (exec_expr_sound' … Hv)); 418 ##[ @2; // ##| napply (bool_of … Hb); ##] 419 ##| napply (step_skip_or_continue_dowhile_false … (exec_expr_sound' … Hv)); 420 ##[ @2; // ##| napply (bool_of … Hb); ##] 421 ##] 422 ##| #ex s1 s2 k'; nwhd; napply step_skip_or_continue_for2; @2; // 423 ##| #k'; nwhd; //; 424 ##] 425 ##| #r; nwhd in ⊢ (?????%); ncases r; 426 ##[ nwhd; nlapply (refl ? (fn_return f)); ncases (fn_return f) in ⊢ (???% → %); //; #H; 427 napply step_return_0; napply H; 428 ##| #ex; ncases (type_eq_dec (fn_return f) Tvoid); //; nwhd in ⊢ (% → ?????%); #Hnotvoid; 429 napply res_bindIO2_OK; #v tr Hv; 430 nwhd; napply (step_return_1 … Hnotvoid (exec_expr_sound' … Hv)); 431 ##] 432 ##| #ex ls; napply res_bindIO2_OK; #v; ncases v; //; #n tr Hv; 433 napply step_switch; napply (exec_expr_sound' … Hv); 434 ##| #l s'; nwhd; //; 435 ##| #l; nwhd in ⊢ (?????%); nlapply (refl ? (find_label l (fn_body f) (call_cont k))); 436 ncases (find_label l (fn_body f) (call_cont k)) in ⊢ (???% → %); //; 437 #sk; ncases sk; #s' k' H; 438 napply (step_goto … H); 439 ##| #l s'; nwhd; //; 440 ##] 441 ##| #f0 vargs k m; nwhd in ⊢ (?????%); ncases f0; 442 ##[ #fn; 443 napply extract_subset_pair_io; #e m1 ealloc Halloc; 444 napply sig_bindIO_OK; #m2 Hbind; 445 nwhd; napply (step_internal_function … Halloc Hbind); 446 ##| #id argtys rty; napply sig_bindIO_OK; #evs Hevs; 447 napply bindIO_OK; #eres; nwhd; 448 napply step_external_function; @; ##[ napply Hevs; ##| napply mk_val_correct; ##] 449 ##] 450 ##| #v k' m'; nwhd in ⊢ (?????%); ncases k'; //; 451 #r f e k; nwhd in ⊢ (?????%); ncases r; 452 ##[ nwhd; //; 453 ##| #x; ncases x; #y; ncases y; #z; ncases z; #pcl b ofs ty; 454 napply opt_bindIO_OK; #m' em'; napply step_returnstate_1; nwhd in em':(??%?); //; 455 ##] 456 ##] 457 nqed. 458 459 nlemma make_initial_state_sound : ∀p. P_res ? (λgs.globalenv Genv ?? p = OK ? (\fst gs) ∧ initial_state p (\snd gs)) (make_initial_state p). 460 #p; ncases p; #fns main vars; 461 nwhd in ⊢ (???%); 462 napply bind_OK; #ge Ege; 463 napply bind_OK; #m Em; 464 napply opt_bind_OK; #x; ncases x; #sp b esb; 465 napply opt_bind_OK; #u ecode; 466 napply opt_bind_OK; #f ef; 467 ncases sp in esb ecode; #esb ecode; nwhd in ecode:(??%%); ##[ ##1,2,3,4,5: ndestruct (ecode); ##] 468 nwhd; @; //; napply (initial_state_intro … Ege Em esb ef); 469 nqed. 470 471 ntheorem exec_steps_sound: ∀ge,n,st. 373 #ge #st cases st; 374 [ #f #s #k #e #m cases s; 375 [ cases k; 376 [ whd in ⊢ (?????%); 377 lapply (refl ? (fn_return f)); cases (fn_return f) in ⊢ (???% → %); 378 //; #H whd; 379 @step_skip_call //; 380 | #s' #k' whd; (* XXX //; *) @step_skip_seq 381 | #ex #s' #k' @step_skip_or_continue_while % //; 382 | #ex #s' #k' 383 @res_bindIO2_OK #v #tr #Hv 384 letin bexpr ≝ (exec_bool_of_val v (typeof ex)); 385 lapply (refl ? bexpr); 386 cases bexpr in ⊢ (???% → %); 387 [ #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb 388 whd in ⊢ (?????%); 389 [ @(step_skip_or_continue_dowhile_true … (exec_expr_sound' … Hv)) 390 [ % // | @(bool_of … Hb) ] 391 | @(step_skip_or_continue_dowhile_false … (exec_expr_sound' … Hv)) 392 [ % // | @(bool_of … Hb) ] 393 ] 394 | #_ //; 395 ] 396 | #ex #s1 #s2 #k' @step_skip_or_continue_for2 % //; 397 | #ex #s1 #s2 #k' @step_skip_for3 398 | #k' @step_skip_break_switch % //; 399 | #r #f' #e' #k' whd in ⊢ (?????%); 400 lapply (refl ? (fn_return f)); cases (fn_return f) in ⊢ (???% → %); 401 //; #H whd; 402 @step_skip_call //; 403 ] 404 | #ex1 #ex2 405 @res_bindIO2_OK #x cases x; #y cases y; #pcl #loc #ofs #tr1 #Hlval 406 @res_bindIO2_OK #v2 #tr2 #Hv2 407 @opt_bindIO_OK #m' #em' 408 whd; @(step_assign … (exec_lvalue_sound' … Hlval) (exec_expr_sound' … Hv2) em') 409 | #lex #fex #args 410 @res_bindIO2_OK #vf #tr1 #Hvf0 lapply (exec_expr_sound' … Hvf0); #Hvf 411 @res_bindIO2_OK #vargs #tr2 #Hvargs0 lapply (P_res_to_P ???? (exec_exprlist_sound …) Hvargs0); #Hvargs 412 @opt_bindIO_OK #fd #efd 413 @bindIO_OK #ety 414 cases lex; whd; 415 [ @(step_call_none … Hvf Hvargs efd ety) 416 | #lhs' 417 @res_bindIO2_OK #x cases x; #y cases y; #pcl #loc #ofs #tr3 #Hlocofs 418 whd; @(step_call_some … (exec_lvalue_sound' … Hlocofs) Hvf Hvargs efd ety) 419 ] 420 | #s1 #s2 whd; (* XXX //; *) @step_seq 421 | #ex #s1 #s2 422 @res_bindIO2_OK #v #tr #Hv 423 letin bexpr ≝ (exec_bool_of_val v (typeof ex)); 424 lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); //; 425 #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb 426 [ @(step_ifthenelse_true … (exec_expr_sound' … Hv)) @(bool_of … Hb) 427 | @(step_ifthenelse_false … (exec_expr_sound' … Hv)) @(bool_of … Hb) 428 ] 429 | #ex #s' 430 @res_bindIO2_OK #v #tr #Hv 431 letin bexpr ≝ (exec_bool_of_val v (typeof ex)); 432 lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); //; 433 #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb 434 [ @(step_while_true … (exec_expr_sound' … Hv)) @(bool_of … Hb) 435 | @(step_while_false … (exec_expr_sound' … Hv)) @(bool_of … Hb) 436 ] 437 | #ex #s' whd; (* XXX //; *) @step_dowhile 438 | #s1 #ex #s2 #s3 whd in ⊢ (?????%); elim (is_Sskip s1); #Hs1 whd in ⊢ (?????%); 439 [ >Hs1 440 @res_bindIO2_OK #v #tr #Hv 441 letin bexpr ≝ (exec_bool_of_val v (typeof ex)); 442 lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); //; 443 #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb 444 [ @(step_for_true … (exec_expr_sound' … Hv)) @(bool_of … Hb) 445 | @(step_for_false … (exec_expr_sound' … Hv)) @(bool_of … Hb) 446 ] 447 | @step_for_start //; 448 ] 449 | whd in ⊢ (?????%); cases k; //; 450 [ #s' #k' whd (* XXX // *) @step_break_seq 451 | #ex #s' #k' whd (* //; *) @step_break_while 452 | #ex #s' #k' whd (* //; *) @step_break_dowhile 453 | #ex #s1 #s2 #k' whd (* //; *) @step_break_for2 454 | #k' @step_skip_break_switch %2 // 455 ] 456 | whd in ⊢ (?????%); cases k; //; 457 [ #s' #k' whd; (* XXX //;*) @step_continue_seq 458 | #ex #s' #k' whd; @step_skip_or_continue_while %2 ; //; 459 | #ex #s' #k' whd; 460 @res_bindIO2_OK #v #tr #Hv 461 letin bexpr ≝ (exec_bool_of_val v (typeof ex)); 462 lapply (refl ? bexpr); cases bexpr in ⊢ (???% → %); //; 463 #b cases b; #eb lapply (exec_bool_of_val_sound … eb); #Hb 464 [ @(step_skip_or_continue_dowhile_true … (exec_expr_sound' … Hv)) 465 [ %2 ; // | @(bool_of … Hb) ] 466 | @(step_skip_or_continue_dowhile_false … (exec_expr_sound' … Hv)) 467 [ %2 ; // | @(bool_of … Hb) ] 468 ] 469 | #ex #s1 #s2 #k' whd; @step_skip_or_continue_for2 %2 ; // 470 | #k' whd; (* XXX //;*) @step_continue_switch 471 ] 472 | #r whd in ⊢ (?????%); cases r; 473 [ whd; lapply (refl ? (fn_return f)); cases (fn_return f) in ⊢ (???% → %); //; #H 474 @step_return_0 @H 475 | #ex cases (type_eq_dec (fn_return f) Tvoid); //; whd in ⊢ (% → ?????%); #Hnotvoid 476 @res_bindIO2_OK #v #tr #Hv 477 whd; @(step_return_1 … Hnotvoid (exec_expr_sound' … Hv)) 478 ] 479 | #ex #ls @res_bindIO2_OK #v cases v; //; #n #tr #Hv 480 @step_switch @(exec_expr_sound' … Hv) 481 | #l #s' whd; @step_label (* XXX //; *) 482 | #l whd in ⊢ (?????%); lapply (refl ? (find_label l (fn_body f) (call_cont k))); 483 cases (find_label l (fn_body f) (call_cont k)) in ⊢ (???% → %); //; 484 #sk cases sk; #s' #k' #H 485 @(step_goto … H) 486 | #l #s' whd; (* XXX //; *) @step_cost 487 ] 488 | #f0 #vargs #k #m whd in ⊢ (?????%); cases f0; 489 [ #fn whd in ⊢ (?????%) 490 lapply (refl ? (exec_alloc_variables empty_env m (fn_params fn@fn_vars fn))) 491 cases (exec_alloc_variables empty_env m (fn_params fn@fn_vars fn)) in ⊢ (???% → %) 492 #en' #m' #ALLOC whd in ⊢ (?????%) 493 @res_bindIO_OK #m2 #BIND 494 whd; @(step_internal_function … (exec_alloc_variables_sound … ALLOC)) 495 @(P_res_to_P … (exec_bind_parameters_sound …) BIND) 496 | #id #argtys #rty @res_bindIO_OK #evs #Hevs 497 @bindIO_OK #eres whd; 498 @step_external_function % [ @(P_res_to_P … (check_eventval_list_sound …) Hevs) 499 | @mk_val_correct ] 500 ] 501 | #v #k' #m' whd in ⊢ (?????%); cases k'; //; 502 #r #f #e #k whd in ⊢ (?????%); cases r; 503 [ whd; @step_returnstate_0 504 | #x cases x; #y cases y; #z cases z; #pcl #b #ofs #ty 505 @opt_bindIO_OK #m' #em' @step_returnstate_1 whd in em':(??%?); //; 506 ] 507 ] 508 qed. 509 510 lemma make_initial_state_sound : ∀p. P_res ? (λgs.globalenv Genv ?? p = OK ? (\fst gs) ∧ initial_state p (\snd gs)) (make_initial_state p). 511 #p cases p; #fns #main #vars 512 whd in ⊢ (???%); 513 @bind_OK #ge #Ege 514 @bind_OK #m #Em 515 @opt_bind_OK #x cases x; #sp #b #esb 516 @opt_bind_OK #u #ecode 517 @opt_bind_OK #f #ef 518 cases sp in esb ecode; #esb #ecode whd in ecode:(??%%); [ 1,2,3,4,5: destruct (ecode); ] 519 whd; % [ whd in ⊢ (???(??%)) // | @(initial_state_intro … Ege Em esb ef) ] 520 qed. 521 522 theorem exec_steps_sound: ∀ge,n,st. 472 523 P_io ??? (λts:trace×state. star (mk_transrel ?? step) ge st (\fst ts) (\snd ts)) 473 524 (exec_steps n ge st). 474 #ge n; nelim n;475 ##[ #st; nwhd in ⊢ (?????%); nelim (is_final_state st); #H; nwhd; //; 476 ##| #n' IH st; nwhd in ⊢ (?????%); nelim (is_final_state st); #H; 477 ##[ nwhd; //;478 ##| napply (P_bindIO2_OK ????????? (exec_step_sound …)); #t s'; ncases s';479 ##[ #f s k e m; ##| #fd args k m; ##| #r k m; ##]480 nwhd in ⊢ (? → ?????(??????%?));481 ncases m; #mc mn mp; #Hstep;482 nwhd in ⊢ (?????(??????%?));483 napply (P_bindIO2_OK ????????? (IH …)); #t' s'' Hsteps;484 nwhd; napply (star_step ????????? Hsteps); ##[ ##2,5,8: napply Hstep; ##| ##3,6,9: // ##]485 ##]486 nqed.525 #ge #n elim n; 526 [ #st whd in ⊢ (?????%); elim (is_final_state st); #H whd; % 527 | #n' #IH #st whd in ⊢ (?????%); elim (is_final_state st); #H 528 [ whd; % 529 | @(P_bindIO2_OK ????????? (exec_step_sound …)) #t #s' cases s'; 530 [ #f #s #k #e #m | #fd #args #k #m | #r #k #m ] 531 whd in ⊢ (? → ?????(??????%?)); 532 cases m; #mc #mn #mp #Hstep 533 whd in ⊢ (?????(??????%?)); 534 @(P_bindIO2_OK ????????? (IH …)) #t' #s'' #Hsteps 535 whd; @(star_step ????????? Hsteps) [ 2,5,8: @Hstep | 3,6,9: // ] 536 ] 537 qed. -
Deliverables/D3.1/C-semantics/Coqlib.ma
r473 r487 19 19 20 20 include "binary/Z.ma". 21 include " datatypes/sums.ma".22 include " datatypes/list.ma".21 include "basics/types.ma". 22 include "basics/list.ma". 23 23 24 24 include "extralib.ma". … … 350 350 (* * Properties of powers of two. *) 351 351 352 nlemma two_power_nat_O : two_power_nat O = one.353 // ; nqed.354 355 nlemma two_power_nat_pos : ∀n:nat. Z_two_power_nat n > 0.356 // ; nqed.357 358 nlemma two_power_nat_two_p:352 lemma two_power_nat_O : two_power_nat O = one. 353 // qed. 354 355 lemma two_power_nat_pos : ∀n:nat. Z_two_power_nat n > 0. 356 // qed. 357 358 lemma two_power_nat_two_p: 359 359 ∀x. Z_two_power_nat x = two_p (Z_of_nat x). 360 #x ; ncases x;361 ##[ //; 362 ##| nnormalize; #p; nelim p; //; #p' H; nrewrite > (nat_succ_pos …); //; 363 ##] nqed.360 #x cases x 361 [ // 362 | normalize #p elim p // #p' #H >(nat_succ_pos …) // 363 ] qed. 364 364 (* 365 365 Lemma two_p_monotone: … … 576 576 (*naxiom align : Z → Z → Z.*) 577 577 578 ndefinition align ≝ λn: Z. λamount: Z.578 definition align ≝ λn: Z. λamount: Z. 579 579 ((n + amount - 1) / amount) * amount. 580 580 (* … … 600 600 (* * Mapping a function over an option type. *) 601 601 602 ndefinition option_map ≝ λA,B.λf:A→B.λv:option A.602 definition option_map ≝ λA,B.λf:A→B.λv:option A. 603 603 match v with [ Some v' ⇒ Some ? (f v') | None ⇒ None ? ]. 604 604 … … 794 794 induction l; simpl; congruence. 795 795 Qed. 796 *) 797 nlemma list_in_map_inv:798 ∀A,B: Type . ∀f: A -> B. ∀l: list A. ∀y: B.796 *) (* 797 lemma list_in_map_inv: 798 ∀A,B: Type[0]. ∀f: A -> B. ∀l: list A. ∀y: B. 799 799 in_list ? y (map ?? f l) → ∃x:A. y = f x ∧ in_list ? x l. 800 800 #A B f l; nelim l; … … 808 808 nelim (IH h H1); #x'; *; #IH1 IH2; @x'; @; /2/; 809 809 ##] 810 ##] nqed. 810 ##] nqed.*) 811 811 (* 812 812 Lemma list_append_map: … … 841 841 (* * [list_disjoint l1 l2] holds iff [l1] and [l2] have no elements 842 842 in common. *) 843 844 ndefinition list_disjoint : ∀A:Type. list A → list A → Prop ≝843 (* 844 definition list_disjoint : ∀A:Type[0]. list A → list A → Prop ≝ 845 845 λA,l1,l2. 846 846 ∀x,y: A. in_list A x l1 → in_list A y l2 → x ≠ y. 847 847 848 nlemma list_disjoint_cons_left:849 ∀A: Type . ∀a: A. ∀l1,l2: list A.848 lemma list_disjoint_cons_left: 849 ∀A: Type[0]. ∀a: A. ∀l1,l2: list A. 850 850 list_disjoint A (a :: l1) l2 → list_disjoint A l1 l2. 851 851 #A;#a;#l1;#l2;nwhd in ⊢ (% → %); #H; … … 874 874 nwhd in ⊢ (% → %); 875 875 #H;#x;#y;#H1;#H2; napply sym_neq; /2/; 876 nqed. 876 nqed.*) 877 877 878 (* 878 879 Lemma list_disjoint_dec: … … 898 899 (* * [list_norepet l] holds iff the list [l] contains no repetitions, 899 900 i.e. no element occurs twice. *) 900 901 ninductive list_norepet (A: Type) : list A → Prop ≝901 (* 902 inductive list_norepet (A: Type[0]) : list A → Prop ≝ 902 903 | list_norepet_nil: 903 904 list_norepet A (nil A) 904 905 | list_norepet_cons: 905 906 ∀hd,tl. 906 ¬(in_list A hd tl) → list_norepet A tl → list_norepet A (hd :: tl). 907 ¬(in_list A hd tl) → list_norepet A tl → list_norepet A (hd :: tl).*) 907 908 (* 908 909 Lemma list_norepet_dec: -
Deliverables/D3.1/C-semantics/CostLabel.ma
r177 r487 1 1 include "AST.ma". 2 2 3 ndefinition costlabel ≝ ident.3 definition costlabel ≝ ident. -
Deliverables/D3.1/C-semantics/Csem.ma
r485 r487 36 36 the null pointer) and the float 0.0 are false. *) 37 37 38 ninductive is_false: val → type → Prop ≝38 inductive is_false: val → type → Prop ≝ 39 39 | is_false_int: ∀sz,sg. 40 40 is_false (Vint zero) (Tint sz sg) … … 44 44 is_false (Vfloat Fzero) (Tfloat sz). 45 45 46 ninductive is_true: val → type → Prop ≝46 inductive is_true: val → type → Prop ≝ 47 47 | is_true_int_int: ∀n,sz,sg. 48 48 n ≠ zero → … … 54 54 is_true (Vfloat f) (Tfloat sz). 55 55 56 ninductive bool_of_val : val → type → val → Prop ≝56 inductive bool_of_val : val → type → val → Prop ≝ 57 57 | bool_of_val_true: ∀v,ty. 58 58 is_true v ty → … … 70 70 promoted [e2] to a float. *) 71 71 72 nlet rec sem_neg (v: val) (ty: type) : option val ≝72 let rec sem_neg (v: val) (ty: type) : option val ≝ 73 73 match ty with 74 74 [ Tint _ _ ⇒ … … 85 85 ]. 86 86 87 nlet rec sem_notint (v: val) : option val ≝87 let rec sem_notint (v: val) : option val ≝ 88 88 match v with 89 89 [ Vint n ⇒ Some ? (Vint (xor n mone)) … … 91 91 ]. 92 92 93 nlet rec sem_notbool (v: val) (ty: type) : option val ≝93 let rec sem_notbool (v: val) (ty: type) : option val ≝ 94 94 match ty with 95 95 [ Tint _ _ ⇒ … … 113 113 ]. 114 114 115 nlet rec sem_add (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝115 let rec sem_add (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝ 116 116 match classify_add t1 t2 with 117 117 [ add_case_ii ⇒ (**r integer addition *) … … 145 145 ]. 146 146 147 nlet rec sem_sub (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝147 let rec sem_sub (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝ 148 148 match classify_sub t1 t2 with 149 149 [ sub_case_ii ⇒ (**r integer subtraction *) … … 179 179 ]. 180 180 181 nlet rec sem_mul (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝181 let rec sem_mul (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝ 182 182 match classify_mul t1 t2 with 183 183 [ mul_case_ii ⇒ … … 197 197 ]. 198 198 199 nlet rec sem_div (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝199 let rec sem_div (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝ 200 200 match classify_div t1 t2 with 201 201 [ div_case_I32unsi ⇒ … … 221 221 ]. 222 222 223 nlet rec sem_mod (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝223 let rec sem_mod (v1:val) (t1:type) (v2: val) (t2:type) : option val ≝ 224 224 match classify_mod t1 t2 with 225 225 [ mod_case_I32unsi ⇒ … … 239 239 ]. 240 240 241 nlet rec sem_and (v1,v2: val) : option val ≝241 let rec sem_and (v1,v2: val) : option val ≝ 242 242 match v1 with 243 243 [ Vint n1 ⇒ match v2 with … … 247 247 ]. 248 248 249 nlet rec sem_or (v1,v2: val) : option val ≝249 let rec sem_or (v1,v2: val) : option val ≝ 250 250 match v1 with 251 251 [ Vint n1 ⇒ match v2 with … … 255 255 ]. 256 256 257 nlet rec sem_xor (v1,v2: val) : option val ≝257 let rec sem_xor (v1,v2: val) : option val ≝ 258 258 match v1 with 259 259 [ Vint n1 ⇒ match v2 with … … 263 263 ]. 264 264 265 nlet rec sem_shl (v1,v2: val): option val ≝265 let rec sem_shl (v1,v2: val): option val ≝ 266 266 match v1 with 267 267 [ Vint n1 ⇒ match v2 with … … 271 271 | _ ⇒ None ? ]. 272 272 273 nlet rec sem_shr (v1: val) (t1: type) (v2: val) (t2: type): option val ≝273 let rec sem_shr (v1: val) (t1: type) (v2: val) (t2: type): option val ≝ 274 274 match classify_shr t1 t2 with 275 275 [ shr_case_I32unsi ⇒ … … 291 291 ]. 292 292 293 nlet rec sem_cmp_mismatch (c: comparison): option val ≝293 let rec sem_cmp_mismatch (c: comparison): option val ≝ 294 294 match c with 295 295 [ Ceq => Some ? Vfalse … … 298 298 ]. 299 299 300 nlet rec sem_cmp_match (c: comparison): option val ≝300 let rec sem_cmp_match (c: comparison): option val ≝ 301 301 match c with 302 302 [ Ceq => Some ? Vtrue … … 305 305 ]. 306 306 307 nlet rec sem_cmp (c:comparison)307 let rec sem_cmp (c:comparison) 308 308 (v1: val) (t1: type) (v2: val) (t2: type) 309 309 (m: mem): option val ≝ … … 349 349 ]. 350 350 351 ndefinition sem_unary_operation351 definition sem_unary_operation 352 352 : unary_operation → val → type → option val ≝ 353 353 λop,v,ty. … … 358 358 ]. 359 359 360 nlet rec sem_binary_operation360 let rec sem_binary_operation 361 361 (op: binary_operation) 362 362 (v1: val) (t1: type) (v2: val) (t2:type) … … 385 385 resulting in value [v2]. *) 386 386 387 nlet rec cast_int_int (sz: intsize) (sg: signedness) (i: int) : int ≝387 let rec cast_int_int (sz: intsize) (sg: signedness) (i: int) : int ≝ 388 388 match sz with 389 389 [ I8 ⇒ match sg with [ Signed ⇒ sign_ext 8 i | Unsigned ⇒ zero_ext 8 i ] … … 392 392 ]. 393 393 394 nlet rec cast_int_float (si : signedness) (i: int) : float ≝394 let rec cast_int_float (si : signedness) (i: int) : float ≝ 395 395 match si with 396 396 [ Signed ⇒ floatofint i … … 398 398 ]. 399 399 400 nlet rec cast_float_int (si : signedness) (f: float) : int ≝400 let rec cast_float_int (si : signedness) (f: float) : int ≝ 401 401 match si with 402 402 [ Signed ⇒ intoffloat f … … 404 404 ]. 405 405 406 nlet rec cast_float_float (sz: floatsize) (f: float) : float ≝406 let rec cast_float_float (sz: floatsize) (f: float) : float ≝ 407 407 match sz with 408 408 [ F32 ⇒ singleoffloat f … … 410 410 ]. 411 411 412 ninductive type_region : type → region → Prop ≝412 inductive type_region : type → region → Prop ≝ 413 413 | type_rgn_pointer : ∀s,t. type_region (Tpointer s t) s 414 414 | type_rgn_array : ∀s,t,n. type_region (Tarray s t n) s … … 416 416 | type_rgn_code : ∀tys,ty. type_region (Tfunction tys ty) Code. 417 417 418 ninductive cast : mem → val → type → type → val → Prop ≝418 inductive cast : mem → val → type → type → val → Prop ≝ 419 419 | cast_ii: ∀m,i,sz2,sz1,si1,si2. (**r int to int *) 420 420 cast m (Vint i) (Tint sz1 si1) (Tint sz2 si2) … … 448 448 and function pointers to their definitions. (See module [Globalenvs].) *) 449 449 450 ndefinition genv ≝ (genv_t Genv) fundef.450 definition genv ≝ (genv_t Genv) fundef. 451 451 452 452 (* * The local environment maps local variables to block references. … … 454 454 block. *) 455 455 456 ndefinition env ≝ (tree_t ? PTree) block. (* map variable -> location *)457 458 ndefinition empty_env: env ≝ (empty …).456 definition env ≝ (tree_t ? PTree) block. (* map variable -> location *) 457 458 definition empty_env: env ≝ (empty …). 459 459 460 460 (* * [load_value_of_type ty m b ofs] computes the value of a datum … … 464 464 reference, the pointer [Vptr b ofs] is returned. *) 465 465 466 nlet rec load_value_of_type (ty: type) (m: mem) (psp:region) (b: block) (ofs: int) : option val ≝466 let rec load_value_of_type (ty: type) (m: mem) (psp:region) (b: block) (ofs: int) : option val ≝ 467 467 match access_mode ty with 468 468 [ By_value chunk ⇒ loadv chunk m (Vptr psp b ofs) … … 476 476 This is allowed only if [ty] indicates an access by value. *) 477 477 478 nlet rec store_value_of_type (ty_dest: type) (m: mem) (psp:region) (loc: block) (ofs: int) (v: val) : option mem ≝478 let rec store_value_of_type (ty_dest: type) (m: mem) (psp:region) (loc: block) (ofs: int) (v: val) : option mem ≝ 479 479 match access_mode ty_dest with 480 480 [ By_value chunk ⇒ storev chunk m (Vptr psp loc ofs) v … … 490 490 and memory state. *) 491 491 492 ninductive alloc_variables: env → mem →492 inductive alloc_variables: env → mem → 493 493 list (ident × type) → 494 494 env → mem → Prop ≝ … … 507 507 [m1] is the initial memory state and [m2] the final memory state. *) 508 508 509 ninductive bind_parameters: env →509 inductive bind_parameters: env → 510 510 mem → list (ident × type) → list val → 511 511 mem → Prop ≝ … … 525 525 (* * Return the list of blocks in the codomain of [e]. *) 526 526 527 ndefinition blocks_of_env : env → list block ≝ λe.527 definition blocks_of_env : env → list block ≝ λe. 528 528 map ?? (λx. snd ?? x) (elements ??? e). 529 529 … … 531 531 of the selector expression. *) 532 532 533 nlet rec select_switch (n: int) (sl: labeled_statements)533 let rec select_switch (n: int) (sl: labeled_statements) 534 534 on sl : labeled_statements ≝ 535 535 match sl with … … 540 540 (* * Turn a labeled statement into a sequence *) 541 541 542 nlet rec seq_of_labeled_statement (sl: labeled_statements) : statement ≝542 let rec seq_of_labeled_statement (sl: labeled_statements) : statement ≝ 543 543 match sl with 544 544 [ LSdefault s ⇒ s … … 562 562 [e] is the current environment and [m] is the current memory state. *) 563 563 564 ninductive eval_expr (ge:genv) (e:env) (m:mem) : expr → val → trace → Prop ≝564 inductive eval_expr (ge:genv) (e:env) (m:mem) : expr → val → trace → Prop ≝ 565 565 | eval_Econst_int: ∀i,ty. 566 566 eval_expr ge e m (Expr (Econst_int i) ty) (Vint i) E0 … … 648 648 eval_lvalue ge e m (Expr (Efield a i) ty) psp l ofs tr. 649 649 650 nlet rec eval_expr_ind (ge:genv) (e:env) (m:mem)650 let rec eval_expr_ind (ge:genv) (e:env) (m:mem) 651 651 (P:∀a,v,tr. eval_expr ge e m a v tr → Prop) 652 652 (eci:∀i,ty. P ??? (eval_Econst_int ge e m i ty)) … … 684 684 ]. 685 685 686 ninverter eval_expr_inv_ind for eval_expr : Prop.687 688 nlet rec eval_lvalue_ind (ge:genv) (e:env) (m:mem)686 inverter eval_expr_inv_ind for eval_expr : Prop. 687 688 let rec eval_lvalue_ind (ge:genv) (e:env) (m:mem) 689 689 (P:∀a,psp,loc,ofs,tr. eval_lvalue ge e m a psp loc ofs tr → Prop) 690 690 (lvl:∀id,l,ty,H. P ????? (eval_Evar_local ge e m id l ty H)) … … 706 706 *) 707 707 708 ndefinition eval_lvalue_inv_ind :708 definition eval_lvalue_inv_ind : 709 709 ∀x1: genv. 710 710 ∀x2: env. … … 755 755 (refl int x7) (refl trace x8))) 756 756 ?)))))))))))))))). 757 ##[ napply (eval_lvalue_ind x1 x2 x3 (λa,psp,loc,ofs,tr,e. ∀e1:eq ? x4 a. ∀e2:eq ? x5 psp. ∀e3:eq ? x6 loc. ∀e4:eq ? x7 ofs. ∀e5:eq ? x8 tr. P a psp loc ofs tr) … Hterm); 758 ##[ napply H1 ##| napply H2 ##| napply H3 ##| napply H4 ##| napply H5 ##]759 ##| ##*: ##skip760 ##] nqed.761 762 nlet rec eval_expr_ind2 (ge:genv) (e:env) (m:mem)757 [ @(eval_lvalue_ind x1 x2 x3 (λa,psp,loc,ofs,tr,e. ∀e1:eq ? x4 a. ∀e2:eq ? x5 psp. ∀e3:eq ? x6 loc. ∀e4:eq ? x7 ofs. ∀e5:eq ? x8 tr. P a psp loc ofs tr) … Hterm) 758 [ @H1 | @H2 | @H3 | @H4 | @H5 ] 759 | *: skip 760 ] qed. 761 762 let rec eval_expr_ind2 (ge:genv) (e:env) (m:mem) 763 763 (P:∀a,v,tr. eval_expr ge e m a v tr → Prop) 764 764 (Q:∀a,psp,loc,ofs,tr. eval_lvalue ge e m a psp loc ofs tr → Prop) … … 834 834 ]. 835 835 836 ndefinition combined_expr_lvalue_ind ≝836 definition combined_expr_lvalue_ind ≝ 837 837 λge,e,m,P,Q,eci,ecF,elv,ead,esz,eun,ebi,ect,ecf,eo1,eo2,ea1,ea2,ecs,eco,lvl,lvg,lde,lfs,lfu. 838 838 conj ?? … … 852 852 expressions [al] to their values [vl]. *) 853 853 854 ninductive eval_exprlist (ge:genv) (e:env) (m:mem) : list expr → list val → trace → Prop ≝854 inductive eval_exprlist (ge:genv) (e:env) (m:mem) : list expr → list val → trace → Prop ≝ 855 855 | eval_Enil: 856 856 eval_exprlist ge e m (nil ?) (nil ?) E0 … … 866 866 (* * Continuations *) 867 867 868 ninductive cont: Type:=868 inductive cont: Type[0] := 869 869 | Kstop: cont 870 870 | Kseq: statement -> cont -> cont … … 887 887 (* * Pop continuation until a call or stop *) 888 888 889 nlet rec call_cont (k: cont) : cont :=889 let rec call_cont (k: cont) : cont := 890 890 match k with 891 891 [ Kseq s k => call_cont k … … 898 898 ]. 899 899 900 ndefinition is_call_cont : cont → Prop ≝ λk.900 definition is_call_cont : cont → Prop ≝ λk. 901 901 match k with 902 902 [ Kstop => True … … 907 907 (* * States *) 908 908 909 ninductive state: Type:=909 inductive state: Type[0] := 910 910 | State: 911 911 ∀f: function. … … 927 927 corresponding to a label *) 928 928 929 nlet rec find_label (lbl: label) (s: statement) (k: cont)929 let rec find_label (lbl: label) (s: statement) (k: cont) 930 930 on s: option (statement × cont) := 931 931 match s with … … 977 977 978 978 (* Strip off outer pointer for use when comparing function types. *) 979 ndefinition fun_typeof ≝979 definition fun_typeof ≝ 980 980 λe. match typeof e with 981 981 [ Tvoid ⇒ Tvoid … … 992 992 (* XXX: note that cost labels in exprs expose a particular eval order. *) 993 993 994 ninductive step (ge:genv) : state → trace → state → Prop ≝994 inductive step (ge:genv) : state → trace → state → Prop ≝ 995 995 996 996 | step_assign: ∀f,a1,a2,k,e,m,psp,loc,ofs,v2,m',tr1,tr2. … … 1172 1172 through the execution of a [break], [continue] or [return] statement. *) 1173 1173 1174 ninductive outcome: Type:=1174 inductive outcome: Type[0] := 1175 1175 | Out_break: outcome (**r terminated by [break] *) 1176 1176 | Out_continue: outcome (**r terminated by [continue] *) … … 1178 1178 | Out_return: option val -> outcome. (**r terminated by [return] *) 1179 1179 1180 ninductive out_normal_or_continue : outcome -> Prop :=1180 inductive out_normal_or_continue : outcome -> Prop := 1181 1181 | Out_normal_or_continue_N: out_normal_or_continue Out_normal 1182 1182 | Out_normal_or_continue_C: out_normal_or_continue Out_continue. 1183 1183 1184 ninductive out_break_or_return : outcome -> outcome -> Prop :=1184 inductive out_break_or_return : outcome -> outcome -> Prop := 1185 1185 | Out_break_or_return_B: out_break_or_return Out_break Out_normal 1186 1186 | Out_break_or_return_R: ∀ov. … … 1207 1207 evaluation. *) 1208 1208 1209 ninductive exec_stmt: env -> mem -> statement -> trace -> mem -> outcome -> Prop :=1209 inductive exec_stmt: env -> mem -> statement -> trace -> mem -> outcome -> Prop := 1210 1210 | exec_Sskip: ∀e,m. 1211 1211 exec_stmt e m Sskip … … 1368 1368 trace of observable events performed during the execution. *) 1369 1369 1370 Co ninductive execinf_stmt: env -> mem -> statement -> traceinf -> Prop :=1370 Coinductive execinf_stmt: env -> mem -> statement -> traceinf -> Prop := 1371 1371 | execinf_Scall_none: ∀e,m,a,al,vf,vargs,f,t. 1372 1372 eval_expr e m a vf -> … … 1475 1475 without arguments and with an empty continuation. *) 1476 1476 1477 ninductive initial_state (p: clight_program): state -> Prop :=1477 inductive initial_state (p: clight_program): state -> Prop := 1478 1478 | initial_state_intro: ∀b,f,ge,m0. 1479 1479 globalenv Genv ?? p = OK ? ge → … … 1485 1485 (* * A final state is a [Returnstate] with an empty continuation. *) 1486 1486 1487 ninductive final_state: state -> int -> Prop :=1487 inductive final_state: state -> int -> Prop := 1488 1488 | final_state_intro: ∀r,m. 1489 1489 final_state (Returnstate (Vint r) Kstop m) r. … … 1494 1494 behavior. *) 1495 1495 1496 ndefinition exec_program : clight_program → program_behavior → Prop ≝ λp,beh.1496 definition exec_program : clight_program → program_behavior → Prop ≝ λp,beh. 1497 1497 ∀ge. globalenv ??? p = OK ? ge → 1498 1498 program_behaves (mk_transrel ?? step) (initial_state p) final_state ge beh. … … 1500 1500 (** Big-step execution of a whole program. *) 1501 1501 1502 ninductive bigstep_program_terminates (p: program): trace -> int -> Prop :=1502 inductive bigstep_program_terminates (p: program): trace -> int -> Prop := 1503 1503 | bigstep_program_terminates_intro: ∀b,f,m1,t,r. 1504 1504 let ge := Genv.globalenv p in … … 1509 1509 bigstep_program_terminates p t r. 1510 1510 1511 ninductive bigstep_program_diverges (p: program): traceinf -> Prop :=1511 inductive bigstep_program_diverges (p: program): traceinf -> Prop := 1512 1512 | bigstep_program_diverges_intro: ∀b,f,t. 1513 1513 let ge := Genv.globalenv p in … … 1532 1532 (eval_funcall_ind2 ge PS PF a b c d e f g h i j k l m n o p q r s t u v w x y). 1533 1533 1534 ninductive outcome_state_match1534 inductive outcome_state_match 1535 1535 (e: env) (m: mem) (f: function) (k: cont): outcome -> state -> Prop := 1536 1536 | osm_normal: -
Deliverables/D3.1/C-semantics/Csyntax.ma
r483 r487 32 32 flag and a bit size: 8, 16 or 32 bits. *) 33 33 34 ninductive signedness : Type≝34 inductive signedness : Type[0] ≝ 35 35 | Signed: signedness 36 36 | Unsigned: signedness. 37 37 38 ninductive intsize : Type≝38 inductive intsize : Type[0] ≝ 39 39 | I8: intsize 40 40 | I16: intsize … … 44 44 and 64-bit (double precision). *) 45 45 46 ninductive floatsize : Type≝46 inductive floatsize : Type[0] ≝ 47 47 | F32: floatsize 48 48 | F64: floatsize. … … 83 83 *) 84 84 85 ninductive type : Type≝85 inductive type : Type[0] ≝ 86 86 | Tvoid: type (**r the [void] type *) 87 87 | Tint: intsize → signedness → type (**r integer types *) … … 94 94 | Tcomp_ptr: region → ident → type (**r pointer to named struct or union *) 95 95 96 with typelist : Type ≝96 with typelist : Type[0] ≝ 97 97 | Tnil: typelist 98 98 | Tcons: type → typelist → typelist 99 99 100 with fieldlist : Type ≝100 with fieldlist : Type[0] ≝ 101 101 | Fnil: fieldlist 102 102 | Fcons: ident → type → fieldlist → fieldlist. 103 103 104 104 (* XXX: no induction scheme! *) 105 nlet rec type_ind105 let rec type_ind 106 106 (P:type → Prop) 107 107 (vo:P Tvoid) … … 127 127 ]. 128 128 129 nlet rec fieldlist_ind129 let rec fieldlist_ind 130 130 (P:fieldlist → Prop) 131 131 (nl:P Fnil) … … 141 141 (* * Arithmetic and logical operators. *) 142 142 143 ninductive unary_operation : Type≝143 inductive unary_operation : Type[0] ≝ 144 144 | Onotbool : unary_operation (**r boolean negation ([!] in C) *) 145 145 | Onotint : unary_operation (**r integer complement ([~] in C) *) 146 146 | Oneg : unary_operation. (**r opposite (unary [-]) *) 147 147 148 ninductive binary_operation : Type≝148 inductive binary_operation : Type[0] ≝ 149 149 | Oadd : binary_operation (**r addition (binary [+]) *) 150 150 | Osub : binary_operation (**r subtraction (binary [-]) *) … … 174 174 *) 175 175 176 ninductive expr : Type≝176 inductive expr : Type[0] ≝ 177 177 | Expr: expr_descr → type → expr 178 178 179 with expr_descr : Type ≝179 with expr_descr : Type[0] ≝ 180 180 | Econst_int: int → expr_descr (**r integer literal *) 181 181 | Econst_float: float → expr_descr (**r float literal *) … … 198 198 (* * Extract the type part of a type-annotated Clight expression. *) 199 199 200 ndefinition typeof : expr → type ≝ λe.200 definition typeof : expr → type ≝ λe. 201 201 match e with [ Expr de te ⇒ te ]. 202 202 … … 208 208 are not supported. *) 209 209 210 ndefinition label ≝ ident.211 212 ninductive statement : Type≝210 definition label ≝ ident. 211 212 inductive statement : Type[0] ≝ 213 213 | Sskip : statement (**r do nothing *) 214 214 | Sassign : expr → expr → statement (**r assignment [lvalue = rvalue] *) … … 227 227 | Scost : costlabel → statement → statement 228 228 229 with labeled_statements : Type ≝ (**r cases of a [switch] *)229 with labeled_statements : Type[0] ≝ (**r cases of a [switch] *) 230 230 | LSdefault: statement → labeled_statements 231 231 | LScase: int → statement → labeled_statements → labeled_statements. 232 232 233 nlet rec statement_ind2233 let rec statement_ind2 234 234 (P:statement → Prop) (Q:labeled_statements → Prop) 235 235 (Ssk:P Sskip) … … 308 308 ]. 309 309 310 ndefinition statement_ind ≝ λP,Ssk,Sas,Sca,Ssq,Sif,Swh,Sdo,Sfo,Sbr,Sco,Sre,Ssw,Sla,Sgo,Scs.310 definition statement_ind ≝ λP,Ssk,Sas,Sca,Ssq,Sif,Swh,Sdo,Sfo,Sbr,Sco,Sre,Ssw,Sla,Sgo,Scs. 311 311 statement_ind2 P (λ_.True) Ssk Sas Sca Ssq Sif Swh Sdo Sfo Sbr Sco Sre Ssw Sla Sgo Scs 312 312 (λ_,_. I) (λ_,_,_,_,_.I). … … 319 319 function (a statement, [fn_body]). *) 320 320 321 nrecord function : Type≝ {321 record function : Type[0] ≝ { 322 322 fn_return: type; 323 323 fn_params: list (ident × type); … … 329 329 external functions ([External]). *) 330 330 331 ninductive fundef : Type≝331 inductive fundef : Type[0] ≝ 332 332 | Internal: function → fundef 333 333 | External: ident → typelist → type → fundef. … … 339 339 data. See module [AST] for more details. *) 340 340 341 ndefinition clight_program : Type≝ program fundef type.341 definition clight_program : Type[0] ≝ program fundef type. 342 342 343 343 (* * * Operations over types *) … … 345 345 (* * The type of a function definition. *) 346 346 347 nlet rec type_of_params (params: list (ident × type)) : typelist ≝347 let rec type_of_params (params: list (ident × type)) : typelist ≝ 348 348 match params with 349 349 [ nil ⇒ Tnil 350 | cons h rem ⇒ match h with [ mk_pair id ty ⇒ Tcons ty (type_of_params rem) ]351 ]. 352 353 ndefinition type_of_function : function → type ≝ λf.350 | cons h rem ⇒ match h with [ pair id ty ⇒ Tcons ty (type_of_params rem) ] 351 ]. 352 353 definition type_of_function : function → type ≝ λf. 354 354 Tfunction (type_of_params (fn_params f)) (fn_return f). 355 355 356 ndefinition type_of_fundef : fundef → type ≝ λf.356 definition type_of_fundef : fundef → type ≝ λf. 357 357 match f with 358 358 [ Internal fd ⇒ type_of_function fd … … 362 362 (* * Natural alignment of a type, in bytes. *) 363 363 (* FIXME: these are old values for 32 bit machines *) 364 nlet rec alignof (t: type) : Z ≝364 let rec alignof (t: type) : Z ≝ 365 365 match t return λ_.Z (* XXX appears to infer nat otherwise *) with 366 366 [ Tvoid ⇒ 1 … … 387 387 388 388 (* XXX: automatic generation? *) 389 nlet rec type_ind2389 let rec type_ind2 390 390 (P:type → Prop) (Q:fieldlist → Prop) 391 391 (vo:P Tvoid) … … 432 432 ]. 433 433 434 nlemma alignof_fields_pos:434 lemma alignof_fields_pos: 435 435 ∀f. alignof_fields f > 0. 436 napply fieldlist_ind;//;437 #i ;#t;#fs';#IH; nlapply (Zmax_r (alignof t) (alignof_fields fs'));438 napply Zlt_to_Zle_to_Zlt; //; nqed.439 440 nlemma alignof_pos:436 @fieldlist_ind //; 437 #i #t #fs' #IH lapply (Zmax_r (alignof t) (alignof_fields fs')); 438 @Zlt_to_Zle_to_Zlt //; qed. 439 440 lemma alignof_pos: 441 441 ∀t. alignof t > 0. 442 #t ;nelim t; nnormalize; //;443 ##[ ##1,2: #z; ncases z; //;444 ##| ##3,4: #i;napplyalignof_fields_pos445 ##] nqed.442 #t elim t; normalize; //; 443 [ 1,2: #z cases z; //; 444 | 3,4: #i @alignof_fields_pos 445 ] qed. 446 446 447 447 (* * Size of a type, in bytes. *) 448 448 449 ndefinition sizeof_pointer : region → Z ≝450 λsp. match sp with [ Data ⇒ 1 | IData ⇒ 1 | PData ⇒ 1 | XData ⇒ 2 | Code ⇒ 2 | Any ⇒ 3 ].451 452 nlet rec sizeof (t: type) : Z ≝449 definition sizeof_pointer : region → Z ≝ 450 λsp. match sp return λ_.Z with [ Data ⇒ 1 | IData ⇒ 1 | PData ⇒ 1 | XData ⇒ 2 | Code ⇒ 2 | Any ⇒ 3 ]. 451 452 let rec sizeof (t: type) : Z ≝ 453 453 match t return λ_.Z with 454 454 [ Tvoid ⇒ 1 … … 475 475 ]. 476 476 (* TODO: needs some Z_times results 477 nlemma sizeof_pos:477 lemma sizeof_pos: 478 478 ∀t. sizeof t > 0. 479 #t0 ;479 #t0 480 480 napply (type_ind2 (λt. sizeof t > 0) 481 481 (λf. sizeof_union f ≥ 0 ∧ ∀pos:Z. pos ≥ 0 → sizeof_struct f pos ≥ 0)); 482 ##[ ##1,4,6,9: //;483 ##| #i;ncases i;#s;//;484 ##| #f;ncases f;//485 ##| #t;#n;#H; nwhd in ⊢ (?%?);482 [ 1,4,6,9: //; 483 | #i cases i;#s //; 484 | #f cases f;// 485 | #t #n #H whd in ⊢ (?%?); 486 486 Proof. 487 487 intro t0. … … 520 520 Open Local Scope string_scope. 521 521 *) 522 nlet rec field_offset_rec (id: ident) (fld: fieldlist) (pos: Z)522 let rec field_offset_rec (id: ident) (fld: fieldlist) (pos: Z) 523 523 on fld : res Z ≝ 524 524 match fld with … … 531 531 ]. 532 532 533 ndefinition field_offset ≝ λid: ident. λfld: fieldlist.533 definition field_offset ≝ λid: ident. λfld: fieldlist. 534 534 field_offset_rec id fld 0. 535 535 536 nlet rec field_type (id: ident) (fld: fieldlist) on fld : res type :=536 let rec field_type (id: ident) (fld: fieldlist) on fld : res type := 537 537 match fld with 538 538 [ Fnil ⇒ Error ? (*MSG "Unknown field " :: CTX id :: nil*) … … 626 626 *) 627 627 628 ninductive mode: Type≝628 inductive mode: Type[0] ≝ 629 629 | By_value: memory_chunk → mode 630 630 | By_reference: mode 631 631 | By_nothing: mode. 632 632 633 ndefinition access_mode : type → mode ≝ λty.633 definition access_mode : type → mode ≝ λty. 634 634 match ty with 635 635 [ Tint i s ⇒ … … 662 662 *) 663 663 664 ninductive classify_add_cases : Type≝664 inductive classify_add_cases : Type[0] ≝ 665 665 | add_case_ii: classify_add_cases (**r int , int *) 666 666 | add_case_ff: classify_add_cases (**r float , float *) … … 669 669 | add_default: classify_add_cases. (**r other *) 670 670 671 ndefinition classify_add ≝ λty1: type. λty2: type.671 definition classify_add ≝ λty1: type. λty2: type. 672 672 (* 673 673 match ty1, ty2 with … … 694 694 ]. 695 695 696 ninductive classify_sub_cases : Type≝696 inductive classify_sub_cases : Type[0] ≝ 697 697 | sub_case_ii: classify_sub_cases (**r int , int *) 698 698 | sub_case_ff: classify_sub_cases (**r float , float *) … … 701 701 | sub_default: classify_sub_cases . (**r other *) 702 702 703 ndefinition classify_sub ≝ λty1: type. λty2: type.703 definition classify_sub ≝ λty1: type. λty2: type. 704 704 (* match ty1, ty2 with 705 705 | Tint _ _ , Tint _ _ ⇒ sub_case_ii … … 732 732 ]. 733 733 734 ninductive classify_mul_cases : Type≝734 inductive classify_mul_cases : Type[0] ≝ 735 735 | mul_case_ii: classify_mul_cases (**r int , int *) 736 736 | mul_case_ff: classify_mul_cases (**r float , float *) 737 737 | mul_default: classify_mul_cases . (**r other *) 738 738 739 ndefinition classify_mul ≝ λty1: type. λty2: type.739 definition classify_mul ≝ λty1: type. λty2: type. 740 740 match ty1 with 741 741 [ Tint _ _ ⇒ match ty2 with [ Tint _ _ ⇒ mul_case_ii | _ ⇒ mul_default ] … … 750 750 end. 751 751 *) 752 ninductive classify_div_cases : Type≝752 inductive classify_div_cases : Type[0] ≝ 753 753 | div_case_I32unsi: classify_div_cases (**r unsigned int32 , int *) 754 754 | div_case_ii: classify_div_cases (**r int , int *) … … 756 756 | div_default: classify_div_cases. (**r other *) 757 757 758 ndefinition classify_32un_aux ≝ λT:Type.λi.λs.λr1:T.λr2:T.758 definition classify_32un_aux ≝ λT:Type[0].λi.λs.λr1:T.λr2:T. 759 759 match i with [ I32 ⇒ 760 760 match s with [ Unsigned ⇒ r1 | _ ⇒ r2 ] 761 761 | _ ⇒ r2 ]. 762 762 763 ndefinition classify_div ≝ λty1: type. λty2: type.763 definition classify_div ≝ λty1: type. λty2: type. 764 764 match ty1 with 765 765 [ Tint i1 s1 ⇒ … … 773 773 ]. 774 774 (* 775 ndefinition classify_div ≝ λty1: type. λty2: type.775 definition classify_div ≝ λty1: type. λty2: type. 776 776 match ty1,ty2 with 777 777 | Tint I32 Unsigned, Tint _ _ ⇒ div_case_I32unsi … … 782 782 end. 783 783 *) 784 ninductive classify_mod_cases : Type≝784 inductive classify_mod_cases : Type[0] ≝ 785 785 | mod_case_I32unsi: classify_mod_cases (**r unsigned I32 , int *) 786 786 | mod_case_ii: classify_mod_cases (**r int , int *) 787 787 | mod_default: classify_mod_cases . (**r other *) 788 788 789 ndefinition classify_mod ≝ λty1:type. λty2:type.789 definition classify_mod ≝ λty1:type. λty2:type. 790 790 match ty1 with 791 791 [ Tint i1 s1 ⇒ … … 806 806 end . 807 807 *) 808 ninductive classify_shr_cases :Type≝808 inductive classify_shr_cases :Type[0] ≝ 809 809 | shr_case_I32unsi: classify_shr_cases (**r unsigned I32 , int *) 810 810 | shr_case_ii :classify_shr_cases (**r int , int *) 811 811 | shr_default : classify_shr_cases . (**r other *) 812 812 813 ndefinition classify_shr ≝ λty1: type. λty2: type.813 definition classify_shr ≝ λty1: type. λty2: type. 814 814 match ty1 with 815 815 [ Tint i1 s1 ⇒ … … 829 829 end. 830 830 *) 831 ninductive classify_cmp_cases : Type≝831 inductive classify_cmp_cases : Type[0] ≝ 832 832 | cmp_case_I32unsi: classify_cmp_cases (**r unsigned I32 , int *) 833 833 | cmp_case_ipip: classify_cmp_cases (**r int|ptr|array , int|ptr|array*) … … 835 835 | cmp_default: classify_cmp_cases . (**r other *) 836 836 837 ndefinition classify_cmp ≝ λty1:type. λty2:type.837 definition classify_cmp ≝ λty1:type. λty2:type. 838 838 match ty1 with 839 839 [ Tint i1 s1 ⇒ … … 875 875 end. 876 876 *) 877 ninductive classify_fun_cases : Type≝877 inductive classify_fun_cases : Type[0] ≝ 878 878 | fun_case_f: typelist → type → classify_fun_cases (**r (pointer to) function *) 879 879 | fun_default: classify_fun_cases . (**r other *) 880 880 881 ndefinition classify_fun ≝ λty: type.881 definition classify_fun ≝ λty: type. 882 882 match ty with 883 883 [ Tfunction args res ⇒ fun_case_f args res … … 891 891 and external functions. *) 892 892 893 ndefinition typ_of_type : type → typ ≝ λt.893 definition typ_of_type : type → typ ≝ λt. 894 894 match t with 895 895 [ Tfloat _ ⇒ ASTfloat … … 897 897 ]. 898 898 899 ndefinition opttyp_of_type : type → option typ ≝ λt.899 definition opttyp_of_type : type → option typ ≝ λt. 900 900 match t with 901 901 [ Tvoid ⇒ None ? … … 904 904 ]. 905 905 906 nlet rec typlist_of_typelist (tl: typelist) : list typ ≝906 let rec typlist_of_typelist (tl: typelist) : list typ ≝ 907 907 match tl with 908 908 [ Tnil ⇒ nil ? … … 910 910 ]. 911 911 912 ndefinition signature_of_type : typelist → type → signature ≝ λargs. λres.912 definition signature_of_type : typelist → type → signature ≝ λargs. λres. 913 913 mk_signature (typlist_of_typelist args) (opttyp_of_type res). 914 914 915 ndefinition external_function915 definition external_function 916 916 : ident → typelist → type → external_function ≝ λid. λtargs. λtres. 917 917 mk_external_function id (signature_of_type targs tres). -
Deliverables/D3.1/C-semantics/Errors.ma
r485 r487 14 14 (* *********************************************************************) 15 15 16 include "datatypes/pairs.ma". 17 include "Plogic/equality.ma". 18 include "Plogic/connectives.ma". 19 include "datatypes/sums.ma". 16 include "basics/types.ma". 17 include "basics/logic.ma". 20 18 21 19 (* * Error reporting and the error monad. *) … … 41 39 or [Error msg] to indicate failure. *) 42 40 43 ninductive res (A: Type) : Type≝41 inductive res (A: Type[0]) : Type[0] ≝ 44 42 | OK: A → res A 45 43 | Error: (* FIXME errmsg →*) res A. … … 50 48 with the following [bind] operation. *) 51 49 52 ndefinition bind ≝ λA,B:Type. λf: res A. λg: A → res B.50 definition bind ≝ λA,B:Type[0]. λf: res A. λg: A → res B. 53 51 match f with 54 52 [ OK x ⇒ g x … … 56 54 ]. 57 55 58 ndefinition bind2 ≝ λA,B,C: Type. λf: res (A × B). λg: A → B → res C.56 definition bind2 ≝ λA,B,C: Type[0]. λf: res (A × B). λg: A → B → res C. 59 57 match f with 60 [ OK v ⇒ match v with [ mk_pair x y ⇒ g x y ]58 [ OK v ⇒ match v with [ pair x y ⇒ g x y ] 61 59 | Error (*msg*) => Error ? (*msg*) 62 60 ]. … … 72 70 notation < "vbox('do' \nbsp 〈ident v1, ident v2〉 ← e; break e')" with precedence 40 for @{'bind2 ${e} (λ${ident v1}.λ${ident v2}.${e'})}. 73 71 notation < "vbox('do' \nbsp 〈ident v1 : ty1, ident v2 : ty2〉 ← e; break e')" with precedence 40 for @{'bind2 ${e} (λ${ident v1} : ${ty1}.λ${ident v2} : ${ty2}.${e'})}. 74 interpretation "error monad pairbind" 'bind2 e f = (bind2 ??? e f).72 interpretation "error monad Prod bind" 'bind2 e f = (bind2 ??? e f). 75 73 (*interpretation "error monad ret" 'ret e = (ret ? e). 76 74 notation "'ret' e" non associative with precedence 45 for @{'ret ${e}}.*) … … 87 85 : error_monad_scope. 88 86 *) 89 nremarkbind_inversion:90 ∀A,B: Type . ∀f: res A. ∀g: A → res B. ∀y: B.87 lemma bind_inversion: 88 ∀A,B: Type[0]. ∀f: res A. ∀g: A → res B. ∀y: B. 91 89 bind ?? f g = OK ? y → 92 90 ∃x. f = OK ? x ∧ g x = OK ? y. 93 #A B f g y; ncases f;94 ##[ #a e; @a; nwhd in e:(??%?); /2/;95 ##| #H; nwhd in H:(??%?); ndestruct (H);96 ##] nqed.97 98 nremarkbind2_inversion:99 ∀A,B,C: Type . ∀f: res (A×B). ∀g: A → B → res C. ∀z: C.91 #A #B #f #g #y cases f; 92 [ #a #e %{a} whd in e:(??%?); /2/; 93 | #H whd in H:(??%?); destruct (H); 94 ] qed. 95 96 lemma bind2_inversion: 97 ∀A,B,C: Type[0]. ∀f: res (A×B). ∀g: A → B → res C. ∀z: C. 100 98 bind2 ??? f g = OK ? z → 101 99 ∃x. ∃y. f = OK ? 〈x, y〉 ∧ g x y = OK ? z. 102 #A B C f g z; ncases f;103 ##[ #ab; ncases ab; #a b e; @a; @b; nwhd in e:(??%?); /2/;104 ##| #H; nwhd in H:(??%?); ndestruct105 ##] nqed.100 #A #B #C #f #g #z cases f; 101 [ #ab cases ab; #a #b #e %{a} %{b} whd in e:(??%?); /2/; 102 | #H whd in H:(??%?); destruct 103 ] qed. 106 104 107 105 (* … … 195 193 196 194 197 ndefinition opt_to_res ≝ λA.λv:option A. match v with [ None ⇒ Error A | Some v ⇒ OK A v ].198 nlemma opt_OK: ∀A,P,e.195 definition opt_to_res ≝ λA.λv:option A. match v with [ None ⇒ Error A | Some v ⇒ OK A v ]. 196 lemma opt_OK: ∀A,P,e. 199 197 (∀v. e = Some ? v → P v) → 200 198 match opt_to_res A e with [ Error ⇒ True | OK v ⇒ P v ]. 201 #A P e; nelim e; /2/;202 nqed.199 #A #P #e elim e; /2/; 200 qed. -
Deliverables/D3.1/C-semantics/Events.ma
r478 r487 22 22 (*include "Floats.ma".*) 23 23 include "Values.ma". 24 include " datatypes/list.ma".24 include "basics/list.ma". 25 25 include "extralib.ma". 26 26 include "CostLabel.ma". … … 37 37 world. *) 38 38 39 ninductive eventval: Type:=39 inductive eventval: Type[0] := 40 40 | EVint: int -> eventval 41 41 | EVfloat: float -> eventval. 42 42 43 ninductive event : Type≝43 inductive event : Type[0] ≝ 44 44 | EVcost: costlabel → event 45 45 | EVextcall: ∀ev_name: ident. ∀ev_args: list eventval. ∀ev_res: eventval. event. … … 48 48 Traces are of two kinds: finite (type [trace]) or infinite (type [traceinf]). *) 49 49 50 ndefinition trace := list event.51 52 ndefinition E0 : trace := nil ?.53 54 ndefinition Echarge : costlabel → trace ≝50 definition trace := list event. 51 52 definition E0 : trace := nil ?. 53 54 definition Echarge : costlabel → trace ≝ 55 55 λlabel. EVcost label :: (nil ?). 56 56 57 ndefinition Eextcall : ident → list eventval → eventval → trace ≝57 definition Eextcall : ident → list eventval → eventval → trace ≝ 58 58 λname: ident. λargs: list eventval. λres: eventval. 59 59 EVextcall name args res :: (nil ?). 60 60 61 ndefinition Eapp : trace → trace → trace ≝ λt1,t2. t1 @ t2.62 63 ncoinductive traceinf : Type:=61 definition Eapp : trace → trace → trace ≝ λt1,t2. t1 @ t2. 62 63 coinductive traceinf : Type[0] := 64 64 | Econsinf: event -> traceinf -> traceinf. 65 65 66 nlet rec Eappinf (t: trace) (T: traceinf) on t : traceinf :=66 let rec Eappinf (t: trace) (T: traceinf) on t : traceinf := 67 67 match t with 68 68 [ nil => T … … 87 87 Infix "***" := Eappinf (at level 60, right associativity). 88 88 *) 89 nlemma E0_left: ∀t. E0 ⧺ t = t.90 //; nqed.91 92 nlemma E0_right: ∀t. t ⧺ E0 = t.93 napply append_nil; nqed.94 95 nlemma Eapp_assoc: ∀t1,t2,t3. (t1 ⧺ t2) ⧺ t3 = t1 ⧺ (t2 ⧺ t3).96 napply associative_append; nqed.97 98 nlemma Eapp_E0_inv: ∀t1,t2. t1 ⧺ t2 = E0 → t1 = E0 ∧ t2 = E0.99 napply nil_append_nil_both; nqed.89 lemma E0_left: ∀t. E0 ⧺ t = t. 90 //; qed. 91 92 lemma E0_right: ∀t. t ⧺ E0 = t. 93 @append_nil qed. 94 95 lemma Eapp_assoc: ∀t1,t2,t3. (t1 ⧺ t2) ⧺ t3 = t1 ⧺ (t2 ⧺ t3). 96 @associative_append qed. 97 98 lemma Eapp_E0_inv: ∀t1,t2. t1 ⧺ t2 = E0 → t1 = E0 ∧ t2 = E0. 99 @nil_append_nil_both qed. 100 100 101 nlemma E0_left_inf: ∀T. E0 ⧻ T = T.102 //; nqed.103 104 nlemma Eappinf_assoc: ∀t1,t2,T. (t1 ⧺ t2) ⧻ T = t1 ⧻ (t2 ⧻ T).105 #t1 ; nelim t1; #t2 T; nnormalize; //; nqed.101 lemma E0_left_inf: ∀T. E0 ⧻ T = T. 102 //; qed. 103 104 lemma Eappinf_assoc: ∀t1,t2,T. (t1 ⧺ t2) ⧻ T = t1 ⧻ (t2 ⧻ T). 105 #t1 elim t1; #t2 #T normalize; //; qed. 106 106 107 107 (* … … 137 137 infinite concatenations of nonempty finite traces. *) 138 138 139 ncoinductive traceinf': Type≝139 coinductive traceinf': Type[0] ≝ 140 140 | Econsinf': ∀t: trace. ∀T: traceinf'. t ≠ E0 → traceinf'. 141 141 142 ndefinition split_traceinf' : ∀t:trace. traceinf' → t ≠ E0 → event × traceinf' ≝142 definition split_traceinf' : ∀t:trace. traceinf' → t ≠ E0 → event × traceinf' ≝ 143 143 λt,T. 144 144 match t return λt0.t0 ≠ E0 → ? with … … 150 150 ]) (refl ? t') 151 151 ]. 152 ##[ *; #NE; napply False_rect_Type0; napply NE; napply refl; 153 ##| @; #E'; nrewrite > E' in E; #E; nwhd in E:(??%%); ndestruct (E);154 ##] nqed.155 156 nlet corec traceinf_of_traceinf' (T': traceinf') : traceinf ≝152 [ *; #NE @False_rect_Type0 @NE @refl 153 | % #E' >E' in E #E whd in E:(??%%); destruct (E); 154 ] qed. 155 156 let corec traceinf_of_traceinf' (T': traceinf') : traceinf ≝ 157 157 match T' with 158 158 [ Econsinf' t T'' NOTEMPTY ⇒ 159 match split_traceinf' t T'' NOTEMPTY with [ mk_pair e tl ⇒159 match split_traceinf' t T'' NOTEMPTY with [ pair e tl ⇒ 160 160 Econsinf e (traceinf_of_traceinf' tl) ] 161 161 ]. 162 162 163 nremarkunroll_traceinf':163 lemma unroll_traceinf': 164 164 ∀T. T = match T with [ Econsinf' t T' NE ⇒ Econsinf' t T' NE ]. 165 #T ; ncases T; //; nqed.166 167 nremarkunroll_traceinf:165 #T cases T; //; qed. 166 167 lemma unroll_traceinf: 168 168 ∀T. T = match T with [ Econsinf t T' ⇒ Econsinf t T' ]. 169 #T ; ncases T; //;170 nqed.171 172 nlemma traceinf_traceinfp_app:169 #T cases T #ev #tr @refl (* XXX //; *) 170 qed. 171 172 lemma traceinf_traceinfp_app: 173 173 ∀t,T,NE. 174 174 traceinf_of_traceinf' (Econsinf' t T NE) = t ⧻ traceinf_of_traceinf' T. 175 #t ; nelim t;176 ##[ #T NE; ncases NE; #NE'; napply False_ind; napply NE'; napply refl; 177 ##| #h t'; ncases t'; ##[ ##2: #h' t''; ##] #IH T NE; 178 nrewrite > (unroll_traceinf (traceinf_of_traceinf' ?));179 nwhd in ⊢ (??%?); //; nrewrite > (IH …); napply refl;180 ##] nqed.175 #t elim t; 176 [ #T #NE cases NE; #NE' @False_ind @NE' @refl 177 | #h #t' cases t'; [ 2: #h' #t'' ] #IH #T #NE 178 >(unroll_traceinf (traceinf_of_traceinf' ?)) 179 whd in ⊢ (??%?); //; >(IH …) @refl 180 ] qed. 181 181 182 182 (* <<< *) … … 187 187 from the operating system. *) 188 188 189 ninductive eventval_match: eventval -> typ -> val -> Prop :=189 inductive eventval_match: eventval -> typ -> val -> Prop := 190 190 | ev_match_int: 191 191 ∀i. eventval_match (EVint i) ASTint (Vint i) … … 193 193 ∀f. eventval_match (EVfloat f) ASTfloat (Vfloat f). 194 194 195 ninductive eventval_list_match: list eventval -> list typ -> list val -> Prop :=195 inductive eventval_list_match: list eventval -> list typ -> list val -> Prop := 196 196 | evl_match_nil: 197 197 eventval_list_match (nil ?) (nil ?) (nil ?) …