X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambda%2Fpreamble.ma;h=c7c736f2ce1d3cad45775d2568fa537914ca2691;hb=f14ac9cceaaa7f1905f0dbc4548d24c4abe940e3;hp=640615d596f2aa2aa2947a23884d7a3ec9b09916;hpb=178820be7648a60af17837727e51fd1f3f2791db;p=helm.git diff --git a/matita/matita/contribs/lambda/preamble.ma b/matita/matita/contribs/lambda/preamble.ma index 640615d59..c7c736f2c 100644 --- a/matita/matita/contribs/lambda/preamble.ma +++ b/matita/matita/contribs/lambda/preamble.ma @@ -28,16 +28,6 @@ notation "⊥" non associative with precedence 90 for @{'false}. -(* relations *) - -definition confluent1: ∀A. relation A → predicate A ≝ λA,R,a0. - ∀a1. R a0 a1 → ∀a2. R a0 a2 → - ∃∃a. R a1 a & R a2 a. - -(* Note: confluent1 would be redundant if \Pi-reduction where enabled *) -definition confluent: ∀A. predicate (relation A) ≝ λA,R. - ∀a0. confluent1 … R a0. - (* arithmetics *) lemma lt_refl_false: ∀n. n < n → ⊥. @@ -100,3 +90,14 @@ 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-.