]> matita.cs.unibo.it Git - helm.git/commitdiff
reworked freescale stuff to put \m as a left parameter
authorEnrico Tassi <enrico.tassi@inria.fr>
Thu, 6 Mar 2008 15:37:52 +0000 (15:37 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Thu, 6 Mar 2008 15:37:52 +0000 (15:37 +0000)
12 files changed:
helm/software/matita/library/freescale/load_write.ma
helm/software/matita/library/freescale/medium_tests_lemmas.ma
helm/software/matita/library/freescale/multivm.ma
helm/software/matita/library/freescale/opcode.ma
helm/software/matita/library/freescale/table_HC05.ma
helm/software/matita/library/freescale/table_HC05_tests.ma [new file with mode: 0644]
helm/software/matita/library/freescale/table_HC08.ma
helm/software/matita/library/freescale/table_HC08_tests.ma [new file with mode: 0644]
helm/software/matita/library/freescale/table_HCS08.ma
helm/software/matita/library/freescale/table_HCS08_tests.ma [new file with mode: 0644]
helm/software/matita/library/freescale/table_RS08.ma
helm/software/matita/library/freescale/table_RS08_tests.ma [new file with mode: 0644]

index 2eee12532fbf9efdd5b71248c1f668a1dca83800..a5fae5375e8f52197421f35ffddded2e7e5e6fdf 100644 (file)
@@ -260,10 +260,12 @@ definition filtered_inc_w16 ≝
 λm:mcu_type.λt:memory_impl.λs:any_status m t.λw:word16.
  get_pc_reg m t (set_pc_reg m t s (succ_w16 w)).
 
-let rec filtered_plus_w16 (m:mcu_type) (t:memory_impl) (s:any_status m t) (w:word16) (n:nat) on n ≝
+definition filtered_plus_w16 := \lambda m:mcu_type.
+let rec filtered_plus_w16  (t:memory_impl) (s:any_status m t) (w:word16) (n:nat) on n ≝
  match n with
   [ O ⇒ w
-  | S n' ⇒ filtered_plus_w16 m t s (filtered_inc_w16 m t s w) n' ].
+  | S n' ⇒ filtered_plus_w16 t s (filtered_inc_w16 m t s w) n' ]
+in filtered_plus_w16.
 
 (* 
    errore1: non esiste traduzione ILL_OP
index e708f57bcc0ff93d0a869a4210b2aaaae4076992..44f770ae34e67d41f50e9e7ca41ba948ee4a7f43 100644 (file)
@@ -127,6 +127,7 @@ axiom MSB16_of_make_word16 :
  ∀bh,bl:byte8.
   MSB_w16 (mk_word16 bh bl) = MSB_b8 bh.
 
+(*
 lemma execute_HCS08_LDHX_maIMM2 :
  ∀t:memory_impl.∀s:any_status HCS08 t.∀bhigh,blow:byte8.
   (get_clk_desc HCS08 t s = None ?) →
@@ -168,7 +169,6 @@ lemma execute_HCS08_LDHX_maIMM2 :
    [ true ⇒ tick_execute HCS08 t s (anyOP HCS08 LDHX) MODE_IMM2 (filtered_inc_w16 HCS08 t s (get_pc_reg HCS08 t s))
    | false ⇒ TickOK (any_status HCS08 t) (set_clk_desc HCS08 t s (Some ? (quintupleT ????? 〈x0,x1〉     (anyOP HCS08 LDHX) MODE_IMM2 〈x0,x3〉 (filtered_inc_w16 HCS08 t s (get_pc_reg HCS08 t s)))))]);
  whd in ⊢ (? ? match % in tick_result return ? with [_⇒?|_⇒?|_⇒?] ?);
-
  letin status1 ≝ (set_clk_desc HCS08 t s (Some ? (quintupleT ????? 〈x0,x1〉 (anyOP HCS08 LDHX) MODE_IMM2 〈x0,x3〉 (filtered_inc_w16 HCS08 t s (get_pc_reg HCS08 t s)))));
  change in ⊢ (? ? % ?) with
   (match 2 with 
@@ -436,3 +436,4 @@ lemma execute_HCS08_STHX_maSP1 :
  [2: elim H; reflexivity ]
  reflexivity.
 qed.
+*)
index c98c1ad35a04e30edc83c6932028f206b47cd62c..2e63c2054a26a598983e7131da91c76ea6733a47 100644 (file)
@@ -1304,12 +1304,14 @@ definition tick ≝
 (* ESECUZIONE *)
 (* ********** *)
 
-let rec execute (m:mcu_type) (t:memory_impl) (s:tick_result (any_status m t)) (n:nat) on n ≝
+definition execute := \lambda m:mcu_type.
+let rec execute  (t:memory_impl) (s:tick_result (any_status m t)) (n:nat) on n ≝
  match s with
   [ TickERR s' error ⇒ TickERR ? s' error
   | TickSUSP s' susp ⇒ TickSUSP ? s' susp
-  | TickOK s'        ⇒ match n with [ O ⇒ TickOK ? s' | S n' ⇒ execute m t (tick m t s') n' ]
-  ].
+  | TickOK s'        ⇒ match n with [ O ⇒ TickOK ? s' | S n' ⇒ execute t (tick m t s') n' ]
+  ]
+in execute.
 
 lemma breakpoint:
  ∀m,u,s,n1,n2. execute m u s (n1 + n2) = execute m u (execute m u s n1) n2.
index e14ea9001edeee6a04c7e09ebdb2c2896366d40f..62199d7d354142d8d35fcb3d816351e423a2ec2c 100644 (file)
@@ -371,56 +371,64 @@ definition eqim ≝
 (* ********************************************* *)
 
 (* su tutta la lista quante volte compare il byte *)
