X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Flib%2Farith.ma;h=9fedf60c0c180fd43f5e2ce90139254ba59dbfca;hb=f6b452b9c9be141740d4058dfbcf81a4b75fd00b;hp=1afc42ebcbccd6d0365a2869eb9035071a03916f;hpb=98c91e19a9cc31c77a0151f5df7f7690813cbd07;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 1afc42ebc..9fedf60c0 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -12,13 +12,51 @@ (* *) (**************************************************************************) +include "ground_2/notation/functions/successor_1.ma". +include "ground_2/notation/functions/predecessor_1.ma". include "arithmetics/nat.ma". include "ground_2/lib/star.ma". (* ARITHMETICAL PROPERTIES **************************************************) +interpretation "nat successor" 'Successor m = (S m). + +interpretation "nat predecessor" 'Predecessor m = (pred m). + +interpretation "nat min" 'and x y = (min x y). + +interpretation "nat max" 'or x y = (max x y). + +(* Iota equations ***********************************************************) + +lemma pred_O: pred 0 = 0. +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-. + (* Note: uses minus_minus_comm, minus_plus_m_m, commutative_plus, plus_minus *) lemma plus_minus_minus_be: ∀x,y,z. y ≤ z → z ≤ x → (x - z) + (z - y) = x - y. #x #z #y #Hzy #Hyx >plus_minus // >commutative_plus >plus_minus // @@ -60,9 +98,13 @@ lemma arith_i: ∀x,y,z. y < x → x+z-y-1 = x-y-1+z. (* Properties ***************************************************************) -axiom eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). - -axiom lt_dec: ∀n1,n2. Decidable (n1 < n2). +lemma eq_nat_dec: ∀n1,n2:nat. Decidable (n1 = n2). +#n1 elim n1 -n1 [| #n1 #IHn1 ] * [2,4: #n2 ] +[1,4: @or_intror #H destruct +| elim (IHn1 n2) -IHn1 /3 width=1 by or_intror, or_introl/ +| /2 width=1 by or_introl/ +] +qed-. lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m. #m #n elim (lt_or_ge m n) /2 width=1 by or3_intro0/ @@ -85,6 +127,18 @@ lemma monotonic_lt_pred: ∀m,n. m < n → O < m → pred m < pred n. @le_S_S_to_le >S_pred /2 width=3 by transitive_lt/ qed. +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. +/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. @@ -104,6 +158,15 @@ qed. (* Inversion & forward lemmas ***********************************************) +lemma plus_inv_O3: ∀x,y. x + y = 0 → x = 0 ∧ y = 0. +/2 width=1 by plus_le_0/ 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-. @@ -115,9 +178,26 @@ lemma lt_zero_false: ∀n. n < 0 → ⊥. #n #H elim (lt_to_not_le … H) -H /2 width=1 by/ qed-. -lemma false_lt_to_le: ∀x,y. (x < y → ⊥) → y ≤ x. -#x #y #H elim (decidable_lt x y) /2 width=1 by not_lt_to_le/ -#Hxy elim (H Hxy) +lemma lt_le_false: ∀x,y. x < y → y ≤ x → ⊥. +/3 width=4 by lt_refl_false, lt_to_le_to_lt/ qed-. + +lemma succ_inv_refl_sn: ∀x. ⫯x = x → ⊥. +#x #H @(lt_le_false x (⫯x)) // +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. @@ -139,10 +219,45 @@ 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-. +(* 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/ +#x #_ #y @(nat_ind_plus … y) -y /2 width=1 by or_intror/ +#y #_ >minus_plus_plus_l +#H lapply (discr_plus_xy_minus_xz … H) -H +#H destruct +qed-. + +lemma zero_eq_plus: ∀x,y. 0 = x + y → 0 = x ∧ 0 = y. +* /2 width=1 by conj/ #x #y normalize #H destruct +qed-. + +lemma lt_S_S_to_lt: ∀x,y. ⫯x < ⫯y → x < y. +/2 width=1 by le_S_S_to_le/ qed-. + +lemma lt_elim: ∀R:relation nat. + (∀n2. R O (⫯n2)) → + (∀n1,n2. R n1 n2 → R (⫯n1) (⫯n2)) → + ∀n2,n1. n1 < n2 → R n1 n2. +#R #IH1 #IH2 #n2 elim n2 -n2 +[ #n1 #H elim (lt_le_false … H) -H // +| #n2 #IH * /4 width=1 by lt_S_S_to_lt/ +] +qed-. + +lemma le_elim: ∀R:relation nat. + (∀n2. R O (n2)) → + (∀n1,n2. R n1 n2 → R (⫯n1) (⫯n2)) → + ∀n1,n2. n1 ≤ n2 → R n1 n2. +#R #IH1 #IH2 #n1 #n2 @(nat_elim2 … n1 n2) -n1 -n2 +/4 width=1 by monotonic_pred/ -IH1 -IH2 +#n1 #H elim (lt_le_false … H) -H // +qed-. + (* Iterators ****************************************************************) -(* Note: see also: lib/arithemetcs/bigops.ma *) -let rec iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝ +(* Note: see also: lib/arithemetics/bigops.ma *) +rec definition iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝ match n with [ O ⇒ nil | S k ⇒ op (iter k B op nil) @@ -150,6 +265,12 @@ let rec iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝ interpretation "iterated function" 'exp op n = (iter n ? op). +lemma iter_O: ∀B:Type[0]. ∀f:B→B.∀b. f^0 b = b. +// qed. + +lemma iter_S: ∀B:Type[0]. ∀f:B→B.∀b,l. f^(S l) b = f (f^l b). +// qed. + lemma iter_SO: ∀B:Type[0]. ∀f:B→B. ∀b,l. f^(l+1) b = f (f^l b). #B #f #b #l >commutative_plus // qed. @@ -165,7 +286,7 @@ qed. (* Trichotomy operator ******************************************************) (* Note: this is "if eqb n1 n2 then a2 else if leb n1 n2 then a1 else a3" *) -let rec tri (A:Type[0]) n1 n2 a1 a2 a3 on n1 : A ≝ +rec definition tri (A:Type[0]) n1 n2 a1 a2 a3 on n1 : A ≝ match n1 with [ O ⇒ match n2 with [ O ⇒ a2 | S n2 ⇒ a1 ] | S n1 ⇒ match n2 with [ O ⇒ a3 | S n2 ⇒ tri A n1 n2 a1 a2 a3 ]