From: Ferruccio Guidi Date: Mon, 16 Dec 2013 20:55:49 +0000 (+0000) Subject: - lsubr moved down one component X-Git-Tag: make_still_working~1015 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=3e9d72c26091f0e157a024ea9bd6f95a95729860 - lsubr moved down one component - extended substitution started (for use in lleq) --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/ygt/zgt.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/ygt/zgt.etc deleted file mode 100644 index ce6c973fd..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/etc/ygt/zgt.etc +++ /dev/null @@ -1,43 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/computation/ltprs.ma". -include "basic_2/dynamic/lsubsv.ma". - -(* "BIG TREE" PROPER PARALLEL COMPUTATION FOR CLOSURES **********************) - -axiom yprs: ∀h. sd h → bi_relation lenv term. - -interpretation "'big tree' parallel computation (closure)" - 'BTPRedStar h g L1 T1 L2 T2 = (yprs h g L1 T1 L2 T2). - -axiom cprs_yprs: ∀h,g,L,T1,T2. L ⊢ T1 ➡* T2 → h ⊢ ⦃L, T1⦄ ≥[g] ⦃L, T2⦄. - -axiom sstas_yprs: ∀h,g,L,T1,T2. ⦃h, L⦄ ⊢ T1 •*[g] T2 → - h ⊢ ⦃L, T1⦄ ≥[g] ⦃L, T2⦄. - -axiom lsubsv_yprs: ∀h,g,L1,L2,T. h ⊢ L2 ¡⊑[g] L1 → h ⊢ ⦃L1, T⦄ ≥[g] ⦃L2, T⦄. - -axiom ltpr_cprs_yprs: ∀h,g,L1,L2,T1,T2. L1 ➡ L2 → L2 ⊢ T1 ➡* T2 → - h ⊢ ⦃L1, T1⦄ ≥[g] ⦃L2, T2⦄. - -axiom ygt: ∀h. sd h → bi_relation lenv term. - -interpretation "'big tree' proper parallel computation (closure)" - 'BTPRedStarProper h g L1 T1 L2 T2 = (ygt h g L1 T1 L2 T2). - -axiom fw_ygt: ∀h,g,L1,L2,T1,T2. ♯{L2, T2} < ♯{L1, T1} → h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄. - -axiom ygt_yprs_trans: ∀h,g,L1,L,L2,T1,T,T2. h ⊢ ⦃L1, T1⦄ >[g] ⦃L, T⦄ → - h ⊢ ⦃L, T⦄ ≥[g] ⦃L2, T2⦄ → h ⊢ ⦃L1, T1⦄ >[g] ⦃L2, T2⦄. diff --git a/matita/matita/contribs/lambdadelta/basic_2/names.txt b/matita/matita/contribs/lambdadelta/basic_2/names.txt index 6f00405d2..1bfe7ed6d 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/names.txt +++ b/matita/matita/contribs/lambdadelta/basic_2/names.txt @@ -79,6 +79,7 @@ r: reduction s: substitution u: supclosure x: extended reduction +y: extended substitution - forth letter (if present) diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extlrsubeq_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extlrsubeq_4.ma new file mode 100644 index 000000000..8a9714dd8 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extlrsubeq_4.ma @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( L1 break ⊑ × [ term 46 d , break term 46 e ] break term 46 L2 )" + non associative with precedence 45 + for @{ 'ExtLRSubEq $L1 $d $e $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extpsubst_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extpsubst_6.ma new file mode 100644 index 000000000..9dde8aa0f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extpsubst_6.ma @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ × [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'ExtPSubst $G $L $T1 $d $e $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma index c7ce48b1f..3c0514a3c 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpr.ma @@ -16,7 +16,7 @@ include "basic_2/notation/relations/pred_4.ma". include "basic_2/grammar/genv.ma". include "basic_2/grammar/cl_shift.ma". include "basic_2/relocation/ldrop_append.ma". -include "basic_2/substitution/lsubr.ma". +include "basic_2/relocation/lsubr.ma". (* CONTEXT-SENSITIVE PARALLEL REDUCTION FOR TERMS ***************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma new file mode 100644 index 000000000..025cd6f65 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma @@ -0,0 +1,288 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/extpsubst_6.ma". +include "basic_2/grammar/genv.ma". +include "basic_2/relocation/ldrop_append.ma". +include "basic_2/relocation/lsuby.ma". + +(* CONTEXT-SENSITIVE EXTENDED PARALLEL SUBSTITUTION FOR TERMS ***************) + +(* activate genv *) +inductive cpy: nat → nat → relation4 genv lenv term term ≝ +| cpy_atom : ∀I,G,L,d,e. cpy d e G L (⓪{I}) (⓪{I}) +| cpy_subst: ∀I,G,L,K,V,W,i,d,e. d ≤ i → i < d + e → + ⇩[0, i] L ≡ K.ⓑ{I}V → ⇧[0, i + 1] V ≡ W → cpy d e G L (#i) W +| cpy_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e. + cpy d e G L V1 V2 → cpy (d + 1) e G (L.ⓑ{I} V2) T1 T2 → + cpy d e G L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) +| cpy_flat : ∀I,G,L,V1,V2,T1,T2,d,e. + cpy d e G L V1 V2 → cpy d e G L T1 T2 → + cpy d e G L (ⓕ{I}V1. T1) (ⓕ{I}V2. T2) +. + +interpretation "context-sensitive extended parallel substritution (term)" + 'ExtPSubst G L T1 d e T2 = (cpy d e G L T1 T2). + +(* Basic properties *********************************************************) + +lemma lsuby_cpy_trans: ∀G,d,e.lsub_trans … (cpy d e G) lsuby. +#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/ +| /3 width=1 by cpy_flat/ +] +qed-. + +lemma cpy_refl: ∀G,T,L,d,e. ⦃G, L⦄ ⊢ T ▶×[d, e] T. +#G #T elim T -T // * /2 width=1 by cpy_bind, cpy_flat/ +qed. + +lemma cpy_full: ∀I,G,K,V,T1,L,d. ⇩[0, d] L ≡ K.ⓑ{I}V → + ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ▶×[d, 1] T2 & ⇧[d, 1] T ≡ T2. +#I #G #K #V #T1 elim T1 -T1 +[ * #i #L #d #HLK + /2 width=4 by lift_sort, lift_gref, ex2_2_intro/ + elim (lt_or_eq_or_gt i d) #Hid + /3 width=4 by lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ + destruct + elim (lift_total V 0 (i+1)) #W #HVW + elim (lift_split … HVW i i) + /3 width=5 by cpy_subst, le_n, ex2_2_intro/ +| * [ #a ] #J #W1 #U1 #IHW1 #IHU1 #L #d #HLK + elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 + [ elim (IHU1 (L.ⓑ{J}W2) (d+1)) -IHU1 + /3 width=9 by cpy_bind, ldrop_ldrop, lift_bind, ex2_2_intro/ + | elim (IHU1 … HLK) -IHU1 -HLK + /3 width=8 by cpy_flat, lift_flat, ex2_2_intro/ + ] +] +qed-. + +lemma cpy_weak: ∀G,L,T1,T2,d1,e1. ⦃G, L⦄ ⊢ T1 ▶×[d1, e1] T2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → + ⦃G, L⦄ ⊢ T1 ▶×[d2, e2] T2. +#G #L #T1 #T2 #d1 #e1 #H elim H -G -L -T1 -T2 -d1 -e1 +[ // +| /3 width=5 by cpy_subst, transitive_le/ +| /4 width=3 by cpy_bind, le_to_lt_to_lt, le_S_S/ +| /3 width=1 by cpy_flat/ +] +qed-. + +lemma cpy_weak_top: ∀G,L,T1,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶×[d, |L| - d] T2. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e +[ // +| #I #G #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW + lapply (ldrop_fwd_length_lt2 … HLK) + /3 width=5 by cpy_subst, lt_to_le_to_lt/ +| normalize /2 width=1 by cpy_bind/ +| /2 width=1 by cpy_flat/ +] +qed-. + +lemma cpy_weak_full: ∀G,L,T1,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶×[0, |L|] T2. +#G #L #T1 #T2 #d #e #HT12 +lapply (cpy_weak … HT12 0 (d + e) ? ?) -HT12 +/2 width=2 by cpy_weak_top/ +qed-. + +lemma cpy_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶×[d, e] T2 → ∀i. d ≤ i → i ≤ d + e → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶×[d, i - d] T & ⦃G, L⦄ ⊢ T ▶×[i, d + e - i] T2. +#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 | -Hdi -Hdj ] + [ >(plus_minus_m_m j d) in ⊢ (%→?); + /3 width=5 by cpy_subst, ex2_intro/ + | #Hij lapply (plus_minus_m_m … Hjde) -Hjde + /3 width=9 by cpy_atom, cpy_subst, ex2_intro/ + ] +| #a #I #G #L #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 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/ +] +qed. + +lemma cpy_split_down: ∀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 + elim (lt_or_ge i j) + [ -Hide -Hjde >(plus_minus_m_m j d) in⦄ ⊢ (% → ?); // -Hdj /3 width=8/ + | -Hdi -Hdj + >(plus_minus_m_m (d+e) j) in Hide; // -Hjde /3 width=4/ + ] +| #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/ + -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/ +] +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. + +(* 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 +] +qed. + +lemma cpy_inv_atom1: ∀L,T2,I,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-. + + +(* 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 +elim (cpy_inv_atom1 … H) -H // +* #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 → + 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/ +qed-. + +lemma cpy_inv_gref1: ∀L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶×[d, e] T2 → T2 = §p. +#L #T2 #p #d #e #H +elim (cpy_inv_atom1 … H) -H // +* #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 → + ∃∃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 +] +qed. + +lemma cpy_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. ⦃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-. + +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/ +] +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 +[ // +| #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct + lapply (le_to_lt_to_lt … Hdi … Hide) -Hdi -Hide shift_append_assoc normalize #H + elim (cpy_inv_bind1 … H) -H + #V0 #T0 #_ #HT10 #H destruct + elim (IH … HT10) -IH -HT10 #L2 #T2 #HL12 #H destruct + >append_length >HL12 -HL12 + @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] // /2 width=3/ (**) (* explicit constructor *) +] +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_cpy.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_cpy.ma new file mode 100644 index 000000000..7e146c21e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_cpy.ma @@ -0,0 +1,132 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/tps_lift.ma". + +(* PARALLEL SUBSTITUTION ON TERMS *******************************************) + +(* Main properties **********************************************************) + +(* Basic_1: was: subst1_confluence_eq *) +theorem tps_conf_eq: ∀L,T0,T1,d1,e1. L ⊢ T0 ▶ [d1, e1] T1 → + ∀T2,d2,e2. L ⊢ T0 ▶ [d2, e2] T2 → + ∃∃T. L ⊢ T1 ▶ [d2, e2] T & L ⊢ T2 ▶ [d1, e1] T. +#L #T0 #T1 #d1 #e1 #H elim H -L -T0 -T1 -d1 -e1 +[ /2 width=3/ +| #L #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #T2 #d2 #e2 #H + elim (tps_inv_lref1 … H) -H + [ #HX destruct /3 width=6/ + | -Hd1 -Hde1 * #K2 #V2 #_ #_ #HLK2 #HVT2 + lapply (ldrop_mono … HLK1 … HLK2) -HLK1 -HLK2 #H destruct + >(lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3/ + ] +| #L #a #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V1) ?) -HT02 /2 width=1/ #HT02 + elim (IHV01 … HV02) -V0 #V #HV1 #HV2 + elim (IHT01 … HT02) -T0 #T #HT1 #HT2 + lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubr_trans … HT2 (L. ⓑ{I} V) ?) -HT2 /3 width=5/ +| #L #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02) -V0 + elim (IHT01 … HT02) -T0 /3 width=5/ +] +qed. + +(* Basic_1: was: subst1_confluence_neq *) +theorem tps_conf_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 ▶ [d1, e1] T1 → + ∀L2,T2,d2,e2. L2 ⊢ T0 ▶ [d2, e2] T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. L2 ⊢ T1 ▶ [d2, e2] T & L1 ⊢ T2 ▶ [d1, e1] T. +#L1 #T0 #T1 #d1 #e1 #H elim H -L1 -T0 -T1 -d1 -e1 +[ /2 width=3/ +| #L1 #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #L2 #T2 #d2 #e2 #H1 #H2 + elim (tps_inv_lref1 … H1) -H1 + [ #H destruct /3 width=6/ + | -HLK1 -HVT1 * #K2 #V2 #Hd2 #Hde2 #_ #_ elim H2 -H2 #Hded + [ -Hd1 -Hde2 + lapply (transitive_le … Hded Hd2) -Hded -Hd2 #H + lapply (lt_to_le_to_lt … Hde1 H) -Hde1 -H #H + elim (lt_refl_false … H) + | -Hd2 -Hde1 + lapply (transitive_le … Hded Hd1) -Hded -Hd1 #H + lapply (lt_to_le_to_lt … Hde2 H) -Hde2 -H #H + elim (lt_refl_false … H) + ] + ] +| #L1 #a #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -V0 #V #HV1 #HV2 + elim (IHT01 … HT02 ?) -T0 + [ -H #T #HT1 #HT2 + lapply (tps_lsubr_trans … HT1 (L2. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubr_trans … HT2 (L1. ⓑ{I} V) ?) -HT2 /2 width=1/ /3 width=5/ + | -HV1 -HV2 >plus_plus_comm_23 >plus_plus_comm_23 in ⊢ (? ? %); elim H -H #H + [ @or_introl | @or_intror ] /2 by monotonic_le_plus_l/ (**) (* /3 / is too slow *) + ] +| #L1 #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -V0 + elim (IHT01 … HT02 H) -T0 -H /3 width=5/ +] +qed. + +(* Note: the constant 1 comes from tps_subst *) +(* Basic_1: was: subst1_trans *) +theorem tps_trans_ge: ∀L,T1,T0,d,e. L ⊢ T1 ▶ [d, e] T0 → + ∀T2. L ⊢ T0 ▶ [d, 1] T2 → 1 ≤ e → + L ⊢ T1 ▶ [d, e] T2. +#L #T1 #T0 #d #e #H elim H -L -T1 -T0 -d -e +[ #L #I #d #e #T2 #H #He + elim (tps_inv_atom1 … H) -H + [ #H destruct // + | * #K #V #i #Hd2i #Hide2 #HLK #HVT2 #H destruct + lapply (lt_to_le_to_lt … (d + e) Hide2 ?) /2 width=4/ + ] +| #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He + lapply (tps_weak … HVT2 0 (i +1) ? ?) -HVT2 /2 width=1/ #HVT2 + <(tps_inv_lift1_eq … HVT2 … HVW) -HVT2 /2 width=4/ +| #L #a #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He + elim (tps_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct + lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 + lapply (IHT10 … HT02 He) -T0 #HT12 + lapply (tps_lsubr_trans … HT12 (L. ⓑ{I} V2) ?) -HT12 /2 width=1/ /3 width=1/ +| #L #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He + elim (tps_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1/ +] +qed. + +theorem tps_trans_down: ∀L,T1,T0,d1,e1. L ⊢ T1 ▶ [d1, e1] T0 → + ∀T2,d2,e2. L ⊢ T0 ▶ [d2, e2] T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 ▶ [d2, e2] T & L ⊢ T ▶ [d1, e1] T2. +#L #T1 #T0 #d1 #e1 #H elim H -L -T1 -T0 -d1 -e1 +[ /2 width=3/ +| #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1 + lapply (transitive_le … Hde2d1 Hdi1) -Hde2d1 #Hde2i1 + lapply (tps_weak … HWT2 0 (i1 + 1) ? ?) -HWT2 normalize /2 width=1/ -Hde2i1 #HWT2 + <(tps_inv_lift1_eq … HWT2 … HVW) -HWT2 /3 width=8/ +| #L #a #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 + elim (IHV10 … HV02 ?) -IHV10 -HV02 // #V + elim (IHT10 … HT02 ?) -T0 /2 width=1/ #T #HT1 #HT2 + lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubr_trans … HT2 (L. ⓑ{I} V2) ?) -HT2 /2 width=1/ /3 width=6/ +| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV10 … HV02 ?) -V0 // + elim (IHT10 … HT02 ?) -T0 // /3 width=6/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma new file mode 100644 index 000000000..1d65d8d63 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma @@ -0,0 +1,294 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/ldrop_ldrop.ma". +include "basic_2/substitution/tps.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-. + +(* Relocation properties ****************************************************) + +(* Basic_1: was: subst1_lift_lt *) +lemma tps_lift_le: ∀K,T1,T2,dt,et. 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 #_ + >(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 + 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_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 + 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/ +] +qed. + +lemma tps_lift_be: ∀K,T1,T2,dt,et. 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 #_ #_ + >(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 #_ + 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/ + | -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/ + ] +| #K #a #I #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 + 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/ +] +qed. + +(* Basic_1: was: subst1_lift_ge *) +lemma tps_lift_ge: ∀K,T1,T2,dt,et. 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 #_ + >(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 + 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 + 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 + 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/ +] +qed. + +(* Basic_1: was: subst1_gen_lift_lt *) +lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. 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/ + ] +| #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 plus_minus // commutative_plus >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 #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 (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/ +] +qed. + +(* Basic_1: was: subst1_gen_lift_ge *) +lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. 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/ + ] +| #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_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 + 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) + ). + + + 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 → + ∀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/ ] -Hdtd #T #HT1 #HTU +lapply (tps_weak … HU2 d e ? ?) -HU2 // [ >commutative_plus (plus_minus_m_m e 1) // /2 width=1/ +qed. + +lemma lsubr_abst_lt: ∀L1,L2,I,V1,V2,e. L1 ⊑ [0, e - 1] L2 → 0 < e → + L1. ⓑ{I}V1 ⊑ [0, e] L2. ⓛV2. +#L1 #L2 #I #V1 #V2 #e #HL12 #He >(plus_minus_m_m e 1) // /2 width=1/ +qed. + +lemma lsubr_skip_lt: ∀L1,L2,d,e. L1 ⊑ [d - 1, e] L2 → 0 < d → + ∀I1,I2,V1,V2. L1. ⓑ{I1} V1 ⊑ [d, e] L2. ⓑ{I2} V2. +#L1 #L2 #d #e #HL12 #Hd >(plus_minus_m_m d 1) // /2 width=1/ +qed. + +lemma lsubr_bind_lt: ∀I,L1,L2,V,e. L1 ⊑ [0, e - 1] L2 → 0 < e → + L1. ⓓV ⊑ [0, e] L2. ⓑ{I}V. +* /2 width=1/ qed. + +lemma lsubr_refl: ∀d,e,L. L ⊑ [d, e] L. +#d elim d -d +[ #e elim e -e // #e #IHe #L elim L -L // /2 width=1/ +| #d #IHd #e #L elim L -L // /2 width=1/ +] +qed. + +lemma TC_lsubr_trans: ∀S,R. lsubr_trans S R → lsubr_trans S (λL. (TC … (R L))). +#S #R #HR #L1 #s1 #s2 #H elim H -s2 +[ /3 width=5/ +| #s #s2 #_ #Hs2 #IHs1 #L2 #d #e #HL12 + lapply (HR … Hs2 … HL12) -HR -Hs2 -HL12 /3 width=3/ +] +qed. + +(* Basic inversion lemmas ***************************************************) + +fact lsubr_inv_atom1_aux: ∀L1,L2,d,e. L1 ⊑ [d, e] L2 → L1 = ⋆ → + L2 = ⋆ ∨ (d = 0 ∧ e = 0). +#L1 #L2 #d #e * -L1 -L2 -d -e +[ /2 width=1/ +| /3 width=1/ +| #L1 #L2 #W #e #_ #H destruct +| #L1 #L2 #I #W1 #W2 #e #_ #H destruct +| #L1 #L2 #I1 #I2 #W1 #W2 #d #e #_ #H destruct +] +qed. + +lemma lsubr_inv_atom1: ∀L2,d,e. ⋆ ⊑ [d, e] L2 → + L2 = ⋆ ∨ (d = 0 ∧ e = 0). +/2 width=3/ qed-. + +fact lsubr_inv_skip1_aux: ∀L1,L2,d,e. L1 ⊑ [d, e] L2 → + ∀I1,K1,V1. L1 = K1.ⓑ{I1}V1 → 0 < d → + ∃∃I2,K2,V2. K1 ⊑ [d - 1, e] K2 & L2 = K2.ⓑ{I2}V2. +#L1 #L2 #d #e * -L1 -L2 -d -e +[ #d #e #I1 #K1 #V1 #H destruct +| #L1 #L2 #I1 #K1 #V1 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #W #e #_ #I1 #K1 #V1 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #W1 #W2 #e #_ #I1 #K1 #V1 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #J1 #J2 #W1 #W2 #d #e #HL12 #I1 #K1 #V1 #H #_ destruct /2 width=5/ +] +qed. + +lemma lsubr_inv_skip1: ∀I1,K1,L2,V1,d,e. K1.ⓑ{I1}V1 ⊑ [d, e] L2 → 0 < d → + ∃∃I2,K2,V2. K1 ⊑ [d - 1, e] K2 & L2 = K2.ⓑ{I2}V2. +/2 width=5/ qed-. + +fact lsubr_inv_atom2_aux: ∀L1,L2,d,e. L1 ⊑ [d, e] L2 → L2 = ⋆ → + L1 = ⋆ ∨ (d = 0 ∧ e = 0). +#L1 #L2 #d #e * -L1 -L2 -d -e +[ /2 width=1/ +| /3 width=1/ +| #L1 #L2 #W #e #_ #H destruct +| #L1 #L2 #I #W1 #W2 #e #_ #H destruct +| #L1 #L2 #I1 #I2 #W1 #W2 #d #e #_ #H destruct +] +qed. + +lemma lsubr_inv_atom2: ∀L1,d,e. L1 ⊑ [d, e] ⋆ → + L1 = ⋆ ∨ (d = 0 ∧ e = 0). +/2 width=3/ qed-. + +fact lsubr_inv_abbr2_aux: ∀L1,L2,d,e. L1 ⊑ [d, e] L2 → + ∀K2,V. L2 = K2.ⓓV → d = 0 → 0 < e → + ∃∃K1. K1 ⊑ [0, e - 1] K2 & L1 = K1.ⓓV. +#L1 #L2 #d #e * -L1 -L2 -d -e +[ #d #e #K1 #V #H destruct +| #L1 #L2 #K1 #V #_ #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #W #e #HL12 #K1 #V #H #_ #_ destruct /2 width=3/ +| #L1 #L2 #I #W1 #W2 #e #_ #K1 #V #H destruct +| #L1 #L2 #I1 #I2 #W1 #W2 #d #e #_ #K1 #V #_ >commutative_plus normalize #H destruct +] +qed. + +lemma lsubr_inv_abbr2: ∀L1,K2,V,e. L1 ⊑ [0, e] K2.ⓓV → 0 < e → + ∃∃K1. K1 ⊑ [0, e - 1] K2 & L1 = K1.ⓓV. +/2 width=5/ qed-. + +fact lsubr_inv_skip2_aux: ∀L1,L2,d,e. L1 ⊑ [d, e] L2 → + ∀I2,K2,V2. L2 = K2.ⓑ{I2}V2 → 0 < d → + ∃∃I1,K1,V1. K1 ⊑ [d - 1, e] K2 & L1 = K1.ⓑ{I1}V1. +#L1 #L2 #d #e * -L1 -L2 -d -e +[ #d #e #I1 #K1 #V1 #H destruct +| #L1 #L2 #I1 #K1 #V1 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #W #e #_ #I1 #K1 #V1 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #I #W1 #W2 #e #_ #I1 #K1 #V1 #_ #H + elim (lt_zero_false … H) +| #L1 #L2 #J1 #J2 #W1 #W2 #d #e #HL12 #I1 #K1 #V1 #H #_ destruct /2 width=5/ +] +qed. + +lemma lsubr_inv_skip2: ∀I2,L1,K2,V2,d,e. L1 ⊑ [d, e] K2.ⓑ{I2}V2 → 0 < d → + ∃∃I1,K1,V1. K1 ⊑ [d - 1, e] K2 & L1 = K1.ⓑ{I1}V1. +/2 width=5/ qed-. + +(* Basic forward lemmas *****************************************************) + +fact lsubr_fwd_length_full1_aux: ∀L1,L2,d,e. L1 ⊑ [d, e] L2 → + d = 0 → e = |L1| → |L1| ≤ |L2|. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize +[ // +| /2 width=1/ +| /3 width=1/ +| /3 width=1/ +| #L1 #L2 #_ #_ #_ #_ #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma lsubr_fwd_length_full1: ∀L1,L2. L1 ⊑ [0, |L1|] L2 → |L1| ≤ |L2|. +/2 width=5/ qed-. + +fact lsubr_fwd_length_full2_aux: ∀L1,L2,d,e. L1 ⊑ [d, e] L2 → + d = 0 → e = |L2| → |L2| ≤ |L1|. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize +[ // +| /2 width=1/ +| /3 width=1/ +| /3 width=1/ +| #L1 #L2 #_ #_ #_ #_ #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma lsubr_fwd_length_full2: ∀L1,L2. L1 ⊑ [0, |L2|] L2 → |L2| ≤ |L1|. +/2 width=5/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsuba.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsuba.ma index eab0ff035..49441a887 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsuba.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsuba.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/notation/relations/lrsubeqa_3.ma". -include "basic_2/substitution/lsubr.ma". +include "basic_2/relocation/lsubr.ma". include "basic_2/static/aaa.ma". (* LOCAL ENVIRONMENT REFINEMENT FOR ATOMIC ARITY ASSIGNMENT *****************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma b/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma index 5e2332c7c..0644771eb 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/static/lsubd.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/notation/relations/lrsubeqd_5.ma". -include "basic_2/substitution/lsubr.ma". +include "basic_2/relocation/lsubr.ma". include "basic_2/static/da.ma". (* LOCAL ENVIRONMENT REFINEMENT FOR DEGREE ASSIGNMENT ***********************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr.ma deleted file mode 100644 index faf254e36..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr.ma +++ /dev/null @@ -1,105 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/lrsubeq_2.ma". -include "basic_2/relocation/ldrop.ma". - -(* RESTRICTED LOCAL ENVIRONMENT REFINEMENT **********************************) - -inductive lsubr: relation lenv ≝ -| lsubr_sort: ∀L. lsubr L (⋆) -| lsubr_bind: ∀I,L1,L2,V. lsubr L1 L2 → lsubr (L1.ⓑ{I}V) (L2.ⓑ{I}V) -| lsubr_abst: ∀L1,L2,V,W. lsubr L1 L2 → lsubr (L1.ⓓⓝW.V) (L2.ⓛW) -. - -interpretation - "local environment refinement (restricted)" - 'LRSubEq L1 L2 = (lsubr L1 L2). - -(* Basic properties *********************************************************) - -lemma lsubr_refl: ∀L. L ⊑ L. -#L elim L -L // /2 width=1/ -qed. - -(* Basic inversion lemmas ***************************************************) - -fact lsubr_inv_atom1_aux: ∀L1,L2. L1 ⊑ L2 → L1 = ⋆ → L2 = ⋆. -#L1 #L2 * -L1 -L2 // -[ #I #L1 #L2 #V #_ #H destruct -| #L1 #L2 #V #W #_ #H destruct -] -qed-. - -lemma lsubr_inv_atom1: ∀L2. ⋆ ⊑ L2 → L2 = ⋆. -/2 width=3 by lsubr_inv_atom1_aux/ qed-. - -fact lsubr_inv_abst1_aux: ∀L1,L2. L1 ⊑ L2 → ∀K1,W. L1 = K1.ⓛW → - L2 = ⋆ ∨ ∃∃K2. K1 ⊑ K2 & L2 = K2.ⓛW. -#L1 #L2 * -L1 -L2 -[ #L #K1 #W #H destruct /2 width=1/ -| #I #L1 #L2 #V #HL12 #K1 #W #H destruct /3 width=3/ -| #L1 #L2 #V1 #V2 #_ #K1 #W #H destruct -] -qed-. - -lemma lsubr_inv_abst1: ∀K1,L2,W. K1.ⓛW ⊑ L2 → - L2 = ⋆ ∨ ∃∃K2. K1 ⊑ K2 & L2 = K2.ⓛW. -/2 width=3 by lsubr_inv_abst1_aux/ qed-. - -fact lsubr_inv_abbr2_aux: ∀L1,L2. L1 ⊑ L2 → ∀K2,W. L2 = K2.ⓓW → - ∃∃K1. K1 ⊑ K2 & L1 = K1.ⓓW. -#L1 #L2 * -L1 -L2 -[ #L #K2 #W #H destruct -| #I #L1 #L2 #V #HL12 #K2 #W #H destruct /2 width=3/ -| #L1 #L2 #V1 #V2 #_ #K2 #W #H destruct -] -qed-. - -lemma lsubr_inv_abbr2: ∀L1,K2,W. L1 ⊑ K2.ⓓW → - ∃∃K1. K1 ⊑ K2 & L1 = K1.ⓓW. -/2 width=3 by lsubr_inv_abbr2_aux/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lsubr_fwd_length: ∀L1,L2. L1 ⊑ L2 → |L2| ≤ |L1|. -#L1 #L2 #H elim H -L1 -L2 // /2 width=1/ -qed-. - -lemma lsubr_fwd_ldrop2_bind: ∀L1,L2. L1 ⊑ L2 → - ∀I,K2,W,i. ⇩[0, i] L2 ≡ K2.ⓑ{I}W → - (∃∃K1. K1 ⊑ K2 & ⇩[0, i] L1 ≡ K1.ⓑ{I}W) ∨ - ∃∃K1,V. K1 ⊑ K2 & ⇩[0, i] L1 ≡ K1.ⓓⓝW.V & I = Abst. -#L1 #L2 #H elim H -L1 -L2 -[ #L #I #K2 #W #i #H - elim (ldrop_inv_atom1 … H) -H #H destruct -| #J #L1 #L2 #V #HL12 #IHL12 #I #K2 #W #i #H - elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ] - [ /3 width=3/ - | elim (IHL12 … HLK2) -IHL12 -HLK2 * /4 width=3/ /4 width=4/ - ] -| #L1 #L2 #V1 #V2 #HL12 #IHL12 #I #K2 #W #i #H - elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK2 destruct [ -IHL12 | -HL12 ] - [ /3 width=4/ - | elim (IHL12 … HLK2) -IHL12 -HLK2 * /4 width=3/ /4 width=4/ - ] -] -qed-. - -lemma lsubr_fwd_ldrop2_abbr: ∀L1,L2. L1 ⊑ L2 → - ∀K2,V,i. ⇩[0, i] L2 ≡ K2.ⓓV → - ∃∃K1. K1 ⊑ K2 & ⇩[0, i] L1 ≡ K1.ⓓV. -#L1 #L2 #HL12 #K2 #V #i #HLK2 elim (lsubr_fwd_ldrop2_bind … HL12 … HLK2) -L2 // * -#K1 #W #_ #_ #H destruct -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr_lsubr.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr_lsubr.ma deleted file mode 100644 index e8b7b5e0f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsubr_lsubr.ma +++ /dev/null @@ -1,53 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/substitution/lsubr.ma". - -(* RESTRICTED LOCAL ENVIRONMENT REFINEMENT **********************************) - -(* Auxiliary inversion lemmas ***********************************************) - -fact lsubr_inv_bind1_aux: ∀L1,L2. L1 ⊑ L2 → ∀I,K1,X. L1 = K1.ⓑ{I}X → - ∨∨ L2 = ⋆ - | ∃∃K2. K1 ⊑ K2 & L2 = K2.ⓑ{I}X - | ∃∃K2,V,W. K1 ⊑ K2 & L2 = K2.ⓛW & - I = Abbr & X = ⓝW.V. -#L1 #L2 * -L1 -L2 -[ #L #J #K1 #X #H destruct /2 width=1/ -| #I #L1 #L2 #V #HL12 #J #K1 #X #H destruct /3 width=3/ -| #L1 #L2 #V #W #HL12 #J #K1 #X #H destruct /3 width=6/ -] -qed-. - -lemma lsubr_inv_bind1: ∀I,K1,L2,X. K1.ⓑ{I}X ⊑ L2 → - ∨∨ L2 = ⋆ - | ∃∃K2. K1 ⊑ K2 & L2 = K2.ⓑ{I}X - | ∃∃K2,V,W. K1 ⊑ K2 & L2 = K2.ⓛW & - I = Abbr & X = ⓝW.V. -/2 width=3 by lsubr_inv_bind1_aux/ qed-. - -(* Main properties **********************************************************) - -theorem lsubr_trans: Transitive … lsubr. -#L1 #L #H elim H -L1 -L -[ #L1 #X #H - lapply (lsubr_inv_atom1 … H) -H // -| #I #L1 #L #V #_ #IHL1 #X #H - elim (lsubr_inv_bind1 … H) -H // * - #L2 [2: #V2 #W2 ] #HL2 #H1 [ #H2 #H3 ] destruct /3 width=1/ -| #L1 #L #V1 #W #_ #IHL1 #X #H - elim (lsubr_inv_abst1 … H) -H // * - #L2 #HL2 #H destruct /3 width=1/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl index 4f3d0f8dd..a5148f07e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl +++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl @@ -208,10 +208,6 @@ table { ] class "yellow" [ { "substitution" * } { - [ { "restricted local env. ref." * } { - [ "lsubr ( ? ⊑ ? )" "lsubr_lsubr" * ] - } - ] [ { "iterated structural successor for closures" * } { [ "fqus ( ⦃?,?,?⦄ ⊃* ⦃?,?,?⦄ )" "fqus_alt" + "fqus_lleq" + "fqus_fqus" * ] [ "fqup ( ⦃?,?,?⦄ ⊃+ ⦃?,?,?⦄ )" "fqup_lleq" + "fqup_fqup" * ] @@ -243,6 +239,18 @@ table { [ "lleq ( ? ⋕[?,?] ? )" "lleq_lleq" * ] } ] + [ { "contxt-sensitive extended substitution" * } { + [ "cpy ( ⦃?,?⦄ ⊢ ? ×▶[?,?] ? )" "cpy_lift" + "cpy_cpy" * ] + } + ] + [ { "local env. ref. for extended substitution" * } { + [ "lsuby ( ? ×⊑ ? )" "lsuby_lsuby" * ] + } + ] + [ { "restricted local env. ref." * } { + [ "lsubr ( ? ⊑ ? )" "lsubr_lsubr" * ] + } + ] [ { "global env. slicing" * } { [ "gdrop ( ⇩[?] ? ≡ ? )" "gdrop_gdrop" * ] }