X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Fynat%2Fynat_le.ma;h=54e946e96ca40a4ace88868863b258d2ce282cd5;hb=1fd63df4c77f5c24024769432ea8492748b4ac79;hp=be13774f5813aa5ca9fcb86b3669f3e3d65f58a8;hpb=23da2aa16489e00889374d81f19cc090faa44582;p=helm.git 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 be13774f5..54e946e96 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_le.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/ynat/ynat_le.ma @@ -46,7 +46,7 @@ fact yle_inv_O2_aux: ∀m:ynat. ∀x:ynat. m ≤ x → x = 0 → m = 0. #m #x * -m -x [ #m #n #Hmn #H destruct /3 width=1 by le_n_O_to_eq, eq_f/ | #m #H destruct -] +] qed-. lemma yle_inv_O2: ∀m:ynat. m ≤ 0 → m = 0. @@ -60,50 +60,88 @@ qed-. lemma yle_inv_Y1: ∀n. ∞ ≤ n → n = ∞. /2 width=3 by yle_inv_Y1_aux/ qed-. +lemma yle_antisym: ∀y,x. x ≤ y → y ≤ x → x = y. +#x #y #H elim H -x -y +/4 width=1 by yle_inv_Y1, yle_inv_inj, le_to_le_to_eq, eq_f/ +qed-. + +(* Basic properties *********************************************************) + +lemma le_O1: ∀n:ynat. 0 ≤ n. +* /2 width=1 by yle_inj/ +qed. + +lemma yle_refl: reflexive … yle. +* /2 width=1 by le_n, yle_inj/ +qed. + +lemma yle_split: ∀x,y:ynat. x ≤ y ∨ y ≤ x. +* /2 width=1 by or_intror/ +#x * /2 width=1 by or_introl/ +#y elim (le_or_ge x y) /3 width=1 by yle_inj, or_introl, or_intror/ +qed-. + (* Inversion lemmas on successor ********************************************) -fact yle_inv_succ1_aux: ∀x,y. x ≤ y → ∀m. x = ⫯m → ∃∃n. m ≤ n & y = ⫯n. +fact yle_inv_succ1_aux: ∀x,y:ynat. 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 yle_inj/ (**) (* explicit constructor *) -| #x #y #H destruct - @(ex2_intro … (∞)) /2 width=1 by yle_Y/ (**) (* explicit constructor *) + #m #Hnm #H destruct /3 width=1 by yle_inj, conj/ +| #x #y #H destruct /2 width=1 by yle_Y, conj/ ] qed-. -lemma yle_inv_succ1: ∀m,y. ⫯m ≤ y → ∃∃n. m ≤ n & y = ⫯n. +lemma yle_inv_succ1: ∀m,y:ynat. ↑m ≤ y → m ≤ ↓y ∧ ↑↓y = y. /2 width=3 by yle_inv_succ1_aux/ qed-. -lemma yle_inv_succ: ∀m,n. ⫯m ≤ ⫯n → m ≤ n. -#m #n #H elim (yle_inv_succ1 … H) -H -#x #Hx #H destruct // +lemma yle_inv_succ: ∀m,n. ↑m ≤ ↑n → m ≤ n. +#m #n #H elim (yle_inv_succ1 … H) -H // qed-. -(* Basic properties *********************************************************) - -lemma yle_refl: reflexive … yle. -* /2 width=1 by le_n, yle_inj/ -qed. +lemma yle_inv_succ2: ∀x,y. x ≤ ↑y → ↓x ≤ y. +#x #y #Hxy elim (ynat_cases x) +[ #H destruct // +| * #m #H destruct /2 width=1 by yle_inv_succ/ +] +qed-. (* Properties on predecessor ************************************************) -lemma yle_pred_sn: ∀m,n. m ≤ n → ⫰m ≤ n. +lemma yle_pred_sn: ∀m,n. m ≤ n → ↓m ≤ n. #m #n * -m -n /3 width=3 by transitive_le, yle_inj/ qed. -lemma yle_refl_pred_sn: ∀x. ⫰x ≤ x. +lemma yle_refl_pred_sn: ∀x. ↓x ≤ x. /2 width=1 by yle_refl, yle_pred_sn/ qed. +lemma yle_pred: ∀m,n. m ≤ n → ↓m ≤ ↓n. +#m #n * -m -n /3 width=1 by yle_inj, monotonic_pred/ +qed. + (* Properties on successor **************************************************) -lemma yle_succ_dx: ∀m,n. m ≤ n → m ≤ ⫯n. +lemma yle_succ: ∀m,n. m ≤ n → ↑m ≤ ↑n. +#m #n * -m -n /3 width=1 by yle_inj, le_S_S/ +qed. + +lemma yle_succ_dx: ∀m,n. m ≤ n → m ≤ ↑n. #m #n * -m -n /3 width=1 by le_S, yle_inj/ qed. -lemma yle_refl_S_dx: ∀x. x ≤ ⫯x. -/2 width=1 by yle_refl, yle_succ_dx/ qed. +lemma yle_refl_S_dx: ∀x. x ≤ ↑x. +/2 width=1 by yle_succ_dx/ qed. + +lemma yle_refl_SP_dx: ∀x. x ≤ ↑↓x. +* // * // +qed. + +lemma yle_succ2: ∀x,y. ↓x ≤ y → x ≤ ↑y. +#x #y #Hxy elim (ynat_cases x) +[ #H destruct // +| * #m #H destruct /2 width=1 by yle_succ/ +] +qed-. (* Main properties **********************************************************) @@ -114,4 +152,4 @@ theorem yle_trans: Transitive … yle. /3 width=3 by transitive_le, yle_inj/ (**) (* full auto too slow *) | #x #z #H lapply (yle_inv_Y1 … H) // ] -qed-. +qed-.