From 6907a0f66a3782ce4273a609203c9b574841c7d1 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Thu, 19 Dec 2013 14:43:18 +0000 Subject: [PATCH] the theory of cpy continues ... --- .../lambdadelta/basic_2/relocation/cpy.ma | 226 ++++++------ .../basic_2/relocation/cpy_lift.ma | 328 ++++++++---------- .../lambdadelta/basic_2/relocation/lsuby.ma | 28 +- 3 files changed, 281 insertions(+), 301 deletions(-) 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 -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma index 1d65d8d63..161b21b01 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma @@ -12,283 +12,243 @@ (* *) (**************************************************************************) -include "basic_2/substitution/ldrop_ldrop.ma". -include "basic_2/substitution/tps.ma". +include "basic_2/relocation/ldrop_ldrop.ma". +include "basic_2/relocation/cpy.ma". -(* PARTIAL SUBSTITUTION ON TERMS ********************************************) - -(* Advanced inversion lemmas ************************************************) - -fact tps_inv_S2_aux: ∀L,T1,T2,d,e1. L ⊢ T1 ▶ [d, e1] T2 → ∀e2. e1 = e2 + 1 → - ∀K,V. ⇩[0, d] L ≡ K. ⓛV → L ⊢ T1 ▶ [d + 1, e2] T2. -#L #T1 #T2 #d #e1 #H elim H -L -T1 -T2 -d -e1 -[ // -| #L #K0 #V0 #W #i #d #e1 #Hdi #Hide1 #HLK0 #HV0 #e2 #He12 #K #V #HLK destruct - elim (lt_or_ge i (d+1)) #HiSd - [ -Hide1 -HV0 - lapply (le_to_le_to_eq … Hdi ?) /2 width=1/ #H destruct - lapply (ldrop_mono … HLK0 … HLK) #H destruct - | -V -Hdi /2 width=4/ - ] -| /4 width=3/ -| /3 width=3/ -] -qed. - -lemma tps_inv_S2: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e + 1] T2 → - ∀K,V. ⇩[0, d] L ≡ K. ⓛV → L ⊢ T1 ▶ [d + 1, e] T2. -/2 width=3/ qed-. - -lemma tps_inv_refl_SO2: ∀L,T1,T2,d. L ⊢ T1 ▶ [d, 1] T2 → - ∀K,V. ⇩[0, d] L ≡ K. ⓛV → T1 = T2. -#L #T1 #T2 #d #HT12 #K #V #HLK -lapply (tps_inv_S2 … T1 T2 … 0 … HLK) -K // -HT12 #HT12 -lapply (tps_inv_refl_O2 … HT12) -HT12 // -qed-. +(* CONTEXT-SENSITIVE EXTENDED PARALLEL SUBSTITUTION FOR TERMS ***************) (* Relocation properties ****************************************************) -(* Basic_1: was: subst1_lift_lt *) -lemma tps_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → +lemma cpy_lift_le: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶×[dt, et] T2 → ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - dt + et ≤ d → - L ⊢ U1 ▶ [dt, et] U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ + dt + et ≤ d → ⦃G, L⦄ ⊢ U1 ▶×[dt, et] U2. +#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et +[ #I #G #K #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdetd +| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdetd lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct - elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=4/ -| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + elim (lift_trans_ge … HVW … HWU2) -W // (lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=5 by cpy_subst/ +| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=6/ | @IHT12 /2 width=6/ ] (**) (* /3 width=6/ is too slow, arith3 needed to avoid crash *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + /4 width=6 by cpy_bind, ldrop_skip, le_S_S/ +| #G #I #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + /3 width=6 by cpy_flat/ ] -qed. +qed-. -lemma tps_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → +lemma cpy_lift_be: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶×[dt, et] T2 → ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - dt ≤ d → d ≤ dt + et → - L ⊢ U1 ▶ [dt, et + e] U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ #_ + dt ≤ d → d ≤ dt + et → ⦃G, L⦄ ⊢ U1 ▶×[dt, et + e] U2. +#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et +[ #I #G #K #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ #_ >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdtd #_ +| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdtd #_ elim (lift_inv_lref1 … H) -H * #Hid #H destruct [ -Hdtd lapply (lt_to_le_to_lt … (dt+et+e) Hidet ?) // -Hidet #Hidete - elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -V #H destruct /2 width=4/ + elim (lift_trans_ge … HVW … HWU2) -W // (lift_mono … HVY … HVW) -V #H destruct /2 width=5 by cpy_subst/ | -Hdti lapply (transitive_le … Hdtd Hid) -Hdtd #Hdti - lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 - lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ + lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=5 by cpy_subst, lt_minus_to_plus_r, transitive_le/ ] -| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdtd #Hddet +| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdtd #Hddet elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=6/ | @IHT12 [3,4: // | skip |5,6: /2 width=1/ | /2 width=1/ ] - ] (**) (* /3 width=6/ is too slow, simplification like tps_lift_le is too slow *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + /4 width=6 by cpy_bind, ldrop_skip, le_S_S/ (**) (* auto a bit slow *) +| #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + /3 width=6 by cpy_flat/ ] -qed. +qed-. -(* Basic_1: was: subst1_lift_ge *) -lemma tps_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → +lemma cpy_lift_ge: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶×[dt, et] T2 → ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - d ≤ dt → - L ⊢ U1 ▶ [dt + e, et] U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ + d ≤ dt → ⦃G, L⦄ ⊢ U1 ▶×[dt + e, et] U2. +#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et +[ #I #G #K #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hddt +| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hddt lapply (transitive_le … Hddt … Hdti) -Hddt #Hid lapply (lift_inv_lref1_ge … H … Hid) -H #H destruct - lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 - lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ -| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt + lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=5 by cpy_subst, lt_minus_to_plus_r, monotonic_le_plus_l/ +| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt + /4 width=5 by cpy_bind, ldrop_skip, le_minus_to_plus/ +| #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=5/ + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + /3 width=5 by cpy_flat/ ] -qed. +qed-. -(* Basic_1: was: subst1_gen_lift_lt *) -lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → +lemma cpy_inv_lift1_le: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶×[dt, et] U2 → ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → dt + et ≤ d → - ∃∃T2. K ⊢ T1 ▶ [dt, et] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et -[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶×[dt, et] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et +[ * #G #L #i #dt #et #K #d #e #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by cpy_atom, lift_sort, ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by cpy_atom, lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by cpy_atom, lift_gref, ex2_intro/ ] -| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdetd +| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdetd lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct elim (ldrop_conf_lt … HLK … HLKV ?) -L // #L #U #HKL #_ #HUV - elim (lift_trans_le … HUV … HVW ?) -V // >minus_plus minus_plus minus_plus minus_plus plus_minus // commutative_plus >plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) + elim (lift_split … HVW d (i-e+1)) -HVW [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hid -Hdie + #V1 #HV1 >plus_minus // commutative_plus >plus_minus /2 width=1 by monotonic_lt_minus_l/ ] -| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet +| #a #I #G #L #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct - elim (IHV12 … HLK … HWV1 ? ?) -V1 // #W2 #HW12 #HWV2 - elim (IHU12 … HTU1 ? ?) -U1 [5: @ldrop_skip // |2: skip |3: >plus_plus_comm_23 >(plus_plus_comm_23 dt) /2 width=1/ |4: /2 width=1/ ] (**) (* 29s without the rewrites *) - /3 width=5/ -| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet + elim (IHV12 … HLK … HWV1) -V1 // #W2 #HW12 #HWV2 + elim (IHU12 … HTU1) -U1 + [5: @ldrop_skip // |2: skip + |3: >plus_plus_comm_23 >(plus_plus_comm_23 dt) /2 width=1 by le_S_S/ + |4: /2 width=1 by le_S_S/ + ] (**) (* 29s without the rewrites *) + /3 width=5 by _//3 width=5 by cpy_bind, lift_bind, ex2_intro/ +| #I #G #L #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct - elim (IHV12 … HLK … HWV1 ? ?) -V1 // - elim (IHU12 … HLK … HTU1 ? ?) -U1 -HLK // /3 width=5/ + elim (IHV12 … HLK … HWV1) -V1 // + elim (IHU12 … HLK … HTU1) -U1 -HLK // + /3 width=5 by cpy_flat, lift_flat, ex2_intro/ ] -qed. +qed-. -(* Basic_1: was: subst1_gen_lift_ge *) -lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → +lemma cpy_inv_lift1_ge: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶×[dt, et] U2 → ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → d + e ≤ dt → - ∃∃T2. K ⊢ T1 ▶ [dt - e, et] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et -[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶×[dt - e, et] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et +[ * #G #L #i #dt #et #K #d #e #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by cpy_atom, lift_sort, ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by cpy_atom, lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by cpy_atom, lift_gref, ex2_intro/ ] -| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdedt +| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdedt lapply (transitive_le … Hdedt … Hdti) #Hdei elim (le_inv_plus_l … Hdedt) -Hdedt #_ #Hedt elim (le_inv_plus_l … Hdei) #Hdie #Hei lapply (lift_inv_lref2_ge … H … Hdei) -H #H destruct lapply (ldrop_conf_ge … HLK … HLKV ?) -L // #HKV - elim (lift_split … HVW d (i - e + 1) ? ? ?) -HVW [4: // |3: /2 width=1/ |2: /3 width=1/ ] -Hdei -Hdie - #V0 #HV10 >plus_minus // plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) -| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd + elim (lift_split … HVW d (i-e+1)) -HVW [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hdei -Hdie + #V0 #HV10 >plus_minus // plus_minus /2 width=1 by monotonic_lt_minus_l/ +| #a #I #G #L #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct elim (le_inv_plus_l … Hdetd) #_ #Hedt - elim (IHV12 … HLK … HWV1 ?) -V1 // #W2 #HW12 #HWV2 - elim (IHU12 … HTU1 ?) -U1 [4: @ldrop_skip // |2: skip |3: /2 width=1/ ] - IHV12 // >IHT12 // -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct >IHV12 // >IHT12 // ] -qed. -(* - Theorem subst0_gen_lift_rev_ge: (t1,v,u2,i,h,d:?) - (subst0 i v t1 (lift h d u2)) -> - (le (plus d h) i) -> - (EX u1 | (subst0 (minus i h) v u1 u2) & - t1 = (lift h d u1) - ). - +qed-. - Theorem subst0_gen_lift_rev_lelt: (t1,v,u2,i,h,d:?) - (subst0 i v t1 (lift h d u2)) -> - (le d i) -> (lt i (plus d h)) -> - (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)). -*) -lemma tps_inv_lift1_ge_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → +lemma cpy_inv_lift1_ge_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶×[dt, et] U2 → ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 ▶ [d, dt + et - (d + e)] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet -elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 -lapply (tps_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 -lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -U -L // commutative_plus /2 width=1 by le_minus_to_plus_r/ ] -Hddt -Hdtde #HU1 +lapply (cpy_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct +elim (cpy_inv_lift1_ge … HU2 … HLK … HTU1) -U -L // commutative_plus /2 width=1/ ] -Hdtd #T #HT1 #HTU -lapply (tps_weak … HU2 d e ? ?) -HU2 // [ >commutative_plus commutative_plus /2 width=1 by le_minus_to_plus_r/ ] -Hdtd #T #HT1 #HTU +lapply (cpy_weak … HU2 d e ? ?) -HU2 // [ >commutative_plus minus_minus_comm >arith_b1 /3 width=4 by ldrop_ldrop_lt, ex2_2_intro/ + ] +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #J2 #K2 #W #i #H #Hdi >plus_plus_comm_23 #Hide + elim (le_inv_plus_l … Hdi) #_ #Hi + lapply (ldrop_inv_ldrop1_lt … H ?) -H // #HLK1 + elim (IHL12 … HLK1) -IHL12 -HLK1 + [2,3: /2 width=1 by lt_plus_to_minus, monotonic_pred/ ] -Hdi -Hide + >minus_minus_comm >arith_b1 /3 width=4 by ldrop_ldrop_lt, ex2_2_intro/ +] +qed-. -- 2.39.2