X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=blobdiff_plain;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Fsyntax%2Fappend.ma;h=b3a07339bbe3bedba4333fccdfde4b04cceb9b10;hp=55e6e3472643150df032765e1e7cc468d0853c83;hb=a77d0bd6a04e94f765d329d47b37d9e04d349b14;hpb=b598b37379baabef24ae511596be7f740cbb0c2e diff --git a/matita/matita/contribs/lambdadelta/basic_2/syntax/append.ma b/matita/matita/contribs/lambdadelta/basic_2/syntax/append.ma index 55e6e3472..b3a07339b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/syntax/append.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/syntax/append.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "ground_2/notation/functions/append_2.ma". include "basic_2/notation/functions/snitem_2.ma". include "basic_2/notation/functions/snbind1_2.ma". include "basic_2/notation/functions/snbind2_3.ma". @@ -28,7 +27,7 @@ rec definition append L K on K ≝ match K with | LBind K I ⇒ (append L K).ⓘ{I} ]. -interpretation "append (local environment)" 'Append L1 L2 = (append L1 L2). +interpretation "append (local environment)" 'plus L1 L2 = (append L1 L2). interpretation "local environment tail binding construction (generic)" 'SnItem I L = (append (LBind LAtom I) L). @@ -49,18 +48,18 @@ interpretation "tail abstraction (local environment)" 'SnAbst L T = (append (LBind LAtom (BPair Abst T)) L). definition d_appendable_sn: predicate (lenv→relation term) ≝ λR. - ∀K,T1,T2. R K T1 T2 → ∀L. R (L@@K) T1 T2. + ∀K,T1,T2. R K T1 T2 → ∀L. R (L+K) T1 T2. (* Basic properties *********************************************************) -lemma append_atom: ∀L. L @@ ⋆ = L. +lemma append_atom: ∀L. (L + ⋆) = L. (**) (* () should be redundant *) // qed. (* Basic_2A1: uses: append_pair *) -lemma append_bind: ∀I,L,K. L@@(K.ⓘ{I}) = (L@@K).ⓘ{I}. +lemma append_bind: ∀I,L,K. L+(K.ⓘ{I}) = (L+K).ⓘ{I}. // qed. -lemma append_atom_sn: ∀L. ⋆@@L = L. +lemma append_atom_sn: ∀L. ⋆ + L = L. #L elim L -L // #L #I >append_bind // qed. @@ -69,25 +68,25 @@ lemma append_assoc: associative … append. #L1 #L2 #L3 elim L3 -L3 // qed. -lemma append_shift: ∀L,K,I. L@@(ⓘ{I}.K) = (L.ⓘ{I})@@K. +lemma append_shift: ∀L,K,I. L+(ⓘ{I}.K) = (L.ⓘ{I})+K. #L #K #I append_bind #H destruct qed-. -lemma append_inv_bind3_sn: ∀I0,L,L0,K. L0.ⓘ{I0} = L @@ K → +lemma append_inv_bind3_sn: ∀I0,L,L0,K. L0.ⓘ{I0} = L + K → ∨∨ ∧∧ L0.ⓘ{I0} = L & ⋆ = K - | ∃∃K0. K = K0.ⓘ{I0} & L0 = L @@ K0. + | ∃∃K0. K = K0.ⓘ{I0} & L0 = L + K0. #I0 #L #L0 * /3 width=1 by or_introl, conj/ #K #I >append_bind #H destruct /3 width=3 by ex2_intro, or_intror/ qed-. -lemma append_inj_sn: ∀K,L1,L2. L1@@K = L2@@K → L1 = L2. +lemma append_inj_sn: ∀K,L1,L2. L1+K = L2+K → L1 = L2. #K elim K -K // #K #I #IH #L1 #L2 >append_bind #H elim (destruct_lbind_lbind_aux … H) -H /2 width=1 by/ (**) (* destruct lemma needed *)