include "basics/lists/list.ma". include "common/Errors.ma". include "common/Globalenvs.ma". include "common/IO.ma". include "common/SmallstepExec.ma". include "RTLabs/syntax.ma". definition genv ≝ genv_t (fundef internal_function). record frame : Type[0] ≝ { func : internal_function ; locals : register_env val ; next : label ; next_ok: present ?? (f_graph func) next ; sp : block ; retdst : option register }. definition adv : ∀l:label. ∀f:frame. present ?? (f_graph (func f)) l → frame ≝ λl,f,p. mk_frame (func f) (locals f) l p (sp f) (retdst f). inductive state : Type[0] ≝ | State : ∀ f : frame. ∀ fs : list frame. ∀ m : mem. state | Callstate : ∀ fd : fundef internal_function. ∀args : list val. ∀ dst : option register. ∀ stk : list frame. ∀ m : mem. state | Returnstate : ∀ rtv : option val. ∀ dst : option register. ∀ stk : list frame. ∀ m : mem. state | Finalstate : ∀ r : int. state . definition build_state ≝ λf.λfs.λm.λn.λp. State (adv n f p) fs m. definition init_locals : list (register × typ) → register_env val ≝ foldr ?? (λidt,en. let 〈id,ty〉 ≝ idt in add RegisterTag val en id Vundef) (empty_map ??). definition reg_store ≝ λreg,v,locals. update RegisterTag val locals reg v. axiom WrongNumberOfParameters : String. let rec params_store (rs:list (register × typ)) (vs:list val) (locals : register_env val) : res (register_env val) ≝ match rs with [ nil ⇒ match vs with [ nil ⇒ OK ? locals | _ ⇒ Error ? (msg WrongNumberOfParameters)] | cons rt rst ⇒ match vs with [ nil ⇒ Error ? (msg WrongNumberOfParameters) | cons v vst ⇒ let 〈r,ty〉 ≝ rt in let locals' ≝ add RegisterTag val locals r v in params_store rst vst locals' ] ]. axiom BadRegister : String. definition reg_retrieve : register_env ? → register → res val ≝ λlocals,reg. opt_to_res … [MSG BadRegister; CTX ? reg] (lookup ?? locals reg). axiom FailedOp : String. axiom MissingSymbol : String. axiom MissingStatement : String. axiom FailedConstant : String. axiom FailedLoad : String. axiom FailedStore : String. axiom BadFunction : String. axiom BadJumpTable : String. axiom BadJumpValue : String. axiom FinalState : String. axiom ReturnMismatch : String. definition eval_statement : genv → state → IO io_out io_in (trace × state) ≝ λge,st. match st return λ_. IO ??? with [ State f fs m ⇒ let s ≝ lookup_present ?? (f_graph (func f)) (next f) (next_ok f) in match s return λs. labels_present ? s → IO ??? with [ St_skip l ⇒ λH. return 〈E0, build_state f fs m l ?〉 | St_cost cl l ⇒ λH. return 〈Echarge cl, build_state f fs m l ?〉 | St_const _ r cst l ⇒ λH. ! v ← opt_to_res … (msg FailedConstant) (eval_constant ? (find_symbol … ge) (sp f) cst); ! locals ← reg_store r v (locals f); return 〈E0, State (mk_frame (func f) locals l ? (sp f) (retdst f)) fs m〉 | St_op1 _ _ op dst src l ⇒ λH. ! v ← reg_retrieve (locals f) src; ! v' ← opt_to_res … (msg FailedOp) (eval_unop ?? op v); ! locals ← reg_store dst v' (locals f); return 〈E0, State (mk_frame (func f) locals l ? (sp f) (retdst f)) fs m〉 | St_op2 _ _ _ op dst src1 src2 l ⇒ λH. ! v1 ← reg_retrieve (locals f) src1; ! v2 ← reg_retrieve (locals f) src2; ! v' ← opt_to_res … (msg FailedOp) (eval_binop ??? op v1 v2); ! locals ← reg_store dst v' (locals f); return 〈E0, State (mk_frame (func f) locals l ? (sp f) (retdst f)) fs m〉 | St_load chunk addr dst l ⇒ λH. ! vaddr ← reg_retrieve (locals f) addr; ! v ← opt_to_res … (msg FailedLoad) (loadv chunk m vaddr); ! locals ← reg_store dst v (locals f); return 〈E0, State (mk_frame (func f) locals l ? (sp f) (retdst f)) fs m〉 | St_store chunk addr src l ⇒ λH. ! vaddr ← reg_retrieve (locals f) addr; ! v ← reg_retrieve (locals f) src; ! m' ← opt_to_res … (msg FailedStore) (storev chunk m vaddr v); return 〈E0, build_state f fs m' l ?〉 | St_call_id id args dst l ⇒ λH. ! b ← opt_to_res … [MSG MissingSymbol; CTX ? id] (find_symbol … ge id); ! fd ← opt_to_res … [MSG BadFunction; CTX ? id] (find_funct_ptr … ge b); ! vs ← m_list_map … (reg_retrieve (locals f)) args; return 〈E0, Callstate fd vs dst (adv l f ?::fs) m〉 | St_call_ptr frs args dst l ⇒ λH. ! fv ← reg_retrieve (locals f) frs; ! fd ← opt_to_res … (msg BadFunction) (find_funct … ge fv); ! vs ← m_list_map … (reg_retrieve (locals f)) args; return 〈E0, Callstate fd vs dst (adv l f ?::fs) m〉 (* | St_tailcall_id id args ⇒ λH. ! b ← opt_to_res … [MSG MissingSymbol; CTX ? id] (find_symbol … ge id); ! fd ← opt_to_res … [MSG BadFunction; CTX ? id] (find_funct_ptr … ge b); ! vs ← m_list_map … (reg_retrieve (locals f)) args; return 〈E0, Callstate fd vs (retdst f) fs (free m (sp f))〉 | St_tailcall_ptr frs args ⇒ λH. ! fv ← reg_retrieve (locals f) frs; ! fd ← opt_to_res … (msg BadFunction) (find_funct … ge fv); ! vs ← m_list_map … (reg_retrieve (locals f)) args; return 〈E0, Callstate fd vs (retdst f) fs (free m (sp f))〉 *) | St_cond src ltrue lfalse ⇒ λH. ! v ← reg_retrieve (locals f) src; ! b ← eval_bool_of_val v; return 〈E0, build_state f fs m (if b then ltrue else lfalse) ?〉 | St_jumptable r ls ⇒ λH. ! v ← reg_retrieve (locals f) r; match v with [ Vint _ i ⇒ match nth_opt ? (nat_of_bitvector ? i) ls return λx. nth_opt ??? = x → ? with [ None ⇒ λ_. Error ? (msg BadJumpTable) | Some l ⇒ λE. return 〈E0, build_state f fs m l ?〉 ] (refl ??) | _ ⇒ Error ? (msg BadJumpValue) ] | St_return ⇒ λH. ! v ← match f_result (func f) with [ None ⇒ return (None ?) | Some rt ⇒ let 〈r,ty〉 ≝ rt in ! v ← reg_retrieve (locals f) r; return (Some ? v) ]; return 〈E0, Returnstate v (retdst f) fs (free m (sp f))〉 ] (f_closed (func f) (next f) s ?) | Callstate fd params dst fs m ⇒ match fd return λ_. IO ??? with [ Internal fn ⇒ ! locals ← params_store (f_params fn) params (init_locals (f_locals fn)); (* CSC: XXX alignment call (concrete_stacksize in OCaml) is missing here *) let 〈m', sp〉 ≝ alloc m 0 (f_stacksize fn) XData in return 〈E0, State (mk_frame fn locals (f_entry fn) ? sp dst) fs m'〉 | External fn ⇒ ! evargs ← err_to_io … (check_eventval_list params (sig_args (ef_sig fn))); ! evres ← do_io (ef_id fn) evargs (proj_sig_res (ef_sig fn)); return 〈Eextcall (ef_id fn) evargs (mk_eventval ? evres), Returnstate (Some ? (mk_val ? evres) (*FIXME should allow None *)) dst fs m〉 ] | Returnstate v dst fs m ⇒ match fs with [ nil ⇒ match v with [ Some v' ⇒ match v' with [ Vint sz r ⇒ match sz return λsz. bvint sz → monad Res ? with [ I32 ⇒ λr. return 〈E0, Finalstate r〉 | _ ⇒ λ_. Error ? (msg ReturnMismatch) ] r | _ ⇒ Error ? (msg ReturnMismatch) ] | _ ⇒ Error ? (msg ReturnMismatch) ] | cons f fs' ⇒ ! locals ← match dst with [ None ⇒ match v with [ None ⇒ OK ? (locals f) | _ ⇒ Error ? (msg ReturnMismatch) ] | Some d ⇒ match v with [ None ⇒ Error ? (msg ReturnMismatch) | Some v' ⇒ reg_store d v' (locals f) ] ]; return 〈E0, State (mk_frame (func f) locals (next f) ? (sp f) (retdst f)) fs' m〉 ] | Finalstate r ⇒ Error ? (msg FinalState) (* Already in final state *) ]. try @(next_ok f) try @lookup_lookup_present try @H [ cases b [ @(proj1 ?? H) | @(proj2 ?? H) ] | whd in H; @(All_nth … H ? E) | cases (f_entry fn) // ] qed. definition RTLabs_is_final : state → option int ≝ λs. match s with [ Finalstate r ⇒ Some ? r | _ ⇒ None ? ]. definition RTLabs_exec : trans_system io_out io_in ≝ mk_trans_system … ? (λ_.RTLabs_is_final) eval_statement. definition make_global : RTLabs_program → genv ≝ λp. globalenv … (λx.[Init_space x]) p. definition make_initial_state : RTLabs_program → res state ≝ λp. let ge ≝ make_global p in do m ← init_mem … (λx.[Init_space x]) p; let main ≝ prog_main ?? p in do b ← opt_to_res ? [MSG MissingSymbol; CTX ? main] (find_symbol … ge main); do f ← opt_to_res ? [MSG BadFunction; CTX ? main] (find_funct_ptr … ge b); OK ? (Callstate f (nil ?) (None ?) (nil ?) m). definition RTLabs_fullexec : fullexec io_out io_in ≝ mk_fullexec … RTLabs_exec make_global make_initial_state. (* Some lemmas about the semantics. *) (* Note parts of frames and states that are preserved. *) inductive frame_rel : frame → frame → Prop ≝ | mk_frame_rel : ∀func,locals,next,next_ok,sp,retdst,locals',next',next_ok'. frame_rel (mk_frame func locals next next_ok sp retdst) (mk_frame func locals' next' next_ok' sp retdst) . inductive state_rel : genv → state → state → Prop ≝ | normal : ∀ge,f,f',fs,m,m'. frame_rel f f' → state_rel ge (State f fs m) (State f' fs m') | to_call : ∀ge,f,fs,m,fd,args,f',dst. frame_rel f f' → ∀b. find_funct_ptr ? ge b = Some ? fd → state_rel ge (State f fs m) (Callstate fd args dst (f'::fs) m) (* | tail_call : ∀ge,f,fs,m,fd,args,dst,m'. ∀b. find_funct_ptr ?? ge b = Some ? fd → state_rel ge (State f fs m) (Callstate fd args dst fs m') *) | from_call : ∀ge,fn,locals,next,nok,sp,fs,m,args,dst,m'. state_rel ge (Callstate (Internal ? fn) args dst fs m) (State (mk_frame fn locals next nok sp dst) fs m') | to_ret : ∀ge,f,fs,m,rtv,dst,m'. state_rel ge (State f fs m) (Returnstate rtv dst fs m') | from_ret : ∀ge,f,fs,rtv,dst,f',m. frame_rel f f' → state_rel ge (Returnstate rtv dst (f::fs) m) (State f' fs m) | finish : ∀ge,r,dst,m. state_rel ge (Returnstate (Some ? (Vint I32 r)) dst [ ] m) (Finalstate r) . lemma bind_ok : ∀A,B,e,f,v. ∀P:B → Type[0]. (∀a. e = OK A a → f a = OK ? v → P v) → (match e with [ OK v ⇒ f v | Error m ⇒ Error ? m ] = OK ? v → P v). #A #B * [ #a #f #v #P #H #E whd in E:(??%?); @(H a) // | #m #f #v #P #H #E whd in E:(??%?); destruct ] qed. lemma eval_preserves : ∀ge,s,tr,s'. eval_statement ge s = Value ??? 〈tr,s'〉 → state_rel ge s s'. #ge * [ * #func #locals #next #next_ok #sp #retdst #fs #m #tr #s' whd in ⊢ (??%? → ?); generalize in ⊢ (??(?%)? → ?); cases (lookup_present LabelTag statement (f_graph func) next next_ok) [ #l #LP whd in ⊢ (??%? → ?); #E destruct % % | #c #l #LP whd in ⊢ (??%? → ?); #E destruct % % | #t #r #c #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #locals' #Eloc #E whd in E:(??%?); destruct % % | #t1 #t2 #op #r1 #r2 #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #v' #Ev' @bind_ok #loc #Eloc #E whd in E:(??%?); destruct % % | #t1 #t2 #t' #op #r1 #r2 #r3 #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v1 #Ev1 @bind_ok #v2 #Ev2 @bind_ok #v' #Ev' @bind_ok #loc #Eloc #E whd in E:(??%?); destruct % % | #ch #r1 #r2 #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #v' #Ev' @bind_ok #loc #Eloc #E whd in E:(??%?); destruct % % | #ch #r1 #r2 #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #v' #Ev' @bind_ok #loc #Eloc #E whd in E:(??%?); destruct % % | #id #rs #or #l #LP whd in ⊢ (??%? → ?); @bind_res_value #b #Eb @bind_ok #fd #Efd @bind_ok #vs #Evs #E whd in E:(??%?); destruct %2 [ % | @b | cases (find_funct_ptr ???) in Efd ⊢ %; normalize [2:#fd'] #E' destruct @refl ] | #r #rs #or #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #fd #Efd @bind_ok #vs #Evs #E whd in E:(??%?); destruct cases (find_funct_find_funct_ptr ?? v ??) [ #b * #Ev' #Efn %2 [ % | @b | @Efn ] | ||| cases (find_funct ???) in Efd ⊢ %; [2:#x] normalize #E' destruct @refl ] (* | #id #rs #LP whd in ⊢ (??%? → ?); @bind_res_value #b #Eb @bind_ok #fd #Efd @bind_ok #vs #Evs #E whd in E:(??%?); destruct %3 [ @b | cases (find_funct_ptr ????) in Efd ⊢ %; [2:#x] normalize #E' destruct @refl ] | #r #rs #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #fd #Efd @bind_ok #vs #Evs #E whd in E:(??%?); destruct cases (find_funct_find_funct_ptr ?? v ??) [ #r * #b * #c * #Ev' #Efn %3 [ @b | @Efn ] | ||| cases (find_funct ????) in Efd ⊢ %; [2:#x] normalize #E' destruct @refl ] *) | #r #l1 #l2 #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #b #Eb #E whd in E:(??%?); destruct % % | #r #ls #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev cases v [ #E whd in E:(??%?); destruct | #sz #i whd in ⊢ (??%? → ?); generalize in ⊢ (??(?%)? → ?); cases (nth_opt ?? ls) in ⊢ (∀e:???%. ??(match % with [_ ⇒ ?|_ ⇒ ?]?)? → ?); [ #e #E whd in E:(??%?); destruct | #l' #e #E whd in E:(??%?); destruct % % ] | *: [ 1,3: #vl ] #E whd in E:(??%?); destruct ] | #LP whd in ⊢ (??%? → ?); @bind_res_value #v cases (f_result func) [ 2: * #r #t whd in ⊢ (??%? → ?); @bind_ok #v0 #Ev0 ] #E whd in E:(??%?); #E' whd in E':(??%?); destruct %4 ] | * #fn #args #retdst #stk #m #tr #s' whd in ⊢ (??%? → ?); [ @bind_res_value #loc #Eloc cases (alloc m O ??) #m' #b whd in ⊢ (??%? → ?); #E destruct %3 | @bindIO_value #evargs #Eargs whd in ⊢ (??%? → ?); #E destruct ] | #v #r * [2: #f #fs ] #m #tr #s' whd in ⊢ (??%? → ?); [ @bind_res_value #loc #Eloc #E whd in E:(??%?); destruct %5 cases f #func #locals #next #next_ok #sp #retdst % | cases v [ normalize #E destruct | * [ 2: * #r normalize #E destruct // | *: normalize #a try #b destruct ] ] ] | #r #tr #s' normalize #E destruct ] qed. (* Show that, in principle at least, we can calculate which function we called on a transition. The proof is a functional inversion similar to eval_preserves, above. *) definition func_block_of_exec : ∀ge,s,t,fd,args,dst,fs,m. eval_statement ge s = Value ??? 〈t,Callstate fd args dst fs m〉 → Σb:block. find_funct_ptr … ge b = Some ? fd. #ge * [ * #func #locals #next #nok #sp #dst #fs #m #tr #fd #args #dst' #fs' #m' whd in ⊢ (??%? → ?); generalize in ⊢ (??(?%)? → ?); cases (lookup_present … nok) (* Function call cases. *) [ 8,9: #x #rs #or #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #fd' #Efd @bind_ok #vs #Evs #E normalize in E; destruct [ %{v} @Efd | cases v in Efd; [ 5: * #b #off #Efd %{b} whd in Efd:(??(???%)?); cases (eq_offset ??) in Efd; [ #E @E | #E normalize in E; destruct ] | *: normalize #A try #B try #C destruct ] ] (* and everything else cannot occur, but we need to work through the definition to find the resulting state to demonstrate that. *) | #l #LP whd in ⊢ (??%? → ?); #E destruct | #c #l #LP whd in ⊢ (??%? → ?); #E destruct | #t #r #c #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #locals' #Eloc #E whd in E:(??%?); destruct | #t1 #t2 #op #r1 #r2 #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #v' #Ev' @bind_ok #loc #Eloc #E whd in E:(??%?); destruct | #t1 #t2 #t' #op #r1 #r2 #r3 #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v1 #Ev1 @bind_ok #v2 #Ev2 @bind_ok #v' #Ev' @bind_ok #loc #Eloc #E whd in E:(??%?); destruct | #ch #r1 #r2 #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #v' #Ev' @bind_ok #loc #Eloc #E whd in E:(??%?); destruct | #ch #r1 #r2 #l #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #v' #Ev' @bind_ok #loc #Eloc #E whd in E:(??%?); destruct | #r #l1 #l2 #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev @bind_ok #b #Eb #E whd in E:(??%?); destruct | #r #ls #LP whd in ⊢ (??%? → ?); @bind_res_value #v #Ev cases v [ #E whd in E:(??%?); destruct | #sz #i whd in ⊢ (??%? → ?); generalize in ⊢ (??(?%)? → ?); cases (nth_opt ?? ls) in ⊢ (∀e:???%. ??(match % with [_ ⇒ ?|_ ⇒ ?]?)? → ?); [ #e #E whd in E:(??%?); destruct | #l' #e #E whd in E:(??%?); destruct ] | *: [1,3: #vl] #E whd in E:(??%?); destruct ] | #LP whd in ⊢ (??%? → ?); @bind_res_value #v cases (f_result func) [ 2: * #r #t whd in ⊢ (??%? → ?); @bind_ok #v0 #Ev0 ] #E whd in E:(??%?); #E' whd in E':(??%?); destruct ] | * #fn #args #retdst #stk #m #tr #fd' #args' #dst' #fs' #m' whd in ⊢ (??%? → ?); [ @bind_res_value #loc #Eloc cases (alloc m O ??) #m' #b whd in ⊢ (??%? → ?); #E destruct | @bindIO_value #evargs #Eargs whd in ⊢ (??%? → ?); #E destruct ] | #v #r * [2: #f #fs ] #m #tr #fd' #args' #dst' #fs' #m' whd in ⊢ (??%? → ?); [ @bind_res_value #loc #Eloc #E whd in E:(??%?); destruct | cases v [ normalize #E destruct | * [ 2: * #r normalize #E destruct | *: normalize #a try #b destruct ] ] ] | #r #tr #fd' #args' #dst' #fs' #m' normalize #E destruct ] qed. lemma initial_state_is_not_final : ∀p,s. make_initial_state p = OK ? s → RTLabs_is_final s = None ?. #p #s whd in ⊢ (??%? → ?); @bind_ok #m #Em @bind_ok #b #Eb @bind_ok #f #Ef #E destruct @refl qed.