]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/freescale/table_HC05_tests.ma
1) \ldots here and there
[helm.git] / helm / software / matita / contribs / ng_assembly / freescale / table_HC05_tests.ma
index a1b2152f4d473ca585c2da728180c8d595b84457..7d78e530f17191ac3c75a931929eda874b73e0f3 100755 (executable)
@@ -42,7 +42,7 @@ 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))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HC05: pseudocodici non implementati come da manuale *)
@@ -55,7 +55,7 @@ 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))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* HC05: modalita' non implementate come da manuale *)
@@ -82,12 +82,12 @@ 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))
  = 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 ??).
+ napply refl_eq.
 nqed.