X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Flib%2Farith.ma;h=f86f7a3bd75b1ca4532acfa39c2829a704335c9f;hb=d2e0a33c75842a10574ef904097803e02571536c;hp=fad49e616661cdfaa9d5f633b7252ba79a7192b8;hpb=4b8544042a6f3c5f5d303d4120c69abbc34ce15b;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 fad49e616..f86f7a3bd 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -23,6 +23,10 @@ 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. @@ -31,6 +35,17 @@ normalize // qed. lemma pred_S: ∀m. pred (S m) = m. // 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 minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m. @@ -109,6 +124,15 @@ 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. @@ -131,6 +155,9 @@ 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-. @@ -145,6 +172,25 @@ 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 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. * // normalize #m #H elim (lt_refl_false m) // qed-. @@ -202,7 +248,7 @@ qed-. (* Iterators ****************************************************************) (* Note: see also: lib/arithemetics/bigops.ma *) -let rec iter (n:nat) (B:Type[0]) (op: B → B) (nil: B) ≝ +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) @@ -231,7 +277,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 ]