]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/contribs/ng_assembly/freescale/table_RS08_tests.ma
1) \ldots here and there
[helm.git] / helm / software / matita / contribs / ng_assembly / freescale / table_RS08_tests.ma
index c98e6070586837d3ccd09d5a898a34d9da9dda53..b47a7ee8720952b1d1e336af6c6b7d0176129b53 100755 (executable)
@@ -38,7 +38,7 @@ nlemma ok_byte_table_RS08 : forall_byte8 (λb.
  (test_not_impl_byte b RS08_not_impl_byte     ⊙ eq_nat (get_byte_count RS08 b 0 opcode_table_RS08) 1) ⊗
  (⊖ (test_not_impl_byte b RS08_not_impl_byte) ⊙ eq_nat (get_byte_count RS08 b 0 opcode_table_RS08) 0))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* RS08: pseudocodici non implementati come da manuale *)
@@ -52,7 +52,7 @@ nlemma ok_pseudo_table_RS08 : forall_opcode (λo.
  (test_not_impl_pseudo o RS08_not_impl_pseudo     ⊙ le_nat 1 (get_pseudo_count RS08 o 0 opcode_table_RS08)) ⊗
  (⊖ (test_not_impl_pseudo o RS08_not_impl_pseudo) ⊙ eq_nat (get_pseudo_count RS08 o 0 opcode_table_RS08) 0))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 (* RS08: modalita' non implementate come da manuale *)
@@ -66,12 +66,12 @@ nlemma ok_mode_table_RS08 : forall_instr_mode (λi.
  (test_not_impl_mode i RS08_not_impl_mode     ⊙ le_nat 1 (get_mode_count RS08 i 0 opcode_table_RS08)) ⊗
  (⊖ (test_not_impl_mode i RS08_not_impl_mode) ⊙ eq_nat (get_mode_count RS08 i 0 opcode_table_RS08) 0))
  = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.
 
 nlemma ok_OpIm_table_RS08 :
  forall_instr_mode (λi:instr_mode.
  forall_opcode     (λop:opcode.
   le_nat (get_OpIm_count RS08 (anyOP RS08 op) i 0 opcode_table_RS08) 1)) = true.
- napply (refl_eq ??).
+ napply refl_eq.
 nqed.