[1882] | 1 | include "joint/blocks.ma". |
---|
[2186] | 2 | include "joint/Traces.ma". |
---|
[2529] | 3 | include "joint/semanticsUtils.ma". |
---|
[1882] | 4 | |
---|
[2595] | 5 | include "common/StatusSimulation.ma". (* for trace_any_any_free *) |
---|
| 6 | |
---|
[2186] | 7 | (* let rec seq_list_labels p g (b : list (joint_step p g)) on b : list label ≝ |
---|
[1949] | 8 | match b with |
---|
| 9 | [ nil ⇒ [ ] |
---|
| 10 | | cons hd tl ⇒ step_labels … hd @ seq_list_labels … tl |
---|
[2186] | 11 | ].*) |
---|
[1949] | 12 | |
---|
[2529] | 13 | definition repeat_eval_seq_no_pc ≝ |
---|
| 14 | λp : evaluation_params.λcurr_id,curr_fn. |
---|
| 15 | m_fold … (eval_seq_no_pc … (ev_genv … p) curr_id curr_fn). |
---|
[1882] | 16 | |
---|
[2595] | 17 | definition taaf_cons : ∀S : abstract_status.∀s1,s2,s3. |
---|
| 18 | as_execute S s1 s2 → |
---|
| 19 | as_classifier … s1 cl_other → |
---|
| 20 | ∀taaf : trace_any_any_free S s2 s3. |
---|
| 21 | (if taaf_non_empty … taaf then ¬as_costed … s2 else True) → |
---|
| 22 | trace_any_any_free S s1 s3 ≝ |
---|
| 23 | λS,s1,s2,s3,H,I,tl. |
---|
| 24 | match tl return λs2,s3,tl. |
---|
| 25 | as_execute … s1 s2 → |
---|
| 26 | if taaf_non_empty … tl then ¬as_costed … s2 else True → |
---|
| 27 | trace_any_any_free S s1 s3 |
---|
| 28 | with |
---|
| 29 | [ taaf_base s2 ⇒ λH.λ_.taaf_step … (taa_base …) H I |
---|
| 30 | | taaf_step s2 s3 s4 taa H' I' ⇒ |
---|
| 31 | λH,J.taaf_step … (taa_step … H I J taa) H' I' |
---|
| 32 | | taaf_step_jump s2 s3 s4 taa H' I' K' ⇒ |
---|
| 33 | λH,J.taaf_step_jump … (taa_step ???? H I J taa) H' I' K' |
---|
| 34 | ] H. |
---|
[1949] | 35 | |
---|
[2595] | 36 | lemma taaf_cons_non_empty : ∀S,s1,s2,s3,H,I,tl,J. |
---|
| 37 | bool_to_Prop (taaf_non_empty … (taaf_cons S s1 s2 s3 H I tl J)). |
---|
| 38 | #S #s1 #s2 #s3 #H #I #tl lapply H -H cases tl |
---|
| 39 | [ #s #H * % |*: #s2 #s3 #s4 #taa #H' #I' [2: #K'] #H #J % ] |
---|
| 40 | qed. |
---|
| 41 | |
---|
| 42 | let rec produce_trace_any_any_free_aux |
---|
[2186] | 43 | (p : evaluation_params) |
---|
[2529] | 44 | (st : state_pc p) |
---|
| 45 | curr_id curr_fn |
---|
[2186] | 46 | (b : list (joint_seq p (globals p))) on b : |
---|
| 47 | ∀l : list (code_point p). |
---|
| 48 | ∀dst : code_point p. |
---|
| 49 | ∀st' : state p. |
---|
[2529] | 50 | fetch_internal_function … (ev_genv p) (pc_block (pc … st)) = |
---|
| 51 | return 〈curr_id, curr_fn〉 → |
---|
| 52 | let src ≝ point_of_pc p (pc … st) in |
---|
[2595] | 53 | (* disambiguation: select 3rd (step list in code) *) |
---|
[2529] | 54 | src ~❨b, l❩~> dst in joint_if_code … curr_fn → |
---|
[2595] | 55 | All ? (no_cost_label …) b → |
---|
[2529] | 56 | repeat_eval_seq_no_pc p curr_id curr_fn b st = return st' → |
---|
[2595] | 57 | Σtaaf : trace_any_any_free (joint_abstract_status p) st |
---|
| 58 | (mk_state_pc ? st' (pc_of_point p (pc_block (pc … st)) dst) (last_pop … st)). |
---|
| 59 | (not_empty … b ↔ bool_to_Prop (taaf_non_empty … taaf)) ≝ |
---|
| 60 | match b |
---|
| 61 | return λb.∀l,dst.?→?→?→?→?→ Σtaaf.(not_empty ? b ↔ bool_to_Prop (taaf_non_empty … taaf)) |
---|
[2186] | 62 | with |
---|
| 63 | [ nil ⇒ |
---|
[2595] | 64 | λl,dst,st',fd_prf,in_code,all_other,EQ1. |
---|
| 65 | «taaf_base (joint_abstract_status p) st |
---|
| 66 | ⌈trace_any_any_free ??? ↦ ?⌉,?» |
---|
[2186] | 67 | | cons hd tl ⇒ |
---|
| 68 | λl. |
---|
[2595] | 69 | match l return λx.∀dst,st'.?→?→?→?→Σtaaf.(True ↔ bool_to_Prop (taaf_non_empty … taaf)) with |
---|
| 70 | [ nil ⇒ λdst,st',fd_prf,in_code.⊥ |
---|
[2186] | 71 | | cons mid rest ⇒ |
---|
[2595] | 72 | λdst,st',fd_prf,in_code,all_other,EQ1. |
---|
[2529] | 73 | let mid_pc ≝ pc_of_point p (pc_block (pc … st)) mid in |
---|
| 74 | match eval_seq_no_pc … (ev_genv p) curr_id curr_fn hd st |
---|
| 75 | return λx.eval_seq_no_pc … (ev_genv p) curr_id curr_fn hd st = x → |
---|
[2595] | 76 | Σtaaf : trace_any_any_free (joint_abstract_status p) st |
---|
| 77 | (mk_state_pc ? st' (pc_of_point p (pc_block (pc … st)) dst) (last_pop … st)). |
---|
| 78 | (True ↔ bool_to_Prop (taaf_non_empty … taaf)) with |
---|
| 79 | [ OK st_mid ⇒ λEQ2. |
---|
| 80 | let tr_tl ≝ produce_trace_any_any_free_aux ? |
---|
| 81 | (mk_state_pc ? st_mid mid_pc (last_pop … st)) |
---|
| 82 | curr_id curr_fn tl rest dst ????? in |
---|
| 83 | «taaf_cons … tr_tl ?,?» |
---|
[2529] | 84 | | _ ⇒ λEQ2.⊥ |
---|
[2186] | 85 | ] (refl …) |
---|
[1949] | 86 | ] |
---|
[2595] | 87 | ]. @hide_prf |
---|
| 88 | [1,2: [2: % [*] generalize in ⊢ (?(???? (match % with [ _ ⇒ ?])) → ?); ] |
---|
| 89 | whd in EQ1 : (??%%); |
---|
| 90 | cases l in in_code; whd in ⊢ (%→?); [2,4: #hd #tl * ] #EQ destruct |
---|
| 91 | >pc_of_point_of_pc cases st // #a #b #c #e >(K ?? e) normalize nodelta * |
---|
[2529] | 92 | | @in_code |
---|
[2595] | 93 | |12: whd in EQ1 : (??%%); >EQ2 in EQ1; whd in ⊢ (??%?→?); #EQ1 destruct(EQ1) |
---|
| 94 | |4: cases tr_tl -tr_tl #tr_tl cases tl in in_code all_other; |
---|
| 95 | [ #_ #_ * #_ cases (taaf_non_empty ????) |
---|
| 96 | [ #ABS cases (ABS I) | #_ % ] |
---|
| 97 | | #hd' #tl' ** #nxt * #EQstmt_at #EQ_nxt cases rest [*] #mid' #rest' * |
---|
| 98 | * #nxt' * #EQstmt_at' #EQ_nxt' #_ * #hd_other * #hd_other' |
---|
| 99 | #_ * #H #_ >(H I) % whd in ⊢ (%→?); |
---|
| 100 | whd in ⊢ (?(??%?)→?); whd in match (as_pc_of ??); |
---|
| 101 | >fetch_statement_eq [2: whd in match point_of_pc; |
---|
| 102 | normalize nodelta >point_of_offset_of_point @EQstmt_at' |3: @fd_prf |*:] |
---|
| 103 | normalize nodelta |
---|
| 104 | >(no_label_uncosted … hd_other') * #ABS @ABS % |
---|
| 105 | ] |
---|
| 106 | |7: % [2: #_ %] * @taaf_cons_non_empty |
---|
[2529] | 107 | ] |
---|
| 108 | change with (! y ← ? ; repeat_eval_seq_no_pc ????? = ?) in EQ1; |
---|
| 109 | >EQ2 in EQ1; >m_return_bind |
---|
| 110 | cases in_code -in_code * #nxt * #EQ_stmt_at #EQ_mid #rest_in_code |
---|
[2595] | 111 | cases all_other -all_other #hd_no_cost #tl_other |
---|
[2529] | 112 | #EQ1 |
---|
| 113 | try assumption |
---|
[2595] | 114 | [2: whd whd in match eval_state; normalize nodelta |
---|
[2529] | 115 | >(fetch_statement_eq … fd_prf EQ_stmt_at) |
---|
[2324] | 116 | >m_return_bind |
---|
[2529] | 117 | whd in match eval_statement_no_pc; normalize nodelta |
---|
| 118 | >EQ2 >m_return_bind |
---|
| 119 | whd in match eval_statement_advance; normalize nodelta |
---|
| 120 | whd in match next; normalize nodelta |
---|
[2324] | 121 | whd in match succ_pc; normalize nodelta |
---|
[2529] | 122 | >EQ_mid % |
---|
[2595] | 123 | |1: whd whd in ⊢ (??%?); |
---|
| 124 | >(fetch_statement_eq … fd_prf EQ_stmt_at) normalize nodelta % |
---|
| 125 | |3: normalize nodelta >point_of_pc_of_point assumption |
---|
[2324] | 126 | ] |
---|
[2595] | 127 | qed. |
---|
| 128 | |
---|
| 129 | definition produce_trace_any_any_free : |
---|
| 130 | ∀p : evaluation_params. |
---|
| 131 | ∀st : state_pc p. |
---|
| 132 | ∀curr_id,curr_fn. |
---|
| 133 | ∀b : list (joint_seq p (globals p)). |
---|
| 134 | ∀l : list (code_point p). |
---|
| 135 | ∀dst : code_point p. |
---|
| 136 | ∀st' : state p. |
---|
| 137 | fetch_internal_function … (ev_genv p) (pc_block (pc … st)) = |
---|
| 138 | return 〈curr_id, curr_fn〉 → |
---|
| 139 | let src ≝ point_of_pc p (pc … st) in |
---|
| 140 | (* disambiguation: select 3rd (step list in code) *) |
---|
| 141 | src ~❨b, l❩~> dst in joint_if_code … curr_fn → |
---|
| 142 | All ? (no_cost_label …) b → |
---|
| 143 | repeat_eval_seq_no_pc p curr_id curr_fn b st = return st' → |
---|
| 144 | trace_any_any_free (joint_abstract_status p) st |
---|
| 145 | (mk_state_pc ? st' (pc_of_point p (pc_block (pc … st)) dst) (last_pop … st)) ≝ |
---|
| 146 | λp,st,curr_id,curr_fn,b,l,dst,st',prf1,prf2,prf3,prf4. |
---|
| 147 | produce_trace_any_any_free_aux p st curr_id curr_fn b l dst st' prf1 prf2 prf3 prf4. |
---|