X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fstatic_2%2Fsyntax%2Fappend_length.ma;h=1e0de051b2573daeb4bb66eba04a2dca1761f337;hb=156d974ad89aa04a086fdf9d332c8b04adf279fd;hp=2222c52ff01876cb6f2f8b675b7e587e92829352;hpb=98e786e1a6bd7b621e37ba7cd4098d4a0a6f8278;p=helm.git 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 2222c52ff..1e0de051b 100644 --- a/matita/matita/contribs/lambdadelta/static_2/syntax/append_length.ma +++ b/matita/matita/contribs/lambdadelta/static_2/syntax/append_length.ma @@ -12,6 +12,8 @@ (* *) (**************************************************************************) +include "ground/arith/nat_plus.ma". +include "ground/arith/wf1_ind_nlt.ma". include "static_2/syntax/lenv_length.ma". include "static_2/syntax/append.ma". @@ -34,31 +36,36 @@ qed. lemma length_inv_succ_dx_ltail: ∀L,n. |L| = ↑n → ∃∃I,K. |K| = n & L = ⓘ[I].K. #Y #n #H elim (length_inv_succ_dx … H) -H #I #L #Hn #HLK destruct -elim (lenv_case_tail … L) [2: * #K #J ] -#H destruct /2 width=4 by ex2_2_intro/ +elim (lenv_case_tail … L) [| * #K #J ] #H destruct +/2 width=4 by ex2_2_intro/ +@(ex2_2_intro … (J) (K.ⓘ[I])) // (**) (* auto fails *) +>ltail_length >length_bind // qed-. (* Basic_2A1: was: length_inv_pos_sn_ltail *) lemma length_inv_succ_sn_ltail: ∀L,n. ↑n = |L| → ∃∃I,K. n = |K| & L = ⓘ[I].K. #Y #n #H elim (length_inv_succ_sn … H) -H #I #L #Hn #HLK destruct -elim (lenv_case_tail … L) [2: * #K #J ] -#H destruct /2 width=4 by ex2_2_intro/ +elim (lenv_case_tail … L) [| * #K #J ] #H destruct +/2 width=4 by ex2_2_intro/ +@(ex2_2_intro … (J) (K.ⓘ[I])) // (**) (* auto fails *) +>ltail_length >length_bind // qed-. (* Inversion lemmas with length for local environments **********************) (* Basic_2A1: was: append_inj_sn *) lemma append_inj_length_sn: ∀K1,K2,L1,L2. L1 + K1 = L2 + K2 → |K1| = |K2| → - L1 = L2 ∧ K1 = K2. + ∧∧ L1 = L2 & K1 = K2. #K1 elim K1 -K1 [ * /2 width=1 by conj/ #K2 #I2 #L1 #L2 #_ >length_atom >length_bind - #H destruct + #H elim (eq_inv_zero_nsucc … H) | #K1 #I1 #IH * [ #L1 #L2 #_ >length_atom >length_bind - #H destruct + #H elim (eq_inv_nsucc_zero … H) | #K2 #I2 #L1 #L2 #H1 >length_bind >length_bind #H2 + lapply (eq_inv_nsucc_bi … H2) -H2 #H2 elim (destruct_lbind_lbind_aux … H1) -H1 #H1 #H3 destruct (**) (* destruct lemma needed *) elim (IH … H1) -IH -H1 /3 width=4 by conj/ ] @@ -68,16 +75,16 @@ qed-. (* Note: lemma 750 *) (* Basic_2A1: was: append_inj_dx *) lemma append_inj_length_dx: ∀K1,K2,L1,L2. L1 + K1 = L2 + K2 → |L1| = |L2| → - L1 = L2 ∧ K1 = K2. + ∧∧ L1 = L2 & K1 = K2. #K1 elim K1 -K1 [ * /2 width=1 by conj/ #K2 #I2 #L1 #L2 >append_atom >append_bind #H destruct - >length_bind >append_length >nplus_succ_dx - #H elim (succ_nplus_refl_sn … H) + >length_bind >append_length #H + elim (succ_nplus_refl_sn (|L2|) (|K2|) ?) // | #K1 #I1 #IH * [ #L1 #L2 >append_bind >append_atom #H destruct - >length_bind >append_length >nplus_succ_dx #H - lapply (nplus_refl_sn … H) -H #H destruct + >length_bind >append_length #H + elim (succ_nplus_refl_sn … H) | #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/ @@ -107,6 +114,6 @@ lemma lenv_ind_tail: ∀Q:predicate lenv. Q (⋆) → (∀I,L. Q L → Q (ⓘ[I].L)) → ∀L. Q L. #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 ] +elim (lenv_case_tail … L) [| * #K #J ] #H destruct /3 width=1 by/ qed-.