(* Replace initialisation of global variables with equivalent Cminor code. *) include "Cminor/syntax.ma". include "common/Globalenvs.ma". (* XXX having to specify the return type in the match is annoying. *) definition init_expr : init_data → option (𝚺t:typ. expr t) ≝ λinit. match init return λ_.option (𝚺t:typ. expr t) with [ Init_int8 i ⇒ Some ? (mk_DPair ?? (ASTint I8 Unsigned) (Cst ? (Ointconst I8 Unsigned i))) | Init_int16 i ⇒ Some ? (mk_DPair ?? (ASTint I16 Unsigned) (Cst ? (Ointconst I16 Unsigned i))) | Init_int32 i ⇒ Some ? (mk_DPair ?? (ASTint I32 Unsigned) (Cst ? (Ointconst I32 Unsigned i))) | Init_float32 f ⇒ Some ? (mk_DPair ?? (ASTfloat F32) (Cst ? (Ofloatconst F32 f))) | Init_float64 f ⇒ Some ? (mk_DPair ?? (ASTfloat F64) (Cst ? (Ofloatconst F64 f))) | Init_space n ⇒ None ? | Init_null ⇒ Some ? (mk_DPair ?? ASTptr (Op1 (ASTint I8 Unsigned) ? (Optrofint ??) (Cst ? (Ointconst I8 Unsigned (zero ?))))) | Init_addrof id off ⇒ Some ? (mk_DPair ?? ASTptr (Cst ? (Oaddrsymbol id off))) ]. (* None of the additional code introduces locals, labels or returns. *) definition stmt_inv : stmt → Prop ≝ stmt_P (λs. stmt_vars (λ_.λ_.False) s ∧ stmt_labels (λ_.False) s ∧ match s with [ St_return _ ⇒ False | _ ⇒ True ]). discriminator option. discriminator DPair. (* Produces a few extra skips - hopefully the compiler will optimise these away. *) definition init_datum : ident → region → init_data → nat (*offset*) → Σs:stmt. stmt_inv s ≝ λid,r,init,off. match init_expr init with [ None ⇒ St_skip | Some x ⇒ match x with [ mk_DPair t e ⇒ St_store t (Cst ? (Oaddrsymbol id off)) e ] ]. % /3/ cases init in p; [ 7: | 8: #a #b | *: #a ] #p normalize in p; destruct;/4 by stmt_labels_mp, conj/ qed. definition init_var : ident → region → list init_data → Σs:stmt. stmt_inv s ≝ λid,r,init. \snd (foldr ?? (λdatum,os. let 〈off,s〉 ≝ os in 〈off + (size_init_data datum), St_seq (init_datum id r datum off) s〉) 〈0, mk_Sig ? (λx.stmt_inv x) St_skip ?〉 init). [ % /3/ | elim init [ % /3/ | #h #t #IH whd in ⊢ (?(???%)); @breakup_pair % [ /3/ | % [ @(pi2 … (init_datum …)) | @IH ] ] ] qed. definition init_vars : list (ident × region × (list init_data)) → Σs:stmt. stmt_inv s ≝ λvars. foldr ? (Σs:stmt. stmt_inv s) (λvar,s. mk_Sig ? (λx.stmt_inv x ) (St_seq s (init_var (\fst (\fst var)) (\snd (\fst var)) (\snd var))) ?) (mk_Sig ? (λx.stmt_inv x ) St_skip ?) vars. [ % [ /3/ | % [ @(pi2 … s) | @(pi2 … (init_var …)) ] ] | /4/ ] qed. lemma no_labels : ∀s. stmt_inv s → labels_of s = [ ]. #s elim s // [ #s1 #s2 #IH1 #IH2 * #_ * #H1 #H2 whd in ⊢ (??%?); >(IH1 H1) >(IH2 H2) @refl | #sz #sg #e #s1 #s2 #IH1 #IH2 * #_ * #H1 #H2 whd in ⊢ (??%?); >(IH1 H1) >(IH2 H2) @refl | #l #s #IH * * * #_ * | #l #s #IH * #_ #H @(IH H) ] qed. (* FIXME: we need to add a new initialisation function instead of augmenting main, because this currently breaks a recursive main function. *) definition add_statement : costlabel → ident → (Σs:stmt. stmt_inv s) → list (ident × (fundef internal_function)) → list (ident × (fundef internal_function)) ≝ λcost,id,s. match s with [ mk_Sig s H ⇒ map ?? (λidf. let 〈id',f'〉 ≝ idf in match ident_eq id id' with [ inl _ ⇒ match f' with [ Internal f ⇒ 〈id, Internal ? (mk_internal_function (f_return f) (f_params f) (f_vars f) (f_distinct f) (f_stacksize f) (St_cost cost (St_seq s (f_body f))) ?)〉 | External f ⇒ 〈id, External ? f〉 (* Error ? *) ] | inr _ ⇒ 〈id',f'〉 ]) ]. % [ % // | % [ % // | % [ @(stmt_P_mp … H) #s * * #V #L #R % [ @(stmt_vars_mp … V) #i #t * | @(stmt_labels_mp … L) #l * | cases s in R ⊢ %; // #x * ] | whd in match (labels_of ?); >(no_labels … H) @(f_inv f) ] ] ] qed. definition empty_vars : list (ident × region × (list init_data)) → list (ident × region × nat) ≝ map ?? (λv. 〈〈\fst (\fst v), \snd (\fst v)〉, size_init_data_list (\snd v)〉). definition replace_init : costlabel → Cminor_program → Cminor_noinit_program ≝ λcost,p. mk_program ?? (empty_vars (prog_vars ?? p)) (add_statement cost (prog_main ?? p) (init_vars (prog_vars ?? p)) (prog_funct ?? p)) (prog_main ?? p).