include "joint/Joint.ma". definition fresh_reg: ∀pars0,globals. joint_internal_function … (rtl_ertl_params pars0 globals) → (joint_internal_function … (rtl_ertl_params pars0 globals)) × register ≝ λpars0,globals,def. let 〈r,runiverse〉 ≝ fresh … (joint_if_runiverse … def) in 〈set_locals ?? (set_runiverse ?? def runiverse) (r::joint_if_locals ?? def), r〉. let rec fresh_regs (pars0:?) (globals: list ident) (def: joint_internal_function … (rtl_ertl_params pars0 globals)) (n: nat) on n ≝ match n with [ O ⇒ 〈def, [ ]〉 | S n' ⇒ let 〈def', regs'〉 ≝ fresh_regs pars0 globals def n' in let 〈def', reg〉 ≝ fresh_reg … def' in 〈def', reg :: regs'〉 ]. lemma fresh_regs_length: ∀pars0,globals.∀def: joint_internal_function … (rtl_ertl_params pars0 globals). ∀n: nat. |(\snd (fresh_regs … def n))| = n. #pars0 #globals #def #n elim n [ % | #m whd in ⊢ (? → ??(??(???%))?); @pair_elim #def' #regs #EQ >EQ normalize // ] qed. definition fresh_regs_strong: ∀pars0,globals. joint_internal_function … (rtl_ertl_params pars0 globals) → ∀n: nat. Σregs: (joint_internal_function … (rtl_ertl_params pars0 globals)) × (list register). |\snd regs| = n ≝ λpars0,globals,def,n.fresh_regs pars0 globals def n. // qed. definition fresh_label: ∀pars0,globals. joint_internal_function … (graph_params pars0 globals) → label × (joint_internal_function … (graph_params pars0 globals)) ≝ λpars0,globals,def. let 〈r,luniverse〉 ≝ fresh … (joint_if_luniverse … def) in 〈r,set_luniverse … def luniverse〉. let rec add_translates (pars1: params1) (globals: list ident) (translate_list: list ?) (start_lbl: label) (dest_lbl: label) (def: joint_internal_function … (graph_params pars1 globals)) on translate_list ≝ match translate_list with [ nil ⇒ add_graph … start_lbl (GOTO … dest_lbl) def | cons trans translate_list ⇒ match translate_list with [ nil ⇒ trans start_lbl dest_lbl def | _ ⇒ let 〈tmp_lbl, def〉 ≝ fresh_label … def in let def ≝ trans start_lbl tmp_lbl def in add_translates pars1 globals translate_list tmp_lbl dest_lbl def]]. definition adds_graph ≝ λpars1:params1.λglobals. λstmt_list: list (label → joint_statement (graph_params_ pars1) globals). add_translates … (map ?? (λf,start_lbl,dest_lbl. add_graph pars1 ? start_lbl (f dest_lbl)) stmt_list).