X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Flibrary%2Fassembly%2Fexadecimal.ma;h=fb1a14ede914299c6b2e9342296e846c39dc54b0;hb=dfbd010fe5a46d049849913bc23000289893ea4f;hp=3a4e35da40384455eefebbd252648dbcfc9ea984;hpb=9655a931c73be820d876da7b23fe54b55fd60c14;p=helm.git diff --git a/matita/library/assembly/exadecimal.ma b/matita/library/assembly/exadecimal.ma index 3a4e35da4..fb1a14ede 100644 --- a/matita/library/assembly/exadecimal.ma +++ b/matita/library/assembly/exadecimal.ma @@ -14,7 +14,7 @@ set "baseuri" "cic:/matita/assembly/exadecimal/". -include "extra.ma". +include "assembly/extra.ma". inductive exadecimal : Type ≝ x0: exadecimal @@ -836,6 +836,13 @@ lemma nat_of_exadecimal_exadecimal_of_nat: ] qed. +lemma exadecimal_of_nat_nat_of_exadecimal: + ∀b.exadecimal_of_nat (nat_of_exadecimal b) = b. + intro; + elim b; + reflexivity. +qed. + lemma plusex_ok: ∀b1,b2,c. match plusex b1 b2 c with @@ -904,9 +911,20 @@ lemma eqex_true_to_eq: ∀b,b'. eqex b b' = true → b=b'. first [ reflexivity | destruct H ]. qed. +lemma eqex_false_to_not_eq: ∀b,b'. eqex b b' = false → b ≠ b'. + intros 2; + elim b 0; + elim b' 0; + simplify; + intro; + try (destruct H); + intro K; + destruct K. +qed. + (* lemma xpred_S: ∀b. xpred (exadecimal_of_nat (S b)) = exadecimal_of_nat b. intros; rewrite > exadecimal_of_nat_mod; rewrite > exadecimal_of_nat_mod in ⊢ (? ? ? %); -*) \ No newline at end of file +*)