-let rec get_byte_count (m:mcu_type) (b:byte8) (c:nat)
+definition get_byte_count := \lambda m:mcu_type.
+let rec get_byte_count  (b:byte8) (c:nat)
  (l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
  match l with
   [ nil ⇒ c
   | cons hd tl ⇒ match thd4T ???? hd with
    [ Byte b' ⇒ match eq_b8 b b' with
-    [ true ⇒ get_byte_count b (S c) tl
-    | false ⇒ get_byte_count b c tl
+    [ true ⇒ get_byte_count b (S c) tl
+    | false ⇒ get_byte_count b c tl
     ]
-   | Word _ ⇒ get_byte_count b c tl
+   | Word _ ⇒ get_byte_count b c tl
    ]
-  ].
+  ]
+in get_byte_count.
 
 (* su tutta la lista quante volte compare la word (0x9E+byte) *)
-let rec get_word_count (m:mcu_type) (b:byte8) (c:nat)
+definition get_word_count := \lambda m:mcu_type.
+let rec get_word_count (b:byte8) (c:nat)
  (l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
  match l with
   [ nil ⇒ c
   | cons hd tl ⇒ match thd4T ???? hd with
-   [ Byte _ ⇒ get_word_count b c tl
+   [ Byte _ ⇒ get_word_count b c tl
    | Word w ⇒ match eq_w16 〈〈x9,xE〉:b〉 w with
-    [ true ⇒ get_word_count b (S c) tl
-    | false ⇒ get_word_count b c tl
+    [ true ⇒ get_word_count b (S c) tl
+    | false ⇒ get_word_count b c tl
     ]
    ]
-  ].
+  ]
+in get_word_count.
 
 (* su tutta la lista quante volte compare lo pseudocodice *)
-let rec get_pseudo_count (m:mcu_type) (o:opcode) (c:nat)
+definition get_pseudo_count := \lambda m:mcu_type.
+let rec get_pseudo_count  (o:opcode) (c:nat)
  (l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
  match l with
   [ nil ⇒ c
   | cons hd tl ⇒ match fst4T ???? hd with
    [ anyOP o' ⇒ match eqop m (anyOP m o) (anyOP m o') with
-    [ true ⇒ get_pseudo_count o (S c) tl
-    | false ⇒ get_pseudo_count o c tl
+    [ true ⇒ get_pseudo_count o (S c) tl
+    | false ⇒ get_pseudo_count o c tl
     ]
    ]
-  ].
+  ]
+in get_pseudo_count.
 
 (* su tutta la lista quante volte compare la modalita' *)
-let rec get_mode_count (m:mcu_type) (i:instr_mode) (c:nat)
+definition get_mode_count := \lambda m:mcu_type.
+let rec get_mode_count  (i:instr_mode) (c:nat)
  (l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
  match l with
   [ nil ⇒ c
   | cons hd tl ⇒ match eqim (snd4T ???? hd) i with
-   [ true ⇒ get_mode_count i (S c) tl
-   | false ⇒ get_mode_count i c tl
+   [ true ⇒ get_mode_count i (S c) tl
+   | false ⇒ get_mode_count i c tl
    ]
-  ].
+  ]
+  in get_mode_count.
 
 (* b e' non implementato? *)
 let rec test_not_impl_byte (b:byte8) (l:list byte8) on l ≝
@@ -453,17 +461,19 @@ let rec test_not_impl_mode (i:instr_mode) (l:list instr_mode) on l ≝
   ].
 
 (* su tutta la lista quante volte compare la coppia opcode,instr_mode *)
-let rec get_OpIm_count (m:mcu_type) (o:any_opcode m) (i:instr_mode) (c:nat)
+definition get_OpIm_count := \lambda  m:mcu_type.
+let rec get_OpIm_count  (o:any_opcode m) (i:instr_mode) (c:nat)
  (l:list (Prod4T (any_opcode m) instr_mode byte8_or_word16 byte8)) on l ≝
  match l with
   [ nil ⇒ c
   | cons hd tl ⇒
    match (eqop m o (fst4T ???? hd)) ⊗
          (eqim i (snd4T ???? hd)) with
-    [ true ⇒ get_OpIm_count o i (S c) tl
-    | false ⇒ get_OpIm_count o i c tl
+    [ true ⇒ get_OpIm_count o i (S c) tl
+    | false ⇒ get_OpIm_count o i c tl
     ] 
-  ].
+  ]
+  in get_OpIm_count.
 
 (* iteratore sugli opcode *)
 definition forall_opcode ≝ λP.
index 265bf693719688eba15c724c730ff70c9e36cd03..e0c9b33e750f87123c006b01dc57fbd3afad4648 100644 (file)
@@ -379,70 +379,3 @@ definition opcode_table_HC05 ≝
  opcode_table_HC05_25 @ opcode_table_HC05_26 @ opcode_table_HC05_27 @ opcode_table_HC05_28 @
  opcode_table_HC05_29 @ opcode_table_HC05_30 @ opcode_table_HC05_31.
 
-(* CORRETTEZZA *)
-
-(* HC05: opcode non implementati come da manuale *)
-definition HC05_not_impl_byte ≝
- [〈x3,x1〉;〈x3,x2〉;〈x3,x5〉;〈x3,xB〉;〈x3,xE〉
- ;〈x4,x1〉;〈x4,x5〉;〈x4,xB〉;〈x4,xE〉
- ;〈x5,x1〉;〈x5,x2〉;〈x5,x5〉;〈x5,xB〉;〈x5,xE〉
- ;〈x6,x1〉;〈x6,x2〉;〈x6,x5〉;〈x6,xB〉;〈x6,xE〉
- ;〈x7,x1〉;〈x7,x2〉;〈x7,x5〉;〈x7,xB〉;〈x7,xE〉
- ;〈x8,x2〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉
- ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,xE〉
- ;〈xA,x7〉;〈xA,xC〉;〈xA,xF〉
- ].
-
-lemma ok_byte_table_HC05 : forall_byte8 (λb.
- (test_not_impl_byte b HC05_not_impl_byte     ⊙ eqb (get_byte_count HC05 b 0 opcode_table_HC05) 1) ⊗
- (⊖ (test_not_impl_byte b HC05_not_impl_byte) ⊙ eqb (get_byte_count HC05 b 0 opcode_table_HC05) 0))
- = true.
- reflexivity.
-qed.
-
-(* HC05: pseudocodici non implementati come da manuale *)
-definition HC05_not_impl_pseudo ≝
- [ AIS ; AIX ; BGE ; BGND ; BGT ; BLE ; BLT ; CBEQA ; CBEQX ; CPHX ; DAA
- ; DBNZ ; DIV ; LDHX ; MOV ; NSA ; PSHA ; PSHH ; PSHX ; PULA ; PULH ; PULX
- ; SHA ; SLA ; STHX ; TAP ; TPA ; TSX ; TXS ].
-
-lemma ok_pseudo_table_HC05 : forall_opcode (λo.
- (test_not_impl_pseudo o HC05_not_impl_pseudo     ⊙ leb 1 (get_pseudo_count HC05 o 0 opcode_table_HC05)) ⊗
- (⊖ (test_not_impl_pseudo o HC05_not_impl_pseudo) ⊙ eqb (get_pseudo_count HC05 o 0 opcode_table_HC05) 0))
- = true.
- reflexivity.
-qed.
-
-(* HC05: modalita' non implementate come da manuale *)
-definition HC05_not_impl_mode ≝
- [ MODE_INHH ; MODE_IMM2 ; MODE_SP1 ; MODE_SP2 ; MODE_DIR1_to_DIR1
- ; MODE_IMM1_to_DIR1 ; MODE_IX0p_to_DIR1 ; MODE_DIR1_to_IX0p
- ; MODE_INHA_and_IMM1 ; MODE_INHX_and_IMM1 ; MODE_IMM1_and_IMM1
- ; MODE_DIR1_and_IMM1 ; MODE_IX0_and_IMM1 ; MODE_IX0p_and_IMM1
- ; MODE_IX1_and_IMM1 ; MODE_IX1p_and_IMM1 ; MODE_SP1_and_IMM1
- ; MODE_TNY x0 ; MODE_TNY x1 ; MODE_TNY x2 ; MODE_TNY x3
- ; MODE_TNY x4 ; MODE_TNY x5 ; MODE_TNY x6 ; MODE_TNY x7
- ; MODE_TNY x8 ; MODE_TNY x9 ; MODE_TNY xA ; MODE_TNY xB
- ; MODE_TNY xC ; MODE_TNY xD ; MODE_TNY xE ; MODE_TNY xF
- ; MODE_SRT t00 ; MODE_SRT t01 ; MODE_SRT t02 ; MODE_SRT t03
- ; MODE_SRT t04 ; MODE_SRT t05 ; MODE_SRT t06 ; MODE_SRT t07
- ; MODE_SRT t08 ; MODE_SRT t09 ; MODE_SRT t0A ; MODE_SRT t0B
- ; MODE_SRT t0C ; MODE_SRT t0D ; MODE_SRT t0E ; MODE_SRT t0F
- ; MODE_SRT t10 ; MODE_SRT t11 ; MODE_SRT t12 ; MODE_SRT t13
- ; MODE_SRT t14 ; MODE_SRT t15 ; MODE_SRT t16 ; MODE_SRT t17
- ; MODE_SRT t18 ; MODE_SRT t19 ; MODE_SRT t1A ; MODE_SRT t1B
- ; MODE_SRT t1C ; MODE_SRT t1D ; MODE_SRT t1E ; MODE_SRT t1F ].
-
-lemma ok_mode_table_HC05 : forall_instr_mode (λi.
- (test_not_impl_mode i HC05_not_impl_mode     ⊙ leb 1 (get_mode_count HC05 i 0 opcode_table_HC05)) ⊗
- (⊖ (test_not_impl_mode i HC05_not_impl_mode) ⊙ eqb (get_mode_count HC05 i 0 opcode_table_HC05) 0))
- = true.
- reflexivity.
-qed.
-
-lemma ok_OpIm_table_HC05 :
- forall_instr_mode (λi:instr_mode.
- forall_opcode     (λop:opcode.
-  leb (get_OpIm_count HC05 (anyOP HC05 op) i 0 opcode_table_HC05) 1)) = true.
- reflexivity.
-qed.
diff --git a/helm/software/matita/library/freescale/table_HC05_tests.ma b/helm/software/matita/library/freescale/table_HC05_tests.ma
new file mode 100644 (file)
index 0000000..5f92c13
--- /dev/null
@@ -0,0 +1,95 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+include "freescale/table_HC05.ma".
+
+(* CORRETTEZZA *)
+
+(* HC05: opcode non implementati come da manuale *)
+definition HC05_not_impl_byte ≝
+ [〈x3,x1〉;〈x3,x2〉;〈x3,x5〉;〈x3,xB〉;〈x3,xE〉
+ ;〈x4,x1〉;〈x4,x5〉;〈x4,xB〉;〈x4,xE〉
+ ;〈x5,x1〉;〈x5,x2〉;〈x5,x5〉;〈x5,xB〉;〈x5,xE〉
+ ;〈x6,x1〉;〈x6,x2〉;〈x6,x5〉;〈x6,xB〉;〈x6,xE〉
+ ;〈x7,x1〉;〈x7,x2〉;〈x7,x5〉;〈x7,xB〉;〈x7,xE〉
+ ;〈x8,x2〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉
+ ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,xE〉
+ ;〈xA,x7〉;〈xA,xC〉;〈xA,xF〉
+ ].
+
+lemma ok_byte_table_HC05 : forall_byte8 (λb.
+ (test_not_impl_byte b HC05_not_impl_byte     ⊙ eqb (get_byte_count HC05 b 0 opcode_table_HC05) 1) ⊗
+ (⊖ (test_not_impl_byte b HC05_not_impl_byte) ⊙ eqb (get_byte_count HC05 b 0 opcode_table_HC05) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HC05: pseudocodici non implementati come da manuale *)
+definition HC05_not_impl_pseudo ≝
+ [ AIS ; AIX ; BGE ; BGND ; BGT ; BLE ; BLT ; CBEQA ; CBEQX ; CPHX ; DAA
+ ; DBNZ ; DIV ; LDHX ; MOV ; NSA ; PSHA ; PSHH ; PSHX ; PULA ; PULH ; PULX
+ ; SHA ; SLA ; STHX ; TAP ; TPA ; TSX ; TXS ].
+
+lemma ok_pseudo_table_HC05 : forall_opcode (λo.
+ (test_not_impl_pseudo o HC05_not_impl_pseudo     ⊙ leb 1 (get_pseudo_count HC05 o 0 opcode_table_HC05)) ⊗
+ (⊖ (test_not_impl_pseudo o HC05_not_impl_pseudo) ⊙ eqb (get_pseudo_count HC05 o 0 opcode_table_HC05) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HC05: modalita' non implementate come da manuale *)
+definition HC05_not_impl_mode ≝
+ [ MODE_INHH ; MODE_IMM2 ; MODE_SP1 ; MODE_SP2 ; MODE_DIR1_to_DIR1
+ ; MODE_IMM1_to_DIR1 ; MODE_IX0p_to_DIR1 ; MODE_DIR1_to_IX0p
+ ; MODE_INHA_and_IMM1 ; MODE_INHX_and_IMM1 ; MODE_IMM1_and_IMM1
+ ; MODE_DIR1_and_IMM1 ; MODE_IX0_and_IMM1 ; MODE_IX0p_and_IMM1
+ ; MODE_IX1_and_IMM1 ; MODE_IX1p_and_IMM1 ; MODE_SP1_and_IMM1
+ ; MODE_TNY x0 ; MODE_TNY x1 ; MODE_TNY x2 ; MODE_TNY x3
+ ; MODE_TNY x4 ; MODE_TNY x5 ; MODE_TNY x6 ; MODE_TNY x7
+ ; MODE_TNY x8 ; MODE_TNY x9 ; MODE_TNY xA ; MODE_TNY xB
+ ; MODE_TNY xC ; MODE_TNY xD ; MODE_TNY xE ; MODE_TNY xF
+ ; MODE_SRT t00 ; MODE_SRT t01 ; MODE_SRT t02 ; MODE_SRT t03
+ ; MODE_SRT t04 ; MODE_SRT t05 ; MODE_SRT t06 ; MODE_SRT t07
+ ; MODE_SRT t08 ; MODE_SRT t09 ; MODE_SRT t0A ; MODE_SRT t0B
+ ; MODE_SRT t0C ; MODE_SRT t0D ; MODE_SRT t0E ; MODE_SRT t0F
+ ; MODE_SRT t10 ; MODE_SRT t11 ; MODE_SRT t12 ; MODE_SRT t13
+ ; MODE_SRT t14 ; MODE_SRT t15 ; MODE_SRT t16 ; MODE_SRT t17
+ ; MODE_SRT t18 ; MODE_SRT t19 ; MODE_SRT t1A ; MODE_SRT t1B
+ ; MODE_SRT t1C ; MODE_SRT t1D ; MODE_SRT t1E ; MODE_SRT t1F ].
+
+lemma ok_mode_table_HC05 : forall_instr_mode (λi.
+ (test_not_impl_mode i HC05_not_impl_mode     ⊙ leb 1 (get_mode_count HC05 i 0 opcode_table_HC05)) ⊗
+ (⊖ (test_not_impl_mode i HC05_not_impl_mode) ⊙ eqb (get_mode_count HC05 i 0 opcode_table_HC05) 0))
+ = true.
+ reflexivity.
+qed.
+
+lemma ok_OpIm_table_HC05 :
+ forall_instr_mode (λi:instr_mode.
+ forall_opcode     (λop:opcode.
+  leb (get_OpIm_count HC05 (anyOP HC05 op) i 0 opcode_table_HC05) 1)) = true.
+ reflexivity.
+qed.
index 9fa67de57d7239db4bdd039d07b586010804f922..7448901115c1671de86af205a07ac86c435668fb 100644 (file)
@@ -476,99 +476,3 @@ opcode_table_HC08_25 @ opcode_table_HC08_26 @ opcode_table_HC08_27 @ opcode_tabl
 opcode_table_HC08_29 @ opcode_table_HC08_30 @ opcode_table_HC08_31 @ opcode_table_HC08_32 @
 opcode_table_HC08_33 @ opcode_table_HC08_34 @ opcode_table_HC08_35.
 
-(* CORRETTEZZA *)
-
-(* HC08: opcode non implementati come da manuale (byte) *)
-definition HC08_not_impl_byte ≝
- [〈x3,x2〉;〈x3,xE〉
- ;〈x8,x2〉;〈x8,xD〉
- ;〈x9,x6〉;〈x9,xE〉
- ;〈xA,xC〉
- ].
-
-lemma ok_byte_table_HC08 : forall_byte8 (λb.
- (test_not_impl_byte b HC08_not_impl_byte     ⊙ eqb (get_byte_count HC08 b 0 opcode_table_HC08) 1) ⊗
- (⊖ (test_not_impl_byte b HC08_not_impl_byte) ⊙ eqb (get_byte_count HC08 b 0 opcode_table_HC08) 0))
- = true.
- reflexivity.
-qed.
-
-(* HC08: opcode non implementati come da manuale (0x9E+byte) *)
-definition HC08_not_impl_word ≝
- [〈x0,x0〉;〈x0,x1〉;〈x0,x2〉;〈x0,x3〉;〈x0,x4〉;〈x0,x5〉;〈x0,x6〉;〈x0,x7〉
- ;〈x0,x8〉;〈x0,x9〉;〈x0,xA〉;〈x0,xB〉;〈x0,xC〉;〈x0,xD〉;〈x0,xE〉;〈x0,xF〉
- ;〈x1,x0〉;〈x1,x1〉;〈x1,x2〉;〈x1,x3〉;〈x1,x4〉;〈x1,x5〉;〈x1,x6〉;〈x1,x7〉
- ;〈x1,x8〉;〈x1,x9〉;〈x1,xA〉;〈x1,xB〉;〈x1,xC〉;〈x1,xD〉;〈x1,xE〉;〈x1,xF〉
- ;〈x2,x0〉;〈x2,x1〉;〈x2,x2〉;〈x2,x3〉;〈x2,x4〉;〈x2,x5〉;〈x2,x6〉;〈x2,x7〉
- ;〈x2,x8〉;〈x2,x9〉;〈x2,xA〉;〈x2,xB〉;〈x2,xC〉;〈x2,xD〉;〈x2,xE〉;〈x2,xF〉
- ;〈x3,x0〉;〈x3,x1〉;〈x3,x2〉;〈x3,x3〉;〈x3,x4〉;〈x3,x5〉;〈x3,x6〉;〈x3,x7〉
- ;〈x3,x8〉;〈x3,x9〉;〈x3,xA〉;〈x3,xB〉;〈x3,xC〉;〈x3,xD〉;〈x3,xE〉;〈x3,xF〉
- ;〈x4,x0〉;〈x4,x1〉;〈x4,x2〉;〈x4,x3〉;〈x4,x4〉;〈x4,x5〉;〈x4,x6〉;〈x4,x7〉
- ;〈x4,x8〉;〈x4,x9〉;〈x4,xA〉;〈x4,xB〉;〈x4,xC〉;〈x4,xD〉;〈x4,xE〉;〈x4,xF〉
- ;〈x5,x0〉;〈x5,x1〉;〈x5,x2〉;〈x5,x3〉;〈x5,x4〉;〈x5,x5〉;〈x5,x6〉;〈x5,x7〉
- ;〈x5,x8〉;〈x5,x9〉;〈x5,xA〉;〈x5,xB〉;〈x5,xC〉;〈x5,xD〉;〈x5,xE〉;〈x5,xF〉
- ;〈x6,x2〉;〈x6,x5〉;〈x6,xE〉
- ;〈x7,x0〉;〈x7,x1〉;〈x7,x2〉;〈x7,x3〉;〈x7,x4〉;〈x7,x5〉;〈x7,x6〉;〈x7,x7〉
- ;〈x7,x8〉;〈x7,x9〉;〈x7,xA〉;〈x7,xB〉;〈x7,xC〉;〈x7,xD〉;〈x7,xE〉;〈x7,xF〉
- ;〈x8,x0〉;〈x8,x1〉;〈x8,x2〉;〈x8,x3〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉
- ;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉;〈x8,xE〉;〈x8,xF〉
- ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,x7〉
- ;〈x9,x8〉;〈x9,x9〉;〈x9,xA〉;〈x9,xB〉;〈x9,xC〉;〈x9,xD〉;〈x9,xE〉;〈x9,xF〉
- ;〈xA,x0〉;〈xA,x1〉;〈xA,x2〉;〈xA,x3〉;〈xA,x4〉;〈xA,x5〉;〈xA,x6〉;〈xA,x7〉
- ;〈xA,x8〉;〈xA,x9〉;〈xA,xA〉;〈xA,xB〉;〈xA,xC〉;〈xA,xD〉;〈xA,xE〉;〈xA,xF〉
- ;〈xB,x0〉;〈xB,x1〉;〈xB,x2〉;〈xB,x3〉;〈xB,x4〉;〈xB,x5〉;〈xB,x6〉;〈xB,x7〉
- ;〈xB,x8〉;〈xB,x9〉;〈xB,xA〉;〈xB,xB〉;〈xB,xC〉;〈xB,xD〉;〈xB,xE〉;〈xB,xF〉
- ;〈xC,x0〉;〈xC,x1〉;〈xC,x2〉;〈xC,x3〉;〈xC,x4〉;〈xC,x5〉;〈xC,x6〉;〈xC,x7〉
- ;〈xC,x8〉;〈xC,x9〉;〈xC,xA〉;〈xC,xB〉;〈xC,xC〉;〈xC,xD〉;〈xC,xE〉;〈xC,xF〉
- ;〈xD,xC〉;〈xD,xD〉
- ;〈xE,xC〉;〈xE,xD〉
- ;〈xF,x0〉;〈xF,x1〉;〈xF,x2〉;〈xF,x3〉;〈xF,x4〉;〈xF,x5〉;〈xF,x6〉;〈xF,x7〉
- ;〈xF,x8〉;〈xF,x9〉;〈xF,xA〉;〈xF,xB〉;〈xF,xC〉;〈xF,xD〉;〈xF,xE〉;〈xF,xF〉
- ].
-
-lemma ok_word_table_HC08 : forall_byte8 (λb.
- (test_not_impl_byte b HC08_not_impl_word     ⊙ eqb (get_word_count HC08 b 0 opcode_table_HC08) 1) ⊗
- (⊖ (test_not_impl_byte b HC08_not_impl_word) ⊙ eqb (get_word_count HC08 b 0 opcode_table_HC08) 0))
- = true.
- reflexivity.
-qed.
-
-(* HC08: pseudocodici non implementati come da manuale *)
-definition HC08_not_impl_pseudo ≝
- [ BGND ; SHA ; SLA ].
-
-lemma ok_pseudo_table_HC08 : forall_opcode (λo.
- (test_not_impl_pseudo o HC08_not_impl_pseudo     ⊙ leb 1 (get_pseudo_count HC08 o 0 opcode_table_HC08)) ⊗
- (⊖ (test_not_impl_pseudo o HC08_not_impl_pseudo) ⊙ eqb (get_pseudo_count HC08 o 0 opcode_table_HC08) 0))
- = true.
- reflexivity.
-qed.
-
-(* HC08: modalita' non implementate come da manuale *)
-definition HC08_not_impl_mode ≝
- [ MODE_TNY x0 ; MODE_TNY x1 ; MODE_TNY x2 ; MODE_TNY x3
- ; MODE_TNY x4 ; MODE_TNY x5 ; MODE_TNY x6 ; MODE_TNY x7
- ; MODE_TNY x8 ; MODE_TNY x9 ; MODE_TNY xA ; MODE_TNY xB
- ; MODE_TNY xC ; MODE_TNY xD ; MODE_TNY xE ; MODE_TNY xF
- ; MODE_SRT t00 ; MODE_SRT t01 ; MODE_SRT t02 ; MODE_SRT t03
- ; MODE_SRT t04 ; MODE_SRT t05 ; MODE_SRT t06 ; MODE_SRT t07
- ; MODE_SRT t08 ; MODE_SRT t09 ; MODE_SRT t0A ; MODE_SRT t0B
- ; MODE_SRT t0C ; MODE_SRT t0D ; MODE_SRT t0E ; MODE_SRT t0F
- ; MODE_SRT t10 ; MODE_SRT t11 ; MODE_SRT t12 ; MODE_SRT t13
- ; MODE_SRT t14 ; MODE_SRT t15 ; MODE_SRT t16 ; MODE_SRT t17
- ; MODE_SRT t18 ; MODE_SRT t19 ; MODE_SRT t1A ; MODE_SRT t1B
- ; MODE_SRT t1C ; MODE_SRT t1D ; MODE_SRT t1E ; MODE_SRT t1F ].
-
-lemma ok_mode_table_HC08 : forall_instr_mode (λi.
- (test_not_impl_mode i HC08_not_impl_mode     ⊙ leb 1 (get_mode_count HC08 i 0 opcode_table_HC08)) ⊗
- (⊖ (test_not_impl_mode i HC08_not_impl_mode) ⊙ eqb (get_mode_count HC08 i 0 opcode_table_HC08) 0))
- = true.
- reflexivity.
-qed.
-
-lemma ok_OpIm_table_HC08 :
- forall_instr_mode (λi:instr_mode.
- forall_opcode     (λop:opcode.
-  leb (get_OpIm_count HC08 (anyOP HC08 op) i 0 opcode_table_HC08) 1)) = true.
- reflexivity.
-qed.
diff --git a/helm/software/matita/library/freescale/table_HC08_tests.ma b/helm/software/matita/library/freescale/table_HC08_tests.ma
new file mode 100644 (file)
index 0000000..be7b06b
--- /dev/null
@@ -0,0 +1,125 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+include "freescale/table_HC08.ma".
+
+(* CORRETTEZZA *)
+
+(* HC08: opcode non implementati come da manuale (byte) *)
+definition HC08_not_impl_byte ≝
+ [〈x3,x2〉;〈x3,xE〉
+ ;〈x8,x2〉;〈x8,xD〉
+ ;〈x9,x6〉;〈x9,xE〉
+ ;〈xA,xC〉
+ ].
+
+lemma ok_byte_table_HC08 : forall_byte8 (λb.
+ (test_not_impl_byte b HC08_not_impl_byte     ⊙ eqb (get_byte_count HC08 b 0 opcode_table_HC08) 1) ⊗
+ (⊖ (test_not_impl_byte b HC08_not_impl_byte) ⊙ eqb (get_byte_count HC08 b 0 opcode_table_HC08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HC08: opcode non implementati come da manuale (0x9E+byte) *)
+definition HC08_not_impl_word ≝
+ [〈x0,x0〉;〈x0,x1〉;〈x0,x2〉;〈x0,x3〉;〈x0,x4〉;〈x0,x5〉;〈x0,x6〉;〈x0,x7〉
+ ;〈x0,x8〉;〈x0,x9〉;〈x0,xA〉;〈x0,xB〉;〈x0,xC〉;〈x0,xD〉;〈x0,xE〉;〈x0,xF〉
+ ;〈x1,x0〉;〈x1,x1〉;〈x1,x2〉;〈x1,x3〉;〈x1,x4〉;〈x1,x5〉;〈x1,x6〉;〈x1,x7〉
+ ;〈x1,x8〉;〈x1,x9〉;〈x1,xA〉;〈x1,xB〉;〈x1,xC〉;〈x1,xD〉;〈x1,xE〉;〈x1,xF〉
+ ;〈x2,x0〉;〈x2,x1〉;〈x2,x2〉;〈x2,x3〉;〈x2,x4〉;〈x2,x5〉;〈x2,x6〉;〈x2,x7〉
+ ;〈x2,x8〉;〈x2,x9〉;〈x2,xA〉;〈x2,xB〉;〈x2,xC〉;〈x2,xD〉;〈x2,xE〉;〈x2,xF〉
+ ;〈x3,x0〉;〈x3,x1〉;〈x3,x2〉;〈x3,x3〉;〈x3,x4〉;〈x3,x5〉;〈x3,x6〉;〈x3,x7〉
+ ;〈x3,x8〉;〈x3,x9〉;〈x3,xA〉;〈x3,xB〉;〈x3,xC〉;〈x3,xD〉;〈x3,xE〉;〈x3,xF〉
+ ;〈x4,x0〉;〈x4,x1〉;〈x4,x2〉;〈x4,x3〉;〈x4,x4〉;〈x4,x5〉;〈x4,x6〉;〈x4,x7〉
+ ;〈x4,x8〉;〈x4,x9〉;〈x4,xA〉;〈x4,xB〉;〈x4,xC〉;〈x4,xD〉;〈x4,xE〉;〈x4,xF〉
+ ;〈x5,x0〉;〈x5,x1〉;〈x5,x2〉;〈x5,x3〉;〈x5,x4〉;〈x5,x5〉;〈x5,x6〉;〈x5,x7〉
+ ;〈x5,x8〉;〈x5,x9〉;〈x5,xA〉;〈x5,xB〉;〈x5,xC〉;〈x5,xD〉;〈x5,xE〉;〈x5,xF〉
+ ;〈x6,x2〉;〈x6,x5〉;〈x6,xE〉
+ ;〈x7,x0〉;〈x7,x1〉;〈x7,x2〉;〈x7,x3〉;〈x7,x4〉;〈x7,x5〉;〈x7,x6〉;〈x7,x7〉
+ ;〈x7,x8〉;〈x7,x9〉;〈x7,xA〉;〈x7,xB〉;〈x7,xC〉;〈x7,xD〉;〈x7,xE〉;〈x7,xF〉
+ ;〈x8,x0〉;〈x8,x1〉;〈x8,x2〉;〈x8,x3〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉
+ ;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉;〈x8,xE〉;〈x8,xF〉
+ ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,x7〉
+ ;〈x9,x8〉;〈x9,x9〉;〈x9,xA〉;〈x9,xB〉;〈x9,xC〉;〈x9,xD〉;〈x9,xE〉;〈x9,xF〉
+ ;〈xA,x0〉;〈xA,x1〉;〈xA,x2〉;〈xA,x3〉;〈xA,x4〉;〈xA,x5〉;〈xA,x6〉;〈xA,x7〉
+ ;〈xA,x8〉;〈xA,x9〉;〈xA,xA〉;〈xA,xB〉;〈xA,xC〉;〈xA,xD〉;〈xA,xE〉;〈xA,xF〉
+ ;〈xB,x0〉;〈xB,x1〉;〈xB,x2〉;〈xB,x3〉;〈xB,x4〉;〈xB,x5〉;〈xB,x6〉;〈xB,x7〉
+ ;〈xB,x8〉;〈xB,x9〉;〈xB,xA〉;〈xB,xB〉;〈xB,xC〉;〈xB,xD〉;〈xB,xE〉;〈xB,xF〉
+ ;〈xC,x0〉;〈xC,x1〉;〈xC,x2〉;〈xC,x3〉;〈xC,x4〉;〈xC,x5〉;〈xC,x6〉;〈xC,x7〉
+ ;〈xC,x8〉;〈xC,x9〉;〈xC,xA〉;〈xC,xB〉;〈xC,xC〉;〈xC,xD〉;〈xC,xE〉;〈xC,xF〉
+ ;〈xD,xC〉;〈xD,xD〉
+ ;〈xE,xC〉;〈xE,xD〉
+ ;〈xF,x0〉;〈xF,x1〉;〈xF,x2〉;〈xF,x3〉;〈xF,x4〉;〈xF,x5〉;〈xF,x6〉;〈xF,x7〉
+ ;〈xF,x8〉;〈xF,x9〉;〈xF,xA〉;〈xF,xB〉;〈xF,xC〉;〈xF,xD〉;〈xF,xE〉;〈xF,xF〉
+ ].
+
+lemma ok_word_table_HC08 : forall_byte8 (λb.
+ (test_not_impl_byte b HC08_not_impl_word     ⊙ eqb (get_word_count HC08 b 0 opcode_table_HC08) 1) ⊗
+ (⊖ (test_not_impl_byte b HC08_not_impl_word) ⊙ eqb (get_word_count HC08 b 0 opcode_table_HC08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HC08: pseudocodici non implementati come da manuale *)
+definition HC08_not_impl_pseudo ≝
+ [ BGND ; SHA ; SLA ].
+
+lemma ok_pseudo_table_HC08 : forall_opcode (λo.
+ (test_not_impl_pseudo o HC08_not_impl_pseudo     ⊙ leb 1 (get_pseudo_count HC08 o 0 opcode_table_HC08)) ⊗
+ (⊖ (test_not_impl_pseudo o HC08_not_impl_pseudo) ⊙ eqb (get_pseudo_count HC08 o 0 opcode_table_HC08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HC08: modalita' non implementate come da manuale *)
+definition HC08_not_impl_mode ≝
+ [ MODE_TNY x0 ; MODE_TNY x1 ; MODE_TNY x2 ; MODE_TNY x3
+ ; MODE_TNY x4 ; MODE_TNY x5 ; MODE_TNY x6 ; MODE_TNY x7
+ ; MODE_TNY x8 ; MODE_TNY x9 ; MODE_TNY xA ; MODE_TNY xB
+ ; MODE_TNY xC ; MODE_TNY xD ; MODE_TNY xE ; MODE_TNY xF
+ ; MODE_SRT t00 ; MODE_SRT t01 ; MODE_SRT t02 ; MODE_SRT t03
+ ; MODE_SRT t04 ; MODE_SRT t05 ; MODE_SRT t06 ; MODE_SRT t07
+ ; MODE_SRT t08 ; MODE_SRT t09 ; MODE_SRT t0A ; MODE_SRT t0B
+ ; MODE_SRT t0C ; MODE_SRT t0D ; MODE_SRT t0E ; MODE_SRT t0F
+ ; MODE_SRT t10 ; MODE_SRT t11 ; MODE_SRT t12 ; MODE_SRT t13
+ ; MODE_SRT t14 ; MODE_SRT t15 ; MODE_SRT t16 ; MODE_SRT t17
+ ; MODE_SRT t18 ; MODE_SRT t19 ; MODE_SRT t1A ; MODE_SRT t1B
+ ; MODE_SRT t1C ; MODE_SRT t1D ; MODE_SRT t1E ; MODE_SRT t1F ].
+
+lemma ok_mode_table_HC08 : forall_instr_mode (λi.
+ (test_not_impl_mode i HC08_not_impl_mode     ⊙ leb 1 (get_mode_count HC08 i 0 opcode_table_HC08)) ⊗
+ (⊖ (test_not_impl_mode i HC08_not_impl_mode) ⊙ eqb (get_mode_count HC08 i 0 opcode_table_HC08) 0))
+ = true.
+ reflexivity.
+qed.
+
+lemma ok_OpIm_table_HC08 :
+ forall_instr_mode (λi:instr_mode.
+ forall_opcode     (λop:opcode.
+  leb (get_OpIm_count HC08 (anyOP HC08 op) i 0 opcode_table_HC08) 1)) = true.
+ reflexivity.
+qed.
+
index f291e6d5833252565d8b93ae2ed11c289166713f..522c98d48efcd2a9cfb70b1818c8d5ef57142920 100644 (file)
@@ -488,94 +488,3 @@ opcode_table_HCS08_25 @ opcode_table_HCS08_26 @ opcode_table_HCS08_27 @ opcode_t
 opcode_table_HCS08_29 @ opcode_table_HCS08_30 @ opcode_table_HCS08_31 @ opcode_table_HCS08_32 @
 opcode_table_HCS08_33 @ opcode_table_HCS08_34 @ opcode_table_HCS08_35.
 
-(* CORRETTEZZA *)
-
-(* HCS08: opcode non implementati come da manuale (byte) *)
-definition HCS08_not_impl_byte ≝
- [〈x8,xD〉
- ;〈x9,xE〉
- ;〈xA,xC〉
- ].
-
-lemma ok_byte_table_HCS08 : forall_byte8 (λb.
- (test_not_impl_byte b HCS08_not_impl_byte     ⊙ eqb (get_byte_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
- (⊖ (test_not_impl_byte b HCS08_not_impl_byte) ⊙ eqb (get_byte_count HCS08 b 0 opcode_table_HCS08) 0))
- = true.
- reflexivity.
-qed.
-
-(* HCS08: opcode non implementati come da manuale (0x9E+byte) *)
-definition HCS08_not_impl_word ≝
- [〈x0,x0〉;〈x0,x1〉;〈x0,x2〉;〈x0,x3〉;〈x0,x4〉;〈x0,x5〉;〈x0,x6〉;〈x0,x7〉
- ;〈x0,x8〉;〈x0,x9〉;〈x0,xA〉;〈x0,xB〉;〈x0,xC〉;〈x0,xD〉;〈x0,xE〉;〈x0,xF〉
- ;〈x1,x0〉;〈x1,x1〉;〈x1,x2〉;〈x1,x3〉;〈x1,x4〉;〈x1,x5〉;〈x1,x6〉;〈x1,x7〉
- ;〈x1,x8〉;〈x1,x9〉;〈x1,xA〉;〈x1,xB〉;〈x1,xC〉;〈x1,xD〉;〈x1,xE〉;〈x1,xF〉
- ;〈x2,x0〉;〈x2,x1〉;〈x2,x2〉;〈x2,x3〉;〈x2,x4〉;〈x2,x5〉;〈x2,x6〉;〈x2,x7〉
- ;〈x2,x8〉;〈x2,x9〉;〈x2,xA〉;〈x2,xB〉;〈x2,xC〉;〈x2,xD〉;〈x2,xE〉;〈x2,xF〉
- ;〈x3,x0〉;〈x3,x1〉;〈x3,x2〉;〈x3,x3〉;〈x3,x4〉;〈x3,x5〉;〈x3,x6〉;〈x3,x7〉
- ;〈x3,x8〉;〈x3,x9〉;〈x3,xA〉;〈x3,xB〉;〈x3,xC〉;〈x3,xD〉;〈x3,xE〉;〈x3,xF〉
- ;〈x4,x0〉;〈x4,x1〉;〈x4,x2〉;〈x4,x3〉;〈x4,x4〉;〈x4,x5〉;〈x4,x6〉;〈x4,x7〉
- ;〈x4,x8〉;〈x4,x9〉;〈x4,xA〉;〈x4,xB〉;〈x4,xC〉;〈x4,xD〉;〈x4,xE〉;〈x4,xF〉
- ;〈x5,x0〉;〈x5,x1〉;〈x5,x2〉;〈x5,x3〉;〈x5,x4〉;〈x5,x5〉;〈x5,x6〉;〈x5,x7〉
- ;〈x5,x8〉;〈x5,x9〉;〈x5,xA〉;〈x5,xB〉;〈x5,xC〉;〈x5,xD〉;〈x5,xE〉;〈x5,xF〉
- ;〈x6,x2〉;〈x6,x5〉;〈x6,xE〉
- ;〈x7,x0〉;〈x7,x1〉;〈x7,x2〉;〈x7,x3〉;〈x7,x4〉;〈x7,x5〉;〈x7,x6〉;〈x7,x7〉
- ;〈x7,x8〉;〈x7,x9〉;〈x7,xA〉;〈x7,xB〉;〈x7,xC〉;〈x7,xD〉;〈x7,xE〉;〈x7,xF〉
- ;〈x8,x0〉;〈x8,x1〉;〈x8,x2〉;〈x8,x3〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉
- ;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉;〈x8,xE〉;〈x8,xF〉
- ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,x7〉
- ;〈x9,x8〉;〈x9,x9〉;〈x9,xA〉;〈x9,xB〉;〈x9,xC〉;〈x9,xD〉;〈x9,xE〉;〈x9,xF〉
- ;〈xA,x0〉;〈xA,x1〉;〈xA,x2〉;〈xA,x3〉;〈xA,x4〉;〈xA,x5〉;〈xA,x6〉;〈xA,x7〉;〈xA,x8〉;〈xA,x9〉;〈xA,xA〉;〈xA,xB〉;〈xA,xC〉;〈xA,xD〉;〈xA,xF〉
- ;〈xB,x0〉;〈xB,x1〉;〈xB,x2〉;〈xB,x3〉;〈xB,x4〉;〈xB,x5〉;〈xB,x6〉;〈xB,x7〉;〈xB,x8〉;〈xB,x9〉;〈xB,xA〉;〈xB,xB〉;〈xB,xC〉;〈xB,xD〉;〈xB,xF〉
- ;〈xC,x0〉;〈xC,x1〉;〈xC,x2〉;〈xC,x3〉;〈xC,x4〉;〈xC,x5〉;〈xC,x6〉;〈xC,x7〉;〈xC,x8〉;〈xC,x9〉;〈xC,xA〉;〈xC,xB〉;〈xC,xC〉;〈xC,xD〉;〈xC,xF〉
- ;〈xD,xC〉;〈xD,xD〉
- ;〈xE,xC〉;〈xE,xD〉
- ;〈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〉
- ].
-
-lemma ok_word_table_HCS08 : forall_byte8 (λb.
- (test_not_impl_byte b HCS08_not_impl_word     ⊙ eqb (get_word_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
- (⊖ (test_not_impl_byte b HCS08_not_impl_word) ⊙ eqb (get_word_count HCS08 b 0 opcode_table_HCS08) 0))
- = true.
- reflexivity.
-qed.
-
-(* HCS08: pseudocodici non implementati come da manuale *)
-definition HCS08_not_impl_pseudo ≝
- [ SHA ; SLA ].
-
-lemma ok_pseudo_table_HCS08 : forall_opcode (λo.
- (test_not_impl_pseudo o HCS08_not_impl_pseudo     ⊙ leb 1 (get_pseudo_count HCS08 o 0 opcode_table_HCS08)) ⊗
- (⊖ (test_not_impl_pseudo o HCS08_not_impl_pseudo) ⊙ eqb (get_pseudo_count HCS08 o 0 opcode_table_HCS08) 0))
- = true.
- reflexivity.
-qed.
-
-(* HCS08: modalita' non implementate come da manuale *)
-definition HCS08_not_impl_mode ≝
- [ MODE_TNY x0 ; MODE_TNY x1 ; MODE_TNY x2 ; MODE_TNY x3
- ; MODE_TNY x4 ; MODE_TNY x5 ; MODE_TNY x6 ; MODE_TNY x7
- ; MODE_TNY x8 ; MODE_TNY x9 ; MODE_TNY xA ; MODE_TNY xB
- ; MODE_TNY xC ; MODE_TNY xD ; MODE_TNY xE ; MODE_TNY xF
- ; MODE_SRT t00 ; MODE_SRT t01 ; MODE_SRT t02 ; MODE_SRT t03
- ; MODE_SRT t04 ; MODE_SRT t05 ; MODE_SRT t06 ; MODE_SRT t07
- ; MODE_SRT t08 ; MODE_SRT t09 ; MODE_SRT t0A ; MODE_SRT t0B
- ; MODE_SRT t0C ; MODE_SRT t0D ; MODE_SRT t0E ; MODE_SRT t0F
- ; MODE_SRT t10 ; MODE_SRT t11 ; MODE_SRT t12 ; MODE_SRT t13
- ; MODE_SRT t14 ; MODE_SRT t15 ; MODE_SRT t16 ; MODE_SRT t17
- ; MODE_SRT t18 ; MODE_SRT t19 ; MODE_SRT t1A ; MODE_SRT t1B
- ; MODE_SRT t1C ; MODE_SRT t1D ; MODE_SRT t1E ; MODE_SRT t1F ].
-
-lemma ok_mode_table_HCS08 : forall_instr_mode (λi.
- (test_not_impl_mode i HCS08_not_impl_mode     ⊙ leb 1 (get_mode_count HCS08 i 0 opcode_table_HCS08)) ⊗
- (⊖ (test_not_impl_mode i HCS08_not_impl_mode) ⊙ eqb (get_mode_count HCS08 i 0 opcode_table_HCS08) 0))
- = true.
- reflexivity.
-qed.
-
-lemma ok_OpIm_table_HCS08 :
- forall_instr_mode (λi:instr_mode.
- forall_opcode     (λop:opcode.
-  leb (get_OpIm_count HCS08 (anyOP HCS08 op) i 0 opcode_table_HCS08) 1)) = true.
- reflexivity.
-qed.
diff --git a/helm/software/matita/library/freescale/table_HCS08_tests.ma b/helm/software/matita/library/freescale/table_HCS08_tests.ma
new file mode 100644 (file)
index 0000000..011d193
--- /dev/null
@@ -0,0 +1,120 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+include "freescale/table_HCS08.ma".
+
+(* CORRETTEZZA *)
+
+(* HCS08: opcode non implementati come da manuale (byte) *)
+definition HCS08_not_impl_byte ≝
+ [〈x8,xD〉
+ ;〈x9,xE〉
+ ;〈xA,xC〉
+ ].
+
+lemma ok_byte_table_HCS08 : forall_byte8 (λb.
+ (test_not_impl_byte b HCS08_not_impl_byte     ⊙ eqb (get_byte_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
+ (⊖ (test_not_impl_byte b HCS08_not_impl_byte) ⊙ eqb (get_byte_count HCS08 b 0 opcode_table_HCS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HCS08: opcode non implementati come da manuale (0x9E+byte) *)
+definition HCS08_not_impl_word ≝
+ [〈x0,x0〉;〈x0,x1〉;〈x0,x2〉;〈x0,x3〉;〈x0,x4〉;〈x0,x5〉;〈x0,x6〉;〈x0,x7〉
+ ;〈x0,x8〉;〈x0,x9〉;〈x0,xA〉;〈x0,xB〉;〈x0,xC〉;〈x0,xD〉;〈x0,xE〉;〈x0,xF〉
+ ;〈x1,x0〉;〈x1,x1〉;〈x1,x2〉;〈x1,x3〉;〈x1,x4〉;〈x1,x5〉;〈x1,x6〉;〈x1,x7〉
+ ;〈x1,x8〉;〈x1,x9〉;〈x1,xA〉;〈x1,xB〉;〈x1,xC〉;〈x1,xD〉;〈x1,xE〉;〈x1,xF〉
+ ;〈x2,x0〉;〈x2,x1〉;〈x2,x2〉;〈x2,x3〉;〈x2,x4〉;〈x2,x5〉;〈x2,x6〉;〈x2,x7〉
+ ;〈x2,x8〉;〈x2,x9〉;〈x2,xA〉;〈x2,xB〉;〈x2,xC〉;〈x2,xD〉;〈x2,xE〉;〈x2,xF〉
+ ;〈x3,x0〉;〈x3,x1〉;〈x3,x2〉;〈x3,x3〉;〈x3,x4〉;〈x3,x5〉;〈x3,x6〉;〈x3,x7〉
+ ;〈x3,x8〉;〈x3,x9〉;〈x3,xA〉;〈x3,xB〉;〈x3,xC〉;〈x3,xD〉;〈x3,xE〉;〈x3,xF〉
+ ;〈x4,x0〉;〈x4,x1〉;〈x4,x2〉;〈x4,x3〉;〈x4,x4〉;〈x4,x5〉;〈x4,x6〉;〈x4,x7〉
+ ;〈x4,x8〉;〈x4,x9〉;〈x4,xA〉;〈x4,xB〉;〈x4,xC〉;〈x4,xD〉;〈x4,xE〉;〈x4,xF〉
+ ;〈x5,x0〉;〈x5,x1〉;〈x5,x2〉;〈x5,x3〉;〈x5,x4〉;〈x5,x5〉;〈x5,x6〉;〈x5,x7〉
+ ;〈x5,x8〉;〈x5,x9〉;〈x5,xA〉;〈x5,xB〉;〈x5,xC〉;〈x5,xD〉;〈x5,xE〉;〈x5,xF〉
+ ;〈x6,x2〉;〈x6,x5〉;〈x6,xE〉
+ ;〈x7,x0〉;〈x7,x1〉;〈x7,x2〉;〈x7,x3〉;〈x7,x4〉;〈x7,x5〉;〈x7,x6〉;〈x7,x7〉
+ ;〈x7,x8〉;〈x7,x9〉;〈x7,xA〉;〈x7,xB〉;〈x7,xC〉;〈x7,xD〉;〈x7,xE〉;〈x7,xF〉
+ ;〈x8,x0〉;〈x8,x1〉;〈x8,x2〉;〈x8,x3〉;〈x8,x4〉;〈x8,x5〉;〈x8,x6〉;〈x8,x7〉
+ ;〈x8,x8〉;〈x8,x9〉;〈x8,xA〉;〈x8,xB〉;〈x8,xC〉;〈x8,xD〉;〈x8,xE〉;〈x8,xF〉
+ ;〈x9,x0〉;〈x9,x1〉;〈x9,x2〉;〈x9,x3〉;〈x9,x4〉;〈x9,x5〉;〈x9,x6〉;〈x9,x7〉
+ ;〈x9,x8〉;〈x9,x9〉;〈x9,xA〉;〈x9,xB〉;〈x9,xC〉;〈x9,xD〉;〈x9,xE〉;〈x9,xF〉
+ ;〈xA,x0〉;〈xA,x1〉;〈xA,x2〉;〈xA,x3〉;〈xA,x4〉;〈xA,x5〉;〈xA,x6〉;〈xA,x7〉;〈xA,x8〉;〈xA,x9〉;〈xA,xA〉;〈xA,xB〉;〈xA,xC〉;〈xA,xD〉;〈xA,xF〉
+ ;〈xB,x0〉;〈xB,x1〉;〈xB,x2〉;〈xB,x3〉;〈xB,x4〉;〈xB,x5〉;〈xB,x6〉;〈xB,x7〉;〈xB,x8〉;〈xB,x9〉;〈xB,xA〉;〈xB,xB〉;〈xB,xC〉;〈xB,xD〉;〈xB,xF〉
+ ;〈xC,x0〉;〈xC,x1〉;〈xC,x2〉;〈xC,x3〉;〈xC,x4〉;〈xC,x5〉;〈xC,x6〉;〈xC,x7〉;〈xC,x8〉;〈xC,x9〉;〈xC,xA〉;〈xC,xB〉;〈xC,xC〉;〈xC,xD〉;〈xC,xF〉
+ ;〈xD,xC〉;〈xD,xD〉
+ ;〈xE,xC〉;〈xE,xD〉
+ ;〈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〉
+ ].
+
+lemma ok_word_table_HCS08 : forall_byte8 (λb.
+ (test_not_impl_byte b HCS08_not_impl_word     ⊙ eqb (get_word_count HCS08 b 0 opcode_table_HCS08) 1) ⊗
+ (⊖ (test_not_impl_byte b HCS08_not_impl_word) ⊙ eqb (get_word_count HCS08 b 0 opcode_table_HCS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HCS08: pseudocodici non implementati come da manuale *)
+definition HCS08_not_impl_pseudo ≝
+ [ SHA ; SLA ].
+
+lemma ok_pseudo_table_HCS08 : forall_opcode (λo.
+ (test_not_impl_pseudo o HCS08_not_impl_pseudo     ⊙ leb 1 (get_pseudo_count HCS08 o 0 opcode_table_HCS08)) ⊗
+ (⊖ (test_not_impl_pseudo o HCS08_not_impl_pseudo) ⊙ eqb (get_pseudo_count HCS08 o 0 opcode_table_HCS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* HCS08: modalita' non implementate come da manuale *)
+definition HCS08_not_impl_mode ≝
+ [ MODE_TNY x0 ; MODE_TNY x1 ; MODE_TNY x2 ; MODE_TNY x3
+ ; MODE_TNY x4 ; MODE_TNY x5 ; MODE_TNY x6 ; MODE_TNY x7
+ ; MODE_TNY x8 ; MODE_TNY x9 ; MODE_TNY xA ; MODE_TNY xB
+ ; MODE_TNY xC ; MODE_TNY xD ; MODE_TNY xE ; MODE_TNY xF
+ ; MODE_SRT t00 ; MODE_SRT t01 ; MODE_SRT t02 ; MODE_SRT t03
+ ; MODE_SRT t04 ; MODE_SRT t05 ; MODE_SRT t06 ; MODE_SRT t07
+ ; MODE_SRT t08 ; MODE_SRT t09 ; MODE_SRT t0A ; MODE_SRT t0B
+ ; MODE_SRT t0C ; MODE_SRT t0D ; MODE_SRT t0E ; MODE_SRT t0F
+ ; MODE_SRT t10 ; MODE_SRT t11 ; MODE_SRT t12 ; MODE_SRT t13
+ ; MODE_SRT t14 ; MODE_SRT t15 ; MODE_SRT t16 ; MODE_SRT t17
+ ; MODE_SRT t18 ; MODE_SRT t19 ; MODE_SRT t1A ; MODE_SRT t1B
+ ; MODE_SRT t1C ; MODE_SRT t1D ; MODE_SRT t1E ; MODE_SRT t1F ].
+
+lemma ok_mode_table_HCS08 : forall_instr_mode (λi.
+ (test_not_impl_mode i HCS08_not_impl_mode     ⊙ leb 1 (get_mode_count HCS08 i 0 opcode_table_HCS08)) ⊗
+ (⊖ (test_not_impl_mode i HCS08_not_impl_mode) ⊙ eqb (get_mode_count HCS08 i 0 opcode_table_HCS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+lemma ok_OpIm_table_HCS08 :
+ forall_instr_mode (λi:instr_mode.
+ forall_opcode     (λop:opcode.
+  leb (get_OpIm_count HCS08 (anyOP HCS08 op) i 0 opcode_table_HCS08) 1)) = true.
+ reflexivity.
+qed.
+
index c7d1a01f144761284fb9b8bd196a7036a3a6dfbc..522d87efeefa1279d95fde0e95e094a1b0e04776 100644 (file)
@@ -397,53 +397,3 @@ opcode_table_RS08_17 @ opcode_table_RS08_18 @ opcode_table_RS08_19 @ opcode_tabl
 opcode_table_RS08_21 @ opcode_table_RS08_22 @ opcode_table_RS08_23 @ opcode_table_RS08_24 @
 opcode_table_RS08_25 @ opcode_table_RS08_26 @ opcode_table_RS08_27.
 
-(* CORRETTEZZA *)
-
-(* RS08: opcode non implementati come da manuale *)
-definition RS08_not_impl_byte ≝
- [〈x3,x2〉;〈x3,x3〉;〈x3,xD〉
- ;〈x4,x0〉;〈x4,x7〉;〈x4,xD〉
- ;〈xA,x3〉;〈xA,x5〉;〈xA,x7〉
- ;〈xB,x3〉;〈xB,x5〉
- ].
-
-lemma ok_byte_table_RS08 : forall_byte8 (λb.
- (test_not_impl_byte b RS08_not_impl_byte     ⊙ eqb (get_byte_count RS08 b 0 opcode_table_RS08) 1) ⊗
- (⊖ (test_not_impl_byte b RS08_not_impl_byte) ⊙ eqb (get_byte_count RS08 b 0 opcode_table_RS08) 0))
- = true.
- reflexivity.
-qed.
-
-(* RS08: pseudocodici non implementati come da manuale *)
-definition RS08_not_impl_pseudo ≝
- [ AIS ; AIX ; ASR ; BGE ; BGT ; BHCC ; BHCS ; BHI ; BIH ; BIL ; BIT ; BLE ; BLS
- ; BLT ; BMC ; BMI ; BMS ; BPL ; BRN ; CBEQX ; CLI ; CPHX ; CPX ; DAA ; DIV
- ; LDHX ; LDX ; MUL ; NEG ; NSA ; PSHA ; PSHH ; PSHX ; PULA ; PULH ; PULX ; RSP  
- ; RTI ; SEI ; STHX ; STX ; SWI ; TAP ; TAX ; TPA ; TST ; TSX ; TXA ; TXS ].
-
-lemma ok_pseudo_table_RS08 : forall_opcode (λo.
- (test_not_impl_pseudo o RS08_not_impl_pseudo     ⊙ leb 1 (get_pseudo_count RS08 o 0 opcode_table_RS08)) ⊗
- (⊖ (test_not_impl_pseudo o RS08_not_impl_pseudo) ⊙ eqb (get_pseudo_count RS08 o 0 opcode_table_RS08) 0))
- = true.
- reflexivity.
-qed.
-
-(* RS08: modalita' non implementate come da manuale *)
-definition RS08_not_impl_mode ≝
- [ MODE_INHX ; MODE_INHH ; MODE_IMM2 ; MODE_IX0 ; MODE_IX1 ; MODE_IX2 ; MODE_SP1 ; MODE_SP2
- ; MODE_IX0p_to_DIR1 ; MODE_DIR1_to_IX0p ; MODE_INHX_and_IMM1 ; MODE_IX0_and_IMM1
- ; MODE_IX0p_and_IMM1 ; MODE_IX1_and_IMM1 ; MODE_IX1p_and_IMM1 ; MODE_SP1_and_IMM1 ].
-
-lemma ok_mode_table_RS08 : forall_instr_mode (λi.
- (test_not_impl_mode i RS08_not_impl_mode     ⊙ leb 1 (get_mode_count RS08 i 0 opcode_table_RS08)) ⊗
- (⊖ (test_not_impl_mode i RS08_not_impl_mode) ⊙ eqb (get_mode_count RS08 i 0 opcode_table_RS08) 0))
- = true.
- reflexivity.
-qed.
-
-lemma ok_OpIm_table_RS08 :
- forall_instr_mode (λi:instr_mode.
- forall_opcode     (λop:opcode.
-  leb (get_OpIm_count RS08 (anyOP RS08 op) i 0 opcode_table_RS08) 1)) = true.
- reflexivity.
-qed.
diff --git a/helm/software/matita/library/freescale/table_RS08_tests.ma b/helm/software/matita/library/freescale/table_RS08_tests.ma
new file mode 100644 (file)
index 0000000..8c1b3cd
--- /dev/null
@@ -0,0 +1,79 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+(* ********************************************************************** *)
+(*                           Progetto FreeScale                           *)
+(*                                                                        *)
+(* Sviluppato da:                                                         *)
+(*   Cosimo Oliboni, oliboni@cs.unibo.it                                  *)
+(*                                                                        *)
+(* Questo materiale fa parte della tesi:                                  *)
+(*   "Formalizzazione Interattiva dei Microcontroller a 8bit FreeScale"   *)
+(*                                                                        *)
+(*                    data ultima modifica 15/11/2007                     *)
+(* ********************************************************************** *)
+
+include "freescale/table_RS08.ma".
+
+(* CORRETTEZZA *)
+
+(* RS08: opcode non implementati come da manuale *)
+definition RS08_not_impl_byte ≝
+ [〈x3,x2〉;〈x3,x3〉;〈x3,xD〉
+ ;〈x4,x0〉;〈x4,x7〉;〈x4,xD〉
+ ;〈xA,x3〉;〈xA,x5〉;〈xA,x7〉
+ ;〈xB,x3〉;〈xB,x5〉
+ ].
+
+lemma ok_byte_table_RS08 : forall_byte8 (λb.
+ (test_not_impl_byte b RS08_not_impl_byte     ⊙ eqb (get_byte_count RS08 b 0 opcode_table_RS08) 1) ⊗
+ (⊖ (test_not_impl_byte b RS08_not_impl_byte) ⊙ eqb (get_byte_count RS08 b 0 opcode_table_RS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* RS08: pseudocodici non implementati come da manuale *)
+definition RS08_not_impl_pseudo ≝
+ [ AIS ; AIX ; ASR ; BGE ; BGT ; BHCC ; BHCS ; BHI ; BIH ; BIL ; BIT ; BLE ; BLS
+ ; BLT ; BMC ; BMI ; BMS ; BPL ; BRN ; CBEQX ; CLI ; CPHX ; CPX ; DAA ; DIV
+ ; LDHX ; LDX ; MUL ; NEG ; NSA ; PSHA ; PSHH ; PSHX ; PULA ; PULH ; PULX ; RSP  
+ ; RTI ; SEI ; STHX ; STX ; SWI ; TAP ; TAX ; TPA ; TST ; TSX ; TXA ; TXS ].
+
+lemma ok_pseudo_table_RS08 : forall_opcode (λo.
+ (test_not_impl_pseudo o RS08_not_impl_pseudo     ⊙ leb 1 (get_pseudo_count RS08 o 0 opcode_table_RS08)) ⊗
+ (⊖ (test_not_impl_pseudo o RS08_not_impl_pseudo) ⊙ eqb (get_pseudo_count RS08 o 0 opcode_table_RS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+(* RS08: modalita' non implementate come da manuale *)
+definition RS08_not_impl_mode ≝
+ [ MODE_INHX ; MODE_INHH ; MODE_IMM2 ; MODE_IX0 ; MODE_IX1 ; MODE_IX2 ; MODE_SP1 ; MODE_SP2
+ ; MODE_IX0p_to_DIR1 ; MODE_DIR1_to_IX0p ; MODE_INHX_and_IMM1 ; MODE_IX0_and_IMM1
+ ; MODE_IX0p_and_IMM1 ; MODE_IX1_and_IMM1 ; MODE_IX1p_and_IMM1 ; MODE_SP1_and_IMM1 ].
+
+lemma ok_mode_table_RS08 : forall_instr_mode (λi.
+ (test_not_impl_mode i RS08_not_impl_mode     ⊙ leb 1 (get_mode_count RS08 i 0 opcode_table_RS08)) ⊗
+ (⊖ (test_not_impl_mode i RS08_not_impl_mode) ⊙ eqb (get_mode_count RS08 i 0 opcode_table_RS08) 0))
+ = true.
+ reflexivity.
+qed.
+
+lemma ok_OpIm_table_RS08 :
+ forall_instr_mode (λi:instr_mode.
+ forall_opcode     (λop:opcode.
+  leb (get_OpIm_count RS08 (anyOP RS08 op) i 0 opcode_table_RS08) 1)) = true.
+ reflexivity.
+qed.
+