X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Femulator%2Fopcodes%2FHCS08_table_tests.ma;h=e51b236a3d688d4ebe5a0e8e331b91296592655f;hb=b082cb1e1fc9dbd471d16d2eb231e883e96e588f;hp=60a40cb73b72562856ea1455ec32c01afe727ac4;hpb=0f13d14b63b012e0ea8ce0d0e71bf808fdd444eb;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/emulator/opcodes/HCS08_table_tests.ma b/helm/software/matita/contribs/ng_assembly/emulator/opcodes/HCS08_table_tests.ma index 60a40cb73..e51b236a3 100755 --- a/helm/software/matita/contribs/ng_assembly/emulator/opcodes/HCS08_table_tests.ma +++ b/helm/software/matita/contribs/ng_assembly/emulator/opcodes/HCS08_table_tests.ma @@ -35,12 +35,13 @@ ndefinition HCS08_not_impl_byte ≝ ]. (* test bytecode non implementati *) -nlemma ok_byte_table_HCS08 : forall_b8 (λb. +(* !!! per brevita... *) +(*nlemma ok_byte_table_HCS08 : forall_b8 (λb. (test_not_impl_byte b HCS08_not_impl_byte ⊙ eq_w16 (get_byte_count HCS08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗ (⊖ (test_not_impl_byte b HCS08_not_impl_byte) ⊙ eq_w16 (get_byte_count HCS08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x0〉〉)) = true. napply refl_eq. -nqed. +nqed.*) (* HCS08: opcode non implementati come da manuale (0x9E+byte) *) ndefinition HCS08_not_impl_word ≝ @@ -72,31 +73,35 @@ ndefinition HCS08_not_impl_word ≝ ]. (* test bytecode non implementati *) -nlemma ok_word_table_HCS08 : forall_b8 (λb. - (test_not_impl_byte b HCS08_not_impl_word ⊙ eq_w16 (get_word_count HCS08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗ - (⊖ (test_not_impl_byte b HCS08_not_impl_word) ⊙ eq_w16 (get_word_count HCS08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x0〉〉)) +(* !!! per brevita... *) +(*nlemma ok_word_table_HCS08 : forall_b8 (λb. + (test_not_impl_byte b HCS08_not_impl_word ⊙ eq_w16 (get_word_count HCS08 〈〈x9,xE〉:b〉 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗ + (⊖ (test_not_impl_byte b HCS08_not_impl_word) ⊙ eq_w16 (get_word_count HCS08 〈〈x9,xE〉:b〉 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x0〉〉)) = true. napply refl_eq. -nqed. +nqed.*) (* tutti op implementati *) -nlemma ok_pseudo_table_HCS08 : +(* !!! per brevita... *) +(*nlemma ok_pseudo_table_HCS08 : forall_op HCS08 (λo. le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_pseudo_count HCS08 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08)) = true. napply refl_eq. -nqed. +nqed.*) (* tutte im implementate *) -nlemma ok_mode_table_HCS08 : +(* !!! per brevita... *) +(*nlemma ok_mode_table_HCS08 : forall_im HCS08 (λi. le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_mode_count HCS08 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08)) = true. napply refl_eq. -nqed. +nqed.*) (* nessuna ripetizione di combinazione op + imm *) -nlemma ok_OpIm_table_HCS08 : +(* !!! per brevita... *) +(*nlemma ok_OpIm_table_HCS08 : forall_im HCS08 (λi. forall_op HCS08 (λo. le_w16 (get_OpIm_count HCS08 o i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x1〉〉)) = true. napply refl_eq. -nqed. +nqed.*)