include "joint/joint_semantics.ma". include "common/StructuredTraces.ma". record evaluation_params : Type[1] ≝ { globals: list ident ; sparams:> sem_params ; ev_genv: genv sparams globals (* ; io_env : state sparams → ∀o:io_out.res (io_in o) *) }. record prog_params : Type[1] ≝ { prog_spars : sem_params ; prog : joint_program prog_spars ; stack_sizes : ident → option ℕ (* ; prog_io : state prog_spars → ∀o.res (io_in o) *) }. lemma map_Exists : ∀A,B,f,P,l.Exists B P (map A B f l) → Exists ? (λx.P (f x)) l. #A #B #f #P #l elim l -l [*] #hd #tl #IH * [ #Phd %1{Phd} | #Ptl %2{(IH Ptl)} ] qed. lemma Exists_In : ∀A,P,l.Exists A P l → ∃x.In A l x ∧ P x. #A #P #l elim l -l [*] #hd #tl #IH * [ #Phd %{hd} %{Phd} %1 % | #Ptl elim (IH Ptl) #x * #H1 #H2 %{x} %{H2} %2{H1} ] qed. lemma In_Exists : ∀A,P,l,x.In A l x → P x → Exists A P l. #A #P #l elim l -l [ #x *] #hd #tl #IH #x * [ #EQ >EQ #H %1{H} | #Intl #Px %2{(IH … Intl Px)} ] qed. lemma Exists_mp : ∀A,P,Q,l.(∀x.P x → Q x) → Exists A P l → Exists A Q l. #A #P #Q #l #H elim l -l [*] #hd #tl #IH * #G [ %1 | %2 ] /2/ qed. definition make_global : prog_params → evaluation_params ≝ λpars. let p ≝ prog pars in let spars ≝ prog_spars pars in let genv ≝ globalenv_noinit ? p in let get_pc_lbl ≝ λid,lbl. ! bl ← block_of_funct_id … spars genv id ; pc_of_label … genv bl lbl in mk_evaluation_params (prog_var_names … p) spars (mk_genv_gen ?? genv ? (stack_sizes pars) get_pc_lbl) (* (prog_io pars) *). #s #H elim (find_symbol_exists … p s ?) [ #bl #EQ >EQ % #ABS destruct(ABS)|*:] @Exists_append_r @(Exists_mp … H) // qed. coercion prog_params_to_ev_params : ∀p : prog_params.evaluation_params ≝ make_global on _p : prog_params to evaluation_params. definition make_initial_state : ∀pars: prog_params.res (state_pc pars) ≝ λpars.let p ≝ prog pars in let sem_globals : evaluation_params ≝ pars in let ge ≝ ev_genv sem_globals in let m0 ≝ alloc_mem … p in let 〈m,spb〉 as H ≝ alloc … m0 0 external_ram_size in let spp : xpointer ≝ «mk_pointer spb (mk_offset (bitvector_of_Z ? external_ram_size)), pi2 … spb» in (* let ispp : xpointer ≝ mk_pointer ispb (mk_offset (bitvector_of_nat ? 47)) in *) let main ≝ prog_main … p in let st0 ≝ mk_state pars (Some ? (empty_framesT …)) empty_is (BBbit false) (empty_regsT … spp) m in (* use exit_pc as ra and call_dest_for_main as dest *) let st0' ≝ mk_state_pc … (set_sp … spp st0) exit_pc exit_pc in ! st0_no_pc ← save_frame ?? sem_globals ID (call_dest_for_main … pars) st0' ; let st0'' ≝ set_no_pc … st0_no_pc st0' in ! bl ← block_of_call … ge (inl … main) st0''; ! 〈i, fn〉 ← fetch_function … ge bl ; match fn with [ Internal ifn ⇒ ! st' ← eval_internal_call … ge i ifn (call_args_for_main … pars) st0'' ; let pc' ≝ pc_of_point … bl (joint_if_entry … ifn) in return mk_state_pc … st' pc' exit_pc | External _ ⇒ Error … [MSG BadMain; CTX ? main] (* External main not supported: why? *) ]. cases m0 in H; #m1 #m2 #m3 #H whd in H:(???%); destruct whd in ⊢(??%?); @Zleb_elim_Type0 // #abs @⊥ @(absurd … (irreflexive_Zlt …)) % #I cases (I OZ) /3 by Zlt_to_Zle_to_Zlt/ qed. definition joint_classify_step : ∀p : evaluation_params.joint_step p (globals … p) → status_class ≝ λp,stmt. match stmt with [ CALL f args dest ⇒ cl_call | COND _ _ ⇒ cl_jump | _ ⇒ cl_other ]. definition joint_classify_final : ∀p : evaluation_params.joint_fin_step p → status_class ≝ λp,stmt. match stmt with [ GOTO _ ⇒ cl_other | RETURN ⇒ cl_return | TAILCALL _ f args ⇒ cl_tailcall ]. definition joint_classify : ∀p : evaluation_params.state_pc p → status_class ≝ λp,st. match fetch_statement … (ev_genv p) (pc … st) with [ OK i_fn_s ⇒ match \snd i_fn_s with [ sequential s _ ⇒ joint_classify_step p s | final s ⇒ joint_classify_final p s | FCOND _ _ _ _ ⇒ cl_jump ] | _ ⇒ cl_other ]. lemma joint_classify_call : ∀p : evaluation_params.∀st. joint_classify p st = cl_call → ∃i_f,f',args,dest,next. fetch_statement … (ev_genv p) (pc … st) = OK ? 〈i_f, sequential … (CALL … f' args dest) next〉. #p #st whd in match joint_classify; normalize nodelta inversion (fetch_statement … (ev_genv p) (pc … st)) [2: #e #_ whd in ⊢ (??%?→?); #ABS destruct(ABS) ] * #i_f * [2,3: [ * [ #lbl | | #fl #f #args ] | #fl #a #ltrue #lfalse ] #_ normalize nodelta normalize in ⊢ (%→?); #ABS destruct ] * [ #c | #f' #args #dest | #a #lbl | #s ] #nxt #fetch normalize nodelta normalize in ⊢ (%→?); #EQ destruct %[|%[|%[|%[|%[| %]]]]] qed. definition joint_after_ret : ∀p:evaluation_params. (Σs : state_pc p.joint_classify p s = cl_call) → state_pc p → Prop ≝ λp,s1,s2. match fetch_statement … (ev_genv p) (pc … s1) with [ OK x ⇒ match \snd x with [ sequential s next ⇒ last_pop … s2 = pc … s1 ∧ pc … s2 = succ_pc p (pc … s1) next | _ ⇒ False (* never happens *) ] | _ ⇒ False (* never happens *) ]. definition joint_call_ident : ∀p:evaluation_params. state_pc p → ident ≝ (* this is a definition without a dummy ident : λp,st. match ? return λx. !〈f, s〉 ← fetch_statement ? p … (ev_genv p) (pc … st) ; match s with [ sequential s next ⇒ match s with [ step_seq s ⇒ match s with [ CALL f' args dest ⇒ function_of_call … (ev_genv p) st f' | _ ⇒ Error … [ ] ] | _ ⇒ Error … [ ] ] | _ ⇒ Error … [ ] ] = x → ? with [ OK v ⇒ λ_.v | Error e ⇒ λABS.⊥ ] (refl …). @hide_prf elim (joint_classify_call … (pi2 … st)) #f *#f' *#args *#dest *#next *#fn *#fd ** #EQ1 #EQ2 #EQ3 lapply ABS -ABS >EQ1 >m_return_bind normalize nodelta >EQ2 #ABS destruct(ABS) qed. *) (* with a dummy ident (which is never used as seen above in the commented script) I think handling of the function is easier *) λp,st. let dummy : ident ≝ an_identifier ? one in (* used where it cannot ever happen *) match fetch_statement … (ev_genv p) (pc … st) with [ OK x ⇒ match \snd x with [ sequential s next ⇒ match s with [ step_seq s ⇒ match s with [ CALL f' args dest ⇒ match (! bl ← block_of_call … (ev_genv p) f' st; fetch_internal_function … (ev_genv p) bl) with [ OK i_f ⇒ \fst i_f | _ ⇒ dummy ] | _ ⇒ dummy ] | _ ⇒ dummy ] | _ ⇒ dummy ] | _ ⇒ dummy ]. definition joint_tailcall_ident : ∀p:evaluation_params. state_pc p → ident ≝ λp,st. let dummy : ident ≝ an_identifier ? one in (* used where it cannot ever happen *) match fetch_statement … (ev_genv p) (pc … st) with [ OK x ⇒ match \snd x with [ final s ⇒ match s with [ TAILCALL _ f' args ⇒ match (! bl ← block_of_call … (ev_genv p) f' st; fetch_internal_function … (ev_genv p) bl) with [ OK i_f ⇒ \fst i_f | _ ⇒ dummy ] | _ ⇒ dummy ] | _ ⇒ dummy ] | _ ⇒ dummy ]. definition pcDeq ≝ mk_DeqSet program_counter eq_program_counter ?. *#p1 #EQ1 * #p2 #EQ2 @eq_program_counter_elim [ #EQ destruct % #H % | * #NEQ % #ABS destruct elim (NEQ ?) % ] qed. (* let rec io_evaluate O I X (env : ∀o.res (I o)) (c : IO O I X) on c : res X ≝ match c with [ Value x ⇒ OK … x | Interact o f ⇒ ! i ← env o ; io_evaluate O I X env (f i) | Wrong e ⇒ Error … e ]. *) definition cost_label_of_stmt : ∀p : evaluation_params.joint_statement p (globals p) → option costlabel ≝ λp,s.match s with [ sequential s _ ⇒ match s with [ COST_LABEL lbl ⇒ Some ? lbl | _ ⇒ None ? ] | _ ⇒ None ? ]. definition joint_label_of_pc ≝ λp : evaluation_params. (λpc. match fetch_statement … (ev_genv p) pc with [ OK fn_stmt ⇒ cost_label_of_stmt … (\snd fn_stmt) | _ ⇒ None ? ]). definition joint_abstract_status : ∀p : prog_params. abstract_status ≝ λp. mk_abstract_status (* as_status ≝ *) (state_pc p) (* as_execute ≝ *) (λs1,s2.eval_state … (ev_genv p) s1 = return s2) (* (* as_init ≝ *) (make_initial_state p) *) (* as_pc ≝ *) pcDeq (* as_pc_of ≝ *) (pc …) (* as_classify ≝ *) (joint_classify p) (* as_label_of_pc ≝ *) (joint_label_of_pc p) (* as_after_return ≝ *) (joint_after_ret p) (* as_result ≝ *) (is_final p (globals ?) (ev_genv p) exit_pc) (* as_call_ident ≝ *) (λst.joint_call_ident p st) (* as_tailcall_ident ≝ *) (λst.joint_tailcall_ident p st). (* alternative definition with integrated prog_params constructor *) definition joint_status : ∀p : sem_params. joint_program p → (ident → option ℕ) → abstract_status ≝ λp,prog,stacksizes.joint_abstract_status (mk_prog_params p prog stacksizes).