X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fsyntax%2Fappend_length.ma;h=2222c52ff01876cb6f2f8b675b7e587e92829352;hp=c76a2a6e7a915eaee2a0ac5e8329e8b52c685fee;hb=98e786e1a6bd7b621e37ba7cd4098d4a0a6f8278;hpb=5d9f7ae4bad2b5926f615141c12942b9a8eb23fb diff --git a/matita/matita/contribs/lambdadelta/static_2/syntax/append_length.ma b/matita/matita/contribs/lambdadelta/static_2/syntax/append_length.ma index c76a2a6e7..2222c52ff 100644 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/append_length.ma +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/append_length.ma @@ -72,12 +72,12 @@ lemma append_inj_length_dx: ∀K1,K2,L1,L2. L1 + K1 = L2 + K2 → |L1| = |L2| #K1 elim K1 -K1 [ * /2 width=1 by conj/ #K2 #I2 #L1 #L2 >append_atom >append_bind #H destruct - >length_bind >append_length >plus_n_Sm - #H elim (plus_xSy_x_false … H) + >length_bind >append_length >nplus_succ_dx + #H elim (succ_nplus_refl_sn … H) | #K1 #I1 #IH * [ #L1 #L2 >append_bind >append_atom #H destruct - >length_bind >append_length >plus_n_Sm #H - lapply (discr_plus_x_xy … H) -H #H destruct + >length_bind >append_length >nplus_succ_dx #H + lapply (nplus_refl_sn … H) -H #H destruct | #K2 #I2 #L1 #L2 >append_bind >append_bind #H1 #H2 elim (destruct_lbind_lbind_aux … H1) -H1 #H1 #H3 destruct (**) (* destruct lemma needed *) elim (IH … H1) -IH -H1 /2 width=1 by conj/ @@ -105,7 +105,7 @@ qed-. (* Basic_2A1: was: lenv_ind_alt *) lemma lenv_ind_tail: ∀Q:predicate lenv. Q (⋆) → (∀I,L. Q L → Q (ⓘ[I].L)) → ∀L. Q L. -#Q #IH1 #IH2 #L @(f_ind … length … L) -L #x #IHx * // +#Q #IH1 #IH2 #L @(wf1_ind_nlt … length … L) -L #x #IHx * // #L #I -IH1 #H destruct elim (lenv_case_tail … L) [2: * #K #J ] #H destruct /3 width=1 by/