X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=matita%2Fmatita%2Fcontribs%2Flambdadelta%2Fbasic_2%2Frelocation%2Fcpy.ma;h=e1b2afd754ef86043369bd39e3b45aa7e0ea7922;hb=6907a0f66a3782ce4273a609203c9b574841c7d1;hp=025cd6f65c6ef6edea32ceba4fc9515e1e52bef5;hpb=e0c1f02de90fdb4e16f322a5bcb39f16c2dc477e;p=helm.git diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma index 025cd6f65..e1b2afd75 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma @@ -14,6 +14,7 @@ include "basic_2/notation/relations/extpsubst_6.ma". include "basic_2/grammar/genv.ma". +include "basic_2/grammar/cl_shift.ma". include "basic_2/relocation/ldrop_append.ma". include "basic_2/relocation/lsuby.ma". @@ -37,12 +38,13 @@ interpretation "context-sensitive extended parallel substritution (term)" (* Basic properties *********************************************************) -lemma lsuby_cpy_trans: ∀G,d,e.lsub_trans … (cpy d e G) lsuby. +lemma lsuby_cpy_trans: ∀G,L1,T1,T2,d,e. ⦃G, L1⦄ ⊢ T1 ▶×[d, e] T2 → + ∀L2. L2 ⊑×[d, e] L1 → ⦃G, L2⦄ ⊢ T1 ▶×[d, e] T2. #G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e [ // | #I #G #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12 - elim (lsuby_fwd_ldrop2_bind … HL12 … HLK1) -HL12 -HLK1 /2 width=5 by cpy_subst/ -| /4 width=1 by lsuby_pair, cpy_bind/ + elim (lsuby_fwd_ldrop2_be … HL12 … HLK1) -HL12 -HLK1 /2 width=5 by cpy_subst/ +| /4 width=1 by lsuby_succ, cpy_bind/ | /3 width=1 by cpy_flat/ ] qed-. @@ -117,154 +119,155 @@ lemma cpy_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → ∀i. elim (IHV12 i) -IHV12 // #V #HV1 #HV2 elim (IHT12 (i + 1)) -IHT12 /2 width=1 by le_S_S/ -Hdi -Hide >arith_c1x #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT1 (L.ⓑ{I} V) ?) -HT1 /3 width=5 by cpy_bind, ex2_intro/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // - -Hdi -Hide /3 width=5/ + lapply (lsuby_cpy_trans … HT1 (L.ⓑ{I} V) ?) -HT1 /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // + -Hdi -Hide /3 width=5 by ex2_intro, cpy_flat/ ] -qed. +qed-. -lemma cpy_split_down: ∀L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → +lemma cpy_split_down: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → ∀i. d ≤ i → i ≤ d + e → ∃∃T. ⦃G, L⦄ ⊢ T1 ▶×[i, d + e - i] T & ⦃G, L⦄ ⊢ T ▶×[d, i - d] T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e -[ /2 width=3/ -| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e +[ /2 width=3 by ex2_intro/ +| #I #G #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde elim (lt_or_ge i j) - [ -Hide -Hjde >(plus_minus_m_m j d) in⦄ ⊢ (% → ?); // -Hdj /3 width=8/ + [ -Hide -Hjde >(plus_minus_m_m j d) in ⊢ (% → ?); // -Hdj /3 width=9 by ex2_intro, cpy_atom, cpy_subst/ | -Hdi -Hdj - >(plus_minus_m_m (d+e) j) in Hide; // -Hjde /3 width=4/ + >(plus_minus_m_m (d+e) j) in Hide; // -Hjde /3 width=5 by ex2_intro, cpy_subst/ ] -| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 - elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i) -IHV12 // #V + elim (IHT12 (i + 1)) -IHT12 /2 width=1 by le_S_S/ -Hdi -Hide >arith_c1x #T #HT1 #HT2 - lapply (cpy_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // - -Hdi -Hide /3 width=5/ + lapply (lsuby_cpy_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // + -Hdi -Hide /3 width=5 by ex2_intro, cpy_flat/ ] -qed. +qed-. -lemma cpy_append: ∀K,T1,T2,d,e. K⦄ ⊢ T1 ▶×[d, e] T2 → - ∀L. L @@ K⦄ ⊢ T1 ▶×[d, e] T2. -#K #T1 #T2 #d #e #H elim H -K -T1 -T2 -d -e // /2 width=1/ -#K #K0 #V #W #i #d #e #Hdi #Hide #HK0 #HVW #L -lapply (ldrop_fwd_ldrop2_length … HK0) #H -@(cpy_subst … (L@@K0) … HVW) // (**) (* /3/ does not work *) -@(ldrop_O1_append_sn_le … HK0) /2 width=2/ -qed. +lemma cpy_append: ∀G,K,T1,T2,d,e. ⦃G, K⦄ ⊢ T1 ▶×[d, e] T2 → + ∀L. ⦃G, L @@ K⦄ ⊢ T1 ▶×[d, e] T2. +#G #K #T1 #T2 #d #e #H elim H -K -T1 -T2 -d -e +/2 width=1 by cpy_atom, cpy_bind, cpy_flat/ +#I #G #K #K0 #V #W #i #d #e #Hdi #Hide #HK0 #HVW #L +lapply (ldrop_fwd_length_lt2 … HK0) #H +@(cpy_subst I … (L@@K0) … HVW) // (**) (* /4/ does not work *) +@(ldrop_O1_append_sn_le … HK0) /2 width=2 by lt_to_le/ +qed-. (* Basic inversion lemmas ***************************************************) -fact cpy_inv_atom1_aux: ∀L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → ∀I. T1 = ⓪{I} → - T2 = ⓪{I} ∨ - ∃∃K,V,i. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2 & - I = LRef i. -#L #T1 #T2 #d #e * -L -T1 -T2 -d -e -[ #L #I #d #e #J #H destruct /2 width=1/ -| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #I #H destruct /3 width=8/ -| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +fact cpy_inv_atom1_aux: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → ∀J. T1 = ⓪{J} → + T2 = ⓪{J} ∨ + ∃∃I,K,V,i. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K.ⓑ{I}V & + ⇧[O, i + 1] V ≡ T2 & + J = LRef i. +#G #L #T1 #T2 #d #e * -G -L -T1 -T2 -d -e +[ #I #G #L #d #e #J #H destruct /2 width=1 by or_introl/ +| #I #G #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #J #H destruct /3 width=9 by ex5_4_intro, or_intror/ +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct ] -qed. +qed-. -lemma cpy_inv_atom1: ∀L,T2,I,d,e. ⦃G, L⦄ ⊢ ⓪{I} ▶×[d, e] T2 → +lemma cpy_inv_atom1: ∀I,G,L,T2,d,e. ⦃G, L⦄ ⊢ ⓪{I} ▶×[d, e] T2 → T2 = ⓪{I} ∨ - ∃∃K,V,i. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2 & - I = LRef i. -/2 width=3/ qed-. + ∃∃J,K,V,i. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K.ⓑ{J}V & + ⇧[O, i + 1] V ≡ T2 & + I = LRef i. +/2 width=4 by cpy_inv_atom1_aux/ qed-. - -(* Basic_1: was: subst1_gen_sort *) -lemma cpy_inv_sort1: ∀L,T2,k,d,e. ⦃G, L⦄ ⊢ ⋆k ▶×[d, e] T2 → T2 = ⋆k. -#L #T2 #k #d #e #H +lemma cpy_inv_sort1: ∀G,L,T2,k,d,e. ⦃G, L⦄ ⊢ ⋆k ▶×[d, e] T2 → T2 = ⋆k. +#G #L #T2 #k #d #e #H elim (cpy_inv_atom1 … H) -H // -* #K #V #i #_ #_ #_ #_ #H destruct +* #I #K #V #i #_ #_ #_ #_ #H destruct qed-. -(* Basic_1: was: subst1_gen_lref *) -lemma cpy_inv_lref1: ∀L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶×[d, e] T2 → +lemma cpy_inv_lref1: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶×[d, e] T2 → T2 = #i ∨ - ∃∃K,V. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2. -#L #T2 #i #d #e #H -elim (cpy_inv_atom1 … H) -H /2 width=1/ -* #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=4/ + ∃∃I,K,V. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K.ⓑ{I}V & + ⇧[O, i + 1] V ≡ T2. +#G #L #T2 #i #d #e #H +elim (cpy_inv_atom1 … H) -H /2 width=1 by or_introl/ +* #I #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=5 by ex4_3_intro, or_intror/ qed-. -lemma cpy_inv_gref1: ∀L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶×[d, e] T2 → T2 = §p. -#L #T2 #p #d #e #H +lemma cpy_inv_gref1: ∀G,L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶×[d, e] T2 → T2 = §p. +#G #L #T2 #p #d #e #H elim (cpy_inv_atom1 … H) -H // -* #K #V #i #_ #_ #_ #_ #H destruct +* #I #K #V #i #_ #_ #_ #_ #H destruct qed-. -fact cpy_inv_bind1_aux: ∀d,e,L,U1,U2. ⦃G, L⦄ ⊢ U1 ▶×[d, e] U2 → - ∀a,I,V1,T1. U1 = ⓑ{a,I} V1. T1 → +fact cpy_inv_bind1_aux: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶×[d, e] U2 → + ∀a,I,V1,T1. U1 = ⓑ{a,I}V1.T1 → ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶×[d, e] V2 & - L. ⓑ{I} V2⦄ ⊢ T1 ▶×[d + 1, e] T2 & - U2 = ⓑ{a,I} V2. T2. -#d #e #L #U1 #U2 * -d -e -L -U1 -U2 -[ #L #k #d #e #a #I #V1 #T1 #H destruct -| #L #K #V #W #i #d #e #_ #_ #_ #_ #a #I #V1 #T1 #H destruct -| #L #b #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #a #I #V #T #H destruct /2 width=5/ -| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #a #I #V #T #H destruct + ⦃G, L. ⓑ{I}V2⦄ ⊢ T1 ▶×[d + 1, e] T2 & + U2 = ⓑ{a,I}V2.T2. +#G #L #U1 #U2 #d #e * -G -L -U1 -U2 -d -e +[ #I #G #L #d #e #b #J #W1 #U1 #H destruct +| #I #G #L #K #V #W #i #d #e #_ #_ #_ #_ #b #J #W1 #U1 #H destruct +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #b #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #b #J #W1 #U1 #H destruct ] -qed. +qed-. -lemma cpy_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. ⦃G, L⦄ ⊢ ⓑ{a,I} V1. T1 ▶×[d, e] U2 → +lemma cpy_inv_bind1: ∀a,I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓑ{a,I} V1. T1 ▶×[d, e] U2 → ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶×[d, e] V2 & - L. ⓑ{I} V2⦄ ⊢ T1 ▶×[d + 1, e] T2 & - U2 = ⓑ{a,I} V2. T2. -/2 width=3/ qed-. + ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ▶×[d + 1, e] T2 & + U2 = ⓑ{a,I}V2.T2. +/2 width=3 by cpy_inv_bind1_aux/ qed-. -fact cpy_inv_flat1_aux: ∀d,e,L,U1,U2. ⦃G, L⦄ ⊢ U1 ▶×[d, e] U2 → - ∀I,V1,T1. U1 = ⓕ{I} V1. T1 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶×[d, e] V2 & ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 & - U2 = ⓕ{I} V2. T2. -#d #e #L #U1 #U2 * -d -e -L -U1 -U2 -[ #L #k #d #e #I #V1 #T1 #H destruct -| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct -| #L #a #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct -| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ +fact cpy_inv_flat1_aux: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶×[d, e] U2 → + ∀I,V1,T1. U1 = ⓕ{I}V1.T1 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶×[d, e] V2 & + ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 & + U2 = ⓕ{I}V2.T2. +#G #L #U1 #U2 #d #e * -G -L -U1 -U2 -d -e +[ #I #G #L #d #e #J #W1 #U1 #H destruct +| #I #G #L #K #V #W #i #d #e #_ #_ #_ #_ #J #W1 #U1 #H destruct +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #W1 #U1 #H destruct +| #I #G #L #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ ] -qed. +qed-. + +lemma cpy_inv_flat1: ∀I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ▶×[d, e] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶×[d, e] V2 & + ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 & + U2 = ⓕ{I}V2.T2. +/2 width=3 by cpy_inv_flat1_aux/ qed-. -lemma cpy_inv_flat1: ∀d,e,L,I,V1,T1,U2. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ▶×[d, e] U2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶×[d, e] V2 & ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 & - U2 = ⓕ{I} V2. T2. -/2 width=3/ qed-. -fact cpy_inv_refl_O2_aux: ∀L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → e = 0 → T1 = T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +fact cpy_inv_refl_O2_aux: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → e = 0 → T1 = T2. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e [ // -| #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct +| #I #G #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct lapply (le_to_lt_to_lt … Hdi … Hide) -Hdi -Hide append_length >HL12 -HL12 - @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] // /2 width=3/ (**) (* explicit constructor *) + @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] (**) (* explicit constructor *) + /2 width=3 by trans_eq/ ] qed-. - -(* Basic_1: removed theorems 25: - subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt - subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans - subst0_lift_lt subst0_lift_ge subst0_lift_ge_S subst0_lift_ge_s - subst0_subst0 subst0_subst0_back subst0_weight_le subst0_weight_lt - subst0_confluence_neq subst0_confluence_eq subst0_tlt_head - subst0_confluence_lift subst0_tlt - subst1_head subst1_gen_head subst1_lift_S subst1_confluence_lift -*)