X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Flib%2Farith.ma;h=dde95e8a7267d844ab844a0861e9ce1a044b480f;hp=1ba525475622dfe3fb6777c4cbe59847a4584a1d;hb=ff612dc35167ec0c145864c9aa8ae5e1ebe20a48;hpb=5832735b721c0bd8567c8f0be761a9136363a2a6 diff --git a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma index 1ba525475..dde95e8a7 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -12,16 +12,20 @@ (* *) (**************************************************************************) -include "ground_2/notation/functions/successor_1.ma". -include "ground_2/notation/functions/predecessor_1.ma". +include "ground_2/notation/functions/uparrow_1.ma". +include "ground_2/notation/functions/downarrow_1.ma". include "arithmetics/nat.ma". -include "ground_2/lib/star.ma". +include "ground_2/lib/relations.ma". (* ARITHMETICAL PROPERTIES **************************************************) -interpretation "nat successor" 'Successor m = (S m). +interpretation "nat successor" 'UpArrow m = (S m). -interpretation "nat predecessor" 'Predecessor m = (pred m). +interpretation "nat predecessor" 'DownArrow m = (pred m). + +interpretation "nat min" 'and x y = (min x y). + +interpretation "nat max" 'or x y = (max x y). (* Iota equations ***********************************************************) @@ -31,8 +35,25 @@ normalize // qed. lemma pred_S: ∀m. pred (S m) = m. // qed. +lemma plus_S1: ∀x,y. ↑(x+y) = (↑x) + y. +// qed. + +lemma max_O1: ∀n. n = (0 ∨ n). +// qed. + +lemma max_O2: ∀n. n = (n ∨ 0). +// qed. + +lemma max_SS: ∀n1,n2. ↑(n1∨n2) = (↑n1 ∨ ↑n2). +#n1 #n2 elim (decidable_le n1 n2) #H normalize +[ >(le_to_leb_true … H) | >(not_le_to_leb_false … H) ] -H // +qed. + (* Equations ****************************************************************) +lemma plus_SO: ∀n. n + 1 = ↑n. +// qed. + lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m. // qed-. @@ -41,6 +62,10 @@ lemma plus_minus_minus_be: ∀x,y,z. y ≤ z → z ≤ x → (x - z) + (z - y) = #x #z #y #Hzy #Hyx >plus_minus // >commutative_plus >plus_minus // qed-. +lemma lt_succ_pred: ∀m,n. n < m → m = ↑↓m. +#m #n #Hm >S_pred /2 width=2 by ltn_to_ltO/ +qed-. + fact plus_minus_minus_be_aux: ∀i,x,y,z. y ≤ z → z ≤ x → i = z - y → x - z + i = x - y. /2 width=1 by plus_minus_minus_be/ qed-. @@ -48,21 +73,21 @@ lemma plus_n_2: ∀n. n + 2 = n + 1 + 1. // qed. lemma le_plus_minus: ∀m,n,p. p ≤ n → m + n - p = m + (n - p). -/2 by plus_minus/ qed. +/2 by plus_minus/ qed-. lemma le_plus_minus_comm: ∀n,m,p. p ≤ m → m + n - p = m - p + n. -/2 by plus_minus/ qed. +/2 by plus_minus/ qed-. lemma minus_minus_comm3: ∀n,x,y,z. n-x-y-z = n-y-z-x. // qed. lemma arith_b1: ∀a,b,c1. c1 ≤ b → a - c1 - (b - c1) = a - b. #a #b #c1 #H >minus_minus_comm >minus_le_minus_minus_comm // -qed. +qed-. lemma arith_b2: ∀a,b,c1,c2. c1 + c2 ≤ b → a - c1 - c2 - (b - c1 - c2) = a - b. #a #b #c1 #c2 #H >minus_plus >minus_plus >minus_plus /2 width=1 by arith_b1/ -qed. +qed-. lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. /3 by monotonic_le_minus_l, le_to_le_to_eq, le_n/ qed. @@ -70,11 +95,29 @@ lemma arith_c1x: ∀x,a,b,c1. x + c1 + a - (b + c1) = x + a - b. lemma arith_h1: ∀a1,a2,b,c1. c1 ≤ a1 → c1 ≤ b → a1 - c1 + a2 - (b - c1) = a1 + a2 - b. #a1 #a2 #b #c1 #H1 #H2 >plus_minus /2 width=1 by arith_b2/ -qed. +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 plus_to_minus_2: ∀m1,m2,n1,n2. n1 ≤ m1 → n2 ≤ m2 → + m1+n2 = m2+n1 → m1-n1 = m2-n2. +/2 width=1 by arith_b1/ 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). @@ -101,17 +144,23 @@ lemma monotonic_le_minus_l2: ∀x1,x2,y,z. x1 ≤ x2 → x1 - y - z ≤ x2 - y - /3 width=1 by monotonic_le_minus_l/ qed. (* Note: this might interfere with nat.ma *) -lemma monotonic_lt_pred: ∀m,n. m < n → O < m → pred m < pred n. +lemma monotonic_lt_pred: ∀m,n. m < n → 0 < m → pred m < pred n. #m #n #Hmn #Hm whd >(S_pred … Hm) @le_S_S_to_le >S_pred /2 width=3 by transitive_lt/ qed. -lemma lt_S_S: ∀x,y. x < y → ⫯x < ⫯y. +lemma lt_S_S: ∀x,y. x < y → ↑x < ↑y. /2 width=1 by le_S_S/ qed. -lemma lt_S: ∀n,m. n < m → n < ⫯m. +lemma lt_S: ∀n,m. n < m → n < ↑m. /2 width=1 by le_S/ qed. +lemma max_S1_le_S: ∀n1,n2,n. (n1 ∨ n2) ≤ n → (↑n1 ∨ n2) ≤ ↑n. +/4 width=2 by to_max, le_maxr, le_S_S, le_S/ qed-. + +lemma max_S2_le_S: ∀n1,n2,n. (n1 ∨ n2) ≤ n → (n1 ∨ ↑n2) ≤ ↑n. +/2 width=1 by max_S1_le_S/ qed-. + lemma arith_j: ∀x,y,z. x-y-1 ≤ x-(y-z)-1. /3 width=1 by monotonic_le_minus_l, monotonic_le_minus_r/ qed. @@ -131,12 +180,6 @@ qed. (* Inversion & forward lemmas ***********************************************) -lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0. -// qed-. - -lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y. -/2 width=1 by monotonic_pred/ qed-. - lemma lt_refl_false: ∀n. n < n → ⊥. #n #H elim (lt_to_not_eq … H) -H /2 width=1 by/ qed-. @@ -148,23 +191,8 @@ qed-. lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥. /3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-. -lemma lt_inv_O1: ∀n. 0 < n → ∃m. ⫯m = n. -* /2 width=2 by ex_intro/ -#H cases (lt_le_false … H) -H // -qed-. - -lemma lt_inv_S1: ∀m,n. ⫯m < n → ∃∃p. m < p & ⫯p = n. -#m * /3 width=3 by lt_S_S_to_lt, ex2_intro/ -#H cases (lt_le_false … H) -H // -qed-. - -lemma lt_inv_gen: ∀y,x. x < y → ∃∃z. x ≤ z & ⫯z = y. -* /3 width=3 by le_S_S_to_le, ex2_intro/ -#x #H elim (lt_le_false … H) -H // -qed-. - -lemma pred_inv_refl: ∀m. pred m = m → m = 0. -* // normalize #m #H elim (lt_refl_false m) // +lemma succ_inv_refl_sn: ∀x. ↑x = x → ⊥. +#x #H @(lt_le_false x (↑x)) // qed-. lemma le_plus_xSy_O_false: ∀x,y. x + S y ≤ 0 → ⊥. @@ -182,6 +210,28 @@ lemma plus_xySz_x_false: ∀z,x,y. x + y + S z = x → ⊥. lemma plus_xSy_x_false: ∀y,x. x + S y = x → ⊥. /2 width=4 by plus_xySz_x_false/ qed-. +lemma pred_inv_refl: ∀m. pred m = m → m = 0. +* // normalize #m #H elim (lt_refl_false m) // +qed-. + +lemma discr_plus_xy_y: ∀x,y. x + y = y → x = 0. +// qed-. + +lemma discr_plus_x_xy: ∀x,y. x = x + y → y = 0. +/2 width=2 by le_plus_minus_comm/ qed-. + +lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y. +/2 width=1 by monotonic_pred/ qed-. + +lemma plus2_inv_le_sn: ∀m1,m2,n1,n2. m1 + n1 = m2 + n2 → m1 ≤ m2 → n2 ≤ n1. +#m1 #m2 #n1 #n2 #H #Hm +lapply (monotonic_le_plus_l n1 … Hm) -Hm >H -H +/2 width=2 by le_plus_to_le/ +qed-. + +lemma lt_S_S_to_lt: ∀x,y. ↑x < ↑y → x < y. +/2 width=1 by le_S_S_to_le/ qed-. + (* Note this should go in nat.ma *) lemma discr_x_minus_xy: ∀x,y. x = x - y → x = 0 ∨ y = 0. #x @(nat_ind_plus … x) -x /2 width=1 by or_introl/ @@ -191,16 +241,55 @@ lemma discr_x_minus_xy: ∀x,y. x = x - y → x = 0 ∨ y = 0. #H destruct qed-. +lemma lt_inv_O1: ∀n. 0 < n → ∃m. ↑m = n. +* /2 width=2 by ex_intro/ +#H cases (lt_le_false … H) -H // +qed-. + +lemma lt_inv_S1: ∀m,n. ↑m < n → ∃∃p. m < p & ↑p = n. +#m * /3 width=3 by lt_S_S_to_lt, ex2_intro/ +#H cases (lt_le_false … H) -H // +qed-. + +lemma lt_inv_gen: ∀y,x. x < y → ∃∃z. x ≤ z & ↑z = y. +* /3 width=3 by le_S_S_to_le, ex2_intro/ +#x #H elim (lt_le_false … H) -H // +qed-. + +lemma plus_inv_O3: ∀x,y. x + y = 0 → x = 0 ∧ y = 0. +/2 width=1 by plus_le_0/ qed-. + +lemma plus_inv_S3_sn: ∀x1,x2,x3. x1+x2 = ↑x3 → + ∨∨ ∧∧ x1 = 0 & x2 = ↑x3 + | ∃∃y1. x1 = ↑y1 & y1 + x2 = x3. +* /3 width=1 by or_introl, conj/ +#x1 #x2 #x3