From: Ferruccio Guidi Date: Wed, 27 Nov 2013 16:20:52 +0000 (+0000) Subject: strict order relation for natural numbers with inifinity X-Git-Tag: make_still_working~1044 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=23da2aa16489e00889374d81f19cc090faa44582;p=helm.git strict order relation for natural numbers with inifinity --- diff --git a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_le.ma b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_le.ma index e706fa4dd..be13774f5 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_le.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_le.ma @@ -26,15 +26,21 @@ interpretation "ynat 'less or equal to'" 'leq x y = (yle x y). (* Basic inversion lemmas ***************************************************) -fact yle_inv_inj_aux: ∀x,y. x ≤ y → ∀m,n. x = yinj m → y = yinj n → m ≤ n. +fact yle_inv_inj2_aux: ∀x,y. x ≤ y → ∀n. y = yinj n → + ∃∃m. m ≤ n & x = yinj m. #x #y * -x -y -[ #x #y #Hxy #m #n #Hx #Hy destruct // -| #x #m #n #_ #Hy destruct +[ #x #y #Hxy #n #Hy destruct /2 width=3 by ex2_intro/ +| #x #n #Hy destruct ] qed-. +lemma yle_inv_inj2: ∀x,n. x ≤ yinj n → ∃∃m. m ≤ n & x = yinj m. +/2 width=3 by yle_inv_inj2_aux/ qed-. + lemma yle_inv_inj: ∀m,n. yinj m ≤ yinj n → m ≤ n. -/2 width=5 by yle_inv_inj_aux/ qed-. +#m #n #H elim (yle_inv_inj2 … H) -H +#x #Hxn #H destruct // +qed-. fact yle_inv_O2_aux: ∀m:ynat. ∀x:ynat. m ≤ x → x = 0 → m = 0. #m #x * -m -x @@ -63,7 +69,7 @@ fact yle_inv_succ1_aux: ∀x,y. x ≤ y → ∀m. x = ⫯m → ∃∃n. m ≤ n #m #Hnm #H destruct @(ex2_intro … m) /2 width=1 by yle_inj/ (**) (* explicit constructor *) | #x #y #H destruct - @(ex2_intro … (∞)) /2 width=1 by yle_Y/ + @(ex2_intro … (∞)) /2 width=1 by yle_Y/ (**) (* explicit constructor *) ] qed-. @@ -106,6 +112,6 @@ theorem yle_trans: Transitive … yle. [ #x #y #Hxy * // #z #H lapply (yle_inv_inj … H) -H /3 width=3 by transitive_le, yle_inj/ (**) (* full auto too slow *) -| #x #z #H lapply ( yle_inv_Y1 … H) // +| #x #z #H lapply (yle_inv_Y1 … H) // ] qed-. diff --git a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_lt.ma b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_lt.ma new file mode 100644 index 000000000..4c63e40a7 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_lt.ma @@ -0,0 +1,137 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/ynat/ynat_le.ma". + +(* NATURAL NUMBERS WITH INFINITY ********************************************) + +(* strict order relation *) +inductive ylt: relation ynat ≝ +| ylt_inj: ∀m,n. m < n → ylt m n +| ylt_Y : ∀m:nat. ylt m (∞) +. + +interpretation "ynat 'less than'" 'lt x y = (ylt x y). + +(* Basic inversion lemmas ***************************************************) + +fact ylt_inv_inj2_aux: ∀x,y. x < y → ∀n. y = yinj n → + ∃∃m. m < n & x = yinj m. +#x #y * -x -y +[ #x #y #Hxy #n #Hy elim (le_inv_S1 … Hxy) -Hxy + #m #Hm #H destruct /3 width=3 by le_S_S, ex2_intro/ +| #x #n #Hy destruct +] +qed-. + +lemma ylt_inv_inj2: ∀x,n. x < yinj n → + ∃∃m. m < n & x = yinj m. +/2 width=3 by ylt_inv_inj2_aux/ qed-. + +lemma ylt_inv_inj: ∀m,n. yinj m < yinj n → m < n. +#m #n #H elim (ylt_inv_inj2 … H) -H +#x #Hx #H destruct // +qed-. + +fact ylt_inv_Y2_aux: ∀x,y. x < y → y = ∞ → ∃m. x = yinj m. +#x #y * -x -y /2 width=2 by ex_intro/ +qed-. + +lemma ylt_inv_Y2: ∀x. x < ∞ → ∃m. x = yinj m. +/2 width=3 by ylt_inv_Y2_aux/ qed-. + +(* Inversion lemmas on successor ********************************************) + +fact ylt_inv_succ1_aux: ∀x,y. x < y → ∀m. x = ⫯m → ∃∃n. m < n & y = ⫯n. +#x #y * -x -y +[ #x #y #Hxy #m #H elim (ysucc_inv_inj_sn … H) -H + #n #H1 #H2 destruct elim (le_inv_S1 … Hxy) -Hxy + #m #Hnm #H destruct + @(ex2_intro … m) /2 width=1 by ylt_inj/ (**) (* explicit constructor *) +| #x #y #H elim (ysucc_inv_inj_sn … H) -H + #m #H #_ destruct + @(ex2_intro … (∞)) /2 width=1 by/ (**) (* explicit constructor *) +] +qed-. + +lemma ylt_inv_succ1: ∀m,y. ⫯m < y → ∃∃n. m < n & y = ⫯n. +/2 width=3 by ylt_inv_succ1_aux/ qed-. + +lemma yle_inv_succ: ∀m,n. ⫯m < ⫯n → m < n. +#m #n #H elim (ylt_inv_succ1 … H) -H +#x #Hx #H destruct // +qed-. + +fact ylt_inv_succ2_aux: ∀x,y. x < y → ∀n. y = ⫯n → x ≤ n. +#x #y * -x -y +[ #x #y #Hxy #m #H elim (ysucc_inv_inj_sn … H) -H + #n #H1 #H2 destruct /3 width=1 by yle_inj, le_S_S_to_le/ +| #x #n #H lapply (ysucc_inv_Y_sn … H) -H // +] +qed-. + +lemma ylt_inv_succ2: ∀m,n. m < ⫯n → m ≤ n. +/2 width=3 by ylt_inv_succ2_aux/ qed-. + +(* inversion and forward lemmas on yle **************************************) + +lemma lt_fwd_le: ∀m:ynat. ∀n:ynat. m < n → m ≤ n. +#m #n * -m -n /3 width=1 by yle_pred_sn, yle_inj, yle_Y/ +qed-. + +lemma ylt_yle_false: ∀m:ynat. ∀n:ynat. m < n → n ≤ m → ⊥. +#m #n * -m -n +[ #m #n #Hmn #H lapply (yle_inv_inj … H) -H + #H elim (lt_refl_false n) /2 width=3 by le_to_lt_to_lt/ +| #m #H lapply (yle_inv_Y1 … H) -H + #H destruct +] +qed-. + +(* Properties on yle ********************************************************) + +lemma yle_to_ylt_or_eq: ∀m:ynat. ∀n:ynat. m ≤ n → m < n ∨ m = n. +#m #n * -m -n +[ #m #n #Hmn elim (le_to_or_lt_eq … Hmn) -Hmn + /3 width=1 by or_introl, ylt_inj/ +| * /2 width=1 by or_introl, ylt_Y/ +] +qed-. + +lemma ylt_yle_trans: ∀x:ynat. ∀y:ynat. ∀z:ynat. y ≤ z → x < y → x < z. +#x #y #z * -y -z +[ #y #z #Hyz #H elim (ylt_inv_inj2 … H) -H + #m #Hm #H destruct /3 width=3 by ylt_inj, lt_to_le_to_lt/ +| #y * // +] +qed-. + +lemma yle_ylt_trans: ∀x:ynat. ∀y:ynat. ∀z:ynat. y < z → x ≤ y → x < z. +#x #y #z * -y -z +[ #y #z #Hyz #H elim (yle_inv_inj2 … H) -H + #m #Hm #H destruct /3 width=3 by ylt_inj, le_to_lt_to_lt/ +| #y #H elim (yle_inv_inj2 … H) -H // +] +qed-. + +(* Main properties **********************************************************) + +theorem ylt_trans: Transitive … ylt. +#x #y * -x -y +[ #x #y #Hxy * // + #z #H lapply (ylt_inv_inj … H) -H + /3 width=3 by transitive_lt, ylt_inj/ (**) (* full auto too slow *) +| #x #z #H elim (ylt_yle_false … H) // +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_succ.ma b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_succ.ma index d0ed1d6bd..707556f30 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_succ.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_succ.ma @@ -53,3 +53,11 @@ qed-. lemma ysucc_inv_inj_dx: ∀m2,n1. ⫯n1 = yinj m2 → ∃∃m1. n1 = yinj m1 & m2 = S m1. /2 width=1 by ysucc_inv_inj_sn/ qed-. + +lemma ysucc_inv_Y_sn: ∀m. ∞ = ⫯m → m = ∞. +* // normalize +#m #H destruct +qed-. + +lemma ysucc_inv_Y_dx: ∀m. ⫯m = ∞ → m = ∞. +/2 width=1 by ysucc_inv_Y_sn/ qed-.