]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/freescale/table_HCS08_tests.ma
freescale porting, work in progress
[helm.git] / helm / software / matita / contribs / ng_assembly / freescale / table_HCS08_tests.ma
index 9880fadb9cdac8497c50264728002e260a1def89..bd4c8cd5fd1e553039230aa6508ca097e2935ebc 100755 (executable)
 (**************************************************************************)
 
 (* ********************************************************************** *)
-(*                           Progetto FreeScale                           *)
+(*                          Progetto FreeScale                            *)
 (*                                                                        *)
-(* Sviluppato da:                                                         *)
-(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*   Sviluppato da: Ing. Cosimo Oliboni, oliboni@cs.unibo.it              *)
+(*   Ultima modifica: 05/08/2009                                          *)
 (*                                                                        *)
-(* Questo materiale fa parte della tesi:                                  *)
-(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
-(*                                                                        *)
-(*                    data ultima modifica 15/11/2007                     *)
 (* ********************************************************************** *)
 
 include "freescale/table_HCS08.ma".
+include "freescale/opcode.ma".
 
 (* ****************** *)
 (* TABELLA DELL'HCS08 *)
@@ -37,11 +34,11 @@ ndefinition HCS08_not_impl_byte ≝
  ;〈xA,xC〉
  ].
 
-nlemma ok_byte_table_HCS08 : forall_byte8 (λb.
- (test_not_impl_byte b HCS08_not_impl_byte     ⊙ eq_nat (get_byte_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
- (⊖ (test_not_impl_byte b HCS08_not_impl_byte) ⊙ eq_nat (get_byte_count HCS08 b 0 opcode_table_HCS08) 0))
+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 ??).
+ napply refl_eq.
 nqed.
 
 (* HCS08: opcode non implementati come da manuale (0x9E+byte) *)
@@ -73,22 +70,22 @@ ndefinition HCS08_not_impl_word ≝
  ;〈xF,x0〉;〈xF,x1〉;〈xF,x2〉;〈xF,x4〉;〈xF,x5〉;〈xF,x6〉;〈xF,x7〉;〈xF,x8〉;〈xF,x9〉;〈xF,xA〉;〈xF,xB〉;〈xF,xC〉;〈xF,xD〉
  ].
 
-nlemma ok_word_table_HCS08 : forall_byte8 (λb.
- (test_not_impl_byte b HCS08_not_impl_word     ⊙ eq_nat (get_word_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
- (⊖ (test_not_impl_byte b HCS08_not_impl_word) ⊙ eq_nat (get_word_count HCS08 b 0 opcode_table_HCS08) 0))
+nlemma ok_word_table_HCS08 : forall_b8 (λb.
+ (test_not_impl_byte b HCS08_not_impl_word     ⊙ eq_w16 (get_word_count HCS08 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 b 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HCS08: pseudocodici non implementati come da manuale *)
 ndefinition HCS08_not_impl_pseudo ≝
  [ SHA ; SLA ].
 
-nlemma ok_pseudo_table_HCS08 : forall_opcode (λo.
- (test_not_impl_pseudo o HCS08_not_impl_pseudo     ⊙ le_nat 1 (get_pseudo_count HCS08 o 0 opcode_table_HCS08)) ⊗
- (⊖ (test_not_impl_pseudo o HCS08_not_impl_pseudo) ⊙ eq_nat (get_pseudo_count HCS08 o 0 opcode_table_HCS08) 0))
+nlemma ok_pseudo_table_HCS08 : forall_op (λo.
+ (test_not_impl_pseudo o HCS08_not_impl_pseudo     ⊙ le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_pseudo_count HCS08 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08)) ⊗
+ (⊖ (test_not_impl_pseudo o HCS08_not_impl_pseudo) ⊙ eq_w16 (get_pseudo_count HCS08 o 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HCS08: modalita' non implementate come da manuale *)
@@ -106,16 +103,16 @@ ndefinition HCS08_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_HCS08 : forall_instr_mode (λi.
- (test_not_impl_mode i HCS08_not_impl_mode     ⊙ le_nat 1 (get_mode_count HCS08 i 0 opcode_table_HCS08)) ⊗
- (⊖ (test_not_impl_mode i HCS08_not_impl_mode) ⊙ eq_nat (get_mode_count HCS08 i 0 opcode_table_HCS08) 0))
+nlemma ok_mode_table_HCS08 : forall_im (λi.
+ (test_not_impl_mode i HCS08_not_impl_mode     ⊙ le_w16 〈〈x0,x0〉:〈x0,x1〉〉 (get_mode_count HCS08 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08)) ⊗
+ (⊖ (test_not_impl_mode i HCS08_not_impl_mode) ⊙ eq_w16 (get_mode_count HCS08 i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x0〉〉))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 nlemma ok_OpIm_table_HCS08 :
- forall_instr_mode (λi:instr_mode.
- forall_opcode     (λop:opcode.
-  le_nat (get_OpIm_count HCS08 (anyOP HCS08 op) i 0 opcode_table_HCS08) 1)) = true.
- napply (refl_eq ??).
+ forall_im (λi:instr_mode.
+ forall_op (λop:opcode.
+  le_w16 (get_OpIm_count HCS08 (anyOP HCS08 op) i 〈〈x0,x0〉:〈x0,x0〉〉 opcode_table_HCS08) 〈〈x0,x0〉:〈x0,x1〉〉)) = true.
+ napply refl_eq.
 nqed.