X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Flibrary%2Fassembly%2Fbyte.ma;h=f2d1dc1566b7194903c2eda58d1eafd853a37ff9;hb=d6ce6da7d0d5572d7551adafe849f18ab8e8f895;hp=f97b13f293a131b4a2c7bb3b4c223aa37779e90d;hpb=61d5edd7f92c89bceae245453dead8cbd2a276b7;p=helm.git diff --git a/matita/library/assembly/byte.ma b/matita/library/assembly/byte.ma index f97b13f29..f2d1dc156 100644 --- a/matita/library/assembly/byte.ma +++ b/matita/library/assembly/byte.ma @@ -42,6 +42,9 @@ coercion cic:/matita/assembly/byte/nat_of_byte.con. definition byte_of_nat ≝ λn. mk_byte (exadecimal_of_nat (n / 16)) (exadecimal_of_nat n). +interpretation "byte_of_nat" 'byte_of_opcode a = + (cic:/matita/assembly/byte/byte_of_nat.con a). + lemma byte_of_nat_nat_of_byte: ∀b. byte_of_nat (nat_of_byte b) = b. intros; elim b; @@ -63,7 +66,7 @@ lemma lt_nat_of_byte_256: ∀b. nat_of_byte b < 256. [ letin Hf ≝ (le_plus ? ? ? ? Hcut K'); clearbody Hf; simplify in Hf:(? ? %); assumption - | autobatch + | apply le_times_r. apply H'. ] qed. @@ -104,7 +107,7 @@ lemma eq_nat_of_byte_n_nat_of_byte_mod_n_256: rewrite > exadecimal_of_nat_mod in ⊢ (? ? ? %); rewrite > divides_to_eq_mod_mod_mod; [ reflexivity - | autobatch + | apply (witness ? ? 16). reflexivity. ] ] qed. @@ -114,7 +117,8 @@ lemma plusbyte_ok: match plusbyte b1 b2 c with [ couple r c' ⇒ b1 + b2 + nat_of_bool c = nat_of_byte r + nat_of_bool c' * 256 ]. - intros; + intros; elim daemon. + (* unfold plusbyte; generalize in match (plusex_ok (bl b1) (bl b2) c); elim (plusex (bl b1) (bl b2) c); @@ -146,6 +150,7 @@ lemma plusbyte_ok: rewrite < associative_plus in ⊢ (? ? (? ? (? % ?)) ?); rewrite > H; clear H; autobatch paramodulation. + *) qed. definition bpred ≝