X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda%2Fpreamble.ma;h=c7c736f2ce1d3cad45775d2568fa537914ca2691;hb=2e700622e2565c6695e8c1264dd4c1207896f28c;hp=53a433d404bb96b5c7435a02ce17a5b10d246ea8;hpb=2b51cf74b9a5f37d0f91780ceae4b8f4d0ee38a1;p=helm.git diff --git a/matita/matita/contribs/lambda/preamble.ma b/matita/matita/contribs/lambda/preamble.ma index 53a433d40..c7c736f2c 100644 --- a/matita/matita/contribs/lambda/preamble.ma +++ b/matita/matita/contribs/lambda/preamble.ma @@ -12,11 +12,15 @@ (* *) (**************************************************************************) -include "arithmetics/nat.ma". +include "basics/star.ma". +include "basics/lists/lstar.ma". +include "arithmetics/exp.ma". include "xoa_notation.ma". include "xoa.ma". +(* logic *) + (* Note: For some reason this cannot be in the standard library *) interpretation "logical false" 'false = False. @@ -24,6 +28,8 @@ notation "⊥" non associative with precedence 90 for @{'false}. +(* arithmetics *) + lemma lt_refl_false: ∀n. n < n → ⊥. #n #H elim (lt_to_not_eq … H) -H /2 width=1/ qed-. @@ -31,3 +37,67 @@ qed-. lemma lt_zero_false: ∀n. n < 0 → ⊥. #n #H elim (lt_to_not_le … H) -H /2 width=1/ qed-. + +lemma plus_lt_false: ∀m,n. m + n < m → ⊥. +#m #n #H elim (lt_to_not_le … H) -H /2 width=1/ +qed-. + +lemma lt_or_eq_or_gt: ∀m,n. ∨∨ m < n | n = m | n < m. +#m #n elim (lt_or_ge m n) /2 width=1/ +#H elim H -m /2 width=1/ +#m #Hm * #H /2 width=1/ /3 width=1/ +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 ≝ + 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 ] + ]. + +lemma tri_lt: ∀A,a1,a2,a3,n2,n1. n1 < n2 → tri A n1 n2 a1 a2 a3 = a1. +#A #a1 #a2 #a3 #n2 elim n2 -n2 +[ #n1 #H elim (lt_zero_false … H) +| #n2 #IH #n1 elim n1 -n1 // /3 width=1/ +] +qed. + +lemma tri_eq: ∀A,a1,a2,a3,n. tri A n n a1 a2 a3 = a2. +#A #a1 #a2 #a3 #n elim n -n normalize // +qed. + +lemma tri_gt: ∀A,a1,a2,a3,n1,n2. n2 < n1 → tri A n1 n2 a1 a2 a3 = a3. +#A #a1 #a2 #a3 #n1 elim n1 -n1 +[ #n2 #H elim (lt_zero_false … H) +| #n1 #IH #n2 elim n2 -n2 // /3 width=1/ +] +qed. + +(* lists *) + +(* Note: notation for nil not involving brackets *) +notation > "◊" + non associative with precedence 90 + for @{'nil}. + +definition map_cons: ∀A. A → list (list A) → list (list A) ≝ λA,a. + map … (cons … a). + +interpretation "map_cons" 'ho_cons a l = (map_cons ? a l). + +notation "hvbox(a ::: break l)" + right associative with precedence 47 + for @{'ho_cons $a $l}. + +(* lstar *) + +(* Note: this cannot be in lib because of the missing xoa quantifier *) +lemma lstar_inv_pos: ∀A,B,R,l,b1,b2. lstar A B R l b1 b2 → 0 < |l| → + ∃∃a,ll,b. a::ll = l & R a b1 b & lstar A B R ll b b2. +#A #B #R #l #b1 #b2 #H @(lstar_ind_l ????????? H) -b1 +[ #H elim (lt_refl_false … H) +| #a #ll #b1 #b #Hb1 #Hb2 #_ #_ /2 width=6/ (**) (* auto fail if we do not remove the inductive premise *) +] +qed-.