]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/freescale/table_HC05_tests.ma
freescale porting, work in progress
[helm.git] / helm / software / matita / contribs / ng_assembly / freescale / table_HC05_tests.ma
index a1b2152f4d473ca585c2da728180c8d595b84457..75bb840b83de827ff6901ecac26235b3d170e694 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_HC05.ma".
+include "freescale/opcode.ma".
 
 (* ***************** *)
 (* TABELLA DELL'HC05 *)
@@ -38,11 +39,11 @@ ndefinition HC05_not_impl_byte ≝
  ;〈xA,x7〉;〈xA,xC〉;〈xA,xF〉
  ].
 
-nlemma ok_byte_table_HC05 : forall_byte8 (λb.
- (test_not_impl_byte b HC05_not_impl_byte     ⊙ eq_nat (get_byte_count HC05 b 0 opcode_table_HC05) 1) ⊗
- (⊖ (test_not_impl_byte b HC05_not_impl_byte) ⊙ eq_nat (get_byte_count HC05 b 0 opcode_table_HC05) 0))
+nlemma ok_byte_table_HC05 : forall_b8 (λb.
+ (test_not_impl_byte b HC05_not_impl_byte     ⊙ eq_w16 (get_byte_count HC05 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC05) 〈〈x0,x0〉:〈x0,x1〉〉) ⊗
+ (⊖ (test_not_impl_byte b HC05_not_impl_byte) ⊙ eq_w16 (get_byte_count HC05 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC05) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HC05: pseudocodici non implementati come da manuale *)
@@ -51,11 +52,11 @@ ndefinition HC05_not_impl_pseudo ≝
  ; DBNZ ; DIV ; LDHX ; MOV ; NSA ; PSHA ; PSHH ; PSHX ; PULA ; PULH ; PULX
  ; SHA ; SLA ; STHX ; TAP ; TPA ; TSX ; TXS ].
 
-nlemma ok_pseudo_table_HC05 : forall_opcode (λo.
- (test_not_impl_pseudo o HC05_not_impl_pseudo     ⊙ le_nat 1 (get_pseudo_count HC05 o 0 opcode_table_HC05)) ⊗
- (⊖ (test_not_impl_pseudo o HC05_not_impl_pseudo) ⊙ eq_nat (get_pseudo_count HC05 o 0 opcode_table_HC05) 0))
+nlemma ok_pseudo_table_HC05 : forall_op (λo.
+ (test_not_impl_pseudo o HC05_not_impl_pseudo     ⊙ le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_pseudo_count HC05 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC05)) ⊗
+ (⊖ (test_not_impl_pseudo o HC05_not_impl_pseudo) ⊙ eq_w16 (get_pseudo_count HC05 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC05) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HC05: modalita' non implementate come da manuale *)
@@ -78,16 +79,16 @@ ndefinition HC05_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_HC05 : forall_instr_mode (λi.
- (test_not_impl_mode i HC05_not_impl_mode     ⊙ le_nat 1 (get_mode_count HC05 i 0 opcode_table_HC05)) ⊗
- (⊖ (test_not_impl_mode i HC05_not_impl_mode) ⊙ eq_nat (get_mode_count HC05 i 0 opcode_table_HC05) 0))
+nlemma ok_mode_table_HC05 : forall_im (λi.
+ (test_not_impl_mode i HC05_not_impl_mode     ⊙ le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_mode_count HC05 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC05)) ⊗
+ (⊖ (test_not_impl_mode i HC05_not_impl_mode) ⊙ eq_w16 (get_mode_count HC05 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC05) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 nlemma ok_OpIm_table_HC05 :
- forall_instr_mode (λi:instr_mode.
- forall_opcode     (λop:opcode.
-  le_nat (get_OpIm_count HC05 (anyOP HC05 op) i 0 opcode_table_HC05) 1)) = true.
- napply (refl_eq ??).
+ forall_im (λi:instr_mode.
+ forall_op (λop:opcode.
+  le_w16 (get_OpIm_count HC05 (anyOP HC05 op) i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HC05) 〈〈x0,x0〉:〈x0,x1〉〉)) = true.
+ napply refl_eq.
 nqed.