X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Fsh.ma;h=824e37e1fd942bd1a0d5fac70af321005a9234a3;hb=c60524dec7ace912c416a90d6b926bee8553250b;hp=0d09320cc871665e968a39e15b0f3434c31d917a;hpb=f10cfe417b6b8ec1c7ac85c6ecf5fb1b3fdf37db;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/sh.ma b/matita/matita/contribs/lambdadelta/basic_2/static/sh.ma index 0d09320cc..824e37e1f 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/sh.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/sh.ma @@ -27,18 +27,18 @@ definition sh_N: sh ≝ mk_sh S …. (* Basic properties *********************************************************) -lemma nexts_le: ∀h,k,l. k ≤ (next h)^l k. -#h #k #l elim l -l // normalize #l #IHl -lapply (next_lt h ((next h)^l k)) #H -lapply (le_to_lt_to_lt … IHl H) -IHl -H /2 width=2 by lt_to_le/ +lemma nexts_le: ∀h,k,d. k ≤ (next h)^d k. +#h #k #d elim d -d // normalize #d #IHd +lapply (next_lt h ((next h)^d k)) #H +lapply (le_to_lt_to_lt … IHd H) -IHd -H /2 width=2 by lt_to_le/ qed. -lemma nexts_lt: ∀h,k,l. k < (next h)^(l+1) k. -#h #k #l >iter_SO -lapply (nexts_le h k l) #H +lemma nexts_lt: ∀h,k,d. k < (next h)^(d+1) k. +#h #k #d >iter_SO +lapply (nexts_le h k d) #H @(le_to_lt_to_lt … H) // qed. -axiom nexts_dec: ∀h,k1,k2. Decidable (∃l. (next h)^l k1 = k2). +axiom nexts_dec: ∀h,k1,k2. Decidable (∃d. (next h)^d k1 = k2). -axiom nexts_inj: ∀h,k,l1,l2. (next h)^l1 k = (next h)^l2 k → l1 = l2. +axiom nexts_inj: ∀h,k,d1,d2. (next h)^d1 k = (next h)^d2 k → d1 = d2.