include "utilities/lists.ma". include "common/Globalenvs.ma". include "Cminor/syntax.ma". include "RTLabs/syntax.ma". definition env ≝ identifier_map SymbolTag register. definition label_env ≝ identifier_map SymbolTag label. definition populate_env : env → universe RegisterTag → list ident → res (list register × env × (universe RegisterTag)) ≝ λen,gen. foldr ?? (λid,rsengen. do 〈rs,en,gen〉 ← rsengen; do 〈r,gen'〉 ← fresh … gen; OK ? 〈r::rs, add ?? en id r, gen'〉) (OK ? 〈[ ], en, gen〉). definition populate_label_env : label_env → universe LabelTag → list ident → res (label_env × (universe LabelTag)) ≝ λen,gen. foldr ?? (λid,engen. do 〈en,gen〉 ← engen; do 〈r,gen'〉 ← fresh … gen; OK ? 〈add ?? en id r, gen'〉) (OK ? 〈en, gen〉). (* Add a statement to the graph, *without* updating the entry label. *) definition fill_in_statement : label → statement → internal_function → internal_function ≝ λl,s,f. mk_internal_function (f_labgen f) (f_reggen f) (f_sig f) (f_result f) (f_params f) (f_locals f) (f_ptrs f) (f_stacksize f) (add ?? (f_graph f) l s) (f_entry f) (f_exit f). (* Add a statement to the graph, making it the entry label. *) definition add_to_graph : label → statement → internal_function → internal_function ≝ λl,s,f. mk_internal_function (f_labgen f) (f_reggen f) (f_sig f) (f_result f) (f_params f) (f_locals f) (f_ptrs f) (f_stacksize f) (add ?? (f_graph f) l s) l (f_exit f). (* Add a statement with a fresh label to the start of the function. The statement is parametrised by the *next* instruction's label. *) definition add_fresh_to_graph : (label → statement) → internal_function → res internal_function ≝ λs,f. do 〈l,g〉 ← fresh … (f_labgen f); let s' ≝ s (f_entry f) in OK ? (mk_internal_function g (f_reggen f) (f_sig f) (f_result f) (f_params f) (f_locals f) (f_ptrs f) (f_stacksize f) (add ?? (f_graph f) l s') l (f_exit f)). (* Generate a fresh label and use it as a dangling entry point, to be filled in later with the loop head. *) definition add_loop_label_to_graph : internal_function → res internal_function ≝ λf. do 〈l,g〉 ← fresh … (f_labgen f); OK ? (mk_internal_function g (f_reggen f) (f_sig f) (f_result f) (f_params f) (f_locals f) (f_ptrs f) (f_stacksize f) (f_graph f) l (f_exit f)). definition fresh_reg : internal_function → res (register × internal_function) ≝ λf. do 〈r,g〉 ← fresh … (f_reggen f); OK ? 〈r, mk_internal_function (f_labgen f) g (f_sig f) (f_result f) (f_params f) (r::(f_locals f)) (f_ptrs f) (f_stacksize f) (f_graph f) (f_entry f) (f_exit f)〉. definition fresh_ptr_reg : internal_function → res (register × internal_function) ≝ λf. do 〈r,g〉 ← fresh … (f_reggen f); OK ? 〈r, mk_internal_function (f_labgen f) g (f_sig f) (f_result f) (f_params f) (r::(f_locals f)) (r::(f_ptrs f)) (f_stacksize f) (f_graph f) (f_entry f) (f_exit f)〉. let rec expr_yields_pointer (e:expr) (ptrs:list ident) : bool ≝ match e with [ Id i ⇒ exists ? (eq_identifier ? i) ptrs | Cst c ⇒ match c with [ Oaddrsymbol _ _ ⇒ true | Oaddrstack _ ⇒ true | _ ⇒ false ] | Op1 op e' ⇒ match op with [ Oid ⇒ expr_yields_pointer e' ptrs | Optrofint _ ⇒ true | _ ⇒ false ] | Op2 op e1 e2 ⇒ match op with [ Oaddp ⇒ true | Osubpi ⇒ true | _ ⇒ false ] | Mem q e' ⇒ match q with [ Mpointer _ ⇒ true | _ ⇒ false ] (* Both branches ought to be the same? *) | Cond e' e1 e2 ⇒ (expr_yields_pointer e1 ptrs) ∨ (expr_yields_pointer e2 ptrs) | Ecost _ e' ⇒ expr_yields_pointer e' ptrs ]. definition choose_reg : env → expr → list ident → internal_function → res (register × internal_function) ≝ λenv,e,ptrs,f. match e with [ Id i ⇒ do r ← opt_to_res … (lookup … env i); OK ? 〈r, f〉 | _ ⇒ if expr_yields_pointer e ptrs then fresh_ptr_reg f else fresh_reg f ]. definition choose_regs : env → list expr → list ident → internal_function → res (list register × internal_function) ≝ λenv,es,ptrs,f. foldr ?? (λe,acc. do 〈rs,f〉 ← acc; do 〈r,f'〉 ← choose_reg env e ptrs f; OK ? 〈r::rs,f'〉) (OK ? 〈[ ], f〉) es. let rec add_expr (env:env) (e:expr) (dst:register) (ptrs:list ident) (f:internal_function) on e: res internal_function ≝ match e with [ Id i ⇒ do r ← opt_to_res … (lookup ?? env i); match register_eq r dst with [ inl _ ⇒ OK ? f | inr _ ⇒ add_fresh_to_graph (St_op1 Oid dst r) f ] | Cst c ⇒ add_fresh_to_graph (St_const dst c) f | Op1 op e' ⇒ do 〈r,f〉 ← choose_reg env e' ptrs f; do f ← add_fresh_to_graph (St_op1 op dst r) f; add_expr env e' r ptrs f | Op2 op e1 e2 ⇒ do 〈r1,f〉 ← choose_reg env e1 ptrs f; do 〈r2,f〉 ← choose_reg env e2 ptrs f; do f ← add_fresh_to_graph (St_op2 op dst r1 r2) f; do f ← add_expr env e2 r2 ptrs f; add_expr env e1 r1 ptrs f | Mem q e' ⇒ add_with_addressing_internal env e' (λm,rs. St_load q m rs dst) ptrs f (add_expr env e') | Cond e' e1 e2 ⇒ let resume_at ≝ f_entry f in do f ← add_expr env e2 dst ptrs f; let lfalse ≝ f_entry f in do f ← add_fresh_to_graph (λ_.St_skip resume_at) f; do f ← add_expr env e1 dst ptrs f; add_branch_internal env e' (f_entry f) lfalse ptrs f (add_expr env e') | Ecost l e' ⇒ do f ← add_expr env e' dst ptrs f; add_fresh_to_graph (St_cost l) f (* Ugh, the termination checker isn't smart enough to notice that calling add_expr with e is OK, so we take it partially applied and define a proper add_ afterwards. *) ] and add_with_addressing_internal (env:env) (e:expr) (s:∀m:addressing. addr m → label → statement) (ptrs:list ident) (f:internal_function) (termination_hack:register → list ident → internal_function → res internal_function) on e : res internal_function ≝ let add_default : unit → res internal_function ≝ λ_. do 〈r, f〉 ← choose_reg env e ptrs f; do f ← add_fresh_to_graph (s (Aindexed zero) [[ r ]]) f; termination_hack r ptrs f in match e with [ Cst c ⇒ match c with [ Oaddrsymbol id i ⇒ add_fresh_to_graph (s (Aglobal id i) [[ ]]) f | Oaddrstack i ⇒ add_fresh_to_graph (s (Ainstack i) [[ ]]) f | _ ⇒ Error ? (* int and float constants are nonsense here *) ] | Op2 op e1 e2 ⇒ match op with [ Oaddp ⇒ let add_generic_addp : unit → res internal_function ≝ λ_. do 〈r1, f〉 ← choose_reg env e1 ptrs f; do 〈r2, f〉 ← choose_reg env e2 ptrs f; do f ← add_fresh_to_graph (s Aindexed2 [[ r1 ; r2 ]]) f; do f ← add_expr env e2 r2 ptrs f; add_expr env e1 r1 ptrs f in match e1 with [ Cst c ⇒ match c with [ Oaddrsymbol id i ⇒ do 〈r, f〉 ← choose_reg env e ptrs f; do f ← add_fresh_to_graph (s (Abased id i) [[ r ]]) f; add_expr env e2 r ptrs f | _ ⇒ add_generic_addp it ] | _ ⇒ add_generic_addp it ] | _ ⇒ add_default it ] | _ ⇒ add_default it ] (* and again *) and add_branch_internal (env:env) (e:expr) (ltrue:label) (lfalse:label) (ptrs:list ident) (f:internal_function) (termination_hack_add_expr : register → list ident → internal_function → res internal_function) on e : res internal_function ≝ match e with [ Id i ⇒ do r ← opt_to_res … (lookup ?? env i); add_fresh_to_graph (λ_. St_cond1 Oid r ltrue lfalse) f | Cst c ⇒ add_fresh_to_graph (λ_. St_condcst c ltrue lfalse) f | Op1 op e' ⇒ do 〈r,f〉 ← choose_reg env e' ptrs f; do f ← add_fresh_to_graph (λ_. St_cond1 op r ltrue lfalse) f; add_expr env e' r ptrs f | Op2 op e1 e2 ⇒ do 〈r1,f〉 ← choose_reg env e1 ptrs f; do 〈r2,f〉 ← choose_reg env e2 ptrs f; do f ← add_fresh_to_graph (λ_. St_cond2 op r1 r2 ltrue lfalse) f; do f ← add_expr env e2 r2 ptrs f; add_expr env e1 r1 ptrs f | _ ⇒ do 〈r,f〉 ← choose_reg env e ptrs f; do f ← add_fresh_to_graph (λ_. St_cond1 Oid r ltrue lfalse) f; termination_hack_add_expr r ptrs f ]. (* See comment above. *) definition add_with_addressing ≝ λenv,e,s,ptrs,f. add_with_addressing_internal env e s ptrs f (add_expr env e). definition add_branch ≝ λenv,e,ltrue,lfalse,ptrs,f. add_branch_internal env e ltrue lfalse ptrs f (add_expr env e). let rec add_exprs (env:env) (es:list expr) (dsts:list register) (ptrs:list ident) (f:internal_function) on es: res internal_function ≝ match es with [ nil ⇒ match dsts with [ nil ⇒ OK ? f | cons _ _ ⇒ Error ? ] | cons e et ⇒ match dsts with [ nil ⇒ Error ? | cons r rt ⇒ do f ← add_exprs env et rt ptrs f; add_expr env e r ptrs f ] ]. let rec add_stmt (env:env) (label_env:label_env) (s:stmt) (exits:list label) (ptrs:list ident) (f:internal_function) on s : res internal_function ≝ match s with [ St_skip ⇒ OK ? f | St_assign x e ⇒ do dst ← opt_to_res … (lookup ?? env x); add_expr env e dst ptrs f | St_store q e1 e2 ⇒ do 〈val_reg, f〉 ← choose_reg env e2 ptrs f; do f ← add_with_addressing env e1 (λm,rs. St_store q m rs val_reg) ptrs f; add_expr env e2 val_reg ptrs f | St_call return_opt_id e args sig ⇒ do return_opt_reg ← match return_opt_id with [ None ⇒ OK ? (None ?) | Some id ⇒ do r ← opt_to_res … (lookup ?? env id); OK ? (Some ? r) ]; do 〈args_regs, f〉 ← choose_regs env args ptrs f; do f ← match e with [ Id id ⇒ add_fresh_to_graph (St_call_id id args_regs return_opt_reg sig) f | _ ⇒ do 〈fnr, f〉 ← choose_reg env e ptrs f; do f ← add_fresh_to_graph (St_call_ptr fnr args_regs return_opt_reg sig) f; add_expr env e fnr ptrs f ]; add_exprs env args args_regs ptrs f | St_tailcall e args sig ⇒ do 〈args_regs, f〉 ← choose_regs env args ptrs f; do f ← match e with [ Id id ⇒ add_fresh_to_graph (λ_. St_tailcall_id id args_regs sig) f | _ ⇒ do 〈fnr, f〉 ← choose_reg env e ptrs f; do f ← add_fresh_to_graph (λ_. St_tailcall_ptr fnr args_regs sig) f; add_expr env e fnr ptrs f ]; add_exprs env args args_regs ptrs f | St_seq s1 s2 ⇒ do f ← add_stmt env label_env s2 exits ptrs f; add_stmt env label_env s1 exits ptrs f | St_ifthenelse e s1 s2 ⇒ let l_next ≝ f_entry f in do f ← add_stmt env label_env s2 exits ptrs f; let l2 ≝ f_entry f in do f ← add_fresh_to_graph ? (* XXX: fails, but works if applied afterwards: λ_. St_skip l_next*) f; do f ← add_stmt env label_env s1 exits ptrs f; add_branch env e (f_entry f) l2 ptrs f | St_loop s ⇒ do f ← add_loop_label_to_graph f; let l_loop ≝ f_entry f in do f ← add_stmt env label_env s exits ptrs f; OK ? (fill_in_statement l_loop (* XXX another odd failure: St_skip (f_entry f)*)? f) | St_block s ⇒ add_stmt env label_env s ((f_entry f)::exits) ptrs f | St_exit n ⇒ do l ← opt_to_res … (nth_opt ? n exits); add_fresh_to_graph (* XXX another: λ_. St_skip l*)? f | St_switch e tab n ⇒ do 〈r,f〉 ← choose_reg env e ptrs f; do l_default ← opt_to_res … (nth_opt ? n exits); do f ← add_fresh_to_graph (* XXX grrrr: λ_. St_skip l_default*)? f; do f ← foldr ?? (λcs,f. do f ← f; let 〈i,n〉 ≝ cs in do 〈cr,f〉 ← fresh_reg … f; do l_case ← opt_to_res … (nth_opt ? n exits); do f ← add_fresh_to_graph (St_cond2 (Ocmpu Ceq) (* signed? *) r cr l_case) f; add_fresh_to_graph (St_const cr (Ointconst i)) f) (OK ? f) tab; add_expr env e r ptrs f | St_return opt_e ⇒ do f ← add_fresh_to_graph (λ_. St_return) f; match opt_e with [ None ⇒ OK ? f | Some e ⇒ match f_result f with [ None ⇒ Error ? | Some r ⇒ add_expr env e r ptrs f ] ] | St_label l s' ⇒ do f ← add_stmt env label_env s' exits ptrs f; do l' ← opt_to_res … (lookup ?? label_env l); OK ? (add_to_graph l' (* XXX again: St_skip (f_entry f)*)? f) | St_goto l ⇒ do l' ← opt_to_res … (lookup ?? label_env l); add_fresh_to_graph (* XXX again: λ_.St_skip l'*)? f | St_cost l s' ⇒ do f ← add_stmt env label_env s' exits ptrs f; add_fresh_to_graph (St_cost l) f | _ ⇒ Error ? (* XXX implement *) ]. [ @(λ_. St_skip l_next) | @(St_skip (f_entry f)) | @(λ_. St_skip l) | @(λ_. St_skip l_default) | @(St_skip (f_entry f)) | @(λ_.St_skip l') ] qed. (* Get labels from a Cminor statement. *) let rec labels_of (s:stmt) : list ident ≝ match s with [ St_seq s1 s2 ⇒ (labels_of s1) @ (labels_of s2) | St_ifthenelse _ s1 s2 ⇒ (labels_of s1) @ (labels_of s2) | St_loop s ⇒ labels_of s | St_block s ⇒ labels_of s | St_label l s ⇒ l::(labels_of s) | St_cost _ s ⇒ labels_of s | _ ⇒ [ ] ]. definition c2ra_function (*: internal_function → internal_function*) ≝ λf. let labgen0 ≝ new_universe LabelTag in let reggen0 ≝ new_universe RegisterTag in let cminor_labels ≝ labels_of (f_body f) in do 〈params, env1, reggen1〉 ← populate_env (empty_map …) reggen0 (f_params f); do 〈locals0, env, reggen2〉 ← populate_env env1 reggen1 (f_vars f); do 〈result, locals, reggen〉 ← match sig_res (f_sig f) with [ None ⇒ OK ? 〈None ?, locals0, reggen2〉 | Some _ ⇒ do 〈r,gen〉 ← fresh … reggen2; OK ? 〈Some ? r, r::locals0, gen〉 ]; do ptrs ← mmap ?? (λid. opt_to_res … (lookup ?? env id)) (f_ptrs f); do 〈label_env, labgen1〉 ← populate_label_env (empty_map …) labgen0 cminor_labels; do 〈l,labgen〉 ← fresh … labgen1; let emptyfn ≝ mk_internal_function labgen reggen (f_sig f) result params locals ptrs (f_stacksize f) (add ?? (empty_map …) l St_return) l l in add_stmt env label_env (f_body f) [ ] (f_ptrs f) emptyfn . definition cminor_to_rtlabs : Cminor_program → res RTLabs_program ≝ transform_partial_program ??? (transf_partial_fundef ?? c2ra_function). include "Cminor/initialisation.ma". definition cminor_init_to_rtlabs : Cminor_program → res RTLabs_program ≝ λp. let p' ≝ replace_init p in cminor_to_rtlabs p.