]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/emulator/opcodes/IP2022_table_tests.ma
freescale porting
[helm.git] / helm / software / matita / contribs / ng_assembly / emulator / opcodes / IP2022_table_tests.ma
index c753b181b29388edf2e2fab5d48ca414bb4f029a..798d25f995aa750d5978d36f838390aa82918f52 100755 (executable)
@@ -34,12 +34,13 @@ ndefinition IP2022_not_impl_byte ≝
  ;〈x6,xB〉;〈x6,xC〉;〈x6,xD〉;〈x6,xE〉;〈x6,xF〉;〈x7,x5〉 ].
 
 (* test bytecode non implementati *)
-nlemma ok_byte_table_IP2022 : forall_b8 (λb.
+(* !!! per brevita... *)
+(*nlemma ok_byte_table_IP2022 : forall_b8 (λb.
  (test_not_impl_byte b IP2022_not_impl_byte     ⊙ eq_w16 (get_byte_count IP2022 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_IP2022) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗
  (⊖ (test_not_impl_byte b IP2022_not_impl_byte) ⊙ eq_w16 (get_byte_count IP2022 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_IP2022) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
  napply refl_eq.
-nqed.
+nqed.*)
 
 (* IP2022: opcode non implementati come da manuale (0x00+byte) *)
 ndefinition IP2022_not_impl_word ≝
@@ -75,31 +76,35 @@ ndefinition IP2022_not_impl_word ≝
  ].
 
 (* test bytecode non implementati *)
-nlemma ok_word_table_IP2022 : forall_b8 (λb.
+(* !!! per brevita... *)
+(*nlemma ok_word_table_IP2022 : forall_b8 (λb.
  (test_not_impl_byte b IP2022_not_impl_word     ⊙ eq_w16 (get_word_count IP2022 〈〈x0,x0〉:b〉 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_IP2022) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗
  (⊖ (test_not_impl_byte b IP2022_not_impl_word) ⊙ eq_w16 (get_word_count IP2022 〈〈x0,x0〉:b〉 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_IP2022) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
  napply refl_eq.
-nqed.
+nqed.*)
 
 (* tutti op implementati *)
-nlemma ok_pseudo_table_IP2022 :
+(* !!! per brevita... *)
+(*nlemma ok_pseudo_table_IP2022 :
  forall_op IP2022 (λo.
   le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_pseudo_count IP2022 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_IP2022)) = true.
  napply refl_eq.
-nqed.
+nqed.*)
 
 (* tutte im implementate *)
-nlemma ok_mode_table_IP2022 :
+(* !!! per brevita... *)
+(*nlemma ok_mode_table_IP2022 :
  forall_im IP2022 (λi.
   le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_mode_count IP2022 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_IP2022)) = true.
  napply refl_eq.
-nqed.
+nqed.*)
 
 (* nessuna ripetizione di combinazione op + imm *)
-nlemma ok_OpIm_table_IP2022 :
+(* !!! per brevita... *)
+(*nlemma ok_OpIm_table_IP2022 :
  forall_im IP2022 (λi.
   forall_op IP2022 (λo.
    le_w16 (get_OpIm_count IP2022 o i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_IP2022) 〈〈x0,x0〉:〈x0,x1〉〉)) = true.
  napply refl_eq.
-nqed.
+nqed.*)