]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/freescale/table_HCS08_tests.ma
1) \ldots here and there
[helm.git] / helm / software / matita / contribs / ng_assembly / freescale / table_HCS08_tests.ma
index f3c4e85e34fb30421e041ee4463b5b934a51fe43..5983746ead0941a724f42bc7356ab983e153872d 100755 (executable)
@@ -37,7 +37,7 @@ 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))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HCS08: opcode non implementati come da manuale (0x9E+byte) *)
@@ -73,7 +73,7 @@ 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))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HCS08: pseudocodici non implementati come da manuale *)
@@ -84,7 +84,7 @@ 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))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HCS08: modalita' non implementate come da manuale *)
@@ -106,12 +106,12 @@ 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))
  = 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 ??).
+ napply refl_eq.
 nqed.