X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fng_assembly%2Fcommon%2Fnat_lemmas.ma;h=64eeb50f122f07b427be0003725cd2c1b136a6a9;hb=45f87a7044ccab3d8f1300bfc33ca93a6ad868cb;hp=4bfc4c11832366232bf43595a755b02d06a98a5d;hpb=38fccc2b774e493a94eedef76342b56079c0e694;p=helm.git diff --git a/helm/software/matita/contribs/ng_assembly/common/nat_lemmas.ma b/helm/software/matita/contribs/ng_assembly/common/nat_lemmas.ma index 4bfc4c118..64eeb50f1 100644 --- a/helm/software/matita/contribs/ng_assembly/common/nat_lemmas.ma +++ b/helm/software/matita/contribs/ng_assembly/common/nat_lemmas.ma @@ -51,26 +51,6 @@ nlemma nat_destruct_S_0 : ∀n:nat.S n = O → False. napply I. nqed. -nlemma symmetric_eqnat : symmetricT nat bool eq_nat. - nnormalize; - #n1; - nelim n1; - ##[ ##1: #n2; - nelim n2; - nnormalize; - ##[ ##1: napply refl_eq - ##| ##2: #n3; #H; napply refl_eq - ##] - ##| ##2: #n2; #H; #n3; - nnormalize; - ncases n3; - nnormalize; - ##[ ##1: napply refl_eq - ##| ##2: #n4; napply (H n4) - ##] - ##] -nqed. - nlemma eq_to_eqnat : ∀n1,n2:nat.n1 = n2 → eq_nat n1 n2 = true. #n1; nelim n1; @@ -88,7 +68,15 @@ nlemma eq_to_eqnat : ∀n1,n2:nat.n1 = n2 → eq_nat n1 n2 = true. napply (H n4 (nat_destruct_S_S … H1)) ##] ##] -nqed. +nqed. + +nlemma neqnat_to_neq : ∀n1,n2:nat.(eq_nat n1 n2 = false → n1 ≠ n2). + #n1; #n2; #H; + napply (not_to_not (n1 = n2) (eq_nat n1 n2 = true) …); + ##[ ##1: napply (eq_to_eqnat n1 n2) + ##| ##2: napply (eqfalse_to_neqtrue … H) + ##] +nqed. nlemma eqnat_to_eq : ∀n1,n2:nat.(eq_nat n1 n2 = true → n1 = n2). #n1; @@ -110,6 +98,31 @@ nlemma eqnat_to_eq : ∀n1,n2:nat.(eq_nat n1 n2 = true → n1 = n2). ##] nqed. +nlemma neq_to_neqnat : ∀n1,n2:nat.n1 ≠ n2 → eq_nat n1 n2 = false. + #n1; #n2; #H; + napply (neqtrue_to_eqfalse (eq_nat n1 n2)); + napply (not_to_not (eq_nat n1 n2 = true) (n1 = n2) ? H); + napply (eqnat_to_eq n1 n2). +nqed. + +nlemma decidable_nat : ∀x,y:nat.decidable (x = y). + #x; #y; nnormalize; + napply (or2_elim (eq_nat x y = true) (eq_nat x y = false) ? (decidable_bexpr ?)); + ##[ ##1: #H; napply (or2_intro1 (x = y) (x ≠ y) (eqnat_to_eq … H)) + ##| ##2: #H; napply (or2_intro2 (x = y) (x ≠ y) (neqnat_to_neq … H)) + ##] +nqed. + +nlemma symmetric_eqnat : symmetricT nat bool eq_nat. + #n1; #n2; + napply (or2_elim (n1 = n2) (n1 ≠ n2) ? (decidable_nat n1 n2)); + ##[ ##1: #H; nrewrite > H; napply refl_eq + ##| ##2: #H; nrewrite > (neq_to_neqnat n1 n2 H); + napply (symmetric_eq ? (eq_nat n2 n1) false); + napply (neq_to_neqnat n2 n1 (symmetric_neq ? n1 n2 H)) + ##] +nqed. + nlemma Sn_p_n_to_S_npn : ∀n1,n2.(S n1) + n2 = S (n1 + n2). #n1; nelim n1;