include "ASM/ASM.ma". include "ASM/Arithmetic.ma". include "ASM/Fetch.ma". include "ASM/Status.ma". include "utilities/extralib.ma". include "ASM/Assembly.ma". (* Internal types *) (* ppc_pc_map: program length × (pseudo program counter ↦ 〈pc, jump_length〉) *) definition ppc_pc_map ≝ ℕ × (BitVectorTrie (ℕ × jump_length) 16). (* The different properties that we want/need to prove at some point *) (* During our iteration, everything not yet seen is None, and vice versa *) definition out_of_program_none ≝ λprefix:list labelled_instruction.λsigma:ppc_pc_map. ∀i.i < 2^16 → (i > |prefix| ↔ bvt_lookup_opt … (bitvector_of_nat ? i) (\snd sigma) = None ?). definition not_jump_default ≝ λprefix:list labelled_instruction.λsigma:ppc_pc_map. ∀i:ℕ.i < |prefix| → ¬is_jump (\snd (nth i ? prefix 〈None ?, Comment []〉)) → \snd (bvt_lookup … (bitvector_of_nat ? i) (\snd sigma) 〈0,short_jump〉) = short_jump. (* Between two policies, jumps cannot decrease *) definition jmpeqb: jump_length → jump_length → bool ≝ λj1.λj2. match j1 with [ short_jump ⇒ match j2 with [ short_jump ⇒ true | _ ⇒ false ] | absolute_jump ⇒ match j2 with [ absolute_jump ⇒ true | _ ⇒ false ] | long_jump ⇒ match j2 with [ long_jump ⇒ true | _ ⇒ false ] ]. lemma jmpeqb_to_eq: ∀j1,j2.jmpeqb j1 j2 → j1 = j2. #j1 #j2 cases j1 cases j2 [1,5,9: / by /] #H cases H qed. definition jmple: jump_length → jump_length → Prop ≝ λj1.λj2. match j1 with [ short_jump ⇒ match j2 with [ short_jump ⇒ False | _ ⇒ True ] | absolute_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 jump_increase ≝ λprefix:list labelled_instruction.λop:ppc_pc_map.λp:ppc_pc_map. ∀i.i ≤ |prefix| → let 〈opc,oj〉 ≝ bvt_lookup … (bitvector_of_nat ? i) (\snd op) 〈0,short_jump〉 in let 〈pc,j〉 ≝ bvt_lookup … (bitvector_of_nat ? i) (\snd p) 〈0,short_jump〉 in jmpleq oj j. (* this is the instruction size as determined by the jump length given *) definition expand_relative_jump_internal_unsafe: jump_length → ([[relative]] → preinstruction [[relative]]) → list instruction ≝ λjmp_len:jump_length.λi. match jmp_len with [ short_jump ⇒ [ RealInstruction (i (RELATIVE (zero 8))) ] | absolute_jump ⇒ [ ] (* this should not happen *) | long_jump ⇒ [ RealInstruction (i (RELATIVE (bitvector_of_nat ? 2))); SJMP (RELATIVE (bitvector_of_nat ? 3)); (* LJMP size? *) LJMP (ADDR16 (zero 16)) ] ]. @I qed. definition strip_target: preinstruction Identifier → ([[relative]] → preinstruction [[relative]]) ⊎ instruction ≝ λi. match i with [ JC _ ⇒ inl … (JC ?) | JNC _ ⇒ inl … (JNC ?) | JB baddr _ ⇒ inl … (JB ? baddr) | JZ _ ⇒ inl … (JZ ?) | JNZ _ ⇒ inl … (JNZ ?) | JBC baddr _ ⇒ inl … (JBC ? baddr) | JNB baddr _ ⇒ inl … (JNB ? baddr) | CJNE addr _ ⇒ inl … (CJNE ? addr) | DJNZ addr _ ⇒ inl … (DJNZ ? addr) | ADD arg1 arg2 ⇒ inr … (ADD ? arg1 arg2) | ADDC arg1 arg2 ⇒ inr … (ADDC ? arg1 arg2) | SUBB arg1 arg2 ⇒ inr … (SUBB ? arg1 arg2) | INC arg ⇒ inr … (INC ? arg) | DEC arg ⇒ inr … (DEC ? arg) | MUL arg1 arg2 ⇒ inr … (MUL ? arg1 arg2) | DIV arg1 arg2 ⇒ inr … (DIV ? arg1 arg2) | DA arg ⇒ inr … (DA ? arg) | ANL arg ⇒ inr … (ANL ? arg) | ORL arg ⇒ inr … (ORL ? arg) | XRL arg ⇒ inr … (XRL ? arg) | CLR arg ⇒ inr … (CLR ? arg) | CPL arg ⇒ inr … (CPL ? arg) | RL arg ⇒ inr … (RL ? arg) | RR arg ⇒ inr … (RR ? arg) | RLC arg ⇒ inr … (RLC ? arg) | RRC arg ⇒ inr … (RRC ? arg) | SWAP arg ⇒ inr … (SWAP ? arg) | MOV arg ⇒ inr … (MOV ? arg) | MOVX arg ⇒ inr … (MOVX ? arg) | SETB arg ⇒ inr … (SETB ? arg) | PUSH arg ⇒ inr … (PUSH ? arg) | POP arg ⇒ inr … (POP ? arg) | XCH arg1 arg2 ⇒ inr … (XCH ? arg1 arg2) | XCHD arg1 arg2 ⇒ inr … (XCHD ? arg1 arg2) | RET ⇒ inr … (RET ?) | RETI ⇒ inr … (RETI ?) | NOP ⇒ inr … (RealInstruction (NOP ?)) ]. definition expand_relative_jump_unsafe: jump_length → preinstruction Identifier → list instruction ≝ λjmp_len:jump_length.λi. match strip_target i with [ inl jmp ⇒ expand_relative_jump_internal_unsafe jmp_len jmp | inr instr ⇒ [ instr ] ]. definition expand_pseudo_instruction_unsafe: jump_length → pseudo_instruction → list instruction ≝ λjmp_len. λi. match i with [ Cost cost ⇒ [ ] | Comment comment ⇒ [ ] | Call call ⇒ match jmp_len with [ short_jump ⇒ [ ] (* this should not happen *) | absolute_jump ⇒ [ ACALL (ADDR11 (zero 11)) ] | long_jump ⇒ [ LCALL (ADDR16 (zero 16)) ] ] | Mov d trgt ⇒ [ 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 ⇒ [ SJMP (RELATIVE (zero 8)) ] | absolute_jump ⇒ [ AJMP (ADDR11 (zero 11)) ] | long_jump ⇒ [ LJMP (ADDR16 (zero 16)) ] ] ]. % qed. definition instruction_size_jmplen: jump_length → pseudo_instruction → ℕ ≝ λjmp_len. λi. let mapped ≝ map ? ? assembly1 (expand_pseudo_instruction_unsafe jmp_len i) in let flattened ≝ flatten ? mapped in let pc_len ≝ length ? flattened in pc_len. definition sigma_compact_unsafe ≝ λprefix:list labelled_instruction.λlabels:label_map.λsigma:ppc_pc_map. ∀n.n < |prefix| → match bvt_lookup_opt … (bitvector_of_nat ? n) (\snd sigma) with [ None ⇒ False | Some x ⇒ let 〈pc,j〉 ≝ x in match bvt_lookup_opt … (bitvector_of_nat ? (S n)) (\snd sigma) with [ None ⇒ False | Some x1 ⇒ let 〈pc1,j1〉 ≝ x1 in pc1 = pc + instruction_size_jmplen j (\snd (nth n ? prefix 〈None ?, Comment []〉)) ] ]. (* new safety condition: sigma corresponds to program and resulting program is compact *) definition sigma_compact ≝ λprogram:list labelled_instruction.λlabels:label_map.λsigma:ppc_pc_map. ∀datalabels.∀n.n < |program| → match bvt_lookup_opt … (bitvector_of_nat ? n) (\snd sigma) with [ None ⇒ False | Some x ⇒ let 〈pc,j〉 ≝ x in match bvt_lookup_opt … (bitvector_of_nat ? (S n)) (\snd sigma) with [ None ⇒ False | Some x1 ⇒ let 〈pc1,j1〉 ≝ x1 in pc1 = pc + instruction_size (λid.bitvector_of_nat ? (lookup_def ?? labels id 0)) datalabels (λppc.bitvector_of_nat ? (\fst (bvt_lookup ?? ppc (\snd sigma) 〈0,short_jump〉))) (λppc.jmpeqb long_jump (\snd (bvt_lookup ?? ppc (\snd sigma) 〈0,short_jump〉))) (bitvector_of_nat ? n) (\snd (nth n ? program 〈None ?, Comment []〉)) ] ]. (* jumps are of the proper size *) definition sigma_safe ≝ λprefix:list labelled_instruction.λlabels:label_map. λold_sigma:ppc_pc_map.λsigma:ppc_pc_map. ∀i.i < |prefix| → let 〈label,instr〉 ≝ nth i ? prefix 〈None ?, Comment [ ]〉 in ∀dest.is_jump_to instr dest → let paddr ≝ lookup_def … labels dest 0 in let 〈j,src,dest〉 ≝ if leb paddr i then (* backward jump *) let 〈pc,j〉 ≝ bvt_lookup … (bitvector_of_nat ? i) (\snd sigma) 〈0,short_jump〉 in let pc_plus_jmp_length ≝ bitvector_of_nat ? (\fst (bvt_lookup … (bitvector_of_nat ? (S i)) (\snd sigma) 〈0,short_jump〉)) in let addr ≝ bitvector_of_nat ? (\fst (bvt_lookup … (bitvector_of_nat ? paddr) (\snd sigma) 〈0,short_jump〉)) in 〈j,pc_plus_jmp_length,addr〉 else let 〈pc,oj〉 ≝ bvt_lookup … (bitvector_of_nat ? i) (\snd old_sigma) 〈0,short_jump〉 in let 〈npc,j〉 ≝ bvt_lookup … (bitvector_of_nat ? i) (\snd sigma) 〈0,short_jump〉 in let pc_plus_jmp_length ≝ bitvector_of_nat ? (\fst (bvt_lookup … (bitvector_of_nat ? (S i)) (\snd old_sigma) 〈0,short_jump〉)) in let addr ≝ bitvector_of_nat ? (\fst (bvt_lookup … (bitvector_of_nat ? paddr) (\snd old_sigma) 〈0,short_jump〉)) in 〈j,pc_plus_jmp_length,addr〉 in match j with [ short_jump ⇒ \fst (short_jump_cond src dest) = true | absolute_jump ⇒ \fst (absolute_jump_cond src dest) = true (*∧ \fst (short_jump_cond src dest) = false*) | long_jump ⇒ True (* do not talk about long jump *) ]. definition sigma_jumps ≝ λprefix:list labelled_instruction.λsigma:ppc_pc_map. ∀i.i < |prefix| → let 〈label,instr〉 ≝ nth i ? prefix 〈None ?, Comment []〉 in (\snd (bvt_lookup … (bitvector_of_nat ? i) (\snd sigma) 〈0,short_jump〉) = short_jump → bool_to_Prop (¬is_call instr)) ∧ (\snd (bvt_lookup … (bitvector_of_nat ? i) (\snd sigma) 〈0,short_jump〉) = absolute_jump → bool_to_Prop (¬is_relative_jump instr)). (* 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 | absolute_jump ⇒ match j2 with [ absolute_jump ⇒ absolute_jump | _ ⇒ long_jump ] | short_jump ⇒ match j2 with [ short_jump ⇒ short_jump | _ ⇒ long_jump ] ]. lemma dec_jmple: ∀x,y:jump_length.Sum (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.Sum (a = b) (a ≠ b). #a #b cases a cases b /2/ %2 @nmk #H destruct (H) qed. (* The function that creates the label-to-address map *) definition create_label_map: ∀program:list labelled_instruction. (Σlabels:label_map. ∀l.occurs_exactly_once ?? l program → (*And (bitvector_of_nat ? (lookup_def ?? labels l 0) = address_of_word_labels program l) ( *)lookup_def ?? labels l 0 < |program|(*)*) ) ≝ λprogram. \fst (create_label_cost_map program). #l #Hl cases (create_label_cost_map_ok program) #_ #X @(X … Hl) qed. (* General note on jump length selection: the jump displacement is added/replaced * AFTER the fetch (and attendant PC increase), but we calculate before the * fetch, which means that we are [jump_length] bytes off and have to compensate. *) definition select_reljump_length: label_map → ppc_pc_map → ppc_pc_map → ℕ → Identifier → ℕ → jump_length ≝ λlabels.λold_sigma.λinc_sigma.λppc.λlbl.λins_len. let paddr ≝ lookup_def ?? labels lbl 0 in let 〈src,dest〉 ≝ if leb paddr ppc then (* backward jump *) let pc ≝ \fst inc_sigma in let addr ≝ \fst (bvt_lookup … (bitvector_of_nat ? paddr) (\snd inc_sigma) 〈0,short_jump〉) in 〈bitvector_of_nat ? (pc+ins_len), bitvector_of_nat ? addr〉 else let pc ≝ \fst (bvt_lookup … (bitvector_of_nat ? ppc) (\snd old_sigma) 〈0,short_jump〉) in let addr ≝ \fst (bvt_lookup … (bitvector_of_nat ? paddr) (\snd old_sigma) 〈0,short_jump〉) in 〈bitvector_of_nat ? (pc+ins_len), bitvector_of_nat ? addr〉 in let 〈sj_possible, disp〉 ≝ short_jump_cond src dest in if sj_possible then short_jump else long_jump. definition select_call_length: label_map → ppc_pc_map → ppc_pc_map → ℕ → Identifier → jump_length ≝ λlabels.λold_sigma.λinc_sigma.λppc.λlbl. let paddr ≝ lookup_def ?? labels lbl 0 in let 〈src,dest〉 ≝ if leb paddr ppc then (* backward jump *) let pc ≝ \fst inc_sigma in let addr ≝ \fst (bvt_lookup … (bitvector_of_nat ? paddr) (\snd inc_sigma) 〈0,short_jump〉) in 〈bitvector_of_nat ? (pc+2), bitvector_of_nat ? addr〉 else let pc ≝ \fst (bvt_lookup … (bitvector_of_nat ? ppc) (\snd old_sigma) 〈0,short_jump〉) in let addr ≝ \fst (bvt_lookup … (bitvector_of_nat ? paddr) (\snd old_sigma) 〈0,short_jump〉) in 〈bitvector_of_nat ? (pc+2), bitvector_of_nat ? addr〉 in let 〈aj_possible, disp〉 ≝ absolute_jump_cond src dest in if aj_possible then absolute_jump else long_jump. definition select_jump_length: label_map → ppc_pc_map → ppc_pc_map → ℕ → Identifier → jump_length ≝ λlabels.λold_sigma.λinc_sigma.λppc.λlbl. let paddr ≝ lookup_def ?? labels lbl 0 in let 〈src,dest〉 ≝ if leb paddr ppc then (* backward jump *) let pc ≝ \fst inc_sigma in let addr ≝ \fst (bvt_lookup … (bitvector_of_nat ? paddr) (\snd inc_sigma) 〈0,short_jump〉) in 〈bitvector_of_nat ? (pc+2), bitvector_of_nat ? addr〉 else let pc ≝ \fst (bvt_lookup … (bitvector_of_nat ? ppc) (\snd old_sigma) 〈0,short_jump〉) in let addr ≝ \fst (bvt_lookup … (bitvector_of_nat ? paddr) (\snd old_sigma) 〈0,short_jump〉) in 〈bitvector_of_nat ? (pc+2), bitvector_of_nat ? addr〉 in let 〈sj_possible, disp〉 ≝ short_jump_cond src dest in if sj_possible then short_jump else select_call_length labels old_sigma inc_sigma ppc lbl. definition destination_of: preinstruction Identifier → option Identifier ≝ λi. match i with [ JC j ⇒ Some ? j | JNC j ⇒ Some ? j | JZ j ⇒ Some ? j | JNZ j ⇒ Some ? j | JB _ j ⇒ Some ? j | JBC _ j ⇒ Some ? j | JNB _ j ⇒ Some ? j | CJNE _ j ⇒ Some ? j | DJNZ _ j ⇒ Some ? j | _ ⇒ None ? ]. definition length_of: preinstruction Identifier → ℕ ≝ λi. match i with [ JC j ⇒ 2 | JNC j ⇒ 2 | JZ j ⇒ 2 | JNZ j ⇒ 2 | JB _ j ⇒ 3 | JBC _ j ⇒ 3 | JNB _ j ⇒ 3 | CJNE _ j ⇒ 3 | DJNZ x j ⇒ match x with [ REGISTER _ ⇒ 2 | _ ⇒ 3 ] | _ ⇒ 0 ]. definition jump_expansion_step_instruction: label_map → ppc_pc_map → ppc_pc_map → ℕ → preinstruction Identifier → option jump_length ≝ λlabels.λold_sigma.λinc_sigma.λppc.λi. let ins_len ≝ length_of i in match destination_of i with [ Some j ⇒ Some ? (select_reljump_length labels old_sigma inc_sigma ppc j ins_len) | None ⇒ None ? ]. (* The first step of the jump expansion: everything to short. *) definition jump_expansion_start: ∀program:(Σl:list labelled_instruction.S (|l|) < 2^16 ∧ is_well_labelled_p l). ∀labels:label_map. Σpolicy:option ppc_pc_map. match policy with [ None ⇒ True | Some p ⇒ And (And (And (And (And (not_jump_default (pi1 ?? program) p) (\fst (bvt_lookup … (bitvector_of_nat ? 0) (\snd p) 〈0,short_jump〉) = 0)) (\fst p = \fst (bvt_lookup … (bitvector_of_nat ? (|program|)) (\snd p) 〈0,short_jump〉))) (sigma_compact_unsafe program labels p)) (∀i.i ≤ |program| → ∃pc. bvt_lookup_opt … (bitvector_of_nat ? i) (\snd p) = Some ? 〈pc,short_jump〉)) (\fst p ≤ 2^16) ] ≝ λprogram.λlabels. let final_policy ≝ foldl_strong (option Identifier × pseudo_instruction) (λprefix.Σpolicy:ppc_pc_map.And (And (And (And (not_jump_default prefix policy) (\fst (bvt_lookup … (bitvector_of_nat ? 0) (\snd policy) 〈0,short_jump〉) = 0)) (\fst policy = \fst (bvt_lookup … (bitvector_of_nat ? (|prefix|)) (\snd policy) 〈0,short_jump〉))) (sigma_compact_unsafe prefix labels policy)) (∀i.i ≤ |prefix| → ∃pc. bvt_lookup_opt … (bitvector_of_nat ? i) (\snd policy) = Some ? 〈pc,short_jump〉)) program (λprefix.λx.λtl.λprf.λp. let 〈pc,sigma〉 ≝ pi1 ?? p in let 〈label,instr〉 ≝ x in let isize ≝ instruction_size_jmplen short_jump instr in 〈pc + isize, bvt_insert … (bitvector_of_nat 16 (S (|prefix|))) 〈pc+isize,short_jump〉 sigma〉 ) 〈0, bvt_insert ?? (bitvector_of_nat 16 0) 〈0,short_jump〉 (Stub ??)〉 in if gtb (\fst (pi1 ?? final_policy)) 2^16 then None ? else Some ? (pi1 ?? final_policy). [ / by I/ | lapply p -p cases final_policy -final_policy #p #Hp #hg @conj [ @Hp | @not_lt_to_le @ltb_false_to_not_lt @hg ] | @conj [ @conj [ @conj [ @conj [ (* not_jump_default *) cases p -p #p cases p -p #pc #sigma #Hp cases x in prf; #lbl #ins #prf #i >append_length lookup_insert_miss [ (* USE[pass]: not_jump_default *) lapply (proj1 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hp))) i Hi) >nth_append_first [ #H #H2 @H @H2 | @Hi ] | @bitvector_of_nat_abs [ @(transitive_lt ??? Hi) @le_S_to_le] [1,2: @(transitive_lt … (proj1 ?? (pi2 ?? program))) @le_S_S >prf >append_length Hi >lookup_insert_miss [ #_ (* USE: everything is short *) elim ((proj2 ?? Hp) (|prefix|) (le_n (|prefix|))) #pc #Hl >(lookup_opt_lookup_hit … Hl 〈0,short_jump〉) @refl | @bitvector_of_nat_abs [ @(transitive_lt … (proj1 ?? (pi2 ?? program))) >prf @le_S_S >append_length @le_plus_n_r | @(transitive_lt … (proj1 ?? (pi2 ?? program))) >prf @le_S_S >append_length lookup_insert_miss [ (* USE[pass]: 0 ↦ 0 *) @(proj2 ?? (proj1 ?? (proj1 ?? (proj1 ?? Hp)))) | @bitvector_of_nat_abs [ @ltb_true_to_lt / by refl/ | @(transitive_lt … (proj1 ?? (pi2 ?? program))) >prf >append_length @le_S_S append_length >(commutative_plus (|prefix|)) >lookup_insert_hit @refl ] | (* policy_compact_unsafe *) #i >append_length lookup_opt_insert_miss [ >lookup_opt_insert_miss [ (* USE[pass]: policy_compact_unsafe *) lapply (proj2 ?? (proj1 ?? Hp) i Hi) lapply (refl ? (bvt_lookup_opt … (bitvector_of_nat ? i) sigma)) cases (bvt_lookup_opt … (bitvector_of_nat ? i) sigma) in ⊢ (???% → %); [ #_ normalize nodelta / by / | #x cases x -x #pci #ji #EQi lapply (refl ? (bvt_lookup_opt … (bitvector_of_nat ? (S i)) sigma)) cases (bvt_lookup_opt … (bitvector_of_nat ? (S i)) sigma) in ⊢ (???% → %); [ #_ normalize nodelta / by / | #x cases x -x #pcSi #jSi #EQSi normalize nodelta >nth_append_first [ / by / | @Hi ] ] ] ] ] [2: lapply (le_S_to_le … Hi) -Hi #Hi] @bitvector_of_nat_abs [1,4: @(transitive_lt … (proj1 ?? (pi2 ?? program))) >prf @le_S_S >append_length prf @le_S_S >append_length lookup_opt_insert_miss [ >Hi >lookup_opt_insert_hit normalize nodelta (* USE: everything is short, fst p = pc *) elim ((proj2 ?? Hp) (|prefix|) (le_n ?)) #pc #Hl lapply (proj2 ?? (proj1 ?? (proj1 ?? Hp))) >Hl >(lookup_opt_lookup_hit … Hl 〈0,short_jump〉) #EQ normalize nodelta >nth_append_second [ EQ @refl | @le_n ] | @bitvector_of_nat_abs [ @(transitive_lt … (proj1 ?? (pi2 ?? program))) >Hi >prf @le_S_S >append_length prf @le_S_S >append_length Hi @le_n ] ] ] ] | (* everything is short *) #i >append_length lookup_opt_insert_miss [ (* USE[pass]: everything is short *) @((proj2 ?? Hp) i (le_S_S_to_le … Hi)) | @bitvector_of_nat_abs [ @(transitive_lt … (proj1 ?? (pi2 ?? program))) >prf >append_length @le_S_S >commutative_plus @le_plus_a @le_S_S_to_le @Hi | @(transitive_lt … (proj1 ?? (pi2 ?? program))) >prf >append_length Hi >lookup_opt_insert_hit @(ex_intro ?? (pc+instruction_size_jmplen short_jump instr)) @refl ] ] | @conj [ @conj [ @conj [ @conj [ #i cases i [ #Hi @⊥ @(absurd … Hi) @not_le_Sn_O | -i #i #Hi #Hj @⊥ @(absurd … Hi) @not_le_Sn_O ] ] ] >lookup_insert_hit @refl | #i cases i [ #Hi @⊥ @(absurd … Hi) @le_to_not_lt @le_n | -i #i #Hi @⊥ @(absurd … Hi) @not_le_Sn_O ] ] | #i cases i [ #Hi >lookup_opt_insert_hit @(ex_intro ?? 0) @refl | -i #i #Hi @⊥ @(absurd … Hi) @not_le_Sn_O ] ] ] qed. (* NOTE: we only compare the first elements here because otherwise the * added = 0 → policy_equal property of jump_expansion_step doesn't hold: * if we have not added anything to the pc, we only know the PC hasn't changed, * there might still have been a short/medium jump change *) definition sigma_pc_equal ≝ λprogram:list labelled_instruction.λp1,p2:ppc_pc_map. (∀n.n ≤ |program| → \fst (bvt_lookup … (bitvector_of_nat 16 n) (\snd p1) 〈0,short_jump〉) = \fst (bvt_lookup … (bitvector_of_nat 16 n) (\snd p2) 〈0,short_jump〉)). definition sigma_jump_equal ≝ λprogram:list labelled_instruction.λp1,p2:ppc_pc_map. (∀n.n < |program| → \snd (bvt_lookup … (bitvector_of_nat 16 n) (\snd p1) 〈0,short_jump〉) = \snd (bvt_lookup … (bitvector_of_nat 16 n) (\snd p2) 〈0,short_jump〉)). definition nec_plus_ultra ≝ λprogram:list labelled_instruction.λp:ppc_pc_map. ¬(∀i.i < |program| → is_jump (\snd (nth i ? program 〈None ?, Comment []〉)) → \snd (bvt_lookup … (bitvector_of_nat 16 i) (\snd p) 〈0,short_jump〉) = long_jump). (*include alias "common/Identifiers.ma".*) include alias "ASM/BitVector.ma". include alias "basics/lists/list.ma". include alias "arithmetics/nat.ma". include alias "basics/logic.ma". lemma jump_length_equal_max: ∀a,b,i. is_jump i → instruction_size_jmplen (max_length a b) i = instruction_size_jmplen a i → (max_length a b) = a. #a #b #i cases i [1: #pi cases pi try (#x #y #H #EQ) try (#x #H #EQ) try (#H #EQ) cases H cases a in EQ; cases b #EQ try % try (normalize in EQ; destruct(EQ) @False) try (lapply EQ @(subaddressing_mode_elim … x) #w #EQ normalize in EQ; destruct(EQ) @False) lapply EQ -EQ cases x * #a1 #a2 @(subaddressing_mode_elim … a2) #w try (#EQ normalize in EQ; destruct(EQ) @False) @(subaddressing_mode_elim … a1) #w #EQ normalize in EQ; destruct(EQ) |2,3,6: #x [3: #y] #H cases H |4,5: #id #_ cases a cases b #H try % normalize in H; destruct(H) ] qed. lemma jump_length_le_max: ∀a,b,i.is_jump i → instruction_size_jmplen a i ≤ instruction_size_jmplen (max_length a b) i. #a #b #i cases i [2,3,6: #x [3: #y] #H cases H |4,5: #id #_ cases a cases b @leb_true_to_le / by refl/ |1: #pi cases pi try (#x #y #H) try (#x #H) try (#H) cases H -H cases a cases b @leb_true_to_le try % try (@(subaddressing_mode_elim … x) #w % @False) cases x * #a1 #a2 @(subaddressing_mode_elim … a2) #w try % @(subaddressing_mode_elim … a1) #w % ] qed. lemma equal_compact_unsafe_compact: ∀program:(Σl.(S (|l|)) < 2^16 ∧ is_well_labelled_p l). ∀old_sigma.∀sigma. sigma_pc_equal program old_sigma sigma → sigma_jump_equal program old_sigma sigma → sigma_jumps program old_sigma → sigma_safe program (create_label_map program) old_sigma sigma → sigma_compact_unsafe program (create_label_map program) sigma → sigma_compact program (create_label_map program) sigma. #program cases program -program #program #Hprogram #old_sigma #sigma #Hpc_equal #Hjump_equal #Hjumps #Hsafe #Hcp_unsafe #dlbl #i #Hi lapply (Hcp_unsafe i Hi) -Hcp_unsafe lapply (Hsafe i Hi) -Hsafe inversion (lookup_opt … (bitvector_of_nat ? i) (\snd sigma)) [ / by / | #x cases x -x #pc #jl #EQ inversion (lookup_opt … (bitvector_of_nat ? (S i)) (\snd sigma)) [ / by / | #x cases x -x #Spc #Sjl #SEQ normalize nodelta #Hsafe #Hcp_unsafe (*CSC: make a lemma here; to shorten the proof, reimplement the safe case so that it also does a pattern matching on the jump_length type *) cut (instruction_size_jmplen jl (\snd (nth i ? program 〈None ?, Comment []〉)) = instruction_size … (bitvector_of_nat ? i) (\snd (nth i ? program 〈None ?, Comment []〉))) [6: #H (lookup_opt_lookup_hit … EQ 〈0,short_jump〉) #jeq #OSeq #Oeq #Hnth_eq #Hsafe lapply (Hsafe id (refl ? id)) -Hsafe normalize nodelta whd in match expand_relative_jump; whd in match expand_relative_jump_internal; normalize nodelta >add_bitvector_of_nat_plus <(plus_n_Sm i 0) (lookup_opt_lookup_hit … SEQ 〈0,short_jump〉) in OSeq; #OSeq >Hcp_unsafe inversion (leb (lookup_def … (create_label_map program) id 0) i) #Hli normalize nodelta [1,3,5,7,9,11,13,15,17,19,21: (* JC JNC JB JNB JBC JZ JNZ CJNE DJNZ Jmp Call *) >(lookup_opt_lookup_hit … EQ 〈0,short_jump〉) cases jl in jeq; normalize nodelta #jeq [2,5,8,11,14,17,20,23,26: lapply (Hjumps i Hi) >Hnth_eq >jeq normalize nodelta #abs cases ((proj2 ?? abs) (refl ? absolute_jump)) (* no absolute reljmps *) |31: lapply (Hjumps i Hi) >Hnth_eq >jeq normalize nodelta #abs cases ((proj1 ?? abs) (refl ? short_jump)) (* no short calls *) |28: >Hnth_eq #Hold cases (short_jump_cond ??) in Hold; #sj_poss #disp #Hold normalize nodelta >Hold normalize nodelta % |1,4,7,10,13,16,19,22,25: >Hnth_eq #Hold cases (short_jump_cond ??) in Hold; #sj_poss #disp #Hold normalize nodelta >Hold try % try (@(subaddressing_mode_elim … x) #w %) cases x * #a1 #a2 @(subaddressing_mode_elim … a2) #w try % @(subaddressing_mode_elim … a1) #w % |3,6,9,12,15,18,21,24,27: >Hnth_eq #_ whd in match (jmpeqb ??); cases (short_jump_cond ??); #sj_poss #disp normalize nodelta cases sj_poss normalize nodelta try % try (@(subaddressing_mode_elim … x) #w %) cases x * #a1 #a2 @(subaddressing_mode_elim … a2) #w try % @(subaddressing_mode_elim … a1) #w % |29,32: >Hnth_eq #Hold cases (absolute_jump_cond ??) in Hold; #aj_poss #disp #Hold >Hold normalize nodelta cases (short_jump_cond ??); [1,2: #sj_poss #disp2 normalize nodelta cases sj_poss % |3,4: @(zero 16) ] |30,33: >Hnth_eq #_ whd in match (jmpeqb ??); cases (absolute_jump_cond ??); #mj_poss #disp2 normalize nodelta cases (short_jump_cond ??); [1,2: #sj_poss #disp normalize nodelta cases sj_poss cases mj_poss % |3,4: @(zero 16) (* where does this come from? *) ] ] |*: cases (lookup … (bitvector_of_nat ? i) (\snd old_sigma) 〈0,short_jump〉) in Oeq jeq; #opc #ojl #Oeq #jeq normalize nodelta cases (lookup … (bitvector_of_nat ? (S i)) (\snd old_sigma) 〈0,short_jump〉) in OSeq; #oSpc #oSjl #OSeq normalize nodelta >jeq >OSeq >Hcp_unsafe -Hcp_unsafe >Hnth_eq lapply (Hpc_equal (lookup_def … (create_label_map program) id 0) ?) [1,3,5,7,9,11,13,15,17,19,21: @(le_S_to_le … (pi2 ?? (create_label_map program) id ?)) cut (i < 2^16) [1,3,5,7,9,11,13,15,17,19,21: @(transitive_lt … (proj1 ?? Hprogram)) @le_S @Hi ] #Hi2 @(proj2 ?? ((proj2 ?? Hprogram) id (bitvector_of_nat ? i) ???)) [1,5,9,13,17,21,25,29,33,37,41: whd in match fetch_pseudo_instruction; normalize nodelta >(nth_safe_nth … 〈None ?, Comment []〉) >nat_of_bitvector_bitvector_of_nat_inverse [1: >Hnth_eq in ⊢ (??%?); |3: >Hnth_eq in ⊢ (??%?); |5: >Hnth_eq in ⊢ (??%?); |7: >Hnth_eq in ⊢ (??%?); |9: >Hnth_eq in ⊢ (??%?); |11: >Hnth_eq in ⊢ (??%?); |13: >Hnth_eq in ⊢ (??%?); |15: >Hnth_eq in ⊢ (??%?); |17: >Hnth_eq in ⊢ (??%?); |19: >Hnth_eq in ⊢ (??%?); |21: >Hnth_eq in ⊢ (??%?); ] [1,2,3,4,5,6,7,8,9,10,11: % |*: assumption] |3,7,11,15,19,23,27,31,35,39,43: >nat_of_bitvector_bitvector_of_nat_inverse assumption |4,8,12,16,20,24,28,32,36,40,44: % ] |*: #Hpc lapply (Hjumps i Hi) >Hnth_eq >(Hjump_equal i Hi) >(lookup_opt_lookup_hit … EQ 〈0,short_jump〉) cases jl normalize nodelta [31: #abs cases ((proj1 ?? abs) (refl ? short_jump)) (* no short calls *) |2,5,8,11,14,17,20,23,26: #abs cases ((proj2 ?? abs) (refl ? absolute_jump)) (* no absolute RJs *) |1,4,7,10,13,16,19,22,25,28: >Hpc cases (short_jump_cond ??); #sj_poss #disp #_ #H normalize nodelta >H normalize nodelta try % try (@(subaddressing_mode_elim … x) #w %) cases x * #a1 #a2 @(subaddressing_mode_elim … a2) #w try % @(subaddressing_mode_elim … a1) #w % |3,6,9,12,15,18,21,24,27: >Hpc #_ #_ cases (short_jump_cond ??); #sj_poss #disp normalize nodelta normalize nodelta whd in match (jmpeqb ??); cases sj_poss try % try (@(subaddressing_mode_elim … x) #w %) cases x * #a1 #a2 @(subaddressing_mode_elim … a2) #w try % @(subaddressing_mode_elim … a1) #w % |29,32: >Hpc #_ #Hold cases (absolute_jump_cond ??) in Hold; #aj_poss #disp #Hold >Hold normalize nodelta cases (short_jump_cond ??) [1,2: #sj_poss #disp2 cases sj_poss normalize nodelta % |3,4: @(zero 16) ] |30,33: >Hnth_eq #_ #_ whd in match (jmpeqb ??); cases (absolute_jump_cond ??); #mj_poss #disp2 normalize nodelta cases (short_jump_cond ??); [1,2: #sj_poss #disp normalize nodelta cases sj_poss cases mj_poss % |3,4: @(zero 16) ] ] ] ] ] ] ] qed. lemma instruction_size_irrelevant: ∀i. ¬is_jump i → ∀j1,j2.instruction_size_jmplen j1 i = instruction_size_jmplen j2 i. #i cases i [2,3,6: #x [3: #y] #Hj #j1 #j2 % |4,5: #x #Hi cases Hi |1: #pi cases pi try (#x #y #Hj #j1 #j2) try (#y #Hj #j1 #j2) try (#Hj #j1 #j2) try % cases Hj ] qed.