X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Femulator%2Fopcodes%2FRS08_table_tests.ma;h=143724ca3b8de10de19ce6f8137080be2dd2c562;hb=e880d6eab5e1700f4a625ddcd7d0fa8f0cce2dcc;hp=73f1c0ad23372656977e38dd8e709fb2ef52ce42;hpb=7f89b4dce54266c281479a14c01edc4bd33993d1;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/emulator/opcodes/RS08_table_tests.ma b/helm/software/matita/contribs/ng_assembly/emulator/opcodes/RS08_table_tests.ma index 73f1c0ad2..143724ca3 100755 --- a/helm/software/matita/contribs/ng_assembly/emulator/opcodes/RS08_table_tests.ma +++ b/helm/software/matita/contribs/ng_assembly/emulator/opcodes/RS08_table_tests.ma @@ -20,7 +20,7 @@ (* *) (* ********************************************************************** *) -include "emulator/opcodes/opcode.ma". +include "emulator/opcodes/pseudo.ma". include "emulator/opcodes/RS08_table.ma". (* ***************** *) @@ -35,36 +35,44 @@ ndefinition RS08_not_impl_byte ≝ ;〈xB,x3〉;〈xB,x5〉 ]. -(* test bytecode non implementati *) -(* !!! per brevita... *) -(*nlemma ok_byte_table_RS08 : forall_b8 (λb. +nlemma ok_byte_table_RS08 : forall_b8 (λb. (test_not_impl_byte b RS08_not_impl_byte ⊙ eq_w16 (get_byte_count RS08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗ (⊖ (test_not_impl_byte b RS08_not_impl_byte) ⊙ eq_w16 (get_byte_count RS08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08) 〈〈x0,x0〉:〈x0,x0〉〉)) = true. napply refl_eq. -nqed.*) +nqed. -(* tutti op implementati *) -(* !!! per brevita... *) -(*nlemma ok_pseudo_table_RS08 : - forall_op RS08 (λo. - le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_pseudo_count RS08 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08)) = true. +(* RS08: pseudocodici non implementati come da manuale *) +ndefinition RS08_not_impl_pseudo ≝ + [ AIS ; AIX ; ASR ; BGE ; BGT ; BHCC ; BHCS ; BHI ; BIH ; BIL ; BIT ; BLE ; BLS + ; BLT ; BMC ; BMI ; BMS ; BPL ; BRN ; CBEQX ; CLI ; CPHX ; CPX ; DAA ; DIV + ; LDHX ; LDX ; MUL ; NEG ; NSA ; PSHA ; PSHH ; PSHX ; PULA ; PULH ; PULX ; RSP + ; RTI ; SEI ; STHX ; STX ; SWI ; TAP ; TAX ; TPA ; TST ; TSX ; TXA ; TXS ]. + +nlemma ok_pseudo_table_RS08 : forall_Freescale_pseudo (λo. + (test_not_impl_pseudo RS08 o RS08_not_impl_pseudo ⊙ le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_pseudo_count RS08 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08)) ⊗ + (⊖ (test_not_impl_pseudo RS08 o RS08_not_impl_pseudo) ⊙ eq_w16 (get_pseudo_count RS08 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08) 〈〈x0,x0〉:〈x0,x0〉〉)) + = true. napply refl_eq. -nqed.*) +nqed. -(* tutte im implementate *) -(* !!! per brevita... *) -(*nlemma ok_mode_table_RS08 : - forall_im RS08 (λi. - le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_mode_count RS08 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08)) = true. +(* RS08: modalita' non implementate come da manuale *) +ndefinition RS08_not_impl_mode ≝ + [ MODE_INHX ; MODE_INHH ; MODE_INHX0ADD ; MODE_INHX1ADD ; MODE_INHX2ADD ; MODE_IMM1EXT + ; MODE_DIR2 ; MODE_IX0 ; MODE_IX1 ; MODE_IX2 ; MODE_SP1 ; MODE_SP2 + ; MODE_IX0p_to_DIR1 ; MODE_DIR1_to_IX0p ; MODE_INHX_and_IMM1 ; MODE_IX0_and_IMM1 + ; MODE_IX0p_and_IMM1 ; MODE_IX1_and_IMM1 ; MODE_IX1p_and_IMM1 ; MODE_SP1_and_IMM1 ]. + +nlemma ok_mode_table_RS08 : forall_Freescale_im (λi. + (test_not_impl_mode RS08 i RS08_not_impl_mode ⊙ le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_mode_count RS08 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08)) ⊗ + (⊖ (test_not_impl_mode RS08 i RS08_not_impl_mode) ⊙ eq_w16 (get_mode_count RS08 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08) 〈〈x0,x0〉:〈x0,x0〉〉)) + = true. napply refl_eq. -nqed.*) +nqed. -(* nessuna ripetizione di combinazione op + imm *) -(* !!! per brevita... *) -(*nlemma ok_OpIm_table_RS08 : - forall_im RS08 (λi. - forall_op RS08 (λo. - le_w16 (get_OpIm_count RS08 o i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08) 〈〈x0,x0〉:〈x0,x1〉〉)) = true. +nlemma ok_PsIm_table_RS08 : + forall_Freescale_im (λi:Freescale_instr_mode. + forall_Freescale_pseudo (λps:Freescale_pseudo. + le_w16 (get_PsIm_count RS08 ps i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_RS08) 〈〈x0,x0〉:〈x0,x1〉〉)) = true. napply refl_eq. -nqed.*) +nqed.