]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/freescale/table_HC08_tests.ma
freescale porting, work in progress
[helm.git] / helm / software / matita / contribs / ng_assembly / freescale / table_HC08_tests.ma
index 76bed4e991d5df953c7800ca05e42d4b62205540..8c85e3c39978d2b08ff2313fc638338ab9d9210a 100755 (executable)
 (* ********************************************************************** *)
 (*                          Progetto FreeScale                            *)
 (*                                                                        *)
-(*   Sviluppato da: Cosimo Oliboni, oliboni@cs.unibo.it                   *)
-(*     Cosimo Oliboni, oliboni@cs.unibo.it                                *)
+(*   Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it              *)
+(*   Ultima modifica: 05/08/2009                                          *)
 (*                                                                        *)
 (* ********************************************************************** *)
 
 include "freescale/table_HC08.ma".
+include "freescale/opcode.ma".
 
 (* ***************** *)
 (* TABELLA DELL'HC08 *)
@@ -34,9 +35,9 @@ ndefinition HC08_not_impl_byte ≝
  ;〈xA,xC〉
  ].
 
-nlemma ok_byte_table_HC08 : forall_byte8 (λb.
- (test_not_impl_byte b HC08_not_impl_byte     ⊙ eq_nat (get_byte_count HC08 b 0 opcode_table_HC08) 1) ⊗
- (⊖ (test_not_impl_byte b HC08_not_impl_byte) ⊙ eq_nat (get_byte_count HC08 b 0 opcode_table_HC08) 0))
+nlemma ok_byte_table_HC08 : forall_b8 (λb.
+ (test_not_impl_byte b HC08_not_impl_byte     ⊙ eq_w16 (get_byte_count HC08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗
+ (⊖ (test_not_impl_byte b HC08_not_impl_byte) ⊙ eq_w16 (get_byte_count HC08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
  napply refl_eq.
 nqed.
@@ -74,9 +75,9 @@ ndefinition HC08_not_impl_word ≝
  ;〈xF,x8〉;〈xF,x9〉;〈xF,xA〉;〈xF,xB〉;〈xF,xC〉;〈xF,xD〉;〈xF,xE〉;〈xF,xF〉
  ].
 
-nlemma ok_word_table_HC08 : forall_byte8 (λb.
- (test_not_impl_byte b HC08_not_impl_word     ⊙ eq_nat (get_word_count HC08 b 0 opcode_table_HC08) 1) ⊗
- (⊖ (test_not_impl_byte b HC08_not_impl_word) ⊙ eq_nat (get_word_count HC08 b 0 opcode_table_HC08) 0))
+nlemma ok_word_table_HC08 : forall_b8 (λb.
+ (test_not_impl_byte b HC08_not_impl_word     ⊙ eq_w16 (get_word_count HC08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗
+ (⊖ (test_not_impl_byte b HC08_not_impl_word) ⊙ eq_w16 (get_word_count HC08 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
  napply refl_eq.
 nqed.
@@ -85,9 +86,9 @@ nqed.
 ndefinition HC08_not_impl_pseudo ≝
  [ BGND ; SHA ; SLA ].
 
-nlemma ok_pseudo_table_HC08 : forall_opcode (λo.
- (test_not_impl_pseudo o HC08_not_impl_pseudo     ⊙ le_nat 1 (get_pseudo_count HC08 o 0 opcode_table_HC08)) ⊗
- (⊖ (test_not_impl_pseudo o HC08_not_impl_pseudo) ⊙ eq_nat (get_pseudo_count HC08 o 0 opcode_table_HC08) 0))
+nlemma ok_pseudo_table_HC08 : forall_op (λo.
+ (test_not_impl_pseudo o HC08_not_impl_pseudo     ⊙ le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_pseudo_count HC08 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08)) ⊗
+ (⊖ (test_not_impl_pseudo o HC08_not_impl_pseudo) ⊙ eq_w16 (get_pseudo_count HC08 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
  napply refl_eq.
 nqed.
@@ -107,16 +108,16 @@ ndefinition HC08_not_impl_mode ≝
  ; MODE_SRT t18 ; MODE_SRT t19 ; MODE_SRT t1A ; MODE_SRT t1B
  ; MODE_SRT t1C ; MODE_SRT t1D ; MODE_SRT t1E ; MODE_SRT t1F ].
 
-nlemma ok_mode_table_HC08 : forall_instr_mode (λi.
- (test_not_impl_mode i HC08_not_impl_mode     ⊙ le_nat 1 (get_mode_count HC08 i 0 opcode_table_HC08)) ⊗
- (⊖ (test_not_impl_mode i HC08_not_impl_mode) ⊙ eq_nat (get_mode_count HC08 i 0 opcode_table_HC08) 0))
+nlemma ok_mode_table_HC08 : forall_im (λi.
+ (test_not_impl_mode i HC08_not_impl_mode     ⊙ le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_mode_count HC08 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08)) ⊗
+ (⊖ (test_not_impl_mode i HC08_not_impl_mode) ⊙ eq_w16 (get_mode_count HC08 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
  napply refl_eq.
 nqed.
 
 nlemma ok_OpIm_table_HC08 :
- forall_instr_mode (λi:instr_mode.
- forall_opcode     (λop:opcode.
-  le_nat (get_OpIm_count HC08 (anyOP HC08 op) i 0 opcode_table_HC08) 1)) = true.
+ forall_im (λi:instr_mode.
+ forall_op (λop:opcode.
+  le_w16 (get_OpIm_count HC08 (anyOP HC08 op) i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC08) 〈〈x0,x0〉:〈x0,x1〉〉)) = true.
  napply refl_eq.
 nqed.