X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Flib%2Farith.ma;h=dad1f938241812ef41928b983c6cb984f5ae5f2d;hb=f4787814123d74c9504e988137c2c13279838257;hp=9fedf60c0c180fd43f5e2ce90139254ba59dbfca;hpb=f6b452b9c9be141740d4058dfbcf81a4b75fd00b;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma index 9fedf60c0..dad1f9382 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -96,6 +96,20 @@ qed. lemma arith_i: ∀x,y,z. y < x → x+z-y-1 = x-y-1+z. /2 width=1 by plus_minus/ qed-. +lemma idempotent_max: ∀n:nat. n = (n ∨ n). +#n normalize >le_to_leb_true // +qed. + +lemma associative_max: associative … max. +#x #y #z normalize +@(leb_elim x y) normalize #Hxy +@(leb_elim y z) normalize #Hyz // +[1,2: >le_to_leb_true /2 width=3 by transitive_le/ +| >not_le_to_leb_false /4 width=3 by lt_to_not_le, not_le_to_lt, transitive_lt/ + >not_le_to_leb_false // +] +qed. + (* Properties ***************************************************************) lemma eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). @@ -158,6 +172,10 @@ qed. (* Inversion & forward lemmas ***********************************************) +lemma max_inv_O3: ∀x,y. (x ∨ y) = 0 → 0 = x ∧ 0 = y. +/4 width=2 by le_maxr, le_maxl, le_n_O_to_eq, conj/ +qed-. + lemma plus_inv_O3: ∀x,y. x + y = 0 → x = 0 ∧ y = 0. /2 width=1 by plus_le_0/ qed-.