include "ASM/Status.ma". include "ASM/Fetch.ma". definition sign_extension: Byte → Word ≝ λc. let b ≝ get_index_v ? 8 c 1 ? in [[ b; b; b; b; b; b; b; b ]] @@ c. normalize; repeat (@ (le_S_S ?)); @ (le_O_n); qed. lemma inclusive_disjunction_true: ∀b, c: bool. (orb b c) = true → b = true ∨ c = true. # b # c elim b [ normalize # H @ or_introl % | normalize /2/ ] qed. lemma conjunction_true: ∀b, c: bool. andb b c = true → b = true ∧ c = true. # b # c elim b normalize [ /2/ | # K destruct ] qed. lemma eq_true_false: false=true → False. # K destruct qed. lemma eq_a_to_eq: ∀a,b. eq_a a b = true → a=b. # a # b cases a cases b normalize // # K cases (eq_true_false K) qed. lemma inclusive_disjunction_b_true: ∀b. orb b true = true. # b cases b % qed. lemma is_a_to_mem_to_is_in: ∀he,a,m,q. is_a he … a = true → mem … eq_a (S m) q he = true → is_in … q a = true. # he # a # m # q elim q [ normalize # _ # K assumption | # m' # t # q' # II # H1 # H2 normalize change in H2: (??%?) with (orb ??) cases (inclusive_disjunction_true … H2) [ # K < (eq_a_to_eq … K) > H1 % | # K > II try assumption cases (is_a t a) normalize % ] ] qed. lemma execute_1_technical: ∀n, m: nat. ∀v: Vector addressing_mode_tag (S n). ∀q: Vector addressing_mode_tag (S m). ∀a: addressing_mode. bool_to_Prop (is_in ? v a) → bool_to_Prop (subvector_with ? ? ? eq_a v q) → bool_to_Prop (is_in ? q a). # n # m # v # q # a elim v [ normalize # K cases K | # n' # he # tl # II whd in ⊢ (% → ? → ?) lapply (refl … (is_in … (he:::tl) a)) cases (is_in … (he:::tl) a) in ⊢ (???% → %) [ # K # _ normalize in K; whd in ⊢ (% → ?) lapply (refl … (subvector_with … eq_a (he:::tl) q)) cases (subvector_with … eq_a (he:::tl) q) in ⊢ (???% → %) [ # K1 # _ change in K1 with ((andb ? (subvector_with …)) = true) cases (conjunction_true … K1) # K3 # K4 cases (inclusive_disjunction_true … K) # K2 [ > (is_a_to_mem_to_is_in … K2 K3) % | @ II [ > K2 % | > K4 % ] ] | # K1 # K2 (normalize in K2) cases K2; ] | # K1 # K2 normalize in K2; cases K2 ] ] qed. include alias "arithmetics/nat.ma". include alias "ASM/BitVectorTrie.ma". definition execute_1_preinstruction: ∀A, M: Type[0]. (A → PreStatus M → Word) → preinstruction A → PreStatus M → PreStatus M ≝ λA, M: Type[0]. λaddr_of: A → (PreStatus M) → Word. λinstr: preinstruction A. λs: PreStatus M. match instr with [ ADD addr1 addr2 ⇒ let 〈result, flags〉 ≝ add_8_with_carry (get_arg_8 ? s false addr1) (get_arg_8 ? s false addr2) false in let cy_flag ≝ get_index' ? O ? flags in let ac_flag ≝ get_index' ? 1 ? flags in let ov_flag ≝ get_index' ? 2 ? flags in let s ≝ set_arg_8 ? s ACC_A result in set_flags ? s cy_flag (Some Bit ac_flag) ov_flag | ADDC addr1 addr2 ⇒ let old_cy_flag ≝ get_cy_flag ? s in let 〈result, flags〉 ≝ add_8_with_carry (get_arg_8 ? s false addr1) (get_arg_8 ? s false addr2) old_cy_flag in let cy_flag ≝ get_index' ? O ? flags in let ac_flag ≝ get_index' ? 1 ? flags in let ov_flag ≝ get_index' ? 2 ? flags in let s ≝ set_arg_8 ? s ACC_A result in set_flags ? s cy_flag (Some Bit ac_flag) ov_flag | SUBB addr1 addr2 ⇒ let old_cy_flag ≝ get_cy_flag ? s in let 〈result, flags〉 ≝ sub_8_with_carry (get_arg_8 ? s false addr1) (get_arg_8 ? s false addr2) old_cy_flag in let cy_flag ≝ get_index' ? O ? flags in let ac_flag ≝ get_index' ? 1 ? flags in let ov_flag ≝ get_index' ? 2 ? flags in let s ≝ set_arg_8 ? s ACC_A result in set_flags ? s cy_flag (Some Bit ac_flag) ov_flag | ANL addr ⇒ match addr with [ inl l ⇒ match l with [ inl l' ⇒ let 〈addr1, addr2〉 ≝ l' in let and_val ≝ conjunction_bv ? (get_arg_8 ? s true addr1) (get_arg_8 ? s true addr2) in set_arg_8 ? s addr1 and_val | inr r ⇒ let 〈addr1, addr2〉 ≝ r in let and_val ≝ conjunction_bv ? (get_arg_8 ? s true addr1) (get_arg_8 ? s true addr2) in set_arg_8 ? s addr1 and_val ] | inr r ⇒ let 〈addr1, addr2〉 ≝ r in let and_val ≝ andb (get_cy_flag ? s) (get_arg_1 ? s addr2 true) in set_flags ? s and_val (None ?) (get_ov_flag ? s) ] | ORL addr ⇒ match addr with [ inl l ⇒ match l with [ inl l' ⇒ let 〈addr1, addr2〉 ≝ l' in let or_val ≝ inclusive_disjunction_bv ? (get_arg_8 ? s true addr1) (get_arg_8 ? s true addr2) in set_arg_8 ? s addr1 or_val | inr r ⇒ let 〈addr1, addr2〉 ≝ r in let or_val ≝ inclusive_disjunction_bv ? (get_arg_8 ? s true addr1) (get_arg_8 ? s true addr2) in set_arg_8 ? s addr1 or_val ] | inr r ⇒ let 〈addr1, addr2〉 ≝ r in let or_val ≝ (get_cy_flag ? s) ∨ (get_arg_1 ? s addr2 true) in set_flags ? s or_val (None ?) (get_ov_flag ? s) ] | XRL addr ⇒ match addr with [ inl l' ⇒ let 〈addr1, addr2〉 ≝ l' in let xor_val ≝ exclusive_disjunction_bv ? (get_arg_8 ? s true addr1) (get_arg_8 ? s true addr2) in set_arg_8 ? s addr1 xor_val | inr r ⇒ let 〈addr1, addr2〉 ≝ r in let xor_val ≝ exclusive_disjunction_bv ? (get_arg_8 ? s true addr1) (get_arg_8 ? s true addr2) in set_arg_8 ? s addr1 xor_val ] | INC addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ acc_a; registr; direct; indirect; dptr ]] x) → ? with [ ACC_A ⇒ λacc_a: True. let 〈 carry, result 〉 ≝ half_add ? (get_arg_8 ? s true ACC_A) (bitvector_of_nat 8 1) in set_arg_8 ? s ACC_A result | REGISTER r ⇒ λregister: True. let 〈 carry, result 〉 ≝ half_add ? (get_arg_8 ? s true (REGISTER r)) (bitvector_of_nat 8 1) in set_arg_8 ? s (REGISTER r) result | DIRECT d ⇒ λdirect: True. let 〈 carry, result 〉 ≝ half_add ? (get_arg_8 ? s true (DIRECT d)) (bitvector_of_nat 8 1) in set_arg_8 ? s (DIRECT d) result | INDIRECT i ⇒ λindirect: True. let 〈 carry, result 〉 ≝ half_add ? (get_arg_8 ? s true (INDIRECT i)) (bitvector_of_nat 8 1) in set_arg_8 ? s (INDIRECT i) result | DPTR ⇒ λdptr: True. let 〈 carry, bl 〉 ≝ half_add ? (get_8051_sfr ? s SFR_DPL) (bitvector_of_nat 8 1) in let 〈 carry, bu 〉 ≝ full_add ? (get_8051_sfr ? s SFR_DPH) (zero 8) carry in let s ≝ set_8051_sfr ? s SFR_DPL bl in set_8051_sfr ? s SFR_DPH bu | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr) | DEC addr ⇒ let 〈 result, flags 〉 ≝ sub_8_with_carry (get_arg_8 ? s true addr) (bitvector_of_nat 8 1) false in set_arg_8 ? s addr result | MUL addr1 addr2 ⇒ let acc_a_nat ≝ nat_of_bitvector 8 (get_8051_sfr ? s SFR_ACC_A) in let acc_b_nat ≝ nat_of_bitvector 8 (get_8051_sfr ? s SFR_ACC_B) in let product ≝ acc_a_nat * acc_b_nat in let ov_flag ≝ product ≥ 256 in let low ≝ bitvector_of_nat 8 (product mod 256) in let high ≝ bitvector_of_nat 8 (product ÷ 256) in let s ≝ set_8051_sfr ? s SFR_ACC_A low in set_8051_sfr ? s SFR_ACC_B high | DIV addr1 addr2 ⇒ let acc_a_nat ≝ nat_of_bitvector 8 (get_8051_sfr ? s SFR_ACC_A) in let acc_b_nat ≝ nat_of_bitvector 8 (get_8051_sfr ? s SFR_ACC_B) in match acc_b_nat with [ O ⇒ set_flags ? s false (None Bit) true | S o ⇒ let q ≝ bitvector_of_nat 8 (acc_a_nat ÷ (S o)) in let r ≝ bitvector_of_nat 8 (acc_a_nat mod 256) in let s ≝ set_8051_sfr ? s SFR_ACC_A q in let s ≝ set_8051_sfr ? s SFR_ACC_B r in set_flags ? s false (None Bit) false ] | DA addr ⇒ let 〈 acc_nu, acc_nl 〉 ≝ split ? 4 4 (get_8051_sfr ? s SFR_ACC_A) in if (gtb (nat_of_bitvector ? acc_nl) 9) ∨ (get_ac_flag ? s) then let 〈result, flags〉 ≝ add_8_with_carry (get_8051_sfr ? s SFR_ACC_A) (bitvector_of_nat 8 6) false in let cy_flag ≝ get_index' ? O ? flags in let 〈acc_nu', acc_nl'〉 ≝ split ? 4 4 result in if (gtb (nat_of_bitvector ? acc_nu') 9) ∨ cy_flag then let 〈 carry, nu 〉 ≝ half_add ? acc_nu' (bitvector_of_nat 4 6) in let new_acc ≝ nu @@ acc_nl' in let s ≝ set_8051_sfr ? s SFR_ACC_A new_acc in set_flags ? s cy_flag (Some ? (get_ac_flag ? s)) (get_ov_flag ? s) else s else s | CLR addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ acc_a; carry; bit_addr ]] x) → ? with [ ACC_A ⇒ λacc_a: True. set_arg_8 ? s ACC_A (zero 8) | CARRY ⇒ λcarry: True. set_arg_1 ? s CARRY false | BIT_ADDR b ⇒ λbit_addr: True. set_arg_1 ? s (BIT_ADDR b) false | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr) | CPL addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ acc_a; carry; bit_addr ]] x) → ? with [ ACC_A ⇒ λacc_a: True. let old_acc ≝ get_8051_sfr ? s SFR_ACC_A in let new_acc ≝ negation_bv ? old_acc in set_8051_sfr ? s SFR_ACC_A new_acc | CARRY ⇒ λcarry: True. let old_cy_flag ≝ get_arg_1 ? s CARRY true in let new_cy_flag ≝ ¬old_cy_flag in set_arg_1 ? s CARRY new_cy_flag | BIT_ADDR b ⇒ λbit_addr: True. let old_bit ≝ get_arg_1 ? s (BIT_ADDR b) true in let new_bit ≝ ¬old_bit in set_arg_1 ? s (BIT_ADDR b) new_bit | _ ⇒ λother: False. ? ] (subaddressing_modein … addr) | SETB b ⇒ set_arg_1 ? s b false | RL _ ⇒ (* DPM: always ACC_A *) let old_acc ≝ get_8051_sfr ? s SFR_ACC_A in let new_acc ≝ rotate_left … 1 old_acc in set_8051_sfr ? s SFR_ACC_A new_acc | RR _ ⇒ (* DPM: always ACC_A *) let old_acc ≝ get_8051_sfr ? s SFR_ACC_A in let new_acc ≝ rotate_right … 1 old_acc in set_8051_sfr ? s SFR_ACC_A new_acc | RLC _ ⇒ (* DPM: always ACC_A *) let old_cy_flag ≝ get_cy_flag ? s in let old_acc ≝ get_8051_sfr ? s SFR_ACC_A in let new_cy_flag ≝ get_index' ? O ? old_acc in let new_acc ≝ shift_left … 1 old_acc old_cy_flag in let s ≝ set_arg_1 ? s CARRY new_cy_flag in set_8051_sfr ? s SFR_ACC_A new_acc | RRC _ ⇒ (* DPM: always ACC_A *) let old_cy_flag ≝ get_cy_flag ? s in let old_acc ≝ get_8051_sfr ? s SFR_ACC_A in let new_cy_flag ≝ get_index' ? 7 ? old_acc in let new_acc ≝ shift_right … 1 old_acc old_cy_flag in let s ≝ set_arg_1 ? s CARRY new_cy_flag in set_8051_sfr ? s SFR_ACC_A new_acc | SWAP _ ⇒ (* DPM: always ACC_A *) let old_acc ≝ get_8051_sfr ? s SFR_ACC_A in let 〈nu,nl〉 ≝ split ? 4 4 old_acc in let new_acc ≝ nl @@ nu in set_8051_sfr ? s SFR_ACC_A new_acc | PUSH addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ direct ]] x) → ? with [ DIRECT d ⇒ λdirect: True. let 〈carry, new_sp〉 ≝ half_add ? (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) in let s ≝ set_8051_sfr ? s SFR_SP new_sp in write_at_stack_pointer ? s d | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr) | POP addr ⇒ let contents ≝ read_at_stack_pointer ? s in let 〈new_sp, flags〉 ≝ sub_8_with_carry (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) false in let s ≝ set_8051_sfr ? s SFR_SP new_sp in set_arg_8 ? s addr contents | XCH addr1 addr2 ⇒ let old_addr ≝ get_arg_8 ? s false addr2 in let old_acc ≝ get_8051_sfr ? s SFR_ACC_A in let s ≝ set_8051_sfr ? s SFR_ACC_A old_addr in set_arg_8 ? s addr2 old_acc | XCHD addr1 addr2 ⇒ let 〈acc_nu, acc_nl〉 ≝ split … 4 4 (get_8051_sfr ? s SFR_ACC_A) in let 〈arg_nu, arg_nl〉 ≝ split … 4 4 (get_arg_8 ? s false addr2) in let new_acc ≝ acc_nu @@ arg_nl in let new_arg ≝ arg_nu @@ acc_nl in let s ≝ set_8051_sfr ? s SFR_ACC_A new_acc in set_arg_8 ? s addr2 new_arg | RET ⇒ let high_bits ≝ read_at_stack_pointer ? s in let 〈new_sp, flags〉 ≝ sub_8_with_carry (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) false in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let low_bits ≝ read_at_stack_pointer ? s in let 〈new_sp, flags〉 ≝ sub_8_with_carry (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) false in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let new_pc ≝ high_bits @@ low_bits in set_program_counter ? s new_pc | RETI ⇒ let high_bits ≝ read_at_stack_pointer ? s in let 〈new_sp, flags〉 ≝ sub_8_with_carry (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) false in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let low_bits ≝ read_at_stack_pointer ? s in let 〈new_sp, flags〉 ≝ sub_8_with_carry (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) false in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let new_pc ≝ high_bits @@ low_bits in set_program_counter ? s new_pc | MOVX addr ⇒ (* DPM: only copies --- doesn't affect I/O *) match addr with [ inl l ⇒ let 〈addr1, addr2〉 ≝ l in set_arg_8 ? s addr1 (get_arg_8 ? s false addr2) | inr r ⇒ let 〈addr1, addr2〉 ≝ r in set_arg_8 ? s addr1 (get_arg_8 ? s false addr2) ] | MOV addr ⇒ match addr with [ inl l ⇒ match l with [ inl l' ⇒ match l' with [ inl l'' ⇒ match l'' with [ inl l''' ⇒ match l''' with [ inl l'''' ⇒ let 〈addr1, addr2〉 ≝ l'''' in set_arg_8 ? s addr1 (get_arg_8 ? s false addr2) | inr r'''' ⇒ let 〈addr1, addr2〉 ≝ r'''' in set_arg_8 ? s addr1 (get_arg_8 ? s false addr2) ] | inr r''' ⇒ let 〈addr1, addr2〉 ≝ r''' in set_arg_8 ? s addr1 (get_arg_8 ? s false addr2) ] | inr r'' ⇒ let 〈addr1, addr2〉 ≝ r'' in set_arg_16 ? s (get_arg_16 ? s addr2) addr1 ] | inr r ⇒ let 〈addr1, addr2〉 ≝ r in set_arg_1 ? s addr1 (get_arg_1 ? s addr2 false) ] | inr r ⇒ let 〈addr1, addr2〉 ≝ r in set_arg_1 ? s addr1 (get_arg_1 ? s addr2 false) ] | JC addr ⇒ if get_cy_flag ? s then set_program_counter ? s (addr_of addr s) else s | JNC addr ⇒ if ¬(get_cy_flag ? s) then set_program_counter ? s (addr_of addr s) else s | JB addr1 addr2 ⇒ if get_arg_1 ? s addr1 false then set_program_counter ? s (addr_of addr2 s) else s | JNB addr1 addr2 ⇒ if ¬(get_arg_1 ? s addr1 false) then set_program_counter ? s (addr_of addr2 s) else s | JBC addr1 addr2 ⇒ let s ≝ set_arg_1 ? s addr1 false in if get_arg_1 ? s addr1 false then set_program_counter ? s (addr_of addr2 s) else s | JZ addr ⇒ if eq_bv ? (get_8051_sfr ? s SFR_ACC_A) (zero 8) then set_program_counter ? s (addr_of addr s) else s | JNZ addr ⇒ if ¬(eq_bv ? (get_8051_sfr ? s SFR_ACC_A) (zero 8)) then set_program_counter ? s (addr_of addr s) else s | CJNE addr1 addr2 ⇒ match addr1 with [ inl l ⇒ let 〈addr1, addr2'〉 ≝ l in let new_cy ≝ ltb (nat_of_bitvector ? (get_arg_8 ? s false addr1)) (nat_of_bitvector ? (get_arg_8 ? s false addr2')) in if ¬(eq_bv ? (get_arg_8 ? s false addr1) (get_arg_8 ? s false addr2')) then let s ≝ set_program_counter ? s (addr_of addr2 s) in set_flags ? s new_cy (None ?) (get_ov_flag ? s) else (set_flags ? s new_cy (None ?) (get_ov_flag ? s)) | inr r' ⇒ let 〈addr1, addr2'〉 ≝ r' in let new_cy ≝ ltb (nat_of_bitvector ? (get_arg_8 ? s false addr1)) (nat_of_bitvector ? (get_arg_8 ? s false addr2')) in if ¬(eq_bv ? (get_arg_8 ? s false addr1) (get_arg_8 ? s false addr2')) then let s ≝ set_program_counter ? s (addr_of addr2 s) in set_flags ? s new_cy (None ?) (get_ov_flag ? s) else (set_flags ? s new_cy (None ?) (get_ov_flag ? s)) ] | DJNZ addr1 addr2 ⇒ let 〈result, flags〉 ≝ sub_8_with_carry (get_arg_8 ? s true addr1) (bitvector_of_nat 8 1) false in let s ≝ set_arg_8 ? s addr1 result in if ¬(eq_bv ? result (zero 8)) then set_program_counter ? s (addr_of addr2 s) else s | NOP ⇒ s ]. try assumption try @I try ( @ (execute_1_technical … (subaddressing_modein …)) @ I ) qed. definition execute_1: Status → Status ≝ λs: Status. let 〈instr_pc, ticks〉 ≝ fetch (code_memory ? s) (program_counter ? s) in let 〈instr, pc〉 ≝ 〈fst … instr_pc, snd … instr_pc〉 in let s ≝ set_clock ? s (clock ? s + ticks) in let s ≝ set_program_counter ? s pc in let s ≝ match instr with [ RealInstruction instr ⇒ execute_1_preinstruction [[ relative ]] ? (λx. λs. match x return λs. bool_to_Prop (is_in ? [[ relative ]] s) → ? with [ RELATIVE r ⇒ λ_. \snd (half_add ? (program_counter ? s) (sign_extension r)) | _ ⇒ λabsd. ⊥ ] (subaddressing_modein … x)) instr s | ACALL addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ addr11 ]] x) → ? with [ ADDR11 a ⇒ λaddr11: True. let 〈carry, new_sp〉 ≝ half_add ? (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let 〈pc_bu, pc_bl〉 ≝ split ? 8 8 (program_counter ? s) in let s ≝ write_at_stack_pointer ? s pc_bl in let 〈carry, new_sp〉 ≝ half_add ? (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let s ≝ write_at_stack_pointer ? s pc_bu in let 〈thr, eig〉 ≝ split ? 3 8 a in let 〈fiv, thr'〉 ≝ split ? 5 3 pc_bu in let new_addr ≝ (fiv @@ thr) @@ pc_bl in set_program_counter ? s new_addr | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr) | LCALL addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ addr16 ]] x) → ? with [ ADDR16 a ⇒ λaddr16: True. let 〈carry, new_sp〉 ≝ half_add ? (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let 〈pc_bu, pc_bl〉 ≝ split ? 8 8 (program_counter ? s) in let s ≝ write_at_stack_pointer ? s pc_bl in let 〈carry, new_sp〉 ≝ half_add ? (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let s ≝ write_at_stack_pointer ? s pc_bu in set_program_counter ? s a | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr) | MOVC addr1 addr2 ⇒ match addr2 return λx. bool_to_Prop (is_in … [[ acc_dptr; acc_pc ]] x) → ? with [ ACC_DPTR ⇒ λacc_dptr: True. let big_acc ≝ (zero 8) @@ (get_8051_sfr ? s SFR_ACC_A) in let dptr ≝ (get_8051_sfr ? s SFR_DPH) @@ (get_8051_sfr ? s SFR_DPL) in let 〈carry, new_addr〉 ≝ half_add ? dptr big_acc in let result ≝ lookup ? ? new_addr (code_memory ? s) (zero ?) in set_8051_sfr ? s SFR_ACC_A result | ACC_PC ⇒ λacc_pc: True. let big_acc ≝ (zero 8) @@ (get_8051_sfr ? s SFR_ACC_A) in let 〈carry,inc_pc〉 ≝ half_add ? (program_counter ? s) (bitvector_of_nat 16 1) in (* DPM: Under specified: does the carry from PC incrementation affect the *) (* addition of the PC with the DPTR? At the moment, no. *) let s ≝ set_program_counter ? s inc_pc in let 〈carry, new_addr〉 ≝ half_add ? inc_pc big_acc in let result ≝ lookup ? ? new_addr (code_memory ? s) (zero ?) in set_8051_sfr ? s SFR_ACC_A result | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr2) | JMP _ ⇒ (* DPM: always indirect_dptr *) let dptr ≝ (get_8051_sfr ? s SFR_DPH) @@ (get_8051_sfr ? s SFR_DPL) in let big_acc ≝ (zero 8) @@ (get_8051_sfr ? s SFR_ACC_A) in let 〈carry, jmp_addr〉 ≝ half_add ? big_acc dptr in let 〈carry, new_pc〉 ≝ half_add ? (program_counter ? s) jmp_addr in set_program_counter ? s new_pc | LJMP addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ addr16 ]] x) → ? with [ ADDR16 a ⇒ λaddr16: True. set_program_counter ? s a | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr) | SJMP addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ relative ]] x) → ? with [ RELATIVE r ⇒ λrelative: True. let 〈carry, new_pc〉 ≝ half_add ? (program_counter ? s) (sign_extension r) in set_program_counter ? s new_pc | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr) | AJMP addr ⇒ match addr return λx. bool_to_Prop (is_in … [[ addr11 ]] x) → ? with [ ADDR11 a ⇒ λaddr11: True. let 〈pc_bu, pc_bl〉 ≝ split ? 8 8 (program_counter ? s) in let 〈nu, nl〉 ≝ split ? 4 4 pc_bu in let bit ≝ get_index' ? O ? nl in let 〈relevant1, relevant2〉 ≝ split ? 3 8 a in let new_addr ≝ (nu @@ (bit ::: relevant1)) @@ relevant2 in let 〈carry, new_pc〉 ≝ half_add ? (program_counter ? s) new_addr in set_program_counter ? s new_pc | _ ⇒ λother: False. ⊥ ] (subaddressing_modein … addr) ] in s. try assumption try ( normalize repeat (@ (le_S_S)) @ (le_O_n) ) try ( @ (execute_1_technical … (subaddressing_modein …)) @ I ) try ( normalize @ I ) qed. let rec index_of_internal (A: Type[0]) (pred: A → bool) (l: list A) (acc: nat) on l: nat ≝ match l with [ nil ⇒ ? | cons hd tl ⇒ if pred hd then acc else index_of_internal A pred tl (S acc) ]. cases not_implemented. qed. definition index_of ≝ λA. λeq. λl. index_of_internal A eq l 0. definition instruction_matches_identifier ≝ λy: Identifier. λx: labelled_instruction. match \fst x with [ None ⇒ false | Some x ⇒ eq_bv ? x y ]. definition address_of_word_labels_code_mem ≝ λcode_mem. λid: Identifier. bitvector_of_nat 16 (index_of ? (instruction_matches_identifier id) code_mem). definition address_of_word_labels ≝ λps: PseudoStatus. λid: Identifier. address_of_word_labels_code_mem (\snd (code_memory ? ps)) id. definition execute_1_pseudo_instruction: (Word → nat) → PseudoStatus → PseudoStatus ≝ λticks_of. λs. let 〈instr, pc〉 ≝ fetch_pseudo_instruction (\snd (code_memory ? s)) (program_counter ? s) in let ticks ≝ ticks_of (program_counter ? s) in let s ≝ set_clock ? s (clock ? s + ticks) in let s ≝ set_program_counter ? s pc in let s ≝ match instr with [ Instruction instr ⇒ execute_1_preinstruction … (λx, y. address_of_word_labels y x) instr s | Comment cmt ⇒ s | Cost cst ⇒ s | Jmp jmp ⇒ set_program_counter ? s (address_of_word_labels s jmp) | Call call ⇒ let a ≝ address_of_word_labels s call in let 〈carry, new_sp〉 ≝ half_add ? (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let 〈pc_bu, pc_bl〉 ≝ split ? 8 8 (program_counter ? s) in let s ≝ write_at_stack_pointer ? s pc_bl in let 〈carry, new_sp〉 ≝ half_add ? (get_8051_sfr ? s SFR_SP) (bitvector_of_nat 8 1) in let s ≝ set_8051_sfr ? s SFR_SP new_sp in let s ≝ write_at_stack_pointer ? s pc_bu in set_program_counter ? s a | Mov dptr ident ⇒ set_arg_16 ? s (get_arg_16 ? s (DATA16 (address_of_word_labels s ident))) dptr ] in s. normalize @ I qed. let rec execute (n: nat) (s: Status) on n: Status ≝ match n with [ O ⇒ s | S o ⇒ execute o (execute_1 s) ]. let rec execute_pseudo_instruction (n: nat) (ticks_of: Word → nat) (s: PseudoStatus) on n: PseudoStatus ≝ match n with [ O ⇒ s | S o ⇒ execute_pseudo_instruction o ticks_of (execute_1_pseudo_instruction ticks_of s) ].