include "joint/blocks.ma". include "joint/Traces.ma". include "joint/semanticsUtils.ma". include "common/StatusSimulation.ma". (* for trace_any_any_free *) (* let rec seq_list_labels p g (b : list (joint_step p g)) on b : list label ≝ match b with [ nil ⇒ [ ] | cons hd tl ⇒ step_labels … hd @ seq_list_labels … tl ].*) definition repeat_eval_seq_no_pc ≝ λp : evaluation_params.λcurr_id. m_fold … (eval_seq_no_pc … (ev_genv … p) curr_id). definition taaf_cons : ∀S : abstract_status.∀s1,s2,s3. as_execute S s1 s2 → as_classifier … s1 cl_other → ∀taaf : trace_any_any_free S s2 s3. (if taaf_non_empty … taaf then ¬as_costed … s2 else True) → trace_any_any_free S s1 s3 ≝ λS,s1,s2,s3,H,I,tl. match tl return λs2,s3,tl. as_execute … s1 s2 → if taaf_non_empty … tl then ¬as_costed … s2 else True → trace_any_any_free S s1 s3 with [ taaf_base s2 ⇒ λH.λ_.taaf_step … (taa_base …) H I | taaf_step s2 s3 s4 taa H' I' ⇒ λH,J.taaf_step … (taa_step … H I J taa) H' I' | taaf_step_jump s2 s3 s4 taa H' I' K' ⇒ λH,J.taaf_step_jump … (taa_step ???? H I J taa) H' I' K' ] H. lemma taaf_cons_non_empty : ∀S,s1,s2,s3,H,I,tl,J. bool_to_Prop (taaf_non_empty … (taaf_cons S s1 s2 s3 H I tl J)). #S #s1 #s2 #s3 #H #I #tl lapply H -H cases tl [ #s #H * % |*: #s2 #s3 #s4 #taa #H' #I' [2: #K'] #H #J % ] qed. lemma produce_step_trace : ∀p : prog_params. ∀st : state_pc p. ∀curr_id,curr_fn. ∀s : joint_seq p (globals p). ∀dst : code_point p. ∀st' : state p. fetch_internal_function … (ev_genv p) (pc_block (pc … st)) = return 〈curr_id, curr_fn〉 → let src ≝ point_of_pc p (pc … st) in step_in_code … (joint_if_code … curr_fn) src s dst → eval_seq_no_pc p (globals p) (ev_genv p) curr_id s st = return st' → as_execute (joint_abstract_status p) st (mk_state_pc ? st' (pc_of_point p (pc_block (pc … st)) dst) (last_pop … st)) ∧ as_classifier (joint_abstract_status p) st cl_other. #p #st#curr_id #curr_fn #s #dst #st' #EQfetch * #nxt * #EQstmt_at #EQdst #EQeval whd in ⊢ (?%%); whd in ⊢ (?(??%?)(??%?)); whd in match eval_statement_no_pc; whd in match fetch_statement; normalize nodelta >EQfetch >m_return_bind >m_return_bind >EQstmt_at >m_return_bind normalize nodelta % [2: % ] whd in ⊢ (??%?); normalize nodelta >EQeval whd in ⊢ (??%%); @eq_f whd in ⊢ (??%?); @eq_f2 [2: %] whd in match succ_pc; normalize nodelta @eq_f @EQdst qed. let rec produce_trace_any_any_free_aux (p : prog_params) (st : state_pc p) curr_id curr_fn (b : list (joint_seq p (globals p))) on b : ∀l : list (code_point p). ∀dst : code_point p. ∀st' : state p. fetch_internal_function … (ev_genv p) (pc_block (pc … st)) = return 〈curr_id, curr_fn〉 → let src ≝ point_of_pc p (pc … st) in seq_list_in_code … (joint_if_code … curr_fn) src b l dst → repeat_eval_seq_no_pc p curr_id b st = return st' → Σtaaf : trace_any_any_free (joint_abstract_status p) st (mk_state_pc ? st' (pc_of_point p (pc_block (pc … st)) dst) (last_pop … st)). (not_empty … b ↔ bool_to_Prop (taaf_non_empty … taaf)) ≝ match b return λb.∀l,dst.?→?→?→?→ Σtaaf.(not_empty ? b ↔ bool_to_Prop (taaf_non_empty … taaf)) with [ nil ⇒ λl,dst,st',fd_prf,in_code,EQ1. «taaf_base (joint_abstract_status p) st ⌈trace_any_any_free ??? ↦ ?⌉,?» | cons hd tl ⇒ λl. match l return λx.∀dst,st'.?→?→?→Σtaaf.(True ↔ bool_to_Prop (taaf_non_empty … taaf)) with [ nil ⇒ λdst,st',fd_prf,in_code.⊥ | cons _ rest ⇒ λdst. let mid ≝ match rest with [ nil ⇒ dst | cons mid _ ⇒ mid ] in λst',fd_prf,in_code,EQ1. let mid_pc ≝ pc_of_point p (pc_block (pc … st)) mid in match eval_seq_no_pc … (ev_genv p) curr_id hd st return λx.eval_seq_no_pc … (ev_genv p) curr_id hd st = x → Σtaaf : trace_any_any_free (joint_abstract_status p) st (mk_state_pc ? st' (pc_of_point p (pc_block (pc … st)) dst) (last_pop … st)). (True ↔ bool_to_Prop (taaf_non_empty … taaf)) with [ OK st_mid ⇒ λEQ2. let tr_tl ≝ produce_trace_any_any_free_aux ? (mk_state_pc ? st_mid mid_pc (last_pop … st)) curr_id curr_fn tl rest dst ???? in «taaf_cons … tr_tl ?,?» | _ ⇒ λEQ2.⊥ ] (refl …) ] ]. @hide_prf [1,2: [2: % [*] generalize in ⊢ (?(???? (match % with [ _ ⇒ ?])) → ?); ] whd in EQ1 : (??%%); cases l in in_code; whd in ⊢ (%→?); [2,4: #hd #tl * #ABS destruct ] * #_ #EQ destruct >pc_of_point_of_pc cases st // #a #b #c #e >(K ?? e) normalize nodelta * | cases in_code #a * #b ** #ABS destruct |11: whd in EQ1 : (??%%); >EQ2 in EQ1; whd in ⊢ (??%?→?); #EQ1 destruct(EQ1) |4: cases tr_tl -tr_tl #tr_tl * #_ #H @if_elim [2: #_ % ] #G lapply (H G) -H -G cases tl in in_code; [ #_ * ] #hd' #tl' * #mid' * #rest' ** #EQ * #nxt * #EQstmt_at #EQ_nxt * #mid'' * #rest'' ** #EQ' * #nxt' * #EQstmt_at' #EQnxt' normalize nodelta -mid_pc destruct #_ #_ % whd in ⊢ (%→?); whd in ⊢ (?(??%?)→?); whd in match (as_pc_of ??); >fetch_statement_eq [2: whd in match point_of_pc; normalize nodelta >point_of_offset_of_point @EQstmt_at' |3: @fd_prf |*:] normalize nodelta * #ABS @ABS % |7: % [2: #_ %] #_ @taaf_cons_non_empty ] change with (! y ← ? ; repeat_eval_seq_no_pc ???? = ?) in EQ1; >EQ2 in EQ1; >m_return_bind cases in_code -in_code #mid' * #rest' ** #EQ1 #step_in cases (step_in) #nxt * #EQ_stmt_at #EQ_mid' #rest_in_code normalize nodelta #EQ3 destruct skip (mid_pc) try assumption [ whd whd in ⊢ (??%?); >(fetch_statement_eq … fd_prf EQ_stmt_at) normalize nodelta % |*: cases tl in rest_in_code; [1,3: * #EQ4 #EQ5 destruct normalize nodelta |*: #hd' #tl' * #mid'' * #rest'' ** #EQ4 #step_in' #rest_in_code' destruct normalize nodelta ] [1,3: @(proj1 … (produce_step_trace … fd_prf … EQ2)) assumption |2: %[%] @point_of_pc_of_point |4: >point_of_pc_of_point %[| %[| %{rest_in_code'} %{step_in'} %]] ] ] qed. definition produce_trace_any_any_free : ∀p : prog_params. ∀st : state_pc p. ∀curr_id,curr_fn. ∀b : list (joint_seq p (globals p)). ∀l : list (code_point p). ∀dst : code_point p. ∀st' : state p. fetch_internal_function … (ev_genv p) (pc_block (pc … st)) = return 〈curr_id, curr_fn〉 → let src ≝ point_of_pc p (pc … st) in seq_list_in_code … (joint_if_code … curr_fn) src b l dst → repeat_eval_seq_no_pc p curr_id b st = return st' → trace_any_any_free (joint_abstract_status p) st (mk_state_pc ? st' (pc_of_point p (pc_block (pc … st)) dst) (last_pop … st)) ≝ λp,st,curr_id,curr_fn,b,l,dst,st',prf1,prf2,prf3. produce_trace_any_any_free_aux p st curr_id curr_fn b l dst st' prf1 prf2 prf3. (* when a seq_list is coerced to a step_block *) definition produce_trace_any_any_free_coerced : ∀p : prog_params. ∀st : state_pc p. ∀curr_id,curr_fn. ∀b : list (joint_seq p (globals p)). ∀l : list (code_point p). ∀dst : code_point p. ∀st' : state p. fetch_internal_function … (ev_genv p) (pc_block (pc … st)) = return 〈curr_id, curr_fn〉 → let src ≝ point_of_pc p (pc … st) in src ~❨b, l❩~> dst in joint_if_code … curr_fn → repeat_eval_seq_no_pc p curr_id b st = return st' → trace_any_any_free (joint_abstract_status p) st (mk_state_pc ? st' (pc_of_point p (pc_block (pc … st)) dst) (last_pop … st)) ≝ λp,st,curr_id,curr_fn,b.?. #l #dst #st' #fd_prf #prf lapply (coerced_step_list_in_code … prf) inversion b normalize nodelta [ #_ #in_code whd in ⊢ (??%%→?); #EQ destruct cases (produce_step_trace … fd_prf … in_code (refl …)) #H #G %2{(taa_base …)} assumption | #hd #tl #_ #EQ