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 ⊢ (? → ??(??%)?) whd in ⊢ (? → ??(??match % with [ _ ⇒ ?])?) cases (fresh_regs pars0 globals def m) normalize nodelta #def' #regs #EQ change in EQ with (|regs| = m)