X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Freduction%2Fcnx.ma;h=e6d1fcf4b77b244f77afdd506a431ba67c053545;hb=c60524dec7ace912c416a90d6b926bee8553250b;hp=0a23bcf039c24d907a909ed77e4993d76474a892;hpb=f10cfe417b6b8ec1c7ac85c6ecf5fb1b3fdf37db;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma index 0a23bcf03..e6d1fcf4b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cnx.ma @@ -29,9 +29,11 @@ interpretation lemma cnx_inv_sort: ∀h,g,G,L,k. ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃⋆k⦄ → deg h g k 0. #h #g #G #L #k #H elim (deg_total h g k) -#l @(nat_ind_plus … l) -l // #l #_ #Hkl -lapply (H (⋆(next h k)) ?) -H /2 width=2 by cpx_st/ -L -l #H destruct -H -e0 (**) (* destruct does not remove some premises *) -lapply (next_lt h k) >e1 -e1 #H elim (lt_refl_false … H) +#d @(nat_ind_plus … d) -d // #d #_ #Hkd +lapply (H (⋆(next h k)) ?) -H /2 width=2 by cpx_st/ -L -d #H +lapply (destruct_tatom_tatom_aux … H) -H #H (**) (* destruct lemma needed *) +lapply (destruct_sort_sort_aux … H) -H #H (**) (* destruct lemma needed *) +lapply (next_lt h k) >H -H #H elim (lt_refl_false … H) qed-. lemma cnx_inv_delta: ∀h,g,I,G,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃#i⦄ → ⊥. @@ -97,13 +99,13 @@ qed-. (* Basic properties *********************************************************) lemma cnx_sort: ∀h,g,G,L,k. deg h g k 0 → ⦃G, L⦄ ⊢ ➡[h, g] 𝐍⦃⋆k⦄. -#h #g #G #L #k #Hk #X #H elim (cpx_inv_sort1 … H) -H // * #l #Hkl #_ -lapply (deg_mono … Hkl Hk) -h -L