]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/emulator/opcodes/HCS08_table_tests.ma
freescale porting
[helm.git] / helm / software / matita / contribs / ng_assembly / emulator / opcodes / HCS08_table_tests.ma
index c7b9c15491cabcdbcb3491f7c64e8a564bde6970..e51b236a3d688d4ebe5a0e8e331b91296592655f 100755 (executable)
@@ -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.
+(* !!! 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.*)