X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fstatic%2Fsd.ma;h=acdc78ac6533de4c23c51443a9e7a335b3dde4eb;hb=c9a1672c725945b47f9ea8af3c23b67cf9026f01;hp=35500d952b1f0c343ed6453146756a9b3ca7b381;hpb=ef49e0e7f5f298c299afdd3cbfdc2404ecb93879;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/sd.ma b/matita/matita/contribs/lambdadelta/basic_2/static/sd.ma index 35500d952..acdc78ac6 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/sd.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/sd.ma @@ -70,8 +70,8 @@ definition sd_SO: ∀h. nat → sd h ≝ λh,k. mk_sd h (deg_SO h k) …. | #K0 #l1 #l2 * [ #l01 ] #H1 * [1,3: #l02 ] #H2 // [ < H2 in H1; -H2 #H lapply (nexts_inj … H) -H #H destruct // - | elim (H1 ?) /2 width=2/ - | elim (H2 ?) /2 width=2/ + | elim H1 /2 width=2/ + | elim H2 /2 width=2/ ] | #k0 #l0 * [ #l #H destruct elim l -l normalize /2 width=1/ @@ -90,32 +90,39 @@ let rec sd_l (h:sh) (k:nat) (l:nat) on l : sd h ≝ ] ]. -(* Basic properties *********************************************************) +(* Basic inversion lemmas ***************************************************) -lemma deg_iter: ∀h,g,k,l1,l2. deg h g k l1 → deg h g ((next h)^l2 k) (l1-l2). -#h #g #k #l1 #l2 @(nat_ind_plus … l2) -l2 [ iter_SO H iter_SO #H -lapply (deg_pred … H) -H <(associative_plus l0 1 1) #H +lapply (deg_inv_pred … H) -H <(associative_plus l0 1 1) #H lapply (IHl … H) -IHl -H // +qed-. + +(* Basic properties *********************************************************) + +lemma deg_iter: ∀h,g,k,l1,l2. deg h g k l1 → deg h g ((next h)^l2 k) (l1-l2). +#h #g #k #l1 #l2 @(nat_ind_plus … l2) -l2 [ iter_SO