X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fground_2%2Flib%2Farith.ma;h=6ce878b44bf0d73d1ba13dfd8e23a19cf008441e;hb=bac74b5cff042d37e1abc9c961a6c41094b8a294;hp=0574e1757927e90a2f848b7ec4e6d550891fb823;hpb=cacd7323994f7621286dbfd93bbf4c50acfbe918;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 0574e1757..6ce878b44 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/lib/arith.ma @@ -57,6 +57,9 @@ lemma plus_SO: ∀n. n + 1 = ↑n. lemma minus_plus_m_m_commutative: ∀n,m:nat. n = m + n - m. // qed-. +lemma plus_minus_m_m_commutative (n) (m): m ≤ n → n = m+(n-m). +/2 width=1 by plus_minus_associative/ qed-. + lemma plus_to_minus_2: ∀m1,m2,n1,n2. n1 ≤ m1 → n2 ≤ m2 → m1+n2 = m2+n1 → m1-n1 = m2-n2. #m1 #m2 #n1 #n2 #H1 #H2 #H @@ -166,6 +169,15 @@ 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 le_dec (n) (m): Decidable (n≤m). +#n elim n -n [ /2 width=1 by or_introl/ ] +#n #IH * [ /3 width=2 by lt_zero_false, or_intror/ ] +#m elim (IH m) -IH +[ /3 width=1 by or_introl, le_S_S/ +| /4 width=1 by or_intror, le_S_S_to_le/ +] +qed-. + lemma succ_inv_refl_sn: ∀x. ↑x = x → ⊥. #x #H @(lt_le_false x (↑x)) // qed-. @@ -330,3 +342,34 @@ lemma tri_gt: ∀A,a1,a2,a3,n1,n2. n2 < n1 → tri A n1 n2 a1 a2 a3 = a3. | #n1 #IH #n2 elim n2 -n2 /3 width=1 by monotonic_lt_pred/ ] qed. + +(* Decidability of predicates ***********************************************) + +lemma dec_lt (R:predicate nat): + (∀n. Decidable … (R n)) → + ∀n. Decidable … (∃∃m. m < n & R m). +#R #HR #n elim n -n [| #n * ] +[ @or_intror * /2 width=2 by lt_zero_false/ +| * /4 width=3 by lt_S, or_introl, ex2_intro/ +| #H0 elim (HR n) -HR + [ /3 width=3 by or_introl, ex2_intro/ + | #Hn @or_intror * #m #Hmn #Hm + elim (le_to_or_lt_eq … Hmn) -Hmn #H destruct [ -Hn | -H0 ] + /4 width=3 by lt_S_S_to_lt, ex2_intro/ + ] +] +qed-. + +lemma dec_min (R:predicate nat): + (∀n. Decidable … (R n)) → ∀n. R n → + ∃∃m. m ≤ n & R m & (∀p. p < m → R p → ⊥). +#R #HR #n +@(nat_elim1 n) -n #n #IH #Hn +elim (dec_lt … HR n) -HR [ -Hn | -IH ] +[ * #p #Hpn #Hp + elim (IH … Hpn Hp) -IH -Hp #m #Hmp #Hm #HNm + @(ex3_intro … Hm HNm) -HNm + /3 width=3 by lt_to_le, le_to_lt_to_lt/ +| /4 width=4 by ex3_intro, ex2_intro/ +] +qed-.