X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fassembly%2Fexadecimal.ma;h=c059b43fa5fe91f079ff05b951157eb6559a9bb8;hb=947ee89dec9e60561dfac3ce7e1842f35f178cb8;hp=fb1a14ede914299c6b2e9342296e846c39dc54b0;hpb=9a3b49d3584873c77a39107cb5ee4f1e6010e43c;p=helm.git diff --git a/helm/software/matita/library/assembly/exadecimal.ma b/helm/software/matita/library/assembly/exadecimal.ma index fb1a14ede..c059b43fa 100644 --- a/helm/software/matita/library/assembly/exadecimal.ma +++ b/helm/software/matita/library/assembly/exadecimal.ma @@ -768,8 +768,7 @@ lemma lt_nat_of_exadecimal_16: ∀b. nat_of_exadecimal b < 16. elim b; simplify; [ - |*: alias id "lt_S_S" = "cic:/matita/algebra/finite_groups/lt_S_S.con". - repeat (apply lt_S_S) + |*: repeat (apply lt_to_lt_S_S) ]; autobatch. qed. @@ -801,7 +800,8 @@ lemma exadecimal_of_nat_mod: change in ⊢ (? ? ? (? (? % ?))) with (n16 \mod 16); rewrite < mod_mod; [ apply H; - autobatch + unfold lt; + autobatch. | autobatch ] qed. @@ -848,11 +848,7 @@ lemma plusex_ok: match plusex b1 b2 c with [ couple r c' ⇒ b1 + b2 + nat_of_bool c = nat_of_exadecimal r + nat_of_bool c' * 16 ]. intros; - elim c; - elim b1; - elim b2; - normalize; - reflexivity. + elim b1; (elim b2; (elim c; normalize; reflexivity)). qed. definition xpred ≝