X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Fynat%2Fynat_lt.ma;h=6cc21d5bcc60bf6018944835931c20bcfae660df;hb=32bdf7f107be22a121fab8225c5fae4eb6b33633;hp=4c63e40a749ecbbc2feb32a3954ee42b4d2b6cfe;hpb=23da2aa16489e00889374d81f19cc090faa44582;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_lt.ma b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_lt.ma index 4c63e40a7..6cc21d5bc 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_lt.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_lt.ma @@ -24,6 +24,16 @@ inductive ylt: relation ynat ≝ interpretation "ynat 'less than'" 'lt x y = (ylt x y). +(* Basic forward lemmas *****************************************************) + +lemma ylt_fwd_gen: ∀x,y. x < y → ∃m. x = yinj m. +#x #y * -x -y /2 width=2 by ex_intro/ +qed-. + +lemma ylt_fwd_le_succ: ∀x,y. x < y → ⫯x ≤ y. +#x #y * -x -y /2 width=1 by yle_inj/ +qed-. + (* Basic inversion lemmas ***************************************************) fact ylt_inv_inj2_aux: ∀x,y. x < y → ∀n. y = yinj n → @@ -44,36 +54,38 @@ lemma ylt_inv_inj: ∀m,n. yinj m < yinj n → m < n. #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/ +lemma ylt_inv_Y1: ∀n. ∞ < n → ⊥. +#n #H elim (ylt_fwd_gen … H) -H +#y #H destruct qed-. -lemma ylt_inv_Y2: ∀x. x < ∞ → ∃m. x = yinj m. -/2 width=3 by ylt_inv_Y2_aux/ qed-. +lemma ylt_inv_O1: ∀n. 0 < n → ⫯⫰n = n. +* // #n #H lapply (ylt_inv_inj … H) -H normalize +/3 width=1 by S_pred, eq_f/ +qed-. (* Inversion lemmas on successor ********************************************) -fact ylt_inv_succ1_aux: ∀x,y. x < y → ∀m. x = ⫯m → ∃∃n. m < n & y = ⫯n. +fact ylt_inv_succ1_aux: ∀x,y. x < y → ∀m. x = ⫯m → m < ⫰y ∧ ⫯⫰y = y. #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 *) + #m #Hnm #H destruct /3 width=1 by ylt_inj, conj/ | #x #y #H elim (ysucc_inv_inj_sn … H) -H - #m #H #_ destruct - @(ex2_intro … (∞)) /2 width=1 by/ (**) (* explicit constructor *) + #m #H #_ destruct /2 width=1 by ylt_Y, conj/ ] qed-. -lemma ylt_inv_succ1: ∀m,y. ⫯m < y → ∃∃n. m < n & y = ⫯n. +lemma ylt_inv_succ1: ∀m,y. ⫯m < y → m < ⫰y ∧ ⫯⫰y = y. /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 // +lemma ylt_inv_succ: ∀m,n. ⫯m < ⫯n → m < n. +#m #n #H elim (ylt_inv_succ1 … H) -H // qed-. -fact ylt_inv_succ2_aux: ∀x,y. x < y → ∀n. y = ⫯n → x ≤ n. +(* Forward lemmas on successor **********************************************) + +fact ylt_fwd_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/ @@ -81,8 +93,8 @@ fact ylt_inv_succ2_aux: ∀x,y. x < y → ∀n. y = ⫯n → x ≤ n. ] qed-. -lemma ylt_inv_succ2: ∀m,n. m < ⫯n → m ≤ n. -/2 width=3 by ylt_inv_succ2_aux/ qed-. +lemma ylt_fwd_succ2: ∀m,n. m < ⫯n → m ≤ n. +/2 width=3 by ylt_fwd_succ2_aux/ qed-. (* inversion and forward lemmas on yle **************************************) @@ -99,9 +111,26 @@ lemma ylt_yle_false: ∀m:ynat. ∀n:ynat. m < n → n ≤ m → ⊥. ] qed-. -(* Properties on yle ********************************************************) +(* Basic properties *********************************************************) -lemma yle_to_ylt_or_eq: ∀m:ynat. ∀n:ynat. m ≤ n → m < n ∨ m = n. +lemma ylt_O: ∀x. ⫯⫰(yinj x) = yinj x → 0 < x. +* /2 width=1 by/ normalize +#H destruct +qed. + +(* Properties on successor **************************************************) + +lemma ylt_O_succ: ∀n. 0 < ⫯n. +* /2 width=1 by ylt_inj/ +qed. + +lemma ylt_succ: ∀m,n. m < n → ⫯m < ⫯n. +#m #n #H elim H -m -n /3 width=1 by ylt_inj, le_S_S/ +qed. + +(* Properties on order ******************************************************) + +lemma yle_split_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/ @@ -109,6 +138,11 @@ lemma yle_to_ylt_or_eq: ∀m:ynat. ∀n:ynat. m ≤ n → m < n ∨ m = n. ] qed-. +lemma ylt_split: ∀m,n:ynat. m < n ∨ n ≤ m.. +#m #n elim (yle_split m n) /2 width=1 by or_intror/ +#H elim (yle_split_eq … H) -H /2 width=1 by or_introl, or_intror/ +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 @@ -134,4 +168,4 @@ theorem ylt_trans: Transitive … ylt. /3 width=3 by transitive_lt, ylt_inj/ (**) (* full auto too slow *) | #x #z #H elim (ylt_yle_false … H) // ] -qed-. +qed-.