]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/emulator/opcodes/IP2022_table_tests.ma
mod change (-x)
[helm.git] / helm / software / matita / contribs / ng_assembly / emulator / opcodes / IP2022_table_tests.ma
old mode 100755 (executable)
new mode 100644 (file)
index 798d25f..3640054
@@ -20,7 +20,7 @@
 (*                                                                        *)
 (* ********************************************************************** *)
 
-include "emulator/opcodes/opcode.ma".
+include "emulator/opcodes/pseudo.ma".
 include "emulator/opcodes/IP2022_table.ma".
 
 (* ******************* *)
@@ -33,8 +33,6 @@ ndefinition IP2022_not_impl_byte ≝
  ;〈x6,x3〉;〈x6,x4〉;〈x6,x5〉;〈x6,x6〉;〈x6,x7〉;〈x6,x8〉;〈x6,x9〉;〈x6,xA〉
  ;〈x6,xB〉;〈x6,xC〉;〈x6,xD〉;〈x6,xE〉;〈x6,xF〉;〈x7,x5〉 ].
 
-(* test bytecode non implementati *)
-(* !!! 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〉〉))
@@ -75,36 +73,30 @@ ndefinition IP2022_not_impl_word ≝
  ;〈xF,x8〉;〈xF,x9〉;〈xF,xA〉;〈xF,xB〉;〈xF,xC〉;〈xF,xD〉;〈xF,xE〉;〈xF,xF〉
  ].
 
-(* test bytecode non implementati *)
-(* !!! per brevita... *)
-(*nlemma ok_word_table_IP2022 : forall_b8 (λb.
+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 *)
-(* !!! per brevita... *)
-(*nlemma ok_pseudo_table_IP2022 :
- forall_op IP2022 (λo.
+(* tutti gli pseudo implementati *)
+nlemma ok_pseudo_table_IP2022 :
+ forall_IP2022_pseudo (λ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 *)
-(* !!! per brevita... *)
-(*nlemma ok_mode_table_IP2022 :
- forall_im IP2022 (λi.
+(* tutte le modalita' implementate *)
+nlemma ok_mode_table_IP2022 :
+ forall_IP2022_im (λ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 *)
-(* !!! 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.
+nlemma ok_PsIm_table_IP2022 :
+ forall_IP2022_im (λi.
+ forall_IP2022_pseudo (λps.
+  le_w16 (get_PsIm_count IP2022 ps i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_IP2022) 〈〈x0,x0〉:〈x0,x1〉〉)) = true.
  napply refl_eq.
-nqed.*)
+nqed.