include "ASM/ASM.ma". include "ASM/Arithmetic.ma". include "ASM/Fetch.ma". include "ASM/Status.ma". include alias "basics/logic.ma". include alias "arithmetics/nat.ma". include "utilities/extralib.ma". include "ASM/Assembly.ma". (* Internal types *) (* label_map: identifier ↦ 〈instruction number, address〉 *) definition label_map ≝ identifier_map ASMTag (ℕ × ℕ). (* jump_expansion_policy: instruction number ↦ 〈pc, [addr, jump_length]〉 *) definition jump_expansion_policy ≝ ℕ × (BitVectorTrie (ℕ × ℕ × jump_length) 16). (* The different properties that we want/need to prove at some point *) (* Anything that's not in the program doesn't end up in the policy *) definition out_of_program_none: list labelled_instruction → jump_expansion_policy → Prop ≝ λprefix.λpolicy. ∀i.i ≥ |prefix| → i < 2^16 → lookup_opt … (bitvector_of_nat ? i) (\snd policy) = None ?. (* If instruction i is a jump, then there will be something in the policy at * position i *) definition is_jump' ≝ λx:preinstruction Identifier. match x with [ JC _ ⇒ True | JNC _ ⇒ True | JZ _ ⇒ True | JNZ _ ⇒ True | JB _ _ ⇒ True | JNB _ _ ⇒ True | JBC _ _ ⇒ True | CJNE _ _ ⇒ True | DJNZ _ _ ⇒ True | _ ⇒ False ]. definition is_jump ≝ λx:labelled_instruction. let 〈label,instr〉 ≝ x in match instr with [ Instruction i ⇒ is_jump' i | Call _ ⇒ True | Jmp _ ⇒ True | _ ⇒ False ]. definition jump_in_policy: list labelled_instruction → jump_expansion_policy → Prop ≝ λprefix.λpolicy. ∀i:ℕ.i < |prefix| → iff (is_jump (nth i ? prefix 〈None ?, Comment []〉)) (∃p:ℕ.∃a:ℕ.∃j:jump_length. lookup_opt … (bitvector_of_nat 16 i) (\snd policy) = Some ? 〈p,a,j〉). definition labels_okay: label_map → jump_expansion_policy → Prop ≝ λlabels.λpolicy. bvt_forall ?? (\snd policy) (λn.λx. let 〈pc,addr_nat,j〉 ≝ x in ∃id:Identifier.\snd (lookup_def … labels id 〈0,pc〉) = addr_nat ). (* Between two policies, jumps cannot decrease *) definition jmple: jump_length → jump_length → Prop ≝ λj1.λj2. match j1 with [ short_jump ⇒ match j2 with [ short_jump ⇒ False | _ ⇒ True ] | medium_jump ⇒ match j2 with [ long_jump ⇒ True | _ ⇒ False ] | long_jump ⇒ False ]. definition jmpleq: jump_length → jump_length → Prop ≝ λj1.λj2.jmple j1 j2 ∨ j1 = j2. definition policy_increase: list labelled_instruction → jump_expansion_policy → jump_expansion_policy → Prop ≝ λprogram.λop.λp. (∀i:ℕ.i < |program| → let 〈pc1,i1,oj〉 ≝ bvt_lookup ?? (bitvector_of_nat ? i) (\snd op) 〈0,0,short_jump〉 in let 〈pc2,i2,j〉 ≝ bvt_lookup ?? (bitvector_of_nat ? i) (\snd p) 〈0,0,short_jump〉 in jmpleq oj j ). (* Policy safety *) definition policy_safe: jump_expansion_policy → Prop ≝ λpolicy. bvt_forall ?? (\snd policy) (λn.λx.let 〈pc_nat,addr_nat,jmp_len〉 ≝ x in match jmp_len with [ short_jump ⇒ if leb pc_nat addr_nat then le (addr_nat - pc_nat) 126 else le (pc_nat - addr_nat) 129 | medium_jump ⇒ let addr ≝ bitvector_of_nat 16 addr_nat in let pc ≝ bitvector_of_nat 16 pc_nat in let 〈fst_5_addr, rest_addr〉 ≝ split bool 5 11 addr in let 〈fst_5_pc, rest_pc〉 ≝ split bool 5 11 pc in eq_bv 5 fst_5_addr fst_5_pc = true | long_jump ⇒ True ] ). (* Definitions and theorems for the jump_length type (itself defined in Assembly) *) definition max_length: jump_length → jump_length → jump_length ≝ λj1.λj2. match j1 with [ long_jump ⇒ long_jump | medium_jump ⇒ match j2 with [ medium_jump ⇒ medium_jump | _ ⇒ long_jump ] | short_jump ⇒ match j2 with [ short_jump ⇒ short_jump | _ ⇒ long_jump ] ]. lemma dec_jmple: ∀x,y:jump_length.jmple x y + ¬(jmple x y). #x #y cases x cases y /3 by inl, inr, nmk, I/ qed. lemma jmpleq_max_length: ∀ol,nl. jmpleq ol (max_length ol nl). #ol #nl cases ol cases nl /2 by or_introl, or_intror, I/ qed. lemma dec_eq_jump_length: ∀a,b:jump_length.(a = b) + (a ≠ b). #a #b cases a cases b /2/ %2 @nmk #H destruct (H) qed. (* Labels *) definition is_label ≝ λx:labelled_instruction.λl:Identifier. let 〈lbl,instr〉 ≝ x in match lbl with [ Some l' ⇒ l' = l | _ ⇒ False ]. lemma label_does_not_occur: ∀i:ℕ.∀p:list labelled_instruction.∀l:Identifier. is_label (nth i ? p 〈None ?, Comment [ ]〉) l → does_not_occur ?? l p = false. #i #p #l generalize in match i; elim p [ #i >nth_nil #H @⊥ @H | #h #t #IH #i cases i -i [ cases h #hi #hp cases hi [ normalize #H @⊥ @H | #l' #Heq whd in ⊢ (??%?); change with (eq_identifier ? l' l) in match (instruction_matches_identifier ????); whd in Heq; >Heq >eq_identifier_refl / by refl/ ] | #i #H whd in match (does_not_occur ????); whd in match (instruction_matches_identifier ????); cases h #hi #hp cases hi normalize nodelta [ @(IH i) @H | #l' @eq_identifier_elim [ normalize / by / | normalize #_ @(IH i) @H ] ] ] ] qed. definition expand_relative_jump_internal_unsafe: jump_length → ([[relative]] → preinstruction [[relative]]) → option (list instruction) ≝ λjmp_len:jump_length.λi. match jmp_len with [ short_jump ⇒ Some ? [ RealInstruction (i (RELATIVE (zero 8))) ] | medium_jump ⇒ None … | long_jump ⇒ Some ? [ RealInstruction (i (RELATIVE (bitvector_of_nat ? 2))); SJMP (RELATIVE (bitvector_of_nat ? 3)); (* LJMP size? *) LJMP (ADDR16 (zero 16)) ] ]. @I qed. definition expand_relative_jump_unsafe: jump_length → preinstruction Identifier → option (list instruction) ≝ λjmp_len:jump_length.λi. match i with [ JC jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (JC ?) | JNC jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (JNC ?) | JB baddr jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (JB ? baddr) | JZ jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (JZ ?) | JNZ jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (JNZ ?) | JBC baddr jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (JBC ? baddr) | JNB baddr jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (JNB ? baddr) | CJNE addr jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (CJNE ? addr) | DJNZ addr jmp ⇒ expand_relative_jump_internal_unsafe jmp_len (DJNZ ? addr) | ADD arg1 arg2 ⇒ Some … [ ADD ? arg1 arg2 ] | ADDC arg1 arg2 ⇒ Some … [ ADDC ? arg1 arg2 ] | SUBB arg1 arg2 ⇒ Some … [ SUBB ? arg1 arg2 ] | INC arg ⇒ Some … [ INC ? arg ] | DEC arg ⇒ Some … [ DEC ? arg ] | MUL arg1 arg2 ⇒ Some … [ MUL ? arg1 arg2 ] | DIV arg1 arg2 ⇒ Some … [ DIV ? arg1 arg2 ] | DA arg ⇒ Some … [ DA ? arg ] | ANL arg ⇒ Some … [ ANL ? arg ] | ORL arg ⇒ Some … [ ORL ? arg ] | XRL arg ⇒ Some … [ XRL ? arg ] | CLR arg ⇒ Some … [ CLR ? arg ] | CPL arg ⇒ Some … [ CPL ? arg ] | RL arg ⇒ Some … [ RL ? arg ] | RR arg ⇒ Some … [ RR ? arg ] | RLC arg ⇒ Some … [ RLC ? arg ] | RRC arg ⇒ Some … [ RRC ? arg ] | SWAP arg ⇒ Some … [ SWAP ? arg ] | MOV arg ⇒ Some … [ MOV ? arg ] | MOVX arg ⇒ Some … [ MOVX ? arg ] | SETB arg ⇒ Some … [ SETB ? arg ] | PUSH arg ⇒ Some … [ PUSH ? arg ] | POP arg ⇒ Some … [ POP ? arg ] | XCH arg1 arg2 ⇒ Some … [ XCH ? arg1 arg2 ] | XCHD arg1 arg2 ⇒ Some … [ XCHD ? arg1 arg2 ] | RET ⇒ Some … [ RET ? ] | RETI ⇒ Some … [ RETI ? ] | NOP ⇒ Some … [ RealInstruction (NOP ?) ] ]. definition expand_pseudo_instruction_unsafe: jump_length → pseudo_instruction → option (list instruction) ≝ λjmp_len. λi. match i with [ Cost cost ⇒ Some ? [ ] | Comment comment ⇒ Some ? [ ] | Call call ⇒ match jmp_len with [ short_jump ⇒ None ? | medium_jump ⇒ Some ? [ ACALL (ADDR11 (zero 11)) ] | long_jump ⇒ Some ? [ LCALL (ADDR16 (zero 16)) ] ] | Mov d trgt ⇒ Some ? [ RealInstruction (MOV ? (inl ? ? (inl ? ? (inr ? ? 〈DPTR, DATA16 (zero 16)〉))))] | Instruction instr ⇒ expand_relative_jump_unsafe jmp_len instr | Jmp jmp ⇒ match jmp_len with [ short_jump ⇒ Some ? [ SJMP (RELATIVE (zero 8)) ] | medium_jump ⇒ Some ? ([ AJMP (ADDR11 (zero 11)) ]) | long_jump ⇒ Some ? [ LJMP (ADDR16 (zero 16)) ] ] ]. @I qed. definition instruction_size: (* ℕ → *) jump_length → pseudo_instruction → ℕ ≝ λjmp_len.λinstr. let ilist ≝ expand_pseudo_instruction_unsafe jmp_len instr in let bv: list (BitVector 8) ≝ match ilist with [ None ⇒ (* this shouldn't happen *) [ ] | Some l ⇒ flatten … (map … assembly1 l) ] in |bv|. definition policy_isize_sum ≝ λprefix:list labelled_instruction.λpolicy:jump_expansion_policy. (\fst policy) = foldl_strong (option Identifier × pseudo_instruction) (λacc.ℕ) prefix (λhd.λx.λtl.λp.λacc. let 〈i1,i2,jl〉 ≝ bvt_lookup … (bitvector_of_nat ? (|hd|)) (\snd policy) 〈0,0,short_jump〉 in acc + (instruction_size jl (\snd x))) 0. (* The function that creates the label-to-address map *) definition create_label_map: ∀program:list labelled_instruction. ∀policy:jump_expansion_policy. (Σlabels:label_map. ∀i:ℕ.lt i (|program|) → ∀l.occurs_exactly_once ?? l program → is_label (nth i ? program 〈None ?, Comment [ ]〉) l → ∃a.lookup … labels l = Some ? 〈i,a〉 ) ≝ λprogram.λpolicy. let 〈final_pc, final_labels〉 ≝ foldl_strong (option Identifier × pseudo_instruction) (λprefix.ℕ × (Σlabels. ∀i:ℕ.lt i (|prefix|) → ∀l.occurs_exactly_once ?? l prefix → is_label (nth i ? prefix 〈None ?, Comment [ ]〉) l → ∃a.lookup … labels l = Some ? 〈i,a〉) ) program (λprefix.λx.λtl.λprf.λacc. let 〈pc,labels〉 ≝ acc in let 〈label,instr〉 ≝ x in let new_labels ≝ match label with [ None ⇒ labels | Some l ⇒ add ? (ℕ×ℕ) labels l 〈|prefix|, pc〉 ] in let 〈i1,i2,jmp_len〉 ≝ bvt_lookup ?? (bitvector_of_nat 16 (|prefix|)) (\snd policy) 〈0,0,short_jump〉 in 〈pc + (instruction_size jmp_len instr), new_labels〉 ) 〈0, empty_map …〉 in final_labels. [#i >append_length >commutative_plus #Hi normalize in Hi; cases (le_to_or_lt_eq … Hi) -Hi; [ #Hi #l normalize nodelta; cases label normalize nodelta [ >occurs_exactly_once_None #Hocc >(nth_append_first ? ? prefix ? ? (le_S_S_to_le ? ? Hi)) #Hl lapply (pi2 … labels) #Hacc elim (Hacc i (le_S_S_to_le … Hi) l Hocc Hl) #addr #Haddr % [ @addr | @Haddr ] | #l' #Hocc #Hl lapply (occurs_exactly_once_Some_stronger ?????? Hocc) -Hocc; @eq_identifier_elim #Heq #Hocc [ normalize in Hocc; >(nth_append_first ? ? prefix ? ? (le_S_S_to_le … Hi)) in Hl; #Hl @⊥ @(absurd … Hocc) | normalize nodelta lapply (pi2 … labels) #Hacc elim (Hacc i (le_S_S_to_le … Hi) l Hocc ?) [ #addr #Haddr % [ @addr | (nth_append_first ? ? prefix ? ? (le_S_S_to_le … Hi)) in Hl; / by / ] ] >(label_does_not_occur i … Hl) normalize nodelta @nmk / by / ] | #Hi #l #Hocc >(injective_S … Hi) >nth_append_second [ H in Hnone; #abs destruct (abs) ] qed. (* these two obviously belong somewhere else *) lemma pi1_eq: ∀A:Type[0].∀P:A->Prop.∀s1:Σa1:A.P a1.∀s2:Σa2:A.P a2. s1 = s2 → (pi1 ?? s1) = (pi1 ?? s2). #A #P #s1 #s2 / by / qed. lemma Some_eq: ∀T:Type[0].∀x,y:T. Some T x = Some T y → x = y. #T #x #y #H @option_destruct_Some @H qed. (* The first step of the jump expansion: everything to short. * The third condition of the dependent type implies jump_in_policy; * I've left it in for convenience of type-checking. *) definition jump_expansion_start: ∀program:(Σl:list labelled_instruction.|l| < 2^16). Σpolicy:jump_expansion_policy. And (And (out_of_program_none (pi1 ?? program) policy) (jump_in_policy (pi1 ?? program) policy)) (∀i.i < |pi1 ?? program| → is_jump (nth i ? (pi1 ?? program) 〈None ?, Comment []〉) → lookup_opt … (bitvector_of_nat 16 i) (\snd policy) = Some ? 〈0,0,short_jump〉) ≝ λprogram. foldl_strong (option Identifier × pseudo_instruction) (λprefix.Σpolicy:jump_expansion_policy. And (And (out_of_program_none prefix policy) (jump_in_policy prefix policy)) (∀i.i < |prefix| → is_jump (nth i ? prefix 〈None ?, Comment []〉) → lookup_opt … (bitvector_of_nat 16 i) (\snd policy) = Some ? 〈0,0,short_jump〉)) program (λprefix.λx.λtl.λprf.λp. let 〈pc,policy〉 ≝ p in let 〈label,instr〉 ≝ x in 〈pc + instruction_size short_jump instr,match instr with [ Instruction i ⇒ match i with [ JC _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | JNC _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | JZ _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | JNZ _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | JB _ _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | JNB _ _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | JBC _ _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | CJNE _ _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | DJNZ _ _ ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | _ ⇒ policy ] | Call c ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | Jmp j ⇒ bvt_insert … (bitvector_of_nat 16 (|prefix|)) 〈0,0,short_jump〉 policy | _ ⇒ policy ]〉 ) 〈0, Stub ? ?〉. [ @conj [ @conj #i >append_length Hi @Hi2 |9,10,11,12,13,14,15,16,17: [1,2,6,7: #x |3,4,5,8,9: #x #id] >lookup_opt_insert_miss [2,4,6,8,10,12,14,16,18: @bitvector_of_nat_abs [1,4,7,10,13,16,19,22,25: @Hi2 |2,5,8,11,14,17,20,23,26: @(transitive_lt … Hi2) @le_S_to_le @Hi |3,6,9,12,15,18,21,24,27: @sym_neq @lt_to_not_eq @le_S_to_le @Hi ] |1,3,5,7,9,11,13,15,17: @(proj1 ?? (proj1 ?? Hp) i ? Hi2) @le_S_to_le @le_S_to_le @Hi ] |46,47,48,49,50,51,52,53,54: [1,2,6,7: #x |3,4,5,8,9: #x #id] >lookup_opt_insert_miss [2,4,6,8,10,12,14,16,18: @bitvector_of_nat_abs [1,4,7,10,13,16,19,22,25: @Hi2 |2,5,8,11,14,17,20,23,26: @(transitive_lt … Hi2) Hi @Hi2 ] |4,5,10,11: #dst normalize nodelta >lookup_opt_insert_miss [2,4,6,8: @bitvector_of_nat_abs [1,4,7,10: @Hi2 |2,5: @(transitive_lt … Hi2) @le_S_to_le @Hi |8,11: @(transitive_lt … Hi2) (nth_append_first ?? prefix ?? (le_S_S_to_le … Hi)) @conj [ #Hj @(ex_intro ?? 0 (ex_intro ?? 0 (ex_intro ?? short_jump ?))) cases p -p #p cases p #pc #pol #Hp cases x #l #ins cases ins normalize nodelta [1: #pi cases pi normalize nodelta [1,2,3,6,7,33,34: #x #y |4,5,8,18,19,20,21,22,23,24,25,26,27,28,29,30,31,32: #x |9,10,11,12,13,14,15,16,17: [1,2,6,7: #x | 3,4,5,8,9: #x #id] >lookup_opt_insert_miss [2,4,6,8,10,12,14,16,18: @bitvector_of_nat_abs [1,4,7,10,13,16,19,22,25: @(transitive_lt … (pi2 ?? program)) >prf >append_length prf >append_length lookup_opt_insert_miss [2,4: @bitvector_of_nat_abs [1,4: @(transitive_lt … (pi2 ?? program)) >prf >append_length prf >append_length lookup_opt_insert_miss [1,3,5,7,9,11,13,15,17: / by / |2,4,6,8,10,12,14,16,18: @bitvector_of_nat_abs [1,4,7,10,13,16,19,22,25: @(transitive_lt … (pi2 ?? program)) >prf >append_length prf >append_length lookup_opt_insert_miss [2,4: @bitvector_of_nat_abs [1,4: @(transitive_lt … (pi2 ?? program)) >prf >append_length prf >append_length (injective_S … Hi) >(nth_append_second ?? prefix ?? (le_n (|prefix|))) (proj1 ?? (proj1 ?? Hp) (|prefix|) (le_n (|prefix|)) ?) [1,3,5,7,9,11,13,15,17,19,21,23,25,27,29,31,33,35,37,39,41,43,45,47,49,51,53,55: #H destruct (H) |2,4,6,8,10,12,14,16,18,20,22,24,26,28,30,32,34,36,38,40,42,44,46,48,50,52,54,56: @(transitive_lt … (pi2 ?? program)) >prf >append_length (proj1 ?? (proj1 ?? Hp) (|prefix|) (le_n (|prefix|)) ?) [1,3,5: #H destruct (H) |2,4,6: @(transitive_lt … (pi2 ?? program)) >prf >append_length append_length (nth_append_first ?? prefix ?? (le_S_S_to_le … Hi)) #H cases p -p #p cases p #pc #pol #Hp cases x #l #ins cases ins normalize nodelta [1: #pi cases pi [1,2,3,4,5,6,7,8,18,19,20,21,22,23,24,25,26,27,28,29,30,31,32,33,34,35,36,37: [1,2,3,6,7,24,25: #x #y |4,5,8,9,10,11,12,13,14,15,16,17,18,19,20,21,22,23: #x] normalize nodelta @((proj2 ?? Hp) i (le_S_S_to_le … Hi) H) |9,10,11,12,13,14,15,16,17: [1,2,6,7: #x | 3,4,5,8,9: #x #id] normalize nodelta >lookup_opt_insert_miss [1,3,5,7,9,11,13,15,17: @((proj2 ?? Hp) i (le_S_S_to_le … Hi) H) |2,4,6,8,10,12,14,16,18,20: @bitvector_of_nat_abs [1,4,7,10,13,16,19,22,25: @(transitive_lt … (pi2 ?? program)) >prf >append_length prf >append_length lookup_opt_insert_miss [1,3: @((proj2 ?? Hp) i (le_S_S_to_le … Hi) H) |2,4: @bitvector_of_nat_abs [1,4: @(transitive_lt … (pi2 ?? program)) >prf >append_length prf >append_length (injective_S … Hi) >(nth_append_second ?? prefix ?? (le_n (|prefix|))) lookup_opt_insert_hit / by / ] |2,3,6: #x [3: #id] #H @⊥ @H |4,5: #id #_ cases p -p #p cases p -p #pc #pol #Hp >lookup_opt_insert_hit / by / ] ] ] | @conj [ @conj [ #i #_ #Hi2 / by refl/ | #i #_ @conj [ >nth_nil #H @⊥ @H | #H elim H -H #p #H elim H -H #a #H elim H -H #j #H normalize in H; destruct (H) ] ] | #i #Hi >nth_nil #H @⊥ @H ] ] qed. definition policy_equal_int ≝ λprogram:list labelled_instruction.λp1,p2:jump_expansion_policy. ∀n:ℕ.n < |program| → let 〈pc1,i1,j1〉 ≝ bvt_lookup … (bitvector_of_nat 16 n) (\snd p1) 〈0,0,short_jump〉 in let 〈pc2,i2,j2〉 ≝ bvt_lookup … (bitvector_of_nat 16 n) (\snd p2) 〈0,0,short_jump〉 in j1 = j2. definition nec_plus_ultra ≝ λprogram:list labelled_instruction.λp:jump_expansion_policy. ¬(∀i.i < |program| → \snd (bvt_lookup … (bitvector_of_nat 16 i) (\snd p) 〈0,0,short_jump〉) = long_jump). lemma geb_to_leb: ∀a,b:ℕ.geb a b = leb b a. #a #b / by refl/ qed. (*lemma foldl_strong_eq: ∀A: Type[0]. ∀P: list A → Type[0]. ∀l: list A. ∀H: ∀prefix. ∀hd. ∀tl. l = prefix @ [hd] @ tl → P prefix → P (prefix @ [hd]). ∀acc: P [ ]. foldl_strong_internal A P l H [ ] l acc (refl …).*) (* One step in the search for a jump expansion fixpoint. *) definition jump_expansion_step: ∀program:(Σl:list labelled_instruction.|l| < 2^16). ∀labels:(Σlm:label_map.∀i:ℕ.lt i (|pi1 ?? program|) → ∀l.occurs_exactly_once ?? l (pi1 ?? program) → is_label (nth i ? (pi1 ?? program) 〈None ?, Comment [ ]〉) l → ∃a.lookup … lm l = Some ? 〈i,a〉). ∀old_policy:(Σpolicy:jump_expansion_policy. out_of_program_none (pi1 ?? program) policy ∧ jump_in_policy (pi1 ?? program) policy ∧ (\fst policy < 2^16) ∧ policy_isize_sum (pi1 ?? program) policy). (Σx:bool × (option jump_expansion_policy). let 〈ch,y〉 ≝ x in match y with [ None ⇒ nec_plus_ultra program old_policy | Some p ⇒ out_of_program_none (pi1 ?? program) p ∧ labels_okay labels p ∧ jump_in_policy (pi1 ?? program) p ∧ policy_increase (pi1 ?? program) old_policy p ∧ policy_safe p ∧ (ch = false → policy_equal_int (pi1 ?? program) old_policy p) ∧ policy_isize_sum (pi1 ?? program) p ∧ \fst p < 2^16 ]) ≝ λprogram.λlabels.λold_policy. let 〈final_changed, final_policy〉 ≝ foldl_strong (option Identifier × pseudo_instruction) (λprefix.Σx:bool × jump_expansion_policy. let 〈changed,policy〉 ≝ x in out_of_program_none prefix policy ∧ labels_okay labels policy ∧ jump_in_policy prefix policy ∧ policy_increase prefix old_policy policy ∧ policy_safe policy ∧ (changed = false → policy_equal_int prefix old_policy policy) ∧ policy_isize_sum prefix policy) program (λprefix.λx.λtl.λprf.λacc. let 〈changed, pol〉 ≝ acc in (* let 〈pc,pol〉 ≝ p in *) let 〈label,instr〉 ≝ x in let add_instr ≝ match instr with [ Instruction i ⇒ jump_expansion_step_instruction labels (\fst pol) i | Call c ⇒ Some ? (select_call_length labels (\fst pol) c) | Jmp j ⇒ Some ? (select_jump_length labels (\fst pol) j) | _ ⇒ None ? ] in let old_length ≝ \snd (bvt_lookup … (bitvector_of_nat 16 (|prefix|)) (\snd old_policy) 〈0,0,short_jump〉) in match add_instr with [ None ⇒ (* i.e. it's not a jump *) let isize ≝ instruction_size short_jump instr in 〈changed, 〈(\fst pol) + isize, (\snd pol)〉〉 | Some z ⇒ let 〈addr,ai〉 ≝ z in let new_length ≝ max_length old_length ai in let isize ≝ instruction_size new_length instr in 〈match dec_eq_jump_length new_length old_length with [ inl _ ⇒ changed | inr _ ⇒ true], 〈(\fst pol) + isize, insert … (bitvector_of_nat 16 (|prefix|)) 〈(\fst pol), addr, new_length〉 (\snd pol)〉〉 ] ) 〈false, 〈0, Stub ? ?〉〉 in if geb (\fst final_policy) 2^16 then 〈final_changed,None ?〉 else 〈final_changed,Some ? final_policy〉. [ @pair_elim #fch #x #Heq <(pair_eq2 ?????? Heq) normalize nodelta @nmk #Hfull lapply p generalize in match (foldl_strong ?????); * #x #H #H2 >H2 in H; normalize nodelta -H2 -x #H @(absurd ((\fst final_policy) ≥ 2^16)) [ @leb_true_to_le (proj2 ?? H) >(proj2 ?? (pi2 ?? old_policy)) cases daemon (* XXX *) ] | normalize nodelta lapply p generalize in match (foldl_strong ?????); * #x #H #H2 >H2 in H; normalize nodelta #H @conj [ @H | @not_le_to_lt @leb_false_to_not_le p #Hpolicy normalize nodelta in Hpolicy; cases (dec_eq_jump_length new_length old_length) #Hlength normalize nodelta @conj [1,3: @conj [1,3: @conj [1,3: @conj [1,3: @conj [1,3: @conj [1,3: (* out_of_policy_none *) #i >append_length lookup_opt_insert_miss [1,3: @(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy))))) i (le_S_to_le … Hi)) @Hi2 |2,4: >eq_bv_sym @bitvector_of_nat_abs [1,4: @(transitive_lt … (pi2 ?? program)) >prf >append_length normalize p3 lapply (refl ? (lookup_def ?? (pi1 ?? labels) id 〈0,\fst pol〉)) cases (lookup_def ?? labels id 〈0,\fst pol〉) in ⊢ (???% → %); #q1 #q2 normalize nodelta #H >(pair_eq1 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) >(pair_eq2 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) lapply (refl ? (leb (\fst pol) q2)) cases (leb (\fst pol) q2) in ⊢ (???% → %); #Hle1 [1,3,5,7,9,11,13,15: lapply (refl ? (leb (q2-\fst pol) 126)) cases (leb (q2-\fst pol) 126) in ⊢ (???% → %); |2,4,6,8,10,12,14,16: lapply (refl ? (leb (\fst pol-q2) 129)) cases (leb (\fst pol-q2) 129) in ⊢ (???% → %); ] #Hle2 normalize nodelta #Hli @(ex_intro ?? id) >H <(pair_eq1 ?????? (Some_eq ??? Hli)) @refl |11,12,13,16,17,48,49,50,53,54: #x #id normalize nodelta whd in match (jump_expansion_step_instruction ???); whd in match (select_reljump_length ???); >p3 lapply (refl ? (lookup_def ?? labels id 〈0,\fst pol〉)) cases (lookup_def ?? labels id 〈0,\fst pol〉) in ⊢ (???% → %); #q1 #q2 normalize nodelta #H >(pair_eq1 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) >(pair_eq2 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) lapply (refl ? (leb (\fst pol) q2)) cases (leb (\fst pol) q2) in ⊢ (???% → %); #Hle1 [1,3,5,7,9,11,13,15,17,19: lapply (refl ? (leb (q2-\fst pol) 126)) cases (leb (q2-\fst pol) 126) in ⊢ (???% → %); |2,4,6,8,10,12,14,16,18,20: lapply (refl ? (leb (\fst pol-q2) 129)) cases (leb (\fst pol-q2) 129) in ⊢ (???% → %); ] #Hle2 normalize nodelta #Hli @(ex_intro ?? id) >H <(pair_eq1 ?????? (Some_eq ??? Hli)) @refl ] |4,5,10,11: #id normalize nodelta whd in match (select_jump_length ???); whd in match (select_call_length ???); >p3 lapply (refl ? (lookup_def ?? labels id 〈0,\fst pol〉)) cases (lookup_def ?? labels id 〈0,\fst pol〉) in ⊢ (???% → %); #q1 #q2 normalize nodelta #H [1,3: cases (leb (\fst pol) q2) [1,3: cases (leb (q2-\fst pol) 126) |2,4: cases (leb (\fst pol-q2) 129) ] [1,3,5,7: normalize nodelta #H2 >(pair_eq1 ?????? (Some_eq ??? H2)) in H; #Hli @(ex_intro ?? id) lapply (proj2 ?? Hl) #H >(pair_eq1 ?????? (pair_eq1 ?????? H)) >(pair_eq2 ?????? (pair_eq1 ?????? H)) >Hli @refl] ] cases (split ? 5 11 (bitvector_of_nat 16 q2)) #n1 #n2 cases (split ? 5 11 (bitvector_of_nat 16 (\fst pol))) #m1 #m2 normalize nodelta cases (eq_bv ? n1 m1) normalize nodelta #H2 >(pair_eq1 ?????? (Some_eq ??? H2)) in H; #H @(ex_intro ?? id) lapply (proj2 ?? Hl) #H2 >(pair_eq1 ?????? (pair_eq1 ?????? H2)) >(pair_eq2 ?????? (pair_eq1 ?????? H2)) >H @refl ] ] ] |2,4: (* jump_in_policy *) #i #Hi cases (le_to_or_lt_eq … Hi) -Hi; [1,3: >append_length (nth_append_first ?? prefix ??(le_S_S_to_le ?? Hi)) @conj [1,3: #Hj lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy)))) i (le_S_S_to_le … Hi)) #Hacc elim (proj1 ?? Hacc Hj) #h #n elim n -n #n #H elim H -H #j #Hj @(ex_intro ?? h (ex_intro ?? n (ex_intro ?? j ?))) whd in match (snd ???); >lookup_opt_insert_miss [1,3: @Hj |2,4: @bitvector_of_nat_abs ] [3,6: @(lt_to_not_eq i (|prefix|)) @(le_S_S_to_le … Hi) |1,4: @(transitive_lt ??? (le_S_S_to_le ?? Hi)) ] @(transitive_lt … (pi2 ?? program)) >prf >append_length normalize prf >append_length normalize append_length (injective_S … Hi) >(nth_append_second ?? prefix ?? (le_n (|prefix|))) append_length >commutative_plus #Hi normalize in Hi; cases (le_to_or_lt_eq … Hi) -Hi; #Hi [1,3: >lookup_insert_miss [1,3: @(proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy))) i (le_S_S_to_le … Hi)) |2,4: @bitvector_of_nat_abs [3,6: @lt_to_not_eq @(le_S_S_to_le … Hi) |1,4: @(transitive_lt ??? (le_S_S_to_le … Hi)) ] @(transitive_lt … (pi2 ?? program)) >prf >append_length normalize (injective_S … Hi) normalize nodelta in Hlength; >lookup_insert_hit normalize nodelta >Hlength @pair_elim #l1 #l2 #Hl @pair_elim #y1 #y2 #Hy >Hl %2 @refl |4: cases daemon (* XXX get rest of proof done first *) ] ] |2,4: (* policy_safe *) @lookup_forall #x cases x -x #x cases x -x #p #a #j #n normalize nodelta #Hl elim (insert_lookup_opt ?? 〈p,a,j〉 ???? Hl) -Hl #Hl [1,3: @(forall_lookup … (proj2 ?? (proj1 ?? (proj1 ?? Hpolicy))) ? n Hl) |2,4: normalize nodelta in p2; cases instr in p2; [2,3,8,9: #x #abs normalize in abs; lapply (jmeq_to_eq ??? abs) #H destruct (H) |6,12: #x #y #abs normalize in abs; lapply (jmeq_to_eq ??? abs) #H destruct (H) |1,7: #pi cases pi normalize nodelta [35,36,37,72,73,74: #abs normalize in abs; lapply (jmeq_to_eq ??? abs) #H destruct (H) |4,5,8,18,19,20,21,22,23,24,25,26,27,28,29,30,31,32,41,42,45,55,56,57,58,59,60,61,62,63,64,65,66,67,68,69: #x #abs normalize in abs; lapply (jmeq_to_eq ??? abs) #H destruct (H) |1,2,3,6,7,33,34,38,39,40,43,44,70,71: #x #y #abs normalize in abs; lapply (jmeq_to_eq ??? abs) #H destruct (H) |9,10,14,15,46,47,51,52: #id >p3 whd in match (jump_expansion_step_instruction ???); whd in match (select_reljump_length ???); cases (lookup_def ?? labels id 〈0,\fst pol〉) #q1 #q2 normalize nodelta >(pair_eq1 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) >(pair_eq2 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) lapply (refl ? (leb (\fst pol) q2)) cases (leb (\fst pol) q2) in ⊢ (???% → %); #Hle1 [1,3,5,7,9,11,13,15: lapply (refl ? (leb (q2-\fst pol) 126)) cases (leb (q2-\fst pol) 126) in ⊢ (???% → %); |2,4,6,8,10,12,14,16: lapply (refl ? (leb (\fst pol-q2) 129)) cases (leb (\fst pol-q2) 129) in ⊢ (???% → %); ] #Hle2 normalize nodelta #Hli <(pair_eq1 ?????? (Some_eq ??? Hli)) >Hle1 >(pair_eq2 ?????? (proj2 ?? Hl)) <(pair_eq2 ?????? (Some_eq ??? Hli)) cases (\snd (lookup ?? (bitvector_of_nat ? (|prefix|)) (\snd old_policy) ?)) [1,7,13,19,25,31,37,43,49,55,61,67,73,79,85,91: @(leb_true_to_le … Hle2) ] normalize @I (* much faster than / by I/, strangely enough *) |11,12,13,16,17,48,49,50,53,54: #x #id >p3 whd in match (jump_expansion_step_instruction ???); whd in match (select_reljump_length ???); cases (lookup_def ?? labels id 〈0,\fst pol〉) #q1 #q2 normalize nodelta >(pair_eq1 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) >(pair_eq2 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) lapply (refl ? (leb (\fst pol) q2)) cases (leb (\fst pol) q2) in ⊢ (???% → %); #Hle1 [1,3,5,7,9,11,13,15,17,19: lapply (refl ? (leb (q2-\fst pol) 126)) cases (leb (q2-\fst pol) 126) in ⊢ (???% → %); |2,4,6,8,10,12,14,16,18,20: lapply (refl ? (leb (\fst pol-q2) 129)) cases (leb (\fst pol-q2) 129) in ⊢ (???% → %); ] #Hle2 normalize nodelta #Hli <(pair_eq1 ?????? (Some_eq ??? Hli)) >Hle1 >(pair_eq2 ?????? (proj2 ?? Hl)) <(pair_eq2 ?????? (Some_eq ??? Hli)) cases (\snd (lookup ?? (bitvector_of_nat ? (|prefix|)) (\snd old_policy) ?)) [1,7,13,19,25,31,37,43,49,55,61,67,73,79,85,91,97,103,109,115: @(leb_true_to_le … Hle2) ] normalize @I (* vide supra *) ] |4,5,10,11: #id >p3 normalize nodelta whd in match (select_jump_length ???); whd in match (select_call_length ???); cases (lookup_def ?? labels id 〈0,\fst pol〉) #q1 #q2 normalize nodelta >(pair_eq1 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) >(pair_eq2 ?????? (pair_eq1 ?????? (proj2 ?? Hl))) [1,3: lapply (refl ? (leb (\fst pol) q2)) cases (leb (\fst pol) q2) in ⊢ (???% → %); #Hle1 [1,3: lapply (refl ? (leb (q2-\fst pol) 126)) cases (leb (q2-\fst pol) 126) in ⊢ (???% → %); |2,4: lapply (refl ? (leb (\fst pol-q2) 129)) cases (leb (\fst pol-q2) 129) in ⊢ (???% → %); ] #Hle2 normalize nodelta ] [2,4,6,8,9,10: lapply (refl ? (split ? 5 11 (bitvector_of_nat ? q2))) cases (split ??? (bitvector_of_nat ? q2)) in ⊢ (???% → %); #x1 #x2 #Hle3 lapply (refl ? (split ? 5 11 (bitvector_of_nat ? (\fst pol)))) cases (split ??? (bitvector_of_nat ? (\fst pol))) in ⊢ (???% → %); #y1 #y2 #Hle4 normalize nodelta lapply (refl ? (eq_bv 5 x1 y1)) cases (eq_bv 5 x1 y1) in ⊢ (???% → %); #Hle5 ] #Hli <(pair_eq1 ?????? (Some_eq ??? Hli)) >(pair_eq2 ?????? (proj2 ?? Hl)) <(pair_eq2 ?????? (Some_eq ??? Hli)) cases (\snd (lookup ?? (bitvector_of_nat ? (|prefix|)) (\snd old_policy) ?)) [2,8,14,20,26,32: >Hle3 @Hle5 |37,40,43,46: >Hle1 @(leb_true_to_le … Hle2) ] normalize @I (* here too *) ] ] ] |2,4: (* changed *) normalize nodelta #Hc [2: destruct (Hc)] #i #Hi cases (le_to_or_lt_eq … Hi) -Hi >append_length >commutative_plus #Hi normalize in Hi; [ >lookup_insert_miss [ @(proj2 ?? (proj1 ?? Hpolicy) Hc i (le_S_S_to_le … Hi)) | @bitvector_of_nat_abs [3: @lt_to_not_eq @(le_S_S_to_le … Hi) |1: @(transitive_lt ??? (le_S_S_to_le … Hi)) ] @(transitive_lt … (pi2 ?? program)) >prf >append_length normalize (injective_S … Hi) >lookup_insert_hit normalize nodelta in Hlength; >Hlength normalize nodelta @pair_elim #l1 #l2 #Hl @pair_elim #y1 #y2 #Hy >Hl @refl ] ] |2,4: (* policy_isize_sum XXX *) cases daemon ] | (* Case where add_instr = None *) normalize nodelta lapply (pi2 ?? acc) >p >p1 normalize nodelta #Hpolicy @conj [ @conj [ @conj [ @conj [ @conj [ @conj [ (* out_of_program_none *) #i >append_length >commutative_plus #Hi normalize in Hi; #Hi2 @(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy))))) i (le_S_to_le ?? Hi) Hi2) | (* labels_okay *) @lookup_forall #x cases x -x #x cases x #p #a #j #lbl #Hl normalize nodelta elim (forall_lookup … (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy)))))) ? lbl Hl) #id #Hid @(ex_intro … id Hid) ] | (* jump_in_policy *) #i >append_length >commutative_plus #Hi normalize in Hi; elim (le_to_or_lt_eq … Hi) -Hi #Hi [ >(nth_append_first ?? prefix ?? (le_S_S_to_le ?? Hi)) @(proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy)))) i (le_S_S_to_le ?? Hi)) | >(injective_S … Hi) @conj [ >(nth_append_second ?? prefix ?? (le_n (|prefix|))) (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy))))) (|prefix|) (le_n (|prefix|)) ?) in H; [ #H destruct (H) | @(transitive_lt … (pi2 ?? program)) >prf >append_length normalize append_length >commutative_plus #Hi normalize in Hi; elim (le_to_or_lt_eq … Hi) -Hi #Hi [ @(proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy))) i (le_S_S_to_le … Hi)) | >(injective_S … Hi) >lookup_opt_lookup_miss [ >lookup_opt_lookup_miss [ normalize nodelta %2 @refl | @(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy))))) (|prefix|) (le_n (|prefix|)) ?) @(transitive_lt … (pi2 ?? program)) >prf >append_length normalize prf [ >append_length normalize (nth_append_second ?? prefix ?? (le_n (|prefix|))) p1 whd in match (nth ????); normalize nodelta in p2; cases instr in p2; [ #pi cases pi [1,2,3,6,7,33,34: #x #y #_ normalize /2 by nmk/ |4,5,8,18,19,20,21,22,23,24,25,26,27,28,29,30,31,32: #x #_ normalize /2 by nmk/ |9,10,14,15: #id (* normalize segfaults here *) normalize nodelta whd in match (jump_expansion_step_instruction ???); #H lapply (jmeq_to_eq ??? H) #H2 destruct (H2) |11,12,13,16,17: #x #id normalize nodelta whd in match (jump_expansion_step_instruction ???); #H lapply (jmeq_to_eq ??? H) #H2 destruct (H2) |35,36,37: #_ normalize /2 by nmk/ ] |2,3: #x #_ normalize /2 by nmk/ |4,5: #id normalize nodelta #H lapply (jmeq_to_eq ??? H) #H2 destruct (H2) |6: #x #id #_ normalize /2 by nmk/ ] ] ] ] ] | (* policy_safe *) @lookup_forall #x cases x -x #x cases x -x #p #a #j #n #Hl @(forall_lookup … (proj2 ?? (proj1 ?? (proj1 ?? Hpolicy))) ? n Hl) ] | (* changed *) #Hc #i >append_length >commutative_plus #Hi normalize in Hi; elim (le_to_or_lt_eq … Hi) -Hi #Hi [ @(proj2 ?? (proj1 ?? Hpolicy) Hc i (le_S_S_to_le … Hi)) | >(injective_S … Hi) >lookup_opt_lookup_miss [ >lookup_opt_lookup_miss [ normalize nodelta @refl | @(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpolicy))))) (|prefix|) (le_n (|prefix|)) ?) @(transitive_lt … (pi2 ?? program)) >prf >append_length normalize prf [ >append_length normalize (nth_append_second ?? prefix ?? (le_n (|prefix|))) p1 whd in match (nth ????); normalize nodelta in p2; cases instr in p2; [ #pi cases pi [1,2,3,6,7,33,34: #x #y #_ normalize /2 by nmk/ |4,5,8,18,19,20,21,22,23,24,25,26,27,28,29,30,31,32: #x #_ normalize /2 by nmk/ |9,10,14,15: #id (* normalize segfaults here *) normalize nodelta whd in match (jump_expansion_step_instruction ???); #H lapply (jmeq_to_eq ??? H) #H2 destruct (H2) |11,12,13,16,17: #x #id normalize nodelta whd in match (jump_expansion_step_instruction ???); #H lapply (jmeq_to_eq ??? H) #H2 destruct (H2) |35,36,37: #_ normalize /2 by nmk/ ] |2,3: #x #_ normalize /2 by nmk/ |4,5: #id normalize nodelta #H lapply (jmeq_to_eq ??? H) #H2 destruct (H2) |6: #x #id #_ normalize /2 by nmk/ ] ] ] ] ] | (* XXX policy_isize_sum *) cases daemon ] | @conj [ @conj [ @conj [ @conj [ @conj [ @conj [ #i #Hi / by refl/ | / by I/ ] | #i #Hi @conj [ >nth_nil #H @⊥ @H | #H elim H #x #H1 elim H1 #y #H2 elim H2 #z #H3 normalize in H3; destruct (H3) ] ] | #i #Hi @⊥ @(absurd (i<0)) [ @Hi | @(not_le_Sn_O) ] ] | / by I/ ] | #_ #i #Hi @⊥ @(absurd (i < 0)) [ @Hi | @not_le_Sn_O ] ] | / by refl/ ] ] qed. (* this might be replaced by a coercion: (∀x.A x → B x) → Σx.A x → Σx.B x *) (* definition weaken_policy: ∀program,op. option (Σp:jump_expansion_policy. And (And (And (And (out_of_program_none program p) (labels_okay (create_label_map program op) p)) (jump_in_policy program p)) (policy_increase program op p)) (policy_safe p)) → option (Σp:jump_expansion_policy.And (out_of_program_none program p) (jump_in_policy program p)) ≝ λprogram.λop.λx. match x return λ_.option (Σp.And (out_of_program_none program p) (jump_in_policy program p)) with [ None ⇒ None ? | Some z ⇒ Some ? (mk_Sig ?? (pi1 ?? z) ?) ]. @conj [ @(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (pi2 ?? z))))) | @(proj2 ?? (proj1 ?? (proj1 ?? (pi2 ?? z)))) ] qed. *) (* This function executes n steps from the starting point. *) (*let rec jump_expansion_internal (program: Σl:list labelled_instruction.lt (|l|) 2^16) (n: ℕ) on n:(Σx:bool × ℕ × (option jump_expansion_policy). let 〈ch,pc,y〉 ≝ x in match y with [ None ⇒ pc > 2^16 | Some p ⇒ And (out_of_program_none program p) (jump_in_policy program p) ]) ≝ match n with [ O ⇒ 〈0,Some ? (pi1 … (jump_expansion_start program))〉 | S m ⇒ let 〈ch,pc,z〉 as p1 ≝ (pi1 ?? (jump_expansion_internal program m)) in match z return λx. z=x → Σa:bool × ℕ × (option jump_expansion_policy).? with [ None ⇒ λp2.〈pc,None ?〉 | Some op ⇒ λp2.pi1 … (jump_expansion_step program (create_label_map program op) «op,?») ] (refl … z) ].*) let rec jump_expansion_internal (program: Σl:list labelled_instruction.lt (length ? l) 2^16) (n: ℕ) on n:(Σx:bool × (option jump_expansion_policy). let 〈c,pol〉 ≝ x in match pol with [ None ⇒ True | Some x ⇒ And (And (And (out_of_program_none program x) (jump_in_policy program x)) (\fst x < 2^16)) (policy_isize_sum program x) ]) ≝ match n with [ O ⇒ 〈true,Some ? (pi1 ?? (jump_expansion_start program))〉 | S m ⇒ let 〈ch,z〉 as p1 ≝ (pi1 ?? (jump_expansion_internal program m)) in match z return λx. z=x → Σa:bool × (option jump_expansion_policy).? with [ None ⇒ λp2.〈false,None ?〉 | Some op ⇒ λp2.if ch then pi1 ?? (jump_expansion_step program (create_label_map program op) «op,?») else (jump_expansion_internal program m) ] (refl … z) ]. [ normalize nodelta @conj [ @conj [ @(proj1 ?? (pi2 ?? (jump_expansion_start program))) | (* XXX *) cases daemon ] | (* XXX *) cases daemon ] | lapply (pi2 ?? (jump_expansion_internal program m)) p2 normalize nodelta / by / | lapply (pi2 ?? (jump_expansion_internal program m)) p2 normalize nodelta / by / | normalize nodelta cases (jump_expansion_step program (create_label_map program op) «op,?») #p cases p -p #p #r cases r normalize nodelta [ #H / by I/ | #j #H @conj [ @conj [ @conj [ @(proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? H))))))) | @(proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? H)))))) ] | @(proj2 ?? H) ] | @(proj2 ?? (proj1 ?? H)) ] ] ] qed. lemma pe_int_refl: ∀program.reflexive ? (policy_equal_int program). #program whd #x whd #n #Hn cases (bvt_lookup … (bitvector_of_nat 16 n) (\snd x) 〈0,0,short_jump〉) #y cases y -y #y #z normalize nodelta @refl qed. lemma pe_int_sym: ∀program.symmetric ? (policy_equal_int program). #program whd #x #y #Hxy whd #n #Hn lapply (Hxy n Hn) cases (bvt_lookup … (bitvector_of_nat ? n) (\snd x) 〈0,0,short_jump〉) #z cases z -z #x1 #x2 #x3 cases (bvt_lookup … (bitvector_of_nat ? n) (\snd y) 〈0,0,short_jump〉) #z cases z -z #y1 #y2 #y3 normalize nodelta // qed. lemma pe_int_trans: ∀program.transitive ? (policy_equal_int program). #program whd #x #y #z whd in match (policy_equal_int ???); whd in match (policy_equal_int ?y ?); whd in match (policy_equal_int ? x z); #Hxy #Hyz #n #Hn lapply (Hxy n Hn) -Hxy lapply (Hyz n Hn) -Hyz cases (bvt_lookup … (bitvector_of_nat ? n) (\snd x) 〈0,0,short_jump〉) #z cases z -z #x1 #x2 #x3 cases (bvt_lookup … (bitvector_of_nat ? n) (\snd y) 〈0,0,short_jump〉) #z cases z -z #y1 #y2 #y3 cases (bvt_lookup … (bitvector_of_nat ? n) (\snd z) 〈0,0,short_jump〉) #z cases z -z #z1 #z2 #z3 normalize nodelta // qed. definition policy_equal ≝ λprogram:list labelled_instruction.λp1,p2:option jump_expansion_policy. match p1 with [ Some x1 ⇒ match p2 with [ Some x2 ⇒ policy_equal_int program x1 x2 | _ ⇒ False ] | None ⇒ p2 = None ? ]. lemma pe_refl: ∀program.reflexive ? (policy_equal program). #program whd #x whd cases x [ // | #y @pe_int_refl ] qed. lemma pe_sym: ∀program.symmetric ? (policy_equal program). #program whd #x #y #Hxy whd cases y in Hxy; [ cases x [ // | #x' #H @⊥ @(absurd ? H) /2 by nmk/ ] | #y' cases x [ #H @⊥ @(absurd ? H) whd in match (policy_equal ???); @nmk #H destruct (H) | #x' #H @pe_int_sym @H ] ] qed. lemma pe_trans: ∀program.transitive ? (policy_equal program). #program whd #x #y #z cases x [ #Hxy #Hyz >Hxy in Hyz; // | #x' cases y [ #H @⊥ @(absurd ? H) /2 by nmk/ | #y' cases z [ #_ #H @⊥ @(absurd ? H) /2 by nmk/ | #z' @pe_int_trans ] ] ] qed. definition step_none: ∀program.∀n. (\snd (pi1 ?? (jump_expansion_internal program n))) = None ? → ∀k.(\snd (pi1 ?? (jump_expansion_internal program (n+k)))) = None ?. #program #n lapply (refl ? (jump_expansion_internal program n)) cases (jump_expansion_internal program n) in ⊢ (???% → %); #x1 cases x1 #p1 #j1 -x1; #H1 #Heqj #Hj #k elim k [ Heqj @Hj | #k' -k Heqj2 normalize nodelta @refl | #x #H #Heqj2 #abs destruct (abs) ] ] qed. lemma pe_step: ∀program:(Σl:list labelled_instruction.|l| < 2^16). ∀n.policy_equal program (\snd (pi1 ?? (jump_expansion_internal program n))) (\snd (pi1 ?? (jump_expansion_internal program (S n)))) → policy_equal program (\snd (pi1 ?? (jump_expansion_internal program (S n)))) (\snd (pi1 ?? (jump_expansion_internal program (S (S n))))). #program #n #Heq cases daemon (* XXX *) qed. (* this is in the stdlib, but commented out, why? *) theorem plus_Sn_m1: ∀n,m:nat. S m + n = m + S n. #n (elim n) normalize /2 by S_pred/ qed. lemma equal_remains_equal: ∀program:(Σl:list labelled_instruction.|l| < 2^16).∀n:ℕ. policy_equal program (\snd (pi1 … (jump_expansion_internal program n))) (\snd (pi1 … (jump_expansion_internal program (S n)))) → ∀k.k ≥ n → policy_equal program (\snd (pi1 … (jump_expansion_internal program n))) (\snd (pi1 … (jump_expansion_internal program k))). #program #n #Heq #k #Hk elim (le_plus_k … Hk); #z #H >H -H -Hk -k; lapply Heq -Heq; lapply n -n; elim z -z; [ #n #Heq associative_plus >(commutative_plus (S d) ?) measure_plus (commutative_plus 2 ?) @le_plus [1,3: @Hind |2,4: / by le_n/ ] ] ] qed. lemma measure_incr_or_equal: ∀program:Σl:list labelled_instruction.|l|<2^16. ∀policy:Σp:jump_expansion_policy. out_of_program_none program p ∧ jump_in_policy program p ∧ \fst p < 2^16 ∧ policy_isize_sum program p. ∀l.|l| ≤ |program| → ∀acc:ℕ. match \snd (jump_expansion_step program (create_label_map program policy) policy) with [ None ⇒ True | Some p ⇒ measure_int l policy acc ≤ measure_int l p acc ]. #program #policy #l elim l -l; [ #Hp #acc cases (jump_expansion_step ???) #pi1 cases pi1 #p #q -pi1; cases q [ // | #x #_ @le_n ] | #h #t #Hind #Hp #acc lapply (refl ? (jump_expansion_step program (create_label_map program policy) policy)) cases (jump_expansion_step ???) in ⊢ (???% → %); #pi1 cases pi1 -pi1 #c #r cases r [ / by I/ | #x normalize nodelta #Hx #Hjeq lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hx)))) (|t|) Hp) whd in match (measure_int ???); whd in match (measure_int ? x ?); cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) (\snd policy) 〈0,0,short_jump〉) #z cases z -z #x1 #x2 #x3 cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) (\snd x) 〈0,0,short_jump〉) #z cases z -z #y1 #y2 #y3 normalize nodelta #Hj cases Hj [ cases x3 cases y3 [1,4,5,7,8,9: #H @⊥ @H |2,3,6: #_ normalize nodelta [1,2: @(transitive_le ? (measure_int t x acc)) |3: @(transitive_le ? (measure_int t x (acc+1))) ] [2,4,5,6: >measure_plus [1,2: @le_plus_n_r] >measure_plus @le_plus / by le_n/] >Hjeq in Hind; #Hind @Hind @(transitive_le … Hp) @le_n_Sn ] | #Heq >Heq cases y3 normalize nodelta [2,3: >measure_plus >measure_plus @le_plus / by le_n/] >Hjeq in Hind; #Hind @Hind @(transitive_le … Hp) @le_n_Sn ] ] ] qed. (* these lemmas seem superfluous, but not sure how *) lemma bla: ∀a,b:ℕ.a + a = b + b → a = b. #a elim a [ normalize #b // | -a #a #Hind #b cases b [ /2 by le_n_O_to_eq/ | -b #b normalize (Hind b (injective_S ?? (injective_S ?? H))) // ] ] qed. lemma sth_not_s: ∀x.x ≠ S x. #x cases x [ // | #y // ] qed. lemma measure_full: ∀program.∀policy. measure_int program policy 0 = 2*|program| → ∀i.i<|program| → (\snd (bvt_lookup ?? (bitvector_of_nat ? i) (\snd policy) 〈0,0,short_jump〉)) = long_jump. #program #policy elim program [ #Hm #i #Hi @⊥ @(absurd … Hi) @not_le_Sn_O | #h #t #Hind #Hm #i #Hi cut (measure_int t policy 0 = 2*|t|) [ whd in match (measure_int ???) in Hm; cases (\snd (bvt_lookup … (bitvector_of_nat ? (|t|)) (\snd policy) 〈0,0,short_jump〉)) in Hm; normalize nodelta [ #H @⊥ @(absurd ? (measure_le t policy)) >H @lt_to_not_le /2 by lt_plus, le_n/ | >measure_plus >commutative_plus #H @⊥ @(absurd ? (measure_le t policy)) <(plus_to_minus … (sym_eq … H)) @lt_to_not_le normalize >(commutative_plus (|t|) 0) plus_n_Sm @le_n | >measure_plus commutative_plus #H lapply (injective_plus_r … H) // ] | #Hmt cases (le_to_or_lt_eq … Hi) -Hi; [ #Hi @(Hind Hmt i (le_S_S_to_le … Hi)) | #Hi >(injective_S … Hi) whd in match (measure_int ???) in Hm; cases (\snd (bvt_lookup … (bitvector_of_nat ? (|t|)) (\snd policy) 〈0,0,short_jump〉)) in Hm; normalize nodelta [ >Hmt normalize (commutative_plus (|t|) (S (|t|))) >plus_n_Sm #H @⊥ @(absurd ? (bla ?? H)) @sth_not_s | >measure_plus >Hmt normalize commutative_plus normalize #H @⊥ @(absurd ? (injective_plus_r … (injective_S ?? H))) @sth_not_s | #_ // ] ]] ] qed. lemma measure_special: ∀program:(Σl:list labelled_instruction.|l| < 2^16). ∀policy:Σp:jump_expansion_policy. out_of_program_none program p ∧ jump_in_policy program p ∧ \fst p < 2^16 ∧ policy_isize_sum program p. match (\snd (pi1 ?? (jump_expansion_step program (create_label_map program policy) policy))) with [ None ⇒ True | Some p ⇒ measure_int program policy 0 = measure_int program p 0 → policy_equal_int program policy p ]. #program #policy lapply (refl ? (pi1 ?? (jump_expansion_step program (create_label_map program policy) policy))) cases (jump_expansion_step program (create_label_map program policy) policy) in ⊢ (???% → %); #p cases p -p #ch #pol normalize nodelta cases pol [ / by I/ | #p normalize nodelta #Hpol #eqpol lapply (le_n (|program|)) @(list_ind ? (λx.|x| ≤ |pi1 ?? program| → measure_int x policy 0 = measure_int x p 0 → policy_equal_int x policy p) ?? (pi1 ?? program)) [ #_ #_ #i #Hi @⊥ @(absurd ? Hi) @not_le_Sn_O | #h #t #Hind #Hp #Hm #i #Hi cases (le_to_or_lt_eq … Hi) -Hi; [ #Hi @Hind [ @(transitive_le … Hp) / by / | whd in match (measure_int ???) in Hm; whd in match (measure_int ? p ?) in Hm; lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpol)))) (|t|) Hp) cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) ? 〈0,0,short_jump〉) in Hm; #x cases x -x #x1 #x2 #x3 cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) ? 〈0,0,short_jump〉); #y cases y -y #y1 #y2 #y3 cases x3 cases y3 normalize nodelta [1: #H #_ @H |2,3: >measure_plus #H #_ @⊥ @(absurd ? (eq_plus_S_to_lt … H)) @le_to_not_lt lapply (measure_incr_or_equal program policy t ? 0) [1,3: @(transitive_le … Hp) @le_n_Sn ] >eqpol / by / |4,7,8: #_ #H elim H #H2 [1,3,5: @⊥ @H2 |2,4,6: destruct (H2) ] |5: >measure_plus >measure_plus >commutative_plus >(commutative_plus ? 1) #H #_ @(injective_plus_r … H) |6: >measure_plus >measure_plus change with (1+1) in match (2); >assoc_plus1 >(commutative_plus 1 (measure_int ???)) #H #_ @⊥ @(absurd ? (eq_plus_S_to_lt … H)) @le_to_not_lt @monotonic_le_plus_l lapply (measure_incr_or_equal program policy t ? 0) [ @(transitive_le … Hp) @le_n_Sn ] >eqpol / by / |9: >measure_plus >measure_plus >commutative_plus >(commutative_plus ? 2) #H #_ @(injective_plus_r … H) ] | @(le_S_S_to_le … Hi) ] | #Hi >(injective_S … Hi) whd in match (measure_int ???) in Hm; whd in match (measure_int ? p ?) in Hm; lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hpol)))) (|t|) Hp) cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) ? 〈0,0,short_jump〉) in Hm; #x cases x -x #x1 #x2 #x3 cases (bvt_lookup ?? (bitvector_of_nat ? (|t|)) ? 〈0,0,short_jump〉); #y cases y -y #y1 #y2 #y3 normalize nodelta cases x3 cases y3 normalize nodelta [1,5,9: #_ #_ // |4,7,8: #_ #H elim H #H2 [1,3,5: @⊥ @H2 |2,4,6: destruct (H2) ] |2,3: >measure_plus #H #_ @⊥ @(absurd ? (eq_plus_S_to_lt … H)) @le_to_not_lt lapply (measure_incr_or_equal program policy t ? 0) [1,3: @(transitive_le … Hp) @le_n_Sn ] >eqpol / by / |6: >measure_plus >measure_plus change with (1+1) in match (2); >assoc_plus1 >(commutative_plus 1 (measure_int ???)) #H #_ @⊥ @(absurd ? (eq_plus_S_to_lt … H)) @le_to_not_lt @monotonic_le_plus_l lapply (measure_incr_or_equal program policy t ? 0) [ @(transitive_le … Hp) @le_n_Sn ] >eqpol / by / ] ] ] qed. lemma le_to_eq_plus: ∀n,z. n ≤ z → ∃k.z = n + k. #n #z elim z [ #H cases (le_to_or_lt_eq … H) [ #H2 @⊥ @(absurd … H2) @not_le_Sn_O | #H2 @(ex_intro … 0) >H2 // ] | #z' #Hind #H cases (le_to_or_lt_eq … H) [ #H' elim (Hind (le_S_S_to_le … H')) #k' #H2 @(ex_intro … (S k')) >H2 >plus_n_Sm // | #H' @(ex_intro … 0) >H' // ] ] qed. lemma measure_zero: ∀l.∀program:Σl:list labelled_instruction.|l| < 2^16. |l| ≤ |program| → measure_int l (jump_expansion_start program) 0 = 0. #l #program elim l [ // | #h #t #Hind #Hp whd in match (measure_int ???); cases (dec_is_jump (nth (|t|) ? program 〈None ?, Comment []〉)) #Hj [ >(lookup_opt_lookup_hit … (proj2 ?? (pi2 ?? (jump_expansion_start program)) (|t|) ? Hj) 〈0,0,short_jump〉) [ normalize nodelta @Hind @le_S_to_le ] @Hp | >(lookup_opt_lookup_miss … (proj1 ?? (jump_not_in_policy program (pi1 ?? (jump_expansion_start program)) (|t|) ?) Hj) 〈0,0,short_jump〉) [ normalize nodelta @Hind @le_S_to_le @Hp | @Hp | % [ @(proj1 ?? (proj1 ?? (pi2 ?? (jump_expansion_start program)))) | @(proj2 ?? (proj1 ?? (pi2 ?? (jump_expansion_start program)))) ] ] ] ] qed. (* the actual computation of the fixpoint *) definition je_fixpoint: ∀program:(Σl:list labelled_instruction.|l| < 2^16). Σp:option jump_expansion_policy.∃n.∀k.n < k → policy_equal program (\snd (pi1 ?? (jump_expansion_internal program k))) p. #program @(\snd (pi1 ?? (jump_expansion_internal program (2*|program|)))) cases (dec_bounded_exists (λk.policy_equal (pi1 ?? program) (\snd (pi1 ?? (jump_expansion_internal program k))) (\snd (pi1 ?? (jump_expansion_internal program (S k))))) ? (2*|program|)) [ #Hex elim Hex -Hex #x #Hx @(ex_intro … x) #k #Hk @pe_trans [ @(\snd (pi1 ?? (jump_expansion_internal program x))) | @pe_sym @equal_remains_equal [ @(proj2 ?? Hx) | @le_S_S_to_le @le_S @Hk ] | @equal_remains_equal [ @(proj2 ?? Hx) | @le_S_S_to_le @le_S @(proj1 ?? Hx) ] ] | #Hnex lapply (not_exists_forall … Hnex) -Hnex; #Hfa @(ex_intro … (2*|program|)) #k #Hk @pe_sym @equal_remains_equal [ lapply (refl ? (jump_expansion_internal program (2*|program|))) cases (jump_expansion_internal program (2*|program|)) in ⊢ (???% → %); #x cases x -x #Fch #Fpol normalize nodelta #HFpol cases Fpol in HFpol; normalize nodelta [ (* if we're at None in 2*|program|, we're at None in S 2*|program| too *) #HFpol #EQ whd in match (jump_expansion_internal ??); >EQ normalize nodelta / by / | #Fp #HFp #EQ whd in match (jump_expansion_internal ??); >EQ normalize nodelta lapply (refl ? (jump_expansion_step program (create_label_map program Fp) «Fp,?»)) [ @HFp | lapply (measure_full program Fp ?) [ @le_to_le_to_eq [ @measure_le | cut (∀x:ℕ.x ≤ 2*|program| → ∃p.(\snd (pi1 ?? (jump_expansion_internal program x)) = Some ? p ∧ x ≤ measure_int program p 0)) [ #x elim x [ #Hx @(ex_intro ?? (jump_expansion_start program)) @conj [ whd in match (jump_expansion_internal ??); @refl | @le_O_n ] | -x #x #Hind #Hx lapply (refl ? (jump_expansion_internal program (S x))) cases (jump_expansion_internal program (S x)) in ⊢ (???% → %); #z cases z -z #Sxch #Sxpol cases Sxpol -Sxpol normalize nodelta [ #H #HeqSxpol @⊥ elim (le_to_eq_plus ?? Hx) #k #Hk @(absurd … (step_none program (S x) ? k)) [ >HeqSxpol / by / | EQ @nmk #H destruct (H) ] | #Sxpol #HSxpol #HeqSxpol @(ex_intro ?? Sxpol) @conj [ @refl | elim (Hind (transitive_le … (le_n_Sn x) Hx)) #xpol #Hxpol @(le_to_lt_to_lt … (proj2 ?? Hxpol)) lapply (measure_incr_or_equal program xpol program (le_n (|program|)) 0) [ cases (jump_expansion_internal program x) in Hxpol; #z cases z -z #xch #xpol normalize nodelta #H #H2 >(proj1 ?? H2) in H; normalize nodelta / by / | lapply (Hfa x Hx) lapply HeqSxpol -HeqSxpol whd in match (jump_expansion_internal program (S x)); lapply (refl ? (jump_expansion_internal program x)) lapply Hxpol -Hxpol cases (jump_expansion_internal program x) in ⊢ (% → ???% → %); #z cases z -z #xch #b normalize nodelta #H #Heq >(proj1 ?? Heq) in H; #H #Heq cases xch in Heq; #Heq normalize nodelta [ lapply (refl ? (jump_expansion_step program (create_label_map (pi1 ?? program) xpol) «xpol,H»)) cases (jump_expansion_step ???) in ⊢ (???% → %); #z cases z -z #a #c normalize nodelta cases c normalize nodelta [ #H1 #Heq #H2 destruct (H2) | #d #H1 #Heq #H2 destruct (H2) #Hfull #H2 elim (le_to_or_lt_eq … H2) [ / by / | #H3 lapply (measure_special program «xpol,H») >Heq normalize nodelta #H4 @⊥ @(absurd … (H4 H3)) @Hfull ] ] | lapply (refl ? (jump_expansion_step program (create_label_map (pi1 ?? program) xpol) «xpol,H»)) cases (jump_expansion_step ???) in ⊢ (???% → %); #z cases z -z #a #c normalize nodelta cases c normalize nodelta [ #H1 #Heq #H2 #H3 #_ @⊥ @(absurd ?? H3) @pe_refl | #d #H1 #Heq #H2 #H3 @⊥ @(absurd ?? H3) @pe_refl ] ] ] ] ] ] | #H elim (H (2*|program|) (le_n ?)) #plp >EQ #Hplp >(Some_eq ??? (proj1 ?? Hplp)) @(proj2 ?? Hplp) ] ] | #Hfull cases (jump_expansion_step program (create_label_map program Fp) «Fp,?») in ⊢ (???% → %); #x cases x -x #Gch #Gpol cases Gpol normalize nodelta [ #H #EQ2 @⊥ @(absurd ?? H) @Hfull | #Gp #HGp #EQ2 cases Fch [ normalize nodelta #i #Hi lapply (refl ? (lookup ?? (bitvector_of_nat 16 i) (\snd Fp) 〈0,0,short_jump〉)) cases (lookup ?? (bitvector_of_nat 16 i) (\snd Fp) 〈0,0,short_jump〉) in ⊢ (???% → %); #x cases x -x #p #a #j normalize nodelta #H lapply (proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? HGp)))) i Hi) lapply (Hfull i Hi) >H #H2 >H2 normalize nodelta cases (lookup ?? (bitvector_of_nat 16 i) (\snd Gp) 〈0,0,short_jump〉) #x cases x -x #p #a #j' cases j' normalize nodelta #H elim H -H #H [1,3: @⊥ @H |2,4: destruct (H) |5,6: @refl ] | normalize nodelta /2 by pe_int_refl/ ] ] ] ] ] | @le_S_S_to_le @le_S @Hk ] | #n cases (jump_expansion_internal program n) cases (jump_expansion_internal program (S n)) #x cases x -x #nch #npol normalize nodelta #Hnpol #x cases x -x #Sch #Spol normalize nodelta #HSpol cases npol in Hnpol; cases Spol in HSpol; [ #Hnpol #HSpol %1 // |2,3: #x #Hnpol #HSpol %2 @nmk whd in match (policy_equal ???); // #H destruct (H) |4: #np #Hnp #Sp #HSp whd in match (policy_equal ???); @dec_bounded_forall #m cases (bvt_lookup ?? (bitvector_of_nat 16 m) ? 〈0,0,short_jump〉) #x cases x -x #x1 #x2 #x3 cases (bvt_lookup ?? (bitvector_of_nat ? m) ? 〈0,0,short_jump〉) #y cases y -y #y1 #y2 #y3 normalize nodelta @dec_eq_jump_length ] ] qed. (* Take a policy of 〈pc, addr, jump_length〉 tuples, and transform it into * a map from pc to jump_length. This cannot be done earlier because the pc * changes between iterations. *) let rec transform_policy (n: nat) (policy: jump_expansion_policy) (acc: BitVectorTrie jump_length 16) on n: BitVectorTrie jump_length 16 ≝ match n with [ O ⇒ acc | S n' ⇒ match lookup_opt … (bitvector_of_nat 16 n') (\snd policy) with [ None ⇒ transform_policy n' policy acc | Some x ⇒ let 〈pc,addr,length〉 ≝ x in transform_policy n' policy (insert … (bitvector_of_nat 16 pc) length acc) ] ]. (* The glue between Policy and Assembly. *) definition jump_expansion': ∀program:preamble × (Σl:list labelled_instruction.|l| < 2^16). ∀lookup_labels.policy_type lookup_labels ≝ λprogram.λlookup_labels.λpc. let policy ≝ transform_policy (|\snd program|) (pi1 … (je_fixpoint (\snd program))) (Stub ??) in «bvt_lookup ? ? pc policy long_jump,?». cases daemon (* XXX later *) qed.