From: Ferruccio Guidi Date: Sun, 25 Oct 2015 18:56:33 +0000 (+0000) Subject: theory of generic slicing almost completed .... X-Git-Tag: make_still_working~678 X-Git-Url: http://matita.cs.unibo.it/gitweb/?p=helm.git;a=commitdiff_plain;h=952ec5aa2e9a54787acb63a5c8d6fdbf9011ab60 theory of generic slicing almost completed .... --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy.etc new file mode 100644 index 000000000..4a27de0a4 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy.etc @@ -0,0 +1,286 @@ +(**************************************************************************) +(* ___ *) +(* ||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/psubst_6.ma". +include "basic_2/grammar/genv.ma". +include "basic_2/substitution/lsuby.ma". + +(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) + +(* activate genv *) +inductive cpy: ynat → ynat → relation4 genv lenv term term ≝ +| cpy_atom : ∀I,G,L,l,m. cpy l m G L (⓪{I}) (⓪{I}) +| cpy_subst: ∀I,G,L,K,V,W,i,l,m. l ≤ i → i < l+m → + ⬇[i] L ≡ K.ⓑ{I}V → ⬆[0, ⫯i] V ≡ W → cpy l m G L (#i) W +| cpy_bind : ∀a,I,G,L,V1,V2,T1,T2,l,m. + cpy l m G L V1 V2 → cpy (⫯l) m G (L.ⓑ{I}V1) T1 T2 → + cpy l m G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) +| cpy_flat : ∀I,G,L,V1,V2,T1,T2,l,m. + cpy l m G L V1 V2 → cpy l m G L T1 T2 → + cpy l m G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) +. + +interpretation "context-sensitive extended ordinary substritution (term)" + 'PSubst G L T1 l m T2 = (cpy l m G L T1 T2). + +(* Basic properties *********************************************************) + +lemma lsuby_cpy_trans: ∀G,l,m. lsub_trans … (cpy l m G) (lsuby l m). +#G #l #m #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -l -m +[ // +| #I #G #L1 #K1 #V #W #i #l #m #Hli #Hilm #HLK1 #HVW #L2 #HL12 + elim (ylt_inv_plus_dx … Hilm) #m0 #H0 #_ + elim (lsuby_drop_trans_be … HL12 … HLK1 … H0) -HL12 -HLK1 -H0 /2 width=5 by cpy_subst/ +| /4 width=1 by lsuby_succ, cpy_bind/ +| /3 width=1 by cpy_flat/ +] +qed-. + +lemma cpy_refl: ∀G,T,L,l,m. ⦃G, L⦄ ⊢ T ▶[l, m] T. +#G #T elim T -T // * /2 width=1 by cpy_bind, cpy_flat/ +qed. + +(* Basic_1: was: subst1_ex *) +lemma cpy_full: ∀I,G,K,V,T1,L,l. ⬇[l] L ≡ K.ⓑ{I}V → + ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ▶[l, 1] T2 & ⬆[l, 1] T ≡ T2. +#I #G #K #V #T1 elim T1 -T1 +[ * #i #L #l #HLK + /2 width=4 by lift_sort, lift_gref, ex2_2_intro/ + elim (ylt_split_eq i l) /3 width=4 by lift_lref_pred, lift_lref_lt, ex2_2_intro/ + #H destruct lapply (drop_fwd_Y2 … HLK) #Hi + elim (lift_total (⫯i) … V 0) /2 width=1 by ylt_succ/ + #W #HVW elim (lift_split … HVW i i … 1) + /4 width=6 by cpy_subst, monotonic_ylt_plus_sn, ex2_2_intro/ +| * [ #a ] #J #W1 #U1 #IHW1 #IHU1 #L #l #HLK + elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 + [ elim (IHU1 (L.ⓑ{J}W1) (⫯l)) -IHU1 + /3 width=9 by cpy_bind, drop_drop, 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,l1,m1. ⦃G, L⦄ ⊢ T1 ▶[l1, m1] T2 → + ∀l2,m2. l2 ≤ l1 → l1 + m1 ≤ l2 + m2 → + ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T2. +#G #L #T1 #T2 #l1 #m1 #H elim H -G -L -T1 -T2 -l1 -m1 // +[ /3 width=5 by cpy_subst, ylt_yle_trans, yle_trans/ +| /4 width=3 by cpy_bind, ylt_yle_trans, yle_succ/ +| /3 width=1 by cpy_flat/ +] +qed-. +(* +lemma cpy_weak_top: ∀G,L,T1,T2,l,m. + ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶[l, ∞] T2. +/2 width=5 by cpy_weak/ qed-. + +lemma cpy_weak_full: ∀G,L,T1,T2,l,m. + ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶[0, ∞] T2. +/2 width=5 by cpy_weak/ qed-. +*) +lemma cpy_split_up: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → + ∀i,m2. i + m2 = l + m → + ∀m1. i ≤ l + m1 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l, m1] T & ⦃G, L⦄ ⊢ T ▶[i, m2] T2. +#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m +[ /2 width=3 by ex2_intro/ +| #I #G #L #K #V #W #i #l #m #Hli #Hilm #HLK #HVW #j #m2 #H2 #m1 #H1 + elim (ylt_split i j) [ -Hilm -H2 | -Hli ] + /4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/ +| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #m2 #H2 #m1 #H1 + elim (IHV12 … H2 … H1) -IHV12 #V + elim (IHT12 (⫯i) … m2 … m1) -IHT12 /2 width=1 by yle_succ/ -H2 -H1 + #T #HT1 #HT2 lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 + /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ +| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #m2 #H2 #m1 #H1 + elim (IHV12 … H2 … H1) -IHV12 elim (IHT12 … H2 … H1) -IHT12 -H2 -H1 + /3 width=5 by ex2_intro, cpy_flat/ +] +qed-. + +lemma cpy_split_down: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → + ∀m1,m2. m = m1 + m2 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l+m2, m1] T & ⦃G, L⦄ ⊢ T ▶[l, m2] T2. +#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m +[ /2 width=3 by ex2_intro/ +| #I #G #L #K #V #W #i #l #m #Hli #Hilm #HLK #HVW #m1 #m2 #H destruct + elim (ylt_split i (l+m2)) [ -Hilm | -Hli ] + /3 width=9 by cpy_subst, ex2_intro/ +| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #m1 #m2 #H destruct + elim (IHV12 m1 m2) -IHV12 // #V + elim (IHT12 m1 m2) -IHT12 // + >yplus_succ1 #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 + /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ +| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #m1 #m2 #H destruct + elim (IHV12 m1 m2) -IHV12 // elim (IHT12 m1 m2) -IHT12 // + /3 width=5 by ex2_intro, cpy_flat/ +] +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma cpy_fwd_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → + ∀T1,l,m. ⬆[l, m] T1 ≡ U1 → + l ≤ lt → l + m ≤ lt + mt → + ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶[l+m, lt+mt-(l+m)] U2 & ⬆[l, m] T2 ≡ U2. +#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt +[ * #i #G #L #lt #mt #T1 #l #m #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ + ] +| #I #G #L #K #V #W #i #lt #mt #Hlti #Hilmt #HLK #HVW #T1 #l #m #H #Hllt #Hlmlmt + elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ -V -Hilmt -Hlmlmt | -Hlti -Hllt ] + [ elim (ylt_yle_false … Hllt) -Hllt /3 width=3 by yle_ylt_trans, ylt_inj/ + | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi + elim (lift_split … HVW l (⫯(i-m)) ? ? ?) [2,3,4: /2 width=1 by yle_succ_dx, le_S_S/ ] -Hlim + #T2 #_ >plus_minus /2 width=1 by yle_inv_inj/ ymax_pre_sn_comm // (**) (* explicit constructor *) + ] +| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #X #l #m #H #Hllt #Hlmlmt + elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HVW1) -V1 -IHW12 // + elim (IHU12 … HTU1) -T1 -IHU12 /2 width=1 by yle_succ/ + yplus_SO2 >yplus_succ1 >yplus_succ1 + /3 width=2 by cpy_bind, lift_bind, ex2_intro/ +| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #X #l #m #H #Hllt #Hlmlmt + elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HVW1) -V1 -IHW12 // elim (IHU12 … HTU1) -T1 -IHU12 + /3 width=2 by cpy_flat, lift_flat, ex2_intro/ +] +qed-. + +lemma cpy_fwd_tw: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ♯{T1} ≤ ♯{T2}. +#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m normalize +/3 width=1 by monotonic_le_plus_l, le_plus/ +qed-. + +(* Basic inversion lemmas ***************************************************) + +fact cpy_inv_atom1_aux: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ∀J. T1 = ⓪{J} → + T2 = ⓪{J} ∨ + ∃∃I,K,V,i. l ≤ yinj i & i < l + m & + ⬇[i] L ≡ K.ⓑ{I}V & + ⬆[O, ⫯i] V ≡ T2 & + J = LRef i. +#G #L #T1 #T2 #l #m * -G -L -T1 -T2 -l -m +[ #I #G #L #l #m #J #H destruct /2 width=1 by or_introl/ +| #I #G #L #K #V #T2 #i #l #m #Hli #Hilm #HLK #HVT2 #J #H destruct /3 width=9 by ex5_4_intro, or_intror/ +| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #H destruct +| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #H destruct +] +qed-. + +lemma cpy_inv_atom1: ∀I,G,L,T2,l,m. ⦃G, L⦄ ⊢ ⓪{I} ▶[l, m] T2 → + T2 = ⓪{I} ∨ + ∃∃J,K,V,i. l ≤ yinj i & i < l + m & + ⬇[i] L ≡ K.ⓑ{J}V & + ⬆[O, ⫯i] V ≡ T2 & + I = LRef i. +/2 width=4 by cpy_inv_atom1_aux/ qed-. + +(* Basic_1: was: subst1_gen_sort *) +lemma cpy_inv_sort1: ∀G,L,T2,k,l,m. ⦃G, L⦄ ⊢ ⋆k ▶[l, m] T2 → T2 = ⋆k. +#G #L #T2 #k #l #m #H +elim (cpy_inv_atom1 … H) -H // +* #I #K #V #i #_ #_ #_ #_ #H destruct +qed-. + +(* Basic_1: was: subst1_gen_lref *) +lemma cpy_inv_lref1: ∀G,L,T2,i,l,m. ⦃G, L⦄ ⊢ #i ▶[l, m] T2 → + T2 = #i ∨ + ∃∃I,K,V. l ≤ i & i < l + m & + ⬇[i] L ≡ K.ⓑ{I}V & + ⬆[O, ⫯i] V ≡ T2. +#G #L #T2 #i #l #m #H +elim (cpy_inv_atom1 … H) -H /2 width=1 by or_introl/ +* #I #K #V #j #Hlj #Hjlm #HLK #HVT2 #H destruct /3 width=5 by ex4_3_intro, or_intror/ +qed-. + +lemma cpy_inv_gref1: ∀G,L,T2,p,l,m. ⦃G, L⦄ ⊢ §p ▶[l, m] T2 → T2 = §p. +#G #L #T2 #p #l #m #H +elim (cpy_inv_atom1 … H) -H // +* #I #K #V #i #_ #_ #_ #_ #H destruct +qed-. + +fact cpy_inv_bind1_aux: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → + ∀a,I,V1,T1. U1 = ⓑ{a,I}V1.T1 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 & + ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ▶[⫯l, m] T2 & + U2 = ⓑ{a,I}V2.T2. +#G #L #U1 #U2 #l #m * -G -L -U1 -U2 -l -m +[ #I #G #L #l #m #b #J #W1 #U1 #H destruct +| #I #G #L #K #V #W #i #l #m #_ #_ #_ #_ #b #J #W1 #U1 #H destruct +| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #HV12 #HT12 #b #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ +| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #b #J #W1 #U1 #H destruct +] +qed-. + +lemma cpy_inv_bind1: ∀a,I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓑ{a,I} V1. T1 ▶[l, m] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 & + ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶[⫯l, m] T2 & + U2 = ⓑ{a,I}V2.T2. +/2 width=3 by cpy_inv_bind1_aux/ qed-. + +fact cpy_inv_flat1_aux: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → + ∀I,V1,T1. U1 = ⓕ{I}V1.T1 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 & + ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 & + U2 = ⓕ{I}V2.T2. +#G #L #U1 #U2 #l #m * -G -L -U1 -U2 -l -m +[ #I #G #L #l #m #J #W1 #U1 #H destruct +| #I #G #L #K #V #W #i #l #m #_ #_ #_ #_ #J #W1 #U1 #H destruct +| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #W1 #U1 #H destruct +| #I #G #L #V1 #V2 #T1 #T2 #l #m #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ +] +qed-. + +lemma cpy_inv_flat1: ∀I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ▶[l, m] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 & + ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 & + U2 = ⓕ{I}V2.T2. +/2 width=3 by cpy_inv_flat1_aux/ qed-. + + +fact cpy_inv_refl_O2_aux: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → m = 0 → T1 = T2. +#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m +[ // +| #I #G #L #K #V #W #i #l #m #Hli #Hilm #_ #_ #H destruct + elim (ylt_yle_false … Hli) -Hli // +| /3 width=1 by eq_f2/ +| /3 width=1 by eq_f2/ +] +qed-. + +lemma cpy_inv_refl_O2: ∀G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▶[l, 0] T2 → T1 = T2. +/2 width=6 by cpy_inv_refl_O2_aux/ qed-. + +(* Basic_1: was: subst1_gen_lift_eq *) +lemma cpy_inv_lift1_eq: ∀G,T1,U1,l,m. ⬆[l, m] T1 ≡ U1 → + ∀L,U2. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → U1 = U2. +#G #T1 #U1 #l #m #HTU1 #L #U2 #HU12 elim (cpy_fwd_up … HU12 … HTU1) -HU12 -HTU1 +/2 width=4 by cpy_inv_refl_O2/ +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/etc_new/cpy/cpy_cpy.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy_cpy.etc new file mode 100644 index 000000000..7a65a2db5 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy_cpy.etc @@ -0,0 +1,123 @@ +(**************************************************************************) +(* ___ *) +(* ||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/cpy_lift.ma". + +(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) + +(* Main properties **********************************************************) + +(* Basic_1: was: subst1_confluence_eq *) +theorem cpy_conf_eq: ∀G,L,T0,T1,l1,m1. ⦃G, L⦄ ⊢ T0 ▶[l1, m1] T1 → + ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶[l2, m2] T2 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L⦄ ⊢ T2 ▶[l1, m1] T. +#G #L #T0 #T1 #l1 #m1 #H elim H -G -L -T0 -T1 -l1 -m1 +[ /2 width=3 by ex2_intro/ +| #I1 #G #L #K1 #V1 #T1 #i0 #l1 #m1 #Hl1 #Hlm1 #HLK1 #HVT1 #T2 #l2 #m2 #H + elim (cpy_inv_lref1 … H) -H + [ #HX destruct /3 width=7 by cpy_subst, ex2_intro/ + | -Hl1 -Hlm1 * #I2 #K2 #V2 #_ #_ #HLK2 #HVT2 + lapply (drop_mono … HLK1 … HLK2) -HLK1 -HLK2 #H destruct + >(lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3 by ex2_intro/ + ] +| #a #I #G #L #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #X #l2 #m2 #HX + elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02) -IHV01 -HV02 #V #HV1 #HV2 + elim (IHT01 … HT02) -T0 #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT1 (L.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/ + lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V2) ?) -HT2 + /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/ +| #I #G #L #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #X #l2 #m2 #HX + elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02) -V0 + elim (IHT01 … HT02) -T0 /3 width=5 by cpy_flat, ex2_intro/ +] +qed-. + +(* Basic_1: was: subst1_confluence_neq *) +theorem cpy_conf_neq: ∀G,L1,T0,T1,l1,m1. ⦃G, L1⦄ ⊢ T0 ▶[l1, m1] T1 → + ∀L2,T2,l2,m2. ⦃G, L2⦄ ⊢ T0 ▶[l2, m2] T2 → + (l1 + m1 ≤ l2 ∨ l2 + m2 ≤ l1) → + ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L1⦄ ⊢ T2 ▶[l1, m1] T. +#G #L1 #T0 #T1 #l1 #m1 #H elim H -G -L1 -T0 -T1 -l1 -m1 +[ /2 width=3 by ex2_intro/ +| #I1 #G #L1 #K1 #V1 #T1 #i0 #l1 #m1 #Hl1 #Hlm1 #HLK1 #HVT1 #L2 #T2 #l2 #m2 #H1 #H2 + elim (cpy_inv_lref1 … H1) -H1 + [ #H destruct /3 width=7 by cpy_subst, ex2_intro/ + | -HLK1 -HVT1 * #I2 #K2 #V2 #Hl2 #Hlm2 #_ #_ elim H2 -H2 #Hlml [ -Hl1 -Hlm2 | -Hl2 -Hlm1 ] + [ elim (ylt_yle_false … Hlm1) -Hlm1 /2 width=3 by yle_trans/ + | elim (ylt_yle_false … Hlm2) -Hlm2 /2 width=3 by yle_trans/ + ] + ] +| #a #I #G #L1 #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #L2 #X #l2 #m2 #HX #H + elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -IHV01 -HV02 #V #HV1 #HV2 + elim (IHT01 … HT02) -T0 + [ -H #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT1 (L2.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/ + lapply (lsuby_cpy_trans … HT2 (L1.ⓑ{I}V2) ?) -HT2 /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/ + | -HV1 -HV2 elim H -H /3 width=1 by yle_succ, or_introl, or_intror/ + ] +| #I #G #L1 #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #L2 #X #l2 #m2 #HX #H + elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -V0 + elim (IHT01 … HT02 H) -T0 -H /3 width=5 by cpy_flat, ex2_intro/ +] +qed-. + +(* Note: the constant 1 comes from cpy_subst *) +(* Basic_1: was: subst1_trans *) +theorem cpy_trans_ge: ∀G,L,T1,T0,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T0 → + ∀T2. ⦃G, L⦄ ⊢ T0 ▶[l, 1] T2 → 1 ≤ m → ⦃G, L⦄ ⊢ T1 ▶[l, m] T2. +#G #L #T1 #T0 #l #m #H elim H -G -L -T1 -T0 -l -m +[ #I #G #L #l #m #T2 #H #Hm + elim (cpy_inv_atom1 … H) -H + [ #H destruct // + | * #J #K #V #i #Hl2i #Hilm2 #HLK #HVT2 #H destruct + lapply (ylt_yle_trans … (l+m) … Hilm2) + /2 width=5 by cpy_subst, monotonic_yle_plus_sn/ + ] +| #I #G #L #K #V #V2 #i #l #m #Hli #Hilm #HLK #HVW #T2 #HVT2 #Hm + lapply (cpy_weak … HVT2 0 (i+1) ? ?) -HVT2 /3 width=1 by yle_plus_dx2_trans, yle_succ/ + >yplus_inj #HVT2 <(cpy_inv_lift1_eq … HVW … HVT2) -HVT2 /2 width=5 by cpy_subst/ +| #a #I #G #L #V1 #V0 #T1 #T0 #l #m #_ #_ #IHV10 #IHT10 #X #H #Hm + elim (cpy_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct + lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02 + lapply (IHT10 … HT02 Hm) -T0 /3 width=1 by cpy_bind/ +| #I #G #L #V1 #V0 #T1 #T0 #l #m #_ #_ #IHV10 #IHT10 #X #H #Hm + elim (cpy_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1 by cpy_flat/ +] +qed-. + +theorem cpy_trans_down: ∀G,L,T1,T0,l1,m1. ⦃G, L⦄ ⊢ T1 ▶[l1, m1] T0 → + ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶[l2, m2] T2 → l2 + m2 ≤ l1 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L⦄ ⊢ T ▶[l1, m1] T2. +#G #L #T1 #T0 #l1 #m1 #H elim H -G -L -T1 -T0 -l1 -m1 +[ /2 width=3 by ex2_intro/ +| #I #G #L #K #V #W #i1 #l1 #m1 #Hli1 #Hilm1 #HLK #HVW #T2 #l2 #m2 #HWT2 #Hlm2l1 + lapply (yle_trans … Hlm2l1 … Hli1) -Hlm2l1 #Hlm2i1 + lapply (cpy_weak … HWT2 0 (i1+1) ? ?) -HWT2 /3 width=1 by yle_succ, yle_pred_sn/ -Hlm2i1 + >yplus_inj #HWT2 <(cpy_inv_lift1_eq … HVW … HWT2) -HWT2 /3 width=9 by cpy_subst, ex2_intro/ +| #a #I #G #L #V1 #V0 #T1 #T0 #l1 #m1 #_ #_ #IHV10 #IHT10 #X #l2 #m2 #HX #lm2l1 + elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02 + elim (IHV10 … HV02) -IHV10 -HV02 // #V + elim (IHT10 … HT02) -T0 /2 width=1 by yle_succ/ #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 /3 width=6 by cpy_bind, lsuby_succ, ex2_intro/ +| #I #G #L #V1 #V0 #T1 #T0 #l1 #m1 #_ #_ #IHV10 #IHT10 #X #l2 #m2 #HX #lm2l1 + elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV10 … HV02) -V0 // + elim (IHT10 … HT02) -T0 /3 width=6 by cpy_flat, ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy_lift.etc new file mode 100644 index 000000000..44a7f3d76 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy_lift.etc @@ -0,0 +1,253 @@ +(**************************************************************************) +(* ___ *) +(* ||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/drop_drop.ma". +include "basic_2/substitution/cpy.ma". + +(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) + +(* Properties on relocation *************************************************) + +(* Basic_1: was: subst1_lift_lt *) +lemma cpy_lift_le: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 → + ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K → + ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 → + lt + mt ≤ l → ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2. +#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt +[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hlmtl + lapply (ylt_yle_trans … Hlmtl … Hilmt) -Hlmtl #Hil + lapply (lift_inv_lref1_lt … H … Hil) -H #H destruct + elim (lift_trans_ge … HVW … HWU2) -W /2 width=1 by ylt_fwd_le_succ1/ + yplus_SO2 >yminus_succ2 #W #HVW #HWU2 + elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H + elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K #Y #_ #HVY + >(lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=5 by cpy_subst/ +| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + /4 width=7 by cpy_bind, drop_skip, yle_succ/ +| #G #I #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl + 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=7 by cpy_flat/ +] +qed-. + +lemma cpy_lift_be: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 → + ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K → + ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 → + lt ≤ l → l ≤ lt + mt → ⦃G, L⦄ ⊢ U1 ▶[lt, mt + m] U2. +#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt +[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_ #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hltl #_ + elim (lift_inv_lref1 … H) -H * #Hil #H destruct + [ -Hltl + lapply (ylt_yle_trans … (lt+mt+m) … Hilmt) // -Hilmt #Hilmtm + elim (lift_trans_ge … HVW … HWU2) -W yplus_SO2 + [2: >yplus_O1 /2 width=1 by ylt_fwd_le_succ1/ ] >yminus_succ2 #W #HVW #HWU2 + elim (drop_trans_le … HLK … HKV) -K /2 width=1 by ylt_fwd_le/ #X #HLK #H + elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K #Y #_ #HVY + >(lift_mono … HVY … HVW) -V #H destruct /2 width=5 by cpy_subst/ + | -Hlti + lapply (yle_trans … Hltl … Hil) -Hltl #Hlti + lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2 + lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hil + /3 width=5 by cpy_subst, drop_inv_gen, monotonic_ylt_plus_dx, yle_plus_dx1_trans/ + ] +| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hltl #Hllmt + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + /4 width=7 by cpy_bind, drop_skip, yle_succ/ +| #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl + 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=7 by cpy_flat/ +] +qed-. + +(* Basic_1: was: subst1_lift_ge *) +lemma cpy_lift_ge: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 → + ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K → + ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 → + l ≤ lt → ⦃G, L⦄ ⊢ U1 ▶[lt+m, mt] U2. +#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt +[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hllt + lapply (yle_trans … Hllt … Hlti) -Hllt #Hil + lapply (lift_inv_lref1_ge … H … Hil) -H #H destruct + lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2 + lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hil + /3 width=5 by cpy_subst, drop_inv_gen, monotonic_ylt_plus_dx, monotonic_yle_plus_dx/ +| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hllt + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + /4 width=6 by cpy_bind, drop_skip, yle_succ/ +| #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hllt + 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 by cpy_flat/ +] +qed-. + +(* Inversion lemmas on relocation *******************************************) + +(* Basic_1: was: subst1_gen_lift_lt *) +lemma cpy_inv_lift1_le: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → + ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → + lt + mt ≤ l → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 & ⬆[l, m] T2 ≡ U2. +#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt +[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ + ] +| #I #G #L #KV #V #W #i #lt #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hlmtl + lapply (ylt_yle_trans … Hlmtl … Hilmt) -Hlmtl #Hil + lapply (lift_inv_lref2_lt … H … Hil) -H #H destruct + elim (drop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV + elim (lift_trans_le … HUV … HVW) -V // >minus_plus yplus_SO2 >ymax_pre_sn /2 width=1 by ylt_fwd_le_succ1/ -Hil + /3 width=5 by cpy_subst, ex2_intro/ +| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl + elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2 + elim (IHU12 … HTU1) -IHU12 -HTU1 + /3 width=6 by cpy_bind, yle_succ, drop_skip, lift_bind, ex2_intro/ +| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl + elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HLK … HVW1) -W1 // + elim (IHU12 … HLK … HTU1) -U1 -HLK + /3 width=5 by cpy_flat, lift_flat, ex2_intro/ +] +qed-. + +lemma cpy_inv_lift1_be: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → + ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → + lt ≤ l → l + m ≤ lt + mt → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, mt-m] T2 & ⬆[l, m] T2 ≡ U2. +#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt +[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_ #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ + ] +| #I #G #L #KV #V #W #i #x #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hltl #Hlmlmt + elim (yle_inv_inj2 … Hlti) -Hlti #lt #Hlti #H destruct + lapply (yle_fwd_plus_yge … Hltl Hlmlmt) #Hmmt + elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ -Hltl -Hilmt | -Hlti -Hlmlmt ] + [ lapply (ylt_yle_trans i l (lt+(mt-m)) ? ?) // + [ >yplus_minus_assoc_inj /2 width=1 by yle_plus1_to_minus_inj2/ ] -Hlmlmt #Hilmtm + elim (drop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV + elim (lift_trans_le … HUV … HVW) -V // + yplus_SO2 >ymax_pre_sn /2 width=1 by ylt_fwd_le_succ1/ -Hil + /4 width=5 by cpy_subst, ex2_intro, yle_inj/ + | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi + lapply (yle_inv_inj … Hmi) -Hmi #Hmi + lapply (yle_trans … Hltl (i-m) ?) // -Hltl #Hltim + lapply (drop_conf_ge … HLK … HLKV ?) -L // #HKV + elim (lift_split … HVW l (i-m+1)) -HVW [2,3,4: /2 width=1 by yle_succ_dx, le_S_S/ ] -Hil -Hlim + #V1 #HV1 >plus_minus // yplus_minus_assoc_inj /3 width=1 by monotonic_ylt_minus_dx, yle_inj/ + ] +| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hltl #Hlmlmt + elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2 + elim (IHU12 … HTU1) -U1 + /3 width=6 by cpy_bind, drop_skip, lift_bind, yle_succ, ex2_intro/ +| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hltl #Hlmlmt + elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HLK … HVW1) -W1 // + elim (IHU12 … HLK … HTU1) -U1 -HLK // + /3 width=5 by cpy_flat, lift_flat, ex2_intro/ +] +qed-. + +(* Basic_1: was: subst1_gen_lift_ge *) +lemma cpy_inv_lift1_ge: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → + ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → + l + m ≤ lt → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt-m, mt] T2 & ⬆[l, m] T2 ≡ U2. +#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt +[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ + ] +| #I #G #L #KV #V #W #i #lt #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hlmlt + lapply (yle_trans … Hlmlt … Hlti) #Hlmi + elim (yle_inv_plus_inj2 … Hlmlt) -Hlmlt #_ #Hmlt + elim (yle_inv_plus_inj2 … Hlmi) #Hlim #Hmi + lapply (yle_inv_inj … Hmi) -Hmi #Hmi + lapply (lift_inv_lref2_ge … H ?) -H // #H destruct + lapply (drop_conf_ge … HLK … HLKV ?) -L // #HKV + elim (lift_split … HVW l (i-m+1)) -HVW [2,3,4: /3 width=1 by yle_succ, yle_pred_sn, le_S_S/ ] -Hlmi -Hlim + #V0 #HV10 >plus_minus // yminus_succ1_inj /3 width=5 by cpy_bind, lift_bind, ex2_intro/ +| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl + elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HLK … HVW1) -W1 // + elim (IHU12 … HLK … HTU1) -U1 -HLK /3 width=5 by cpy_flat, lift_flat, ex2_intro/ +] +qed-. + +(* Advanced inversion lemmas on relocation ***********************************) + +lemma cpy_inv_lift1_ge_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → + ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → + l ≤ lt → lt ≤ l + m → l + m ≤ lt + mt → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[l, lt + mt - (l + m)] T2 & ⬆[l, m] T2 ≡ U2. +#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hllt #Hltlm #Hlmlmt +elim (cpy_split_up … HU12 (l + m)) -HU12 // -Hlmlmt #U #HU1 #HU2 +lapply (cpy_weak … HU1 l m ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hllt -Hltlm #HU1 +lapply (cpy_inv_lift1_eq … HTU1 … HU1) -HU1 #HU1 destruct +elim (cpy_inv_lift1_ge … HU2 … HLK … HTU1) -U -L /2 width=3 by ex2_intro/ +qed-. + +lemma cpy_inv_lift1_be_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → + ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → + lt ≤ l → lt + mt ≤ l + m → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, l-lt] T2 & ⬆[l, m] T2 ≡ U2. +#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hlmtlm +lapply (cpy_weak … HU12 lt (l+m-lt) ? ?) -HU12 // +[ >ymax_pre_sn_comm /2 width=1 by yle_plus_dx1_trans/ ] -Hlmtlm #HU12 +elim (cpy_inv_lift1_be … HU12 … HLK … HTU1) -U1 -L /2 width=3 by ex2_intro/ +qed-. + +lemma cpy_inv_lift1_le_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → + ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → + lt ≤ l → l ≤ lt + mt → lt + mt ≤ l + m → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, l - lt] T2 & ⬆[l, m] T2 ≡ U2. +#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hllmt #Hlmtlm +elim (cpy_split_up … HU12 l) -HU12 // #U #HU1 #HU2 +elim (cpy_inv_lift1_le … HU1 … HLK … HTU1) -U1 +[2: >ymax_pre_sn_comm // ] -Hltl #T #HT1 #HTU +lapply (cpy_weak … HU2 l m ? ?) -HU2 // +[ >ymax_pre_sn_comm // ] -Hllmt -Hlmtlm #HU2 +lapply (cpy_inv_lift1_eq … HTU … HU2) -L #H destruct /2 width=3 by ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy_nlift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy_nlift.etc new file mode 100644 index 000000000..3a6450724 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/cpy/cpy_nlift.etc @@ -0,0 +1,65 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lift_neg.ma". +include "basic_2/substitution/lift_lift.ma". +include "basic_2/substitution/cpy.ma". + +(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) + +(* Inversion lemmas on negated relocation ***********************************) + +lemma cpy_fwd_nlift2_ge: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → + ∀i. l ≤ i → (∀T2. ⬆[i, 1] T2 ≡ U2 → ⊥) → + (∀T1. ⬆[i, 1] T1 ≡ U1 → ⊥) ∨ + ∃∃I,K,W,j. l ≤ yinj j & j < i & ⬇[j]L ≡ K.ⓑ{I}W & + (∀V. ⬆[⫰(i-j), 1] V ≡ W → ⊥) & (∀T1. ⬆[j, 1] T1 ≡ U1 → ⊥). +#G #L #U1 #U2 #l #m #H elim H -G -L -U1 -U2 -l -m +[ /3 width=2 by or_introl/ +| #I #G #L #K #V #W #j #l #m #Hlj #Hjlm #HLK #HVW #i #Hli #HnW + elim (ylt_split j i) #Hij + [ @or_intror @(ex5_4_intro … HLK) // -HLK + [ #X #HXV elim (lift_trans_le … HXV … HVW ?) -V // #Y #HXY + yplus_SO2 >ymax_pre_sn /2 width=2 by ylt_fwd_le_succ1/ + | -HnW /3 width=7 by lift_inv_lref2_be, ylt_inj/ + ] + | elim (lift_split … HVW i j) -HVW // + #X #_ #H elim HnW -HnW // + ] +| #a #I #G #L #W1 #W2 #U1 #U2 #l #m #_ #_ #IHW12 #IHU12 #i #Hli #H elim (nlift_inv_bind … H) -H + [ #HnW2 elim (IHW12 … HnW2) -IHW12 -HnW2 -IHU12 // + [ /4 width=9 by nlift_bind_sn, or_introl/ + | * /5 width=9 by nlift_bind_sn, ex5_4_intro, or_intror/ + ] + | #HnU2 elim (IHU12 … HnU2) -IHU12 -HnU2 -IHW12 /2 width=1 by yle_succ/ + [ /4 width=9 by nlift_bind_dx, or_introl/ + | * #J #K #W #j #Hlj elim (yle_inv_succ1 … Hlj) -Hlj #Hlj #Hj + yminus_succ #Hji #HLK #HnW + lapply (drop_inv_drop1_lt … HLK ?) /2 width=1 by ylt_O/ -HLK #HLK + minus_plus_plus_l + #L #HL1 #HL2 elim (lt_or_ge (|L1|) (m2-m1)) #H0 + [ elim (drop_inv_O1_gt … HL1 H0) -HL1 #H1 #H2 destruct + elim (drop_inv_atom1 … HL2) -HL2 #H #_ destruct + @(ex2_intro … (⋆)) [ @drop_O1_ge normalize // ] + @drop_atom #H destruct + | elim (drop_O1_pair … HL1 H0 I V) -HL1 -H0 /3 width=5 by drop_drop, ex2_intro/ + ] + ] +| #I #L1 #L2 #V1 #V2 #l #m2 #_ #HV21 #IHL12 #m1 #Hm12 elim (IHL12 … Hm12) -IHL12 + #L #HL1 #HL2 elim (lift_split … HV21 l m1) -HV21 /3 width=5 by drop_skip, ex2_intro/ +] +qed-. + +(* Basic_2A1: includes: drop_split *) +lemma drops_split_trans: ∀L1,L2,t. ⬇*[t] L1 ≡ L2 → ∀t1,t2. t1 ⊚ t2 ≡ t → + ∃∃L. ⬇*[t1] L1 ≡ L & ⬇*[t2] L ≡ L2. +#L1 #L2 #t #H elim H -L1 -L2 -t +[ #t1 #t2 #H elim (after_inv_empty3 … H) -H + /2 width=3 by ex2_intro, drops_atom/ +| #I #L1 #L2 #V #t #HL12 #IHL12 #t1 #t2 #H elim (after_inv_false3 … H) -H * + [ #tl1 #tl2 #H1 #H2 #Ht destruct elim (IHL12 … Ht) -t + #tl #H1 #H2 + | #tl1 #H #Ht destruct elim (IHL12 … Ht) -t + /3 width=5 by ex2_intro, drops_drop/ + ] +| #I #L1 #L2 #V1 #V2 #t #_ #HV21 #IHL12 #t1 #t2 #H elim (after_inv_true3 … H) -H + #tl1 #tl2 #H1 #H2 #Ht destruct elim (lifts_split_trans … HV21 … Ht) -HV21 + elim (IHL12 … Ht) -t /3 width=5 by ex2_intro, drops_skip/ +] +qed-. + +lemma drop_inv_refl: ∀L,l,m. ⬇[Ⓕ, l, m] L ≡ L → m = 0. +/2 width=5 by drop_inv_length_eq/ qed-. + +fact drop_inv_FT_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → + ∀I,K,V. L2 = K.ⓑ{I}V → s = Ⓣ → l = 0 → + ⬇[Ⓕ, l, m] L1 ≡ K.ⓑ{I}V. +#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m +[ #l #m #_ #J #K #W #H destruct +| #I #L #V #J #K #W #H destruct // +| #I #L1 #L2 #V #m #_ #IHL12 #J #K #W #H1 #H2 destruct + /3 width=1 by drop_drop/ +| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #_ #J #K #W #_ #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma drop_inv_FT: ∀I,L,K,V,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] L ≡ +K.ⓑ{I}V. +/2 width=5 by drop_inv_FT_aux/ qed. + +lemma drop_inv_gen: ∀I,L,K,V,s,m. ⬇[s, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] L ≡ +K.ⓑ{I}V. +#I #L #K #V * /2 width=1 by drop_inv_FT/ +qed-. + +lemma drop_inv_T: ∀I,L,K,V,s,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[s, 0, m] L +≡ K.ⓑ{I}V. +#I #L #K #V * /2 width=1 by drop_inv_FT/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops_append.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops_append.etc new file mode 100644 index 000000000..d134fe768 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops_append.etc @@ -0,0 +1,65 @@ +(**************************************************************************) +(* ___ *) +(* ||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/grammar/lenv_append.ma". +include "basic_2/substitution/drop.ma". + +(* DROPPING *****************************************************************) + +(* Properties on append for local environments ******************************) + +fact drop_O1_append_sn_le_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → + l = 0 → m ≤ |L1| → + ∀L. ⬇[s, 0, m] L @@ L1 ≡ L @@ L2. +#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m // +[ #l #m #_ #_ #H >(yle_inv_O2 … H) -m // +| /4 width=1 by drop_drop, yle_inv_succ/ +| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #_ #H elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma drop_O1_append_sn_le: ∀L1,L2,s,m. ⬇[s, yinj 0, m] L1 ≡ L2 → m ≤ |L1| → + ∀L. ⬇[s, 0, m] L @@ L1 ≡ L @@ L2. +/2 width=3 by drop_O1_append_sn_le_aux/ qed. + +(* Inversion lemmas on append for local environments ************************) + +lemma drop_O1_inv_append1_ge: ∀K,L1,L2,s,m. ⬇[s, 0, m] L1 @@ L2 ≡ K → + ∀m0. |L2| + m0 = m → ⬇[s, 0, m0] L1 ≡ K. +#K #L1 #L2 elim L2 -L2 +[ #s #m #H #m0 >yplus_O1 #H0 destruct // +| #L2 #I #V #IHL2 #s #m #H #m0 >yplus_succ1 + #H0 elim (drop_inv_O1_pair1 … H) -H * #Hm #HL12 destruct + [ elim (ysucc_inv_O_dx … Hm) + | /2 width=3 by/ + ] +] +qed-. + +lemma drop_O1_inv_append1_le: ∀K,L1,L2,s,m. ⬇[s, 0, m] L1 @@ L2 ≡ K → m ≤ |L2| → + ∀K2. ⬇[s, 0, m] L2 ≡ K2 → K = L1 @@ K2. +#K #L1 #L2 elim L2 -L2 +[ #s #m #H1 #H2 #K2 #H3 lapply (yle_inv_O2 … H2) -H2 + #H2 elim (drop_inv_atom1 … H3) -H3 #H3 #_ destruct + >(drop_inv_O2 … H1) -H1 // +| #L2 #I #V #IHL2 #s #m @(ynat_ind … m) -m [ -IHL2 || -IHL2 ] + [ #H1 #_ #K2 #H2 + lapply (drop_inv_O2 … H1) -H1 #H1 + lapply (drop_inv_O2 … H2) -H2 #H2 destruct // + | /3 width=7 by drop_inv_drop1, yle_inv_succ/ + | #_ #H lapply (yle_inv_Y1 … H) -H + #H elim (ylt_yle_false (|L2.ⓑ{I}V|) (∞)) // + ] +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops_length.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops_length.etc new file mode 100644 index 000000000..8953abb9a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/drops_length.etc @@ -0,0 +1,152 @@ +include "basic_2/grammar/lenv_length.ma". + +lemma drop_inv_O1_gt: ∀L,K,m,s. ⬇[s, 0, m] L ≡ K → |L| < m → + s = Ⓣ ∧ K = ⋆. +#L elim L -L [| #L #Z #X #IHL ] #K #m #s #H normalize in ⊢ (?%?→?); #H1m +[ elim (drop_inv_atom1 … H) -H elim s -s /2 width=1 by conj/ + #_ #Hs lapply (Hs ?) // -Hs #H destruct elim (ylt_yle_false … H1m) -H1m // +| elim (drop_inv_O1_pair1 … H) -H * #H2m #HLK destruct + [ elim (ylt_yle_false … H1m) -H1m // + | elim (IHL … HLK) -IHL -HLK /2 width=1 by ylt_pred, conj/ + ] +] +qed-. + +lemma drop_O1_le: ∀s,m,L. m ≤ |L| → ∃K. ⬇[s, 0, m] L ≡ K. +#s #m @(ynat_ind … m) -m /2 width=2 by ex_intro/ +[ #m #IHm * + [ #H elim (ylt_yle_false … H) -H // + | #L #I #V #H elim (IHm L) -IHm /3 width=2 by drop_drop, yle_inv_succ, ex_intro/ + ] +| #L #H elim (ylt_yle_false … H) -H // +] +qed-. + +lemma drop_O1_lt: ∀s,L,m. m < |L| → ∃∃I,K,V. ⬇[s, 0, m] L ≡ K.ⓑ{I}V. +#s #L elim L -L +[ #m #H elim (ylt_yle_false … H) -H // +| #L #I #V #IHL #m @(ynat_ind … m) -m /2 width=4 by drop_pair, ex1_3_intro/ + [ #m #_#H elim (IHL m) -IHL /3 width=4 by drop_drop, ylt_inv_succ, ex1_3_intro/ + | #H elim (ylt_yle_false … H) -H // + ] +] +qed-. + +lemma drop_O1_pair: ∀L,K,m,s. ⬇[s, 0, m] L ≡ K → m ≤ |L| → ∀I,V. + ∃∃J,W. ⬇[s, 0, m] L.ⓑ{I}V ≡ K.ⓑ{J}W. +#L elim L -L [| #L #Z #X #IHL ] #K #m #s #H #Hm #I #V +[ elim (drop_inv_atom1 … H) -H #H >(yle_inv_O2 … Hm) -m + #Hs destruct /2 width=3 by ex1_2_intro/ +| elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK destruct /2 width=3 by ex1_2_intro/ + elim (IHL … HLK … Z X) -IHL -HLK + /3 width=3 by yle_pred, drop_drop_lt, ex1_2_intro/ +] +qed-. + +lemma drop_O1_ge: ∀L,m. |L| ≤ m → ⬇[Ⓣ, 0, m] L ≡ ⋆. +#L elim L -L [ #m #_ @drop_atom #H destruct ] +#L #I #V #IHL #m @(ynat_ind … m) -m // +[ #H elim (ylt_yle_false … H) -H /2 width=1 by ylt_inj/ +| /4 width=1 by drop_drop, yle_inv_succ/ +] +qed. + +lemma drop_O1_eq: ∀L,s. ⬇[s, 0, |L|] L ≡ ⋆. +#L elim L -L /2 width=1 by drop_drop/ +qed. + +lemma drop_fwd_length_ge: ∀L1,L2,l,m,s. ⬇[s, l, m] L1 ≡ L2 → |L1| ≤ l → |L2| = |L1|. +#L1 #L2 #l #m #s #H elim H -L1 -L2 -l -m // +[ #I #L1 #L2 #V #m #_ #_ #H elim (ylt_yle_false … H) -H // +| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #IH #H + lapply (yle_inv_succ … H) -H #H + >length_pair >length_pair /3 width=1 by eq_f/ +] +qed-. + +lemma drop_fwd_length_le_le: ∀L1,L2,l,m,s. ⬇[s, l, m] L1 ≡ L2 → + ∀l0. l + m + l0 = |L1| → |L2| = l + l0. +#L1 #L2 #l #m #s #H elim H -L1 -L2 -l -m // +[ #l #m #Hm #l0 #H elim (yplus_inv_O … H) -H + #H #H0 elim (yplus_inv_O … H) -H + #H1 #_ destruct // +| #I #L1 #L2 #V #m #_ >yplus_O1 >yplus_O1 #IH #l0 + /3 width=1 by ysucc_inv_inj/ +| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #IHL12 #l0 >yplus_succ1 >yplus_succ1 #H + lapply (ysucc_inv_inj … H) -H #Hl1 + >yplus_succ1 /3 width=1 by eq_f/ +] +qed-. + +lemma drop_fwd_length_le_ge: ∀L1,L2,l,m,s. ⬇[s, l, m] L1 ≡ L2 → l ≤ |L1| → |L1| ≤ l + m → |L2| = l. +#L1 #L2 #l #m #s #H elim H -L1 -L2 -l -m +[ #l #m #_ #H #_ /2 width=1 by yle_inv_O2/ +| #I #L #V #_ #H elim (ylt_yle_false … H) -H // +| #I #L1 #L2 #V #m #_ >yplus_O1 >yplus_O1 + /3 width=1 by yle_inv_succ/ +| #I #L1 #L2 #V1 #v2 #l #m #_ #_ #IH + >yplus_SO2 >yplus_SO2 + /4 width=1 by yle_inv_succ, eq_f/ +] +qed-. + +lemma drop_fwd_length: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L1| = |L2| + m. +#L1 #L2 #l #m #H elim H -L1 -L2 -l -m // +#l #m #H >H -m // +qed-. + +lemma drop_fwd_length_le2: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → m ≤ |L1|. +#L1 #L2 #l #m #H lapply (drop_fwd_length … H) -H // +qed-. + +lemma drop_fwd_length_le4: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L2| ≤ |L1|. +#L1 #L2 #l #m #H lapply (drop_fwd_length … H) -H // +qed-. + +lemma drop_fwd_length_lt2: ∀L1,I2,K2,V2,l,m. + ⬇[Ⓕ, l, m] L1 ≡ K2. ⓑ{I2} V2 → m < |L1|. +#L1 #I2 #K2 #V2 #l #m #H +lapply (drop_fwd_Y2 … H) #Hm +lapply (drop_fwd_length … H) -l #H <(yplus_O2 m) >H -L1 +/2 width=1 by monotonic_ylt_plus_sn/ +qed-. + +lemma drop_fwd_length_lt4: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → 0 < m → |L2| < |L1|. +#L1 #L2 #l #m #H +lapply (drop_fwd_Y2 … H) #Hm +lapply (drop_fwd_length … H) -l +/2 width=1 by monotonic_ylt_plus_sn/ +qed-. + +lemma drop_fwd_length_eq1: ∀L1,L2,K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 → + |L1| = |L2| → |K1| = |K2|. +#L1 #L2 #K1 #K2 #l #m #HLK1 #HLK2 #HL12 +lapply (drop_fwd_Y2 … HLK1) #Hm +lapply (drop_fwd_length … HLK1) -HLK1 +lapply (drop_fwd_length … HLK2) -HLK2 +#H #H0 >H in HL12; -H >H0 -H0 #H +@(yplus_inv_monotonic_dx … H) -H // (**) (* auto fails *) +qed-. + +lemma drop_fwd_length_eq2: ∀L1,L2,K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 → + |K1| = |K2| → |L1| = |L2|. +#L1 #L2 #K1 #K2 #l #m #HLK1 #HLK2 #HL12 +lapply (drop_fwd_length … HLK1) -HLK1 +lapply (drop_fwd_length … HLK2) -HLK2 // +qed-. + +lemma drop_inv_length_eq: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → + |L1| = |L2| → m = 0. +#L1 #L2 #l #m #H #HL12 lapply (drop_fwd_length … H) -H +>HL12 -L1 #H elim (discr_yplus_x_xy … H) -H // +#H elim (ylt_yle_false (|L2|) (∞)) // +qed-. + +lemma drop_fwd_be: ∀L,K,s,l,m,i. ⬇[s, l, m] L ≡ K → |K| ≤ i → i < l → |L| ≤ i. +#L #K #s #l #m #i #HLK #HK #Hl elim (ylt_split i (|L|)) // +#HL elim (drop_O1_lt (Ⓕ) … HL) #I #K0 #V #HLK0 -HL +elim (ylt_inv_plus_sn … Hl) -Hl #l0 #H0 +elim (drop_conf_lt … HLK … HLK0 … H0) -HLK -HLK0 -H0 +#K1 #V1 #HK1 #_ #_ lapply (drop_fwd_length_lt2 … HK1) -I -K1 -V1 +#H elim (ylt_yle_false … H) -H // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/rdropstar_3.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/rdropstar_3.etc new file mode 100644 index 000000000..8a6f2f4c1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/drops/rdropstar_3.etc @@ -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 m ] break term 46 L1 ≡ break term 46 L2 )" + non associative with precedence 45 + for @{ 'RDropStar $m $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/gget.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/gget.etc new file mode 100644 index 000000000..1ea9ce7ff --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/gget.etc @@ -0,0 +1,81 @@ +(**************************************************************************) +(* ___ *) +(* ||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/rdrop_3.ma". +include "basic_2/grammar/genv.ma". + +(* GLOBAL ENVIRONMENT READING ***********************************************) + +inductive gget (m:nat): relation genv ≝ +| gget_gt: ∀G. |G| ≤ m → gget m G (⋆) +| gget_eq: ∀G. |G| = m + 1 → gget m G G +| gget_lt: ∀I,G1,G2,V. m < |G1| → gget m G1 G2 → gget m (G1. ⓑ{I} V) G2 +. + +interpretation "global reading" + 'RDrop m G1 G2 = (gget m G1 G2). + +(* basic inversion lemmas ***************************************************) + +lemma gget_inv_gt: ∀G1,G2,m. ⬇[m] G1 ≡ G2 → |G1| ≤ m → G2 = ⋆. +#G1 #G2 #m * -G1 -G2 // +[ #G #H >H -H >commutative_plus #H (**) (* lemma needed here *) + lapply (le_plus_to_le_r … 0 H) -H #H + lapply (le_n_O_to_eq … H) -H #H destruct +| #I #G1 #G2 #V #H1 #_ #H2 + lapply (le_to_lt_to_lt … H2 H1) -H2 -H1 normalize in ⊢ (? % ? → ?); >commutative_plus #H + lapply (lt_plus_to_lt_l … 0 H) -H #H + elim (lt_zero_false … H) +] +qed-. + +lemma gget_inv_eq: ∀G1,G2,m. ⬇[m] G1 ≡ G2 → |G1| = m + 1 → G1 = G2. +#G1 #G2 #m * -G1 -G2 // +[ #G #H1 #H2 >H2 in H1; -H2 >commutative_plus #H (**) (* lemma needed here *) + lapply (le_plus_to_le_r … 0 H) -H #H + lapply (le_n_O_to_eq … H) -H #H destruct +| #I #G1 #G2 #V #H1 #_ normalize #H2 + <(injective_plus_l … H2) in H1; -H2 #H + elim (lt_refl_false … H) +] +qed-. + +fact gget_inv_lt_aux: ∀I,G,G1,G2,V,m. ⬇[m] G ≡ G2 → G = G1. ⓑ{I} V → + m < |G1| → ⬇[m] G1 ≡ G2. +#I #G #G1 #G2 #V #m * -G -G2 +[ #G #H1 #H destruct #H2 + lapply (le_to_lt_to_lt … H1 H2) -H1 -H2 normalize in ⊢ (? % ? → ?); >commutative_plus #H + lapply (lt_plus_to_lt_l … 0 H) -H #H + elim (lt_zero_false … H) +| #G #H1 #H2 destruct >(injective_plus_l … H1) -H1 #H + elim (lt_refl_false … H) +| #J #G #G2 #W #_ #HG2 #H destruct // +] +qed-. + +lemma gget_inv_lt: ∀I,G1,G2,V,m. + ⬇[m] G1. ⓑ{I} V ≡ G2 → m < |G1| → ⬇[m] G1 ≡ G2. +/2 width=5 by gget_inv_lt_aux/ qed-. + +(* Basic properties *********************************************************) + +lemma gget_total: ∀m,G1. ∃G2. ⬇[m] G1 ≡ G2. +#m #G1 elim G1 -G1 /3 width=2 by gget_gt, ex_intro/ +#I #V #G1 * #G2 #HG12 +elim (lt_or_eq_or_gt m (|G1|)) #Hm +[ /3 width=2 by gget_lt, ex_intro/ +| destruct /3 width=2 by gget_eq, ex_intro/ +| @ex_intro [2: @gget_gt normalize /2 width=1 by/ | skip ] (**) (* explicit constructor *) +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/gget_gget.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/gget_gget.etc new file mode 100644 index 000000000..7bd39f09b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/gget_gget.etc @@ -0,0 +1,40 @@ +(**************************************************************************) +(* ___ *) +(* ||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/gget.ma". + +(* GLOBAL ENVIRONMENT READING ***********************************************) + +(* Main properties **********************************************************) + +theorem gget_mono: ∀G,G1,m. ⬇[m] G ≡ G1 → ∀G2. ⬇[m] G ≡ G2 → G1 = G2. +#G #G1 #m #H elim H -G -G1 +[ #G #Hm #G2 #H + >(gget_inv_gt … H Hm) -H -Hm // +| #G #Hm #G2 #H + >(gget_inv_eq … H Hm) -H -Hm // +| #I #G #G1 #V #Hm #_ #IHG1 #G2 #H + lapply (gget_inv_lt … H Hm) -H -Hm /2 width=1 by/ +] +qed-. + +lemma gget_dec: ∀G1,G2,m. Decidable (⬇[m] G1 ≡ G2). +#G1 #G2 #m +elim (gget_total m G1) #G #HG1 +elim (eq_genv_dec G G2) #HG2 +[ destruct /2 width=1 by or_introl/ +| @or_intror #HG12 + lapply (gget_mono … HG1 … HG12) -HG1 -HG12 /2 width=1 by/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/rdrop_3.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/rdrop_3.etc new file mode 100644 index 000000000..a761d47c6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/gget/rdrop_3.etc @@ -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 m ] break term 46 L1 ≡ break term 46 L2 )" + non associative with precedence 45 + for @{ 'RDrop $m $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts.etc new file mode 100644 index 000000000..ebb4a2456 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts.etc @@ -0,0 +1,29 @@ +(* Basic_1: was: lift_r *) +lemma lift_refl: ∀T,l. ⬆[l, 0] T ≡ T. +#T elim T -T +[ * #i // #l elim (lt_or_ge i l) /2 width=1 by lift_lref_lt, +lift_lref_ge/ +| * /2 width=1 by lift_bind, lift_flat/ +] +qed. + +lemma lift_total: ∀T1,l,m. ∃T2. ⬆[l,m] T1 ≡ T2. +#T1 elim T1 -T1 +[ * #i /2 width=2 by lift_sort, lift_gref, ex_intro/ + #l #m elim (lt_or_ge i l) /3 width=2 by lift_lref_lt, lift_lref_ge, +ex_intro/ +| * [ #a ] #I #V1 #T1 #IHV1 #IHT1 #l #m + elim (IHV1 l m) -IHV1 #V2 #HV12 + [ elim (IHT1 (l+1) m) -IHT1 /3 width=2 by lift_bind, ex_intro/ + | elim (IHT1 l m) -IHT1 /3 width=2 by lift_flat, ex_intro/ + ] +] +qed. + +lemma liftv_total: ∀l,m. ∀T1s:list term. ∃T2s. ⬆[l, m] T1s ≡ T2s. +#l #m #T1s elim T1s -T1s +[ /2 width=2 by liftv_nil, ex_intro/ +| #T1 #T1s * #T2s #HT12s + elim (lift_total T1 l m) /3 width=2 by liftv_cons, ex_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts_neg.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts_neg.etc new file mode 100644 index 000000000..f4b8a70c1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lifts/lifts_neg.etc @@ -0,0 +1,75 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lift.ma". + +(* BASIC TERM RELOCATION ****************************************************) + +(* Properties on negated basic relocation ***********************************) + +lemma nlift_lref_be_SO: ∀X,j. j < ∞ → ⬆[j, 1] X ≡ #j → ⊥. +#X #j #Hj #H elim (lift_inv_lref2 … H) -H * +[ #H elim (ylt_yle_false … H) -H // +| #i #Hij #_ #H1 #H2 destruct + elim (ylt_inv_plus_Y … Hj) -Hj #Hi #_ + elim (ylt_yle_false … Hij) -Hij /2 width=1 by monotonic_ylt_plus_sn/ +] +qed-. + +lemma nlift_bind_sn: ∀W,l,m. (∀V. ⬆[l, m] V ≡ W → ⊥) → + ∀a,I,U. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥). +#W #l #m #HW #a #I #U #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/ +qed-. + +lemma nlift_bind_dx: ∀U,l,m. (∀T. ⬆[⫯l, m] T ≡ U → ⊥) → + ∀a,I,W. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥). +#U #l #m #HU #a #I #W #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/ +qed-. + +lemma nlift_flat_sn: ∀W,l,m. (∀V. ⬆[l, m] V ≡ W → ⊥) → + ∀I,U. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥). +#W #l #m #HW #I #U #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/ +qed-. + +lemma nlift_flat_dx: ∀U,l,m. (∀T. ⬆[l, m] T ≡ U → ⊥) → + ∀I,W. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥). +#U #l #m #HU #I #W #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/ +qed-. + +(* Inversion lemmas on negated basic relocation *****************************) + +lemma nlift_inv_lref_be_SO: ∀i,j. (∀X. ⬆[i, 1] X ≡ #j → ⊥) → j = i ∧ j < ∞. +#i #j elim (ylt_split_eq i j) #Hij #H destruct +[ elim (H (#⫰j)) -H /2 width=1 by lift_lref_pred/ +| elim (yle_split_eq i (∞)) /2 width=1 by conj/ #H0 destruct + elim (H (#∞)) -H /2 width=1 by lift_lref_plus, ylt_Y/ +| elim (H (#j)) -H /2 width=1 by lift_lref_lt/ +] +qed-. + +lemma nlift_inv_bind: ∀a,I,W,U,l,m. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥) → + (∀V. ⬆[l, m] V ≡ W → ⊥) ∨ (∀T. ⬆[⫯l, m] T ≡ U → ⊥). +#a #I #W #U #l #m #H elim (is_lift_dec W l m) +[ * /4 width=2 by lift_bind, or_intror/ +| /4 width=2 by ex_intro, or_introl/ +] +qed-. + +lemma nlift_inv_flat: ∀I,W,U,l,m. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥) → + (∀V. ⬆[l, m] V ≡ W → ⊥) ∨ (∀T. ⬆[l, m] T ≡ U → ⊥). +#I #W #U #l #m #H elim (is_lift_dec W l m) +[ * /4 width=2 by lift_flat, or_intror/ +| /4 width=2 by ex_intro, or_introl/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsuby/lsuby.etc b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsuby/lsuby.etc new file mode 100644 index 000000000..e2ee7c7c2 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc_new/lsuby/lsuby.etc @@ -0,0 +1,236 @@ +(**************************************************************************) +(* ___ *) +(* ||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_4.ma". +include "basic_2/substitution/drop.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR EXTENDED SUBSTITUTION *******************) + +inductive lsuby: relation4 ynat ynat lenv lenv ≝ +| lsuby_atom: ∀L,l,m. lsuby l m L (⋆) +| lsuby_zero: ∀I1,I2,L1,L2,V1,V2. + lsuby 0 0 L1 L2 → lsuby 0 0 (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2) +| lsuby_pair: ∀I1,I2,L1,L2,V,m. lsuby 0 m L1 L2 → + lsuby 0 (⫯m) (L1.ⓑ{I1}V) (L2.ⓑ{I2}V) +| lsuby_succ: ∀I1,I2,L1,L2,V1,V2,l,m. + lsuby l m L1 L2 → lsuby (⫯l) m (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2) +. + +interpretation + "local environment refinement (extended substitution)" + 'LRSubEq L1 l m L2 = (lsuby l m L1 L2). + +(* Basic properties *********************************************************) + +lemma lsuby_pair_lt: ∀I1,I2,L1,L2,V,m. L1 ⊆[0, ⫰m] L2 → 0 < m → + L1.ⓑ{I1}V ⊆[0, m] L2.ⓑ{I2}V. +#I1 #I2 #L1 #L2 #V #m #HL12 #Hm <(ylt_inv_O1 … Hm) /2 width=1 by lsuby_pair/ +qed. + +lemma lsuby_succ_lt: ∀I1,I2,L1,L2,V1,V2,l,m. L1 ⊆[⫰l, m] L2 → 0 < l → + L1.ⓑ{I1}V1 ⊆[l, m] L2. ⓑ{I2}V2. +#I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #Hl <(ylt_inv_O1 … Hl) /2 width=1 by lsuby_succ/ +qed. + +lemma lsuby_pair_O_Y: ∀L1,L2. L1 ⊆[0, ∞] L2 → + ∀I1,I2,V. L1.ⓑ{I1}V ⊆[0,∞] L2.ⓑ{I2}V. +#L1 #L2 #HL12 #I1 #I2 #V lapply (lsuby_pair I1 I2 … V … HL12) -HL12 // +qed. + +lemma lsuby_refl: ∀L,l,m. L ⊆[l, m] L. +#L elim L -L // +#L #I #V #IHL #l elim (ynat_cases … l) [| * #x ] +#Hl destruct /2 width=1 by lsuby_succ/ +#m elim (ynat_cases … m) [| * #x ] +#Hm destruct /2 width=1 by lsuby_zero, lsuby_pair/ +qed. + +lemma lsuby_O2: ∀L2,L1,l. |L2| ≤ |L1| → L1 ⊆[l, 0] L2. +#L2 elim L2 -L2 // #L2 #I2 #V2 #IHL2 * +[ #l #H elim (ylt_yle_false … H) -H // +| #L1 #I1 #V1 #l + #H lapply (yle_inv_succ … H) -H #HL12 + elim (ynat_cases l) /3 width=1 by lsuby_zero/ + * /3 width=1 by lsuby_succ/ +] +qed. + +lemma lsuby_sym: ∀l,m,L1,L2. L1 ⊆[l, m] L2 → |L1| = |L2| → L2 ⊆[l, m] L1. +#l #m #L1 #L2 #H elim H -l -m -L1 -L2 +[ #L1 #l #m #H >(length_inv_zero_dx … H) -L1 // +| /2 width=1 by lsuby_O2/ +| #I1 #I2 #L1 #L2 #V #m #_ #IHL12 #H lapply (ysucc_inv_inj … H) -H + /3 width=1 by lsuby_pair/ +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL12 #H lapply (ysucc_inv_inj … H) -H + /3 width=1 by lsuby_succ/ +] +qed-. + +(* Basic inversion lemmas ***************************************************) + +fact lsuby_inv_atom1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → L1 = ⋆ → L2 = ⋆. +#L1 #L2 #l #m * -L1 -L2 -l -m // +[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #H destruct +| #I1 #I2 #L1 #L2 #V #m #_ #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #H destruct +] +qed-. + +lemma lsuby_inv_atom1: ∀L2,l,m. ⋆ ⊆[l, m] L2 → L2 = ⋆. +/2 width=5 by lsuby_inv_atom1_aux/ qed-. + +fact lsuby_inv_zero1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → + ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → l = 0 → m = 0 → + L2 = ⋆ ∨ + ∃∃J2,K2,W2. K1 ⊆[0, 0] K2 & L2 = K2.ⓑ{J2}W2. +#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/ +[ #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J1 #K1 #W1 #H #_ #_ destruct + /3 width=5 by ex2_3_intro, or_intror/ +| #I1 #I2 #L1 #L2 #V #m #_ #J1 #K1 #W1 #_ #_ #H + elim (ysucc_inv_O_dx … H) +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J1 #K1 #W1 #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma lsuby_inv_zero1: ∀I1,K1,L2,V1. K1.ⓑ{I1}V1 ⊆[0, 0] L2 → + L2 = ⋆ ∨ + ∃∃I2,K2,V2. K1 ⊆[0, 0] K2 & L2 = K2.ⓑ{I2}V2. +/2 width=9 by lsuby_inv_zero1_aux/ qed-. + +fact lsuby_inv_pair1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → + ∀J1,K1,W. L1 = K1.ⓑ{J1}W → l = 0 → 0 < m → + L2 = ⋆ ∨ + ∃∃J2,K2. K1 ⊆[0, ⫰m] K2 & L2 = K2.ⓑ{J2}W. +#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/ +[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W #_ #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #m #HL12 #J1 #K1 #W #H #_ #_ destruct + /3 width=4 by ex2_2_intro, or_intror/ +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J1 #K1 #W #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma lsuby_inv_pair1: ∀I1,K1,L2,V,m. K1.ⓑ{I1}V ⊆[0, m] L2 → 0 < m → + L2 = ⋆ ∨ + ∃∃I2,K2. K1 ⊆[0, ⫰m] K2 & L2 = K2.ⓑ{I2}V. +/2 width=6 by lsuby_inv_pair1_aux/ qed-. + +fact lsuby_inv_succ1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → + ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < l → + L2 = ⋆ ∨ + ∃∃J2,K2,W2. K1 ⊆[⫰l, m] K2 & L2 = K2.ⓑ{J2}W2. +#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/ +[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W1 #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #m #_ #J1 #K1 #W1 #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #J1 #K1 #W1 #H #_ destruct + /3 width=5 by ex2_3_intro, or_intror/ +] +qed-. + +lemma lsuby_inv_succ1: ∀I1,K1,L2,V1,l,m. K1.ⓑ{I1}V1 ⊆[l, m] L2 → 0 < l → + L2 = ⋆ ∨ + ∃∃I2,K2,V2. K1 ⊆[⫰l, m] K2 & L2 = K2.ⓑ{I2}V2. +/2 width=5 by lsuby_inv_succ1_aux/ qed-. + +fact lsuby_inv_zero2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → + ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → l = 0 → m = 0 → + ∃∃J1,K1,W1. K1 ⊆[0, 0] K2 & L1 = K1.ⓑ{J1}W1. +#L1 #L2 #l #m * -L1 -L2 -l -m +[ #L1 #l #m #J2 #K2 #W1 #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J2 #K2 #W2 #H #_ #_ destruct + /2 width=5 by ex2_3_intro/ +| #I1 #I2 #L1 #L2 #V #m #_ #J2 #K2 #W2 #_ #_ #H + elim (ysucc_inv_O_dx … H) +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J2 #K2 #W2 #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma lsuby_inv_zero2: ∀I2,K2,L1,V2. L1 ⊆[0, 0] K2.ⓑ{I2}V2 → + ∃∃I1,K1,V1. K1 ⊆[0, 0] K2 & L1 = K1.ⓑ{I1}V1. +/2 width=9 by lsuby_inv_zero2_aux/ qed-. + +fact lsuby_inv_pair2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → + ∀J2,K2,W. L2 = K2.ⓑ{J2}W → l = 0 → 0 < m → + ∃∃J1,K1. K1 ⊆[0, ⫰m] K2 & L1 = K1.ⓑ{J1}W. +#L1 #L2 #l #m * -L1 -L2 -l -m +[ #L1 #l #m #J2 #K2 #W #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W #_ #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #m #HL12 #J2 #K2 #W #H #_ #_ destruct + /2 width=4 by ex2_2_intro/ +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J2 #K2 #W #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma lsuby_inv_pair2: ∀I2,K2,L1,V,m. L1 ⊆[0, m] K2.ⓑ{I2}V → 0 < m → + ∃∃I1,K1. K1 ⊆[0, ⫰m] K2 & L1 = K1.ⓑ{I1}V. +/2 width=6 by lsuby_inv_pair2_aux/ qed-. + +fact lsuby_inv_succ2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → + ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → 0 < l → + ∃∃J1,K1,W1. K1 ⊆[⫰l, m] K2 & L1 = K1.ⓑ{J1}W1. +#L1 #L2 #l #m * -L1 -L2 -l -m +[ #L1 #l #m #J2 #K2 #W2 #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W2 #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #m #_ #J2 #K1 #W2 #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #J2 #K2 #W2 #H #_ destruct + /2 width=5 by ex2_3_intro/ +] +qed-. + +lemma lsuby_inv_succ2: ∀I2,K2,L1,V2,l,m. L1 ⊆[l, m] K2.ⓑ{I2}V2 → 0 < l → + ∃∃I1,K1,V1. K1 ⊆[⫰l, m] K2 & L1 = K1.ⓑ{I1}V1. +/2 width=5 by lsuby_inv_succ2_aux/ qed-. + +(* Basic forward lemmas *****************************************************) + +lemma lsuby_fwd_length: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → |L2| ≤ |L1|. +#L1 #L2 #l #m #H elim H -L1 -L2 -l -m /2 width=1 by yle_succ/ +qed-. + +(* Properties on basic slicing **********************************************) + +lemma lsuby_drop_trans_be: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → + ∀I2,K2,W,s,i. ⬇[s, 0, i] L2 ≡ K2.ⓑ{I2}W → + l ≤ i → ∀m0. i + ⫯m0 = l + m → + ∃∃I1,K1. K1 ⊆[0, m0] K2 & ⬇[s, 0, i] L1 ≡ K1.ⓑ{I1}W. +#L1 #L2 #l #m #H elim H -L1 -L2 -l -m +[ #L1 #l #m #J2 #K2 #W #s #i #H + elim (drop_inv_atom1 … H) -H #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J2 #K2 #W #s #i #_ #_ #m0 + >yplus_O2 >yplus_succ2 #H elim (ysucc_inv_O_dx … H) +| #I1 #I2 #L1 #L2 #V #m #HL12 #IHL12 #J2 #K2 #W #s #i #H #_ #m0 + >yplus_succ2 >yplus_succ2 #H0 lapply (ysucc_inv_inj … H0) -H0 + elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ] + [ destruct -I2 /2 width=4 by drop_pair, ex2_2_intro/ + | lapply (ylt_inv_O1 … Hi) + #H yplus_succ1 >yplus_succ2 #H lapply (ysucc_inv_inj … H) -H + (drop_fwd_length … HLK) in H; -L + #H elim (discr_yplus_xy_x … H) -H /2 width=2 by ysucc_inv_O_sn/ + #H elim (ylt_yle_false (|K|) (∞)) // +] +/2 width=4 by discr_tpair_xy_y, discr_tpair_xy_x/ +qed-. + +lemma fqu_inv_eq: ∀G,L1,L2,T. ⦃G, L1, T⦄ ⊐ ⦃G, L2, T⦄ → |L1| = |L2| → ⊥. +#G #L1 #L2 #T #H #H0 @(fqu_inv_eq_aux … H … H0) // (**) (* full auto fails *) +qed-. + +(* Advanced eliminators *****************************************************) + +lemma fqu_wf_ind: ∀R:relation3 …. ( + ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G2 L2 T2) → + R G1 L1 T1 + ) → ∀G1,L1,T1. R G1 L1 T1. +#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=1 by fqu_fwd_fw/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq.ma new file mode 100644 index 000000000..914af3702 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq.ma @@ -0,0 +1,64 @@ +(**************************************************************************) +(* ___ *) +(* ||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/suptermopt_6.ma". +include "basic_2/substitution/fqu.ma". + +(* OPTIONAL SUPCLOSURE ******************************************************) + +(* activate genv *) +inductive fquq: tri_relation genv lenv term ≝ +| fquq_lref_O : ∀I,G,L,V. fquq G (L.ⓑ{I}V) (#0) G L V +| fquq_pair_sn: ∀I,G,L,V,T. fquq G L (②{I}V.T) G L V +| fquq_bind_dx: ∀a,I,G,L,V,T. fquq G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T +| fquq_flat_dx: ∀I,G, L,V,T. fquq G L (ⓕ{I}V.T) G L T +| fquq_drop : ∀G,L,K,T,U,m. + ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → fquq G L U G K T +. + +interpretation + "optional structural successor (closure)" + 'SupTermOpt G1 L1 T1 G2 L2 T2 = (fquq G1 L1 T1 G2 L2 T2). + +(* Basic properties *********************************************************) + +lemma fquq_refl: tri_reflexive … fquq. +/2 width=3 by fquq_drop/ qed. + +lemma fqu_fquq: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 /2 width=3 by fquq_drop/ +qed. + +(* Basic forward lemmas *****************************************************) + +lemma fquq_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 /2 width=1 by lt_to_le/ +#G1 #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 +lapply (drop_fwd_lw … HLK1) -HLK1 +lapply (lift_fwd_tw … HTU1) -HTU1 +/2 width=1 by le_plus, le_n/ +qed-. + +fact fquq_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → + ∀i. T1 = #i → |L2| ≤ |L1|. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // +[ #a #I #G #L #V #T #j #H destruct +| #G1 #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #i #H destruct + /2 width=3 by drop_fwd_length_le4/ +] +qed-. + +lemma fquq_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐⸮ ⦃G2, L2, T2⦄ → |L2| ≤ |L1|. +/2 width=7 by fquq_fwd_length_lref1_aux/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq_alt.ma new file mode 100644 index 000000000..7d7f05bd9 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/fquq_alt.ma @@ -0,0 +1,58 @@ +(**************************************************************************) +(* ___ *) +(* ||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/suptermoptalt_6.ma". +include "basic_2/substitution/fquq.ma". + +(* OPTIONAL SUPCLOSURE ******************************************************) + +(* alternative definition of fquq *) +definition fquqa: tri_relation genv lenv term ≝ tri_RC … fqu. + +interpretation + "optional structural successor (closure) alternative" + 'SupTermOptAlt G1 L1 T1 G2 L2 T2 = (fquqa G1 L1 T1 G2 L2 T2). + +(* Basic properties *********************************************************) + +lemma fquqa_refl: tri_reflexive … fquqa. +// qed. + +lemma fquqa_drop: ∀G,L,K,T,U,m. + ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → ⦃G, L, U⦄ ⊐⊐⸮ ⦃G, K, T⦄. +#G #L #K #T #U #m @(ynat_ind … m) -m /3 width=3 by fqu_drop, or_introl/ +#HLK #HTU >(drop_inv_O2 … HLK) -L >(lift_inv_O2 … HTU) -T // +qed. + +(* Main properties **********************************************************) + +theorem fquq_fquqa: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⊐⸮ ⦃G2, L2, T2⦄. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 +/2 width=3 by fquqa_drop, fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, or_introl/ +qed. + +(* Main inversion properties ************************************************) + +theorem fquqa_inv_fquq: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⊐⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄. +#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H /2 width=1 by fqu_fquq/ +* #H1 #H2 #H3 destruct // +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma fquq_inv_gen: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → + ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ ∨ (∧∧ G1 = G2 & L1 = L2 & T1 = T2). +#G1 #G2 #L1 #L2 #T1 #T2 #H elim (fquq_fquqa … H) -H [| * ] +/2 width=1 by or_introl/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drop_lreq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drop_lreq.ma new file mode 100644 index 000000000..85b93239d --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drop_lreq.ma @@ -0,0 +1,97 @@ +(**************************************************************************) +(* ___ *) +(* ||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/grammar/lreq_lreq.ma". +include "basic_2/substitution/drop.ma". + +(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************) + +definition dedropable_sn: predicate (relation lenv) ≝ + λR. ∀L1,K1,s,l,m. ⬇[s, l, m] L1 ≡ K1 → ∀K2. R K1 K2 → + ∃∃L2. R L1 L2 & ⬇[s, l, m] L2 ≡ K2 & L1 ⩬[l, m] L2. + +(* Properties on equivalence ************************************************) + +lemma lreq_drop_trans_be: ∀L1,L2,l,m. L1 ⩬[l, m] L2 → + ∀I,K2,W,s,i. ⬇[s, 0, i] L2 ≡ K2.ⓑ{I}W → + l ≤ i → ∀m0. i + ⫯m0 = l + m → + ∃∃K1. K1 ⩬[0, m0] K2 & ⬇[s, 0, i] L1 ≡ K1.ⓑ{I}W. +#L1 #L2 #l #m #H elim H -L1 -L2 -l -m +[ #l #m #J #K2 #W #s #i #H + elim (drop_inv_atom1 … H) -H #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J #K2 #W #s #i #_ #_ #m0 + >yplus_succ2 #H elim (ysucc_inv_O_dx … H) +| #I #L1 #L2 #V #m #HL12 #IHL12 #J #K2 #W #s #i #H #_ >yplus_O1 #m0 #H0 + elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ] + [ destruct + /2 width=3 by drop_pair, ex2_intro/ + | lapply (ylt_inv_O1 … Hi) #H yplus_succ1 #H lapply (ysucc_inv_inj … H) -H <(yplus_O1 m) + #H0 elim (IHL12 … HLK1 … H0) -IHL12 -HLK1 -H0 // + /3 width=3 by drop_drop_lt, ex2_intro/ + ] +| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL12 #J #K2 #W #s #i #HLK2 #Hli #m0 + elim (yle_inv_succ1 … Hli) -Hli + #Hli #Hi yplus_succ1 >yplus_succ1 #H lapply (ysucc_inv_inj … H) -H + #H0 lapply (drop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O1/ + #HLK1 elim (IHL12 … HLK1 … H0) -IHL12 -HLK1 -H0 + /4 width=3 by ylt_O1, drop_drop_lt, ex2_intro/ +] +qed-. + +lemma lreq_drop_conf_be: ∀L1,L2,l,m. L1 ⩬[l, m] L2 → + ∀I,K1,W,s,i. ⬇[s, 0, i] L1 ≡ K1.ⓑ{I}W → + l ≤ i → ∀m0. i + ⫯m0 = l + m → + ∃∃K2. K1 ⩬[0, m0] K2 & ⬇[s, 0, i] L2 ≡ K2.ⓑ{I}W. +#L1 #L2 #l #m #HL12 #I #K1 #W #s #i #HLK1 #Hli #m0 #H0 +elim (lreq_drop_trans_be … (lreq_sym … HL12) … HLK1 … H0) // -L1 -Hli -H0 +/3 width=3 by lreq_sym, ex2_intro/ +qed-. + +lemma drop_O1_ex: ∀K2,i,L1. |L1| = |K2| + i → + ∃∃L2. L1 ⩬[0, i] L2 & ⬇[i] L2 ≡ K2. +#K2 #i @(ynat_ind … i) -i +[ /3 width=3 by lreq_O2, ex2_intro/ +| #i #IHi #Y >yplus_succ2 #Hi + elim (drop_O1_lt (Ⓕ) Y 0) [2: >Hi // ] + #I #L1 #V #H lapply (drop_inv_O2 … H) -H #H destruct + >length_pair in Hi; #H lapply (ysucc_inv_inj … H) -H + #HL1K2 elim (IHi L1) -IHi // -HL1K2 + /3 width=5 by lreq_pair, drop_drop, ex2_intro/ +| #L1 >yplus_Y2 #H elim (ylt_yle_false (|L1|) (∞)) // +] +qed-. + +lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R). +#R #HR #L1 #K1 #s #l #m #HLK1 #K2 #H elim H -K2 +[ #K2 #HK12 elim (HR … HLK1 … HK12) -HR -K1 + /3 width=4 by inj, ex3_intro/ +| #K #K2 #_ #HK2 * #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2) -HR -K + /3 width=6 by lreq_trans, step, ex3_intro/ +] +qed-. + +(* Inversion lemmas on equivalence ******************************************) + +lemma drop_O1_inj: ∀i,L1,L2,K. ⬇[i] L1 ≡ K → ⬇[i] L2 ≡ K → L1 ⩬[i, ∞] L2. +#i @(ynat_ind … i) -i +[ #L1 #L2 #K #H <(drop_inv_O2 … H) -K #H <(drop_inv_O2 … H) -L1 // +| #i #IHi * [2: #L1 #I1 #V1 ] * [2,4: #L2 #I2 #V2 ] #K #HLK1 #HLK2 // + lapply (drop_fwd_length … HLK1) + <(drop_fwd_length … HLK2) [ /4 width=5 by drop_inv_drop1, lreq_succ/ ] + #H [ elim (ysucc_inv_O_sn … H) | elim (ysucc_inv_O_dx … H) ] +| #L1 #L2 #K #H1 lapply (drop_fwd_Y2 … H1) -H1 + #H elim (ylt_yle_false … H) // +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma index 065135c9a..1016ec0d5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops.ma @@ -12,113 +12,191 @@ (* *) (**************************************************************************) -include "basic_2/notation/relations/rdropstar_3.ma". include "basic_2/notation/relations/rdropstar_4.ma". -include "basic_2/substitution/drop.ma". -include "basic_2/multiple/mr2_minus.ma". -include "basic_2/multiple/lifts_vector.ma". +include "basic_2/grammar/lenv.ma". +include "basic_2/relocation/lifts.ma". -(* ITERATED LOCAL ENVIRONMENT SLICING ***************************************) +(* GENERAL SLICING FOR LOCAL ENVIRONMENTS ***********************************) -inductive drops (s:bool): list2 ynat nat → relation lenv ≝ -| drops_nil : ∀L. drops s (◊) L L -| drops_cons: ∀L1,L,L2,cs,l,m. - drops s cs L1 L → ⬇[s, l, m] L ≡ L2 → drops s ({l, m} @ cs) L1 L2 +(* Basic_1: includes: drop_skip_bind drop1_skip_bind *) +(* Basic_2A1: includes: drop_atom drop_pair drop_drop drop_skip + drop_refl_atom_O2 drop_drop_lt drop_skip_lt +*) +inductive drops (s:bool): trace → relation lenv ≝ +| drops_atom: ∀t. (s = Ⓕ → 𝐈⦃t⦄) → drops s (t) (⋆) (⋆) +| drops_drop: ∀I,L1,L2,V,t. drops s t L1 L2 → drops s (Ⓕ@t) (L1.ⓑ{I}V) L2 +| drops_skip: ∀I,L1,L2,V1,V2,t. + drops s t L1 L2 → ⬆*[t] V2 ≡ V1 → + drops s (Ⓣ@t) (L1.ⓑ{I}V1) (L2.ⓑ{I}V2) . -interpretation "iterated slicing (local environment) abstract" - 'RDropStar s cs T1 T2 = (drops s cs T1 T2). -(* -interpretation "iterated slicing (local environment) general" - 'RDropStar des T1 T2 = (drops true des T1 T2). -*) +interpretation "general slicing (local environment)" + 'RDropStar s t L1 L2 = (drops s t L1 L2). definition d_liftable1: relation2 lenv term → predicate bool ≝ - λR,s. ∀K,T. R K T → ∀L,l,m. ⬇[s, l, m] L ≡ K → - ∀U. ⬆[l, m] T ≡ U → R L U. + λR,s. ∀L,K,t. ⬇*[s, t] L ≡ K → + ∀T,U. ⬆*[t] T ≡ U → R K T → R L U. + +definition d_liftable2: predicate (lenv → relation term) ≝ + λR. ∀K,T1,T2. R K T1 T2 → ∀L,s,t. ⬇*[s, t] L ≡ K → + ∀U1. ⬆*[t] T1 ≡ U1 → + ∃∃U2. ⬆*[t] T2 ≡ U2 & R L U1 U2. + +definition d_deliftable2_sn: predicate (lenv → relation term) ≝ + λR. ∀L,U1,U2. R L U1 U2 → ∀K,s,t. ⬇*[s, t] L ≡ K → + ∀T1. ⬆*[t] T1 ≡ U1 → + ∃∃T2. ⬆*[t] T2 ≡ U2 & R K T1 T2. + +definition dropable_sn: predicate (relation lenv) ≝ + λR. ∀L1,K1,s,t. ⬇*[s, t] L1 ≡ K1 → ∀L2. R L1 L2 → + ∃∃K2. R K1 K2 & ⬇*[s, t] L2 ≡ K2. +(* +definition dropable_dx: predicate (relation lenv) ≝ + λR. ∀L1,L2. R L1 L2 → ∀K2,s,m. ⬇[s, 0, m] L2 ≡ K2 → + ∃∃K1. ⬇[s, 0, m] L1 ≡ K1 & R K1 K2. +*) +(* Basic inversion lemmas ***************************************************) -definition d_liftables1: relation2 lenv term → predicate bool ≝ - λR,s. ∀L,K,cs. ⬇*[s, cs] L ≡ K → - ∀T,U. ⬆*[cs] T ≡ U → R K T → R L U. +fact drops_inv_atom1_aux: ∀X,Y,s,t. ⬇*[s, t] X ≡ Y → X = ⋆ → + Y = ⋆ ∧ (s = Ⓕ → 𝐈⦃t⦄). +#X #Y #s #t * -X -Y -t +[ /3 width=1 by conj/ +| #I #L1 #L2 #V #t #_ #H destruct +| #I #L1 #L2 #V1 #V2 #t #_ #_ #H destruct +] +qed-. -definition d_liftables1_all: relation2 lenv term → predicate bool ≝ - λR,s. ∀L,K,cs. ⬇*[s, cs] L ≡ K → - ∀Ts,Us. ⬆*[cs] Ts ≡ Us → - all … (R K) Ts → all … (R L) Us. +(* Basic_1: includes: drop_gen_sort *) +(* Basic_2A1: includes: drop_inv_atom1 *) +lemma drops_inv_atom1: ∀Y,s,t. ⬇*[s, t] ⋆ ≡ Y → Y = ⋆ ∧ (s = Ⓕ → 𝐈⦃t⦄). +/2 width=3 by drops_inv_atom1_aux/ qed-. + +fact drops_inv_drop1_aux: ∀X,Y,s,t. ⬇*[s, t] X ≡ Y → ∀I,K,V,tl. X = K.ⓑ{I}V → t = Ⓕ@tl → + ⬇*[s, tl] K ≡ Y. +#X #Y #s #t * -X -Y -t +[ #t #Ht #J #K #W #tl #H destruct +| #I #L1 #L2 #V #t #HL #J #K #W #tl #H1 #H2 destruct // +| #I #L1 #L2 #V1 #V2 #t #_ #_ #J #K #W #tl #_ #H2 destruct +] +qed-. -(* Basic inversion lemmas ***************************************************) +(* Basic_1: includes: drop_gen_drop *) +(* Basic_2A1: includes: drop_inv_drop1_lt drop_inv_drop1 *) +lemma drops_inv_drop1: ∀I,K,Y,V,s,t. ⬇*[s, Ⓕ@t] K.ⓑ{I}V ≡ Y → ⬇*[s, t] K ≡ Y. +/2 width=7 by drops_inv_drop1_aux/ qed-. -fact drops_inv_nil_aux: ∀L1,L2,s,cs. ⬇*[s, cs] L1 ≡ L2 → cs = ◊ → L1 = L2. -#L1 #L2 #s #cs * -L1 -L2 -cs // -#L1 #L #L2 #l #m #cs #_ #_ #H destruct -qed-. -(* Basic_1: was: drop1_gen_pnil *) -lemma drops_inv_nil: ∀L1,L2,s. ⬇*[s, ◊] L1 ≡ L2 → L1 = L2. -/2 width=4 by drops_inv_nil_aux/ qed-. - -fact drops_inv_cons_aux: ∀L1,L2,s,cs. ⬇*[s, cs] L1 ≡ L2 → - ∀l,m,tl. cs = {l, m} @ tl → - ∃∃L. ⬇*[s, tl] L1 ≡ L & ⬇[s, l, m] L ≡ L2. -#L1 #L2 #s #cs * -L1 -L2 -cs -[ #L #l #m #tl #H destruct -| #L1 #L #L2 #cs #l #m #HT1 #HT2 #l0 #m0 #tl #H destruct - /2 width=3 by ex2_intro/ +fact drops_inv_skip1_aux: ∀X,Y,s,t. ⬇*[s, t] X ≡ Y → ∀I,K1,V1,tl. X = K1.ⓑ{I}V1 → t = Ⓣ@tl → + ∃∃K2,V2. ⬇*[s, tl] K1 ≡ K2 & ⬆*[tl] V2 ≡ V1 & Y = K2.ⓑ{I}V2. +#X #Y #s #t * -X -Y -t +[ #t #Ht #J #K1 #W1 #tl #H destruct +| #I #L1 #L2 #V #t #_ #J #K1 #W1 #tl #_ #H2 destruct +| #I #L1 #L2 #V1 #V2 #t #HL #HV #J #K1 #W1 #tl #H1 #H2 destruct + /2 width=5 by ex3_2_intro/ ] qed-. -(* Basic_1: was: drop1_gen_pcons *) -lemma drops_inv_cons: ∀L1,L2,s,l,m,cs. ⬇*[s, {l, m} @ cs] L1 ≡ L2 → - ∃∃L. ⬇*[s, cs] L1 ≡ L & ⬇[s, l, m] L ≡ L2. -/2 width=3 by drops_inv_cons_aux/ qed-. - -lemma drops_inv_skip2: ∀I,s,cs,cs2,i. cs ▭ i ≡ cs2 → - ∀L1,K2,V2. ⬇*[s, cs2] L1 ≡ K2. ⓑ{I} V2 → - ∃∃K1,V1,cs1. cs + 1 ▭ i + 1 ≡ cs1 + 1 & - ⬇*[s, cs1] K1 ≡ K2 & - ⬆*[cs1] V2 ≡ V1 & - L1 = K1. ⓑ{I} V1. -#I #s #cs #cs2 #i #H elim H -cs -cs2 -i -[ #i #L1 #K2 #V2 #H - >(drops_inv_nil … H) -L1 /2 width=7 by lifts_nil, minuss_nil, ex4_3_intro, drops_nil/ -| #cs #cs2 #l #m #i #Hil #_ #IHcs2 #L1 #K2 #V2 #H - elim (drops_inv_cons … H) -H #L #HL1 #H - elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ #K #V pluss_SO2 >pluss_SO2 - >yminus_succ2 >ylt_inv_O1 /2 width=1 by ylt_to_minus/ commutative_plus (**) (* (lifts_inv_nil … HV12) -HV12 // -| #L1 #L #L2 #cs #l #m #_ #HL2 #IHL #V1 #V2 #H #I - elim (lifts_inv_cons … H) -H /3 width=5 by drop_skip, drops_cons/ -]. +(* Basic_2A1: includes: drop_FT *) +lemma drops_FT: ∀L1,L2,t. ⬇*[Ⓕ, t] L1 ≡ L2 → ⬇*[Ⓣ, t] L1 ≡ L2. +#L1 #L2 #t #H elim H -L1 -L2 -t +/3 width=1 by drops_atom, drops_drop, drops_skip/ qed. -lemma d1_liftable_liftables: ∀R,s. d_liftable1 R s → d_liftables1 R s. -#R #s #HR #L #K #cs #H elim H -L -K -cs -[ #L #T #U #H #HT <(lifts_inv_nil … H) -H // -| #L1 #L #L2 #cs #l #m #_ #HL2 #IHL #T2 #T1 #H #HLT2 - elim (lifts_inv_cons … H) -H /3 width=10 by/ +(* Basic_2A1: includes: drop_gen *) +lemma drops_gen: ∀L1,L2,s,t. ⬇*[Ⓕ, t] L1 ≡ L2 → ⬇*[s, t] L1 ≡ L2. +#L1 #L2 * /2 width=1 by drops_FT/ +qed-. + +(* Basic_2A1: includes: drop_T *) +lemma drops_T: ∀L1,L2,s,t. ⬇*[s, t] L1 ≡ L2 → ⬇*[Ⓣ, t] L1 ≡ L2. +#L1 #L2 * /2 width=1 by drops_FT/ +qed-. + +(* Basic forward lemmas *****************************************************) + +fact drops_fwd_drop2_aux: ∀X,Y,s,t2. ⬇*[s, t2] X ≡ Y → ∀I,K,V. Y = K.ⓑ{I}V → + ∃∃t1,t. 𝐈⦃t1⦄ & t2 ⊚ Ⓕ@t1 ≡ t & ⬇*[s, t] X ≡ K. +#X #Y #s #t2 #H elim H -X -Y -t2 +[ #t2 #Ht2 #J #K #W #H destruct +| #I #L1 #L2 #V #t2 #_ #IHL #J #K #W #H elim (IHL … H) -IHL + /3 width=5 by after_false, ex3_2_intro, drops_drop/ +| #I #L1 #L2 #V1 #V2 #t2 #HL #_ #_ #J #K #W #H destruct + elim (isid_after_dx t2) /3 width=5 by after_true, ex3_2_intro, drops_drop/ ] -qed. +qed-. -lemma d1_liftables_liftables_all: ∀R,s. d_liftables1 R s → d_liftables1_all R s. -#R #s #HR #L #K #cs #HLK #Ts #Us #H elim H -Ts -Us normalize // -#Ts #Us #T #U #HTU #_ #IHTUs * /3 width=7 by conj/ -qed. +(* Basic_1: includes: drop_S *) +(* Basic_2A1: includes: drop_fwd_drop2 *) +lemma drops_fwd_drop2: ∀I,X,K,V,s,t2. ⬇*[s, t2] X ≡ K.ⓑ{I}V → + ∃∃t1,t. 𝐈⦃t1⦄ & t2 ⊚ Ⓕ@t1 ≡ t & ⬇*[s, t] X ≡ K. +/2 width=5 by drops_fwd_drop2_aux/ qed-. + +fact drops_after_fwd_drop2_aux: ∀X,Y,s,t2. ⬇*[s, t2] X ≡ Y → ∀I,K,V. Y = K.ⓑ{I}V → + ∀t1,t. 𝐈⦃t1⦄ → t2 ⊚ Ⓕ@t1 ≡ t → ⬇*[s, t] X ≡ K. +#X #Y #s #t2 #H elim H -X -Y -t2 +[ #t2 #Ht2 #J #K #W #H destruct +| #I #L1 #L2 #V #t2 #_ #IHL #J #K #W #H #t1 #t #Ht1 #Ht elim (after_inv_false1 … Ht) -Ht + /3 width=3 by drops_drop/ +| #I #L1 #L2 #V1 #V2 #t2 #HL #_ #_ #J #K #W #H #t1 #t #Ht1 #Ht elim (after_inv_true1 … Ht) -Ht + #u1 #u #b #H1 #H2 #Hu destruct >(after_isid_inv_dx … Hu) -Hu /2 width=1 by drops_drop/ +] +qed-. + +lemma drops_after_fwd_drop2: ∀I,X,K,V,s,t2. ⬇*[s, t2] X ≡ K.ⓑ{I}V → + ∀t1,t. 𝐈⦃t1⦄ → t2 ⊚ Ⓕ@t1 ≡ t → ⬇*[s, t] X ≡ K. +/2 width=9 by drops_after_fwd_drop2_aux/ qed-. -(* Basic_1: removed theorems 1: drop1_getl_trans *) +(* Basic_1: includes: drop_gen_refl *) +(* Basic_2A1: includes: drop_inv_O2 *) +lemma drops_fwd_isid: ∀L1,L2,s,t. ⬇*[s, t] L1 ≡ L2 → 𝐈⦃t⦄ → L1 = L2. +#L1 #L2 #s #t #H elim H -L1 -L2 -t // +[ #I #L1 #L2 #V #t #_ #_ #H elim (isid_inv_false … H) +| /5 width=3 by isid_inv_true, lifts_fwd_isid, eq_f3, sym_eq/ +] +qed-. + +(* Basic_2A1: removed theorems 13: + drops_inv_nil drops_inv_cons d1_liftable_liftables + drop_inv_O1_pair1 drop_inv_pair1 drop_inv_O1_pair2 + drop_inv_Y1 drop_Y1 drop_O_Y drop_fwd_Y2 + drop_fwd_length_minus2 drop_fwd_length_minus4 +*) +(* Basic_1: removed theorems 53: + drop1_gen_pnil drop1_gen_pcons drop1_getl_trans + drop_ctail drop_skip_flat + cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf + drop_clear drop_clear_O drop_clear_S + clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r + clear_gen_all clear_clear clear_mono clear_trans clear_ctail clear_cle + getl_ctail_clen getl_gen_tail clear_getl_trans getl_clear_trans + getl_clear_bind getl_clear_conf getl_dec getl_drop getl_drop_conf_lt + getl_drop_conf_ge getl_conf_ge_drop getl_drop_conf_rev + drop_getl_trans_lt drop_getl_trans_le drop_getl_trans_ge + getl_drop_trans getl_flt getl_gen_all getl_gen_sort getl_gen_O + getl_gen_S getl_gen_2 getl_gen_flat getl_gen_bind getl_conf_le + getl_trans getl_refl getl_head getl_flat getl_ctail getl_mono +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_drops.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_drops.ma index c2b0b30d1..34c446f64 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_drops.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_drops.ma @@ -12,14 +12,78 @@ (* *) (**************************************************************************) -include "basic_2/multiple/drops_drop.ma". +include "basic_2/relocation/lifts_lifts.ma". +include "basic_2/relocation/drops.ma". -(* ITERATED LOCAL ENVIRONMENT SLICING ***************************************) +(* GENERAL SLICING FOR LOCAL ENVIRONMENTS ***********************************) (* Main properties **********************************************************) +(* Basic_2A1: includes: drop_conf_ge drop_conf_be drop_conf_le *) +theorem drops_conf: ∀L1,L,s1,t1. ⬇*[s1, t1] L1 ≡ L → + ∀L2,s2,t. ⬇*[s2, t] L1 ≡ L2 → + ∀t2. t1 ⊚ t2 ≡ t → ⬇*[s2, t2] L ≡ L2. +#L1 #L #s1 #t1 #H elim H -L1 -L -t1 +[ #t1 #_ #L2 #s2 #t #H #t2 #Ht12 elim (drops_inv_atom1 … H) -s1 -H + #H #Ht destruct @drops_atom + #H elim (after_inv_isid3 … Ht12) -Ht12 /2 width=1 by/ +| #I #K1 #K #V1 #t1 #_ #IH #L2 #s2 #t #H12 #t2 #Ht elim (after_inv_false1 … Ht) -Ht + #u #H #Hu destruct /3 width=3 by drops_inv_drop1/ +| #I #K1 #K #V1 #V #t1 #_ #HV1 #IH #L2 #s2 #t #H #t2 #Ht elim (after_inv_true1 … Ht) -Ht + #u2 #u * #H1 #H2 #Hu destruct + [ elim (drops_inv_skip1 … H) -H /3 width=6 by drops_skip, lifts_div/ + | /4 width=3 by drops_inv_drop1, drops_drop/ + ] +] +qed-. + (* Basic_1: was: drop1_trans *) -theorem drops_trans: ∀L,L2,s,cs2. ⬇*[s, cs2] L ≡ L2 → ∀L1,cs1. ⬇*[s, cs1] L1 ≡ L → - ⬇*[s, cs2 @@ cs1] L1 ≡ L2. -#L #L2 #s #cs2 #H elim H -L -L2 -cs2 /3 width=3 by drops_cons/ +(* Basic_2A1: includes: drop_trans_ge drop_trans_le drop_trans_ge_comm + drops_drop_trans +*) +theorem drops_trans: ∀L1,L,s1,t1. ⬇*[s1, t1] L1 ≡ L → + ∀L2,s2,t2. ⬇*[s2, t2] L ≡ L2 → + ∀t. t1 ⊚ t2 ≡ t → ⬇*[s1∨s2, t] L1 ≡ L2. +#L1 #L #s1 #t1 #H elim H -L1 -L -t1 +[ #t1 #Ht1 #L2 #s2 #t2 #H #t #Ht elim (drops_inv_atom1 … H) -H + #H #Ht2 destruct @drops_atom #H elim (orb_false_r … H) -H + #H1 #H2 >(after_isid_inv_sn … Ht) -Ht /2 width=1 by/ +| #I #K1 #K #V1 #t1 #_ #IH #L #s2 #t2 #HKL #t #Ht elim (after_inv_false1 … Ht) -Ht + /3 width=3 by drops_drop/ +| #I #K1 #K #V1 #V #t1 #_ #HV1 #IH #L #s2 #t2 #H #t #Ht elim (after_inv_true1 … Ht) -Ht + #u2 #u * #H1 #H2 #Hu destruct + [ elim (drops_inv_skip1 … H) -H /3 width=6 by drops_skip, lifts_trans/ + | /4 width=3 by drops_inv_drop1, drops_drop/ + ] +] +qed-. + +(* Advanced properties ******************************************************) + +(* Basic_2A1: includes: drop_mono *) +lemma drops_mono: ∀L,L1,s1,t. ⬇*[s1, t] L ≡ L1 → + ∀L2,s2. ⬇*[s2, t] L ≡ L2 → L1 = L2. +#L #L1 #s1 #t elim (isid_after_dx t) +/3 width=8 by drops_conf, drops_fwd_isid/ +qed-. + +(* Basic_2A1: includes: drop_conf_lt *) +lemma drops_conf_skip1: ∀L,L2,s2,t. ⬇*[s2, t] L ≡ L2 → + ∀I,K1,V1,s1,t1. ⬇*[s1, t1] L ≡ K1.ⓑ{I}V1 → + ∀t2. t1 ⊚ Ⓣ@t2 ≡ t → + ∃∃K2,V2. L2 = K2.ⓑ{I}V2 & + ⬇*[s2, t2] K1 ≡ K2 & ⬆*[t2] V2 ≡ V1. +#L #L2 #s2 #t #H2 #I #K1 #V1 #s1 #t1 #H1 #t2 #Ht lapply (drops_conf … H1 … H2 … Ht) -L -Ht +#H elim (drops_inv_skip1 … H) -H /2 width=5 by ex3_2_intro/ +qed-. + +(* Basic_2A1: includes: drop_trans_lt *) +lemma drops_trans_skip2: ∀L1,L,s1,t1. ⬇*[s1, t1] L1 ≡ L → + ∀I,K2,V2,s2,t2. ⬇*[s2, t2] L ≡ K2.ⓑ{I}V2 → + ∀t. t1 ⊚ t2 ≡ Ⓣ@t → + ∃∃K1,V1. L1 = K1.ⓑ{I}V1 & + ⬇*[s1∨s2, t] K1 ≡ K2 & ⬆*[t] V2 ≡ V1. +#L1 #L #s1 #t1 #H1 #I #K2 #V2 #s2 #t2 #H2 #t #Ht +lapply (drops_trans … H1 … H2 … Ht) -L -Ht +#H elim (drops_inv_skip2 … H) -H /2 width=5 by ex3_2_intro/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lstar.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lstar.ma new file mode 100644 index 000000000..5851e809e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_lstar.ma @@ -0,0 +1,82 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "ground_2/lib/lstar.ma". +include "basic_2/relocation/drops.ma". + +(* GENERAL SLICING FOR LOCAL ENVIRONMENTS ***********************************) + +(* Properties on reflexive and transitive closure ***************************) + +(* Basic_2A1: was: d_liftable_LTC *) +lemma d2_liftable_LTC: ∀R. d_liftable2 R → d_liftable2 (LTC … R). +#R #HR #K #T1 #T2 #H elim H -T2 +[ #T2 #HT12 #L #s #t #HLK #U1 #HTU1 + elim (HR … HT12 … HLK … HTU1) /3 width=3 by inj, ex2_intro/ +| #T #T2 #_ #HT2 #IHT1 #L #s #t #HLK #U1 #HTU1 + elim (IHT1 … HLK … HTU1) -T1 #U #HTU #HU1 + elim (HR … HT2 … HLK … HTU) -HR -K -T /3 width=5 by step, ex2_intro/ +] +qed-. + +(* Basic_2A1: was: d_deliftable_sn_LTC *) +lemma d2_deliftable_sn_LTC: ∀R. d_deliftable2_sn R → d_deliftable2_sn (LTC … R). +#R #HR #L #U1 #U2 #H elim H -U2 +[ #U2 #HU12 #K #s #t #HLK #T1 #HTU1 + elim (HR … HU12 … HLK … HTU1) -HR -L -U1 /3 width=3 by inj, ex2_intro/ +| #U #U2 #_ #HU2 #IHU1 #K #s #t #HLK #T1 #HTU1 + elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1 + elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by step, ex2_intro/ +] +qed-. + +lemma dropable_sn_TC: ∀R. dropable_sn R → dropable_sn (TC … R). +#R #HR #L1 #K1 #s #t #HLK1 #L2 #H elim H -L2 +[ #L2 #HL12 elim (HR … HLK1 … HL12) -HR -L1 + /3 width=3 by inj, ex2_intro/ +| #L #L2 #_ #HL2 * #K #HK1 #HLK elim (HR … HLK … HL2) -HR -L + /3 width=3 by step, ex2_intro/ +] +qed-. +(* +lemma dropable_dx_TC: ∀R. dropable_dx R → dropable_dx (TC … R). +#R #HR #L1 #L2 #H elim H -L2 +[ #L2 #HL12 #K2 #s #m #HLK2 elim (HR … HL12 … HLK2) -HR -L2 + /3 width=3 by inj, ex2_intro/ +| #L #L2 #_ #HL2 #IHL1 #K2 #s #m #HLK2 elim (HR … HL2 … HLK2) -HR -L2 + #K #HLK #HK2 elim (IHL1 … HLK) -L + /3 width=5 by step, ex2_intro/ +] +qed-. +*) +(* Basic_2A1: was: d_liftable_llstar *) +lemma d2_liftable_llstar: ∀R. d_liftable2 R → ∀d. d_liftable2 (llstar … R d). +#R #HR #d #K #T1 #T2 #H @(lstar_ind_r … d T2 H) -d -T2 +[ #L #s #t #_ #U1 #HTU1 -HR -K -s /2 width=3 by ex2_intro/ +| #d #T #T2 #_ #HT2 #IHT1 #L #s #t #HLK #U1 #HTU1 + elim (IHT1 … HLK … HTU1) -T1 #U #HTU #HU1 + elim (HR … HT2 … HLK … HTU) -T /3 width=5 by lstar_dx, ex2_intro/ +] +qed-. + +(* Basic_2A1: was: d_deliftable_sn_llstar *) +lemma d2_deliftable_sn_llstar: ∀R. d_deliftable2_sn R → + ∀d. d_deliftable2_sn (llstar … R d). +#R #HR #d #L #U1 #U2 #H @(lstar_ind_r … d U2 H) -d -U2 +[ /2 width=3 by lstar_O, ex2_intro/ +| #d #U #U2 #_ #HU2 #IHU1 #K #s #t #HLK #T1 #HTU1 + elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1 + elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by lstar_dx, ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_vector.ma new file mode 100644 index 000000000..63bc71da6 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_vector.ma @@ -0,0 +1,31 @@ +(**************************************************************************) +(* ___ *) +(* ||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/relocation/lifts_vector.ma". +include "basic_2/relocation/drops.ma". + +(* GENERAL SLICING FOR LOCAL ENVIRONMENTS ***********************************) + +definition d_liftable1_all: relation2 lenv term → predicate bool ≝ + λR,s. ∀L,K,t. ⬇*[s, t] L ≡ K → + ∀Ts,Us. ⬆*[t] Ts ≡ Us → + all … (R K) Ts → all … (R L) Us. + +(* Properties on general relocation for term vectors ************************) + +(* Basic_2A1: was: d1_liftables_liftables_all *) +lemma d1_liftable_liftable_all: ∀R,s. d_liftable1 R s → d_liftable1_all R s. +#R #s #HR #L #K #t #HLK #Ts #Us #H elim H -Ts -Us normalize // +#Ts #Us #T #U #HTU #_ #IHTUs * /3 width=7 by conj/ +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_weight.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_weight.ma new file mode 100644 index 000000000..0db32e2a2 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/drops_weight.ma @@ -0,0 +1,50 @@ +(**************************************************************************) +(* ___ *) +(* ||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/grammar/cl_restricted_weight.ma". +include "basic_2/relocation/lifts_weight.ma". +include "basic_2/relocation/drops.ma". + +(* GENERAL SLICING FOR LOCAL ENVIRONMENTS ***********************************) + +(* Forward lemmas on weight for local environments **************************) + +(* Basic_2A1: includes: drop_fwd_lw *) +lemma drops_fwd_lw: ∀L1,L2,s,t. ⬇*[s, t] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}. +#L1 #L2 #s #t #H elim H -L1 -L2 -t // +[ /2 width=3 by transitive_le/ +| #I #L1 #L2 #V1 #V2 #t #_ #HV21 #IHL12 normalize + >(lifts_fwd_tw … HV21) -HV21 /2 width=1 by monotonic_le_plus_l/ +] +qed-. + +(* Basic_2A1: includes: drop_fwd_lw_lt *) +(* Note: 𝐈⦃t⦄ → ⊥ is ∥l∥ < |L| *) +lemma drops_fwd_lw_lt: ∀L1,L2,t. ⬇*[Ⓕ, t] L1 ≡ L2 → + (𝐈⦃t⦄ → ⊥) → ♯{L2} < ♯{L1}. +#L1 #L2 #t #H elim H -L1 -L2 -t +[ #t #Ht #Hnt elim Hnt -Hnt /2 width=1 by/ +| /3 width=3 by drops_fwd_lw, le_to_lt_to_lt/ +| #I #L1 #L2 #V1 #V2 #t #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I + >(lifts_fwd_tw … HV21) -V2 /4 width=1 by monotonic_lt_plus_l/ +] +qed-. + +(* Forward lemmas on restricted weight for closures *************************) + +(* Basic_2A1: includes: drop_fwd_rfw *) +lemma drops_pair2_fwd_rfw: ∀I,L,K,V,s,t. ⬇*[s, t] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}. +#I #L #K #V #s #t #HLK lapply (drops_fwd_lw … HLK) -HLK +normalize in ⊢ (%→?→?%%); /3 width=3 by le_to_lt_to_lt/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy.ma deleted file mode 100644 index cd16f5042..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy.ma +++ /dev/null @@ -1,297 +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 "ground_2/ynat/ynat_max.ma". -include "basic_2/notation/relations/psubst_6.ma". -include "basic_2/grammar/genv.ma". -include "basic_2/substitution/lsuby.ma". - -(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) - -(* activate genv *) -inductive cpy: ynat → ynat → relation4 genv lenv term term ≝ -| cpy_atom : ∀I,G,L,l,m. cpy l m G L (⓪{I}) (⓪{I}) -| cpy_subst: ∀I,G,L,K,V,W,i,l,m. l ≤ yinj i → i < l+m → - ⬇[i] L ≡ K.ⓑ{I}V → ⬆[0, i+1] V ≡ W → cpy l m G L (#i) W -| cpy_bind : ∀a,I,G,L,V1,V2,T1,T2,l,m. - cpy l m G L V1 V2 → cpy (⫯l) m G (L.ⓑ{I}V1) T1 T2 → - cpy l m G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) -| cpy_flat : ∀I,G,L,V1,V2,T1,T2,l,m. - cpy l m G L V1 V2 → cpy l m G L T1 T2 → - cpy l m G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) -. - -interpretation "context-sensitive extended ordinary substritution (term)" - 'PSubst G L T1 l m T2 = (cpy l m G L T1 T2). - -(* Basic properties *********************************************************) - -lemma lsuby_cpy_trans: ∀G,l,m. lsub_trans … (cpy l m G) (lsuby l m). -#G #l #m #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -l -m -[ // -| #I #G #L1 #K1 #V #W #i #l #m #Hli #Hilm #HLK1 #HVW #L2 #HL12 - elim (lsuby_drop_trans_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-. - -lemma cpy_refl: ∀G,T,L,l,m. ⦃G, L⦄ ⊢ T ▶[l, m] T. -#G #T elim T -T // * /2 width=1 by cpy_bind, cpy_flat/ -qed. - -(* Basic_1: was: subst1_ex *) -lemma cpy_full: ∀I,G,K,V,T1,L,l. ⬇[l] L ≡ K.ⓑ{I}V → - ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ▶[l, 1] T2 & ⬆[l, 1] T ≡ T2. -#I #G #K #V #T1 elim T1 -T1 -[ * #i #L #l #HLK - /2 width=4 by lift_sort, lift_gref, ex2_2_intro/ - elim (lt_or_eq_or_gt i l) #Hil - /4 width=4 by lift_lref_ge_minus, lift_lref_lt, ex2_2_intro, ylt_inj, yle_inj/ (**) (* was /3 width=4/ without inj *) - destruct - elim (lift_total V 0 (i+1)) #W #HVW - elim (lift_split … HVW i i) - /4 width=5 by cpy_subst, ylt_inj, ex2_2_intro/ -| * [ #a ] #J #W1 #U1 #IHW1 #IHU1 #L #l #HLK - elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 - [ elim (IHU1 (L.ⓑ{J}W1) (l+1)) -IHU1 - /3 width=9 by cpy_bind, drop_drop, 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,l1,m1. ⦃G, L⦄ ⊢ T1 ▶[l1, m1] T2 → - ∀l2,m2. l2 ≤ l1 → l1 + m1 ≤ l2 + m2 → - ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T2. -#G #L #T1 #T2 #l1 #m1 #H elim H -G -L -T1 -T2 -l1 -m1 // -[ /3 width=5 by cpy_subst, ylt_yle_trans, yle_trans/ -| /4 width=3 by cpy_bind, ylt_yle_trans, yle_succ/ -| /3 width=1 by cpy_flat/ -] -qed-. - -lemma cpy_weak_top: ∀G,L,T1,T2,l,m. - ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶[l, |L| - l] T2. -#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m // -[ #I #G #L #K #V #W #i #l #m #Hli #_ #HLK #HVW - lapply (drop_fwd_length_lt2 … HLK) - /4 width=5 by cpy_subst, ylt_yle_trans, ylt_inj/ -| #a #I #G #L #V1 #V2 normalize in match (|L.ⓑ{I}V2|); (**) (* |?| does not work *) - yplus_SO2 >yminus_succ - /2 width=1 by cpy_bind/ -| /2 width=1 by cpy_flat/ -] -qed-. - -lemma cpy_weak_full: ∀G,L,T1,T2,l,m. - ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ⦃G, L⦄ ⊢ T1 ▶[0, |L|] T2. -#G #L #T1 #T2 #l #m #HT12 -lapply (cpy_weak … HT12 0 (l + m) ? ?) -HT12 -/2 width=2 by cpy_weak_top/ -qed-. - -lemma cpy_split_up: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ∀i. i ≤ l + m → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l, i-l] T & ⦃G, L⦄ ⊢ T ▶[i, l+m-i] T2. -#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m -[ /2 width=3 by ex2_intro/ -| #I #G #L #K #V #W #i #l #m #Hli #Hilm #HLK #HVW #j #Hjlm - elim (ylt_split i j) [ -Hilm -Hjlm | -Hli ] - /4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/ -| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #Hilm - elim (IHV12 i) -IHV12 // #V - elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hilm - >yplus_SO2 >yplus_succ1 #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 - /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ -| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #Hilm - elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hilm - /3 width=5 by ex2_intro, cpy_flat/ -] -qed-. - -lemma cpy_split_down: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ∀i. i ≤ l + m → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[i, l+m-i] T & ⦃G, L⦄ ⊢ T ▶[l, i-l] T2. -#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m -[ /2 width=3 by ex2_intro/ -| #I #G #L #K #V #W #i #l #m #Hli #Hilm #HLK #HVW #j #Hjlm - elim (ylt_split i j) [ -Hilm -Hjlm | -Hli ] - /4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/ -| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #Hilm - elim (IHV12 i) -IHV12 // #V - elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hilm - >yplus_SO2 >yplus_succ1 #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 - /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ -| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #i #Hilm - elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hilm - /3 width=5 by ex2_intro, cpy_flat/ -] -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma cpy_fwd_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → - ∀T1,l,m. ⬆[l, m] T1 ≡ U1 → - l ≤ lt → l + m ≤ lt + mt → - ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶[l+m, lt+mt-(l+m)] U2 & ⬆[l, m] T2 ≡ U2. -#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt -[ * #i #G #L #lt #mt #T1 #l #m #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ - | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ - ] -| #I #G #L #K #V #W #i #lt #mt #Hlti #Hilmt #HLK #HVW #T1 #l #m #H #Hllt #Hlmlmt - elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ -V -Hilmt -Hlmlmt | -Hlti -Hllt ] - [ elim (ylt_yle_false … Hllt) -Hllt /3 width=3 by yle_ylt_trans, ylt_inj/ - | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi - elim (lift_split … HVW l (i-m+1) ? ? ?) [2,3,4: /2 width=1 by yle_succ_dx, le_S_S/ ] -Hlim - #T2 #_ >plus_minus /2 width=1 by yle_inv_inj/ ymax_pre_sn_comm // (**) (* explicit constructor *) - ] -| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #X #l #m #H #Hllt #Hlmlmt - elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HVW1) -V1 -IHW12 // - elim (IHU12 … HTU1) -T1 -IHU12 /2 width=1 by yle_succ/ - yplus_SO2 >yplus_succ1 >yplus_succ1 - /3 width=2 by cpy_bind, lift_bind, ex2_intro/ -| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #X #l #m #H #Hllt #Hlmlmt - elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HVW1) -V1 -IHW12 // elim (IHU12 … HTU1) -T1 -IHU12 - /3 width=2 by cpy_flat, lift_flat, ex2_intro/ -] -qed-. - -lemma cpy_fwd_tw: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ♯{T1} ≤ ♯{T2}. -#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m normalize -/3 width=1 by monotonic_le_plus_l, le_plus/ -qed-. - -(* Basic inversion lemmas ***************************************************) - -fact cpy_inv_atom1_aux: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → ∀J. T1 = ⓪{J} → - T2 = ⓪{J} ∨ - ∃∃I,K,V,i. l ≤ yinj i & i < l + m & - ⬇[i] L ≡ K.ⓑ{I}V & - ⬆[O, i+1] V ≡ T2 & - J = LRef i. -#G #L #T1 #T2 #l #m * -G -L -T1 -T2 -l -m -[ #I #G #L #l #m #J #H destruct /2 width=1 by or_introl/ -| #I #G #L #K #V #T2 #i #l #m #Hli #Hilm #HLK #HVT2 #J #H destruct /3 width=9 by ex5_4_intro, or_intror/ -| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #H destruct -| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #H destruct -] -qed-. - -lemma cpy_inv_atom1: ∀I,G,L,T2,l,m. ⦃G, L⦄ ⊢ ⓪{I} ▶[l, m] T2 → - T2 = ⓪{I} ∨ - ∃∃J,K,V,i. l ≤ yinj i & i < l + m & - ⬇[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: ∀G,L,T2,k,l,m. ⦃G, L⦄ ⊢ ⋆k ▶[l, m] T2 → T2 = ⋆k. -#G #L #T2 #k #l #m #H -elim (cpy_inv_atom1 … H) -H // -* #I #K #V #i #_ #_ #_ #_ #H destruct -qed-. - -(* Basic_1: was: subst1_gen_lref *) -lemma cpy_inv_lref1: ∀G,L,T2,i,l,m. ⦃G, L⦄ ⊢ #i ▶[l, m] T2 → - T2 = #i ∨ - ∃∃I,K,V. l ≤ i & i < l + m & - ⬇[i] L ≡ K.ⓑ{I}V & - ⬆[O, i+1] V ≡ T2. -#G #L #T2 #i #l #m #H -elim (cpy_inv_atom1 … H) -H /2 width=1 by or_introl/ -* #I #K #V #j #Hlj #Hjlm #HLK #HVT2 #H destruct /3 width=5 by ex4_3_intro, or_intror/ -qed-. - -lemma cpy_inv_gref1: ∀G,L,T2,p,l,m. ⦃G, L⦄ ⊢ §p ▶[l, m] T2 → T2 = §p. -#G #L #T2 #p #l #m #H -elim (cpy_inv_atom1 … H) -H // -* #I #K #V #i #_ #_ #_ #_ #H destruct -qed-. - -fact cpy_inv_bind1_aux: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → - ∀a,I,V1,T1. U1 = ⓑ{a,I}V1.T1 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 & - ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ▶[⫯l, m] T2 & - U2 = ⓑ{a,I}V2.T2. -#G #L #U1 #U2 #l #m * -G -L -U1 -U2 -l -m -[ #I #G #L #l #m #b #J #W1 #U1 #H destruct -| #I #G #L #K #V #W #i #l #m #_ #_ #_ #_ #b #J #W1 #U1 #H destruct -| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #HV12 #HT12 #b #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ -| #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #b #J #W1 #U1 #H destruct -] -qed-. - -lemma cpy_inv_bind1: ∀a,I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓑ{a,I} V1. T1 ▶[l, m] U2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 & - ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶[⫯l, m] T2 & - U2 = ⓑ{a,I}V2.T2. -/2 width=3 by cpy_inv_bind1_aux/ qed-. - -fact cpy_inv_flat1_aux: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → - ∀I,V1,T1. U1 = ⓕ{I}V1.T1 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 & - ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 & - U2 = ⓕ{I}V2.T2. -#G #L #U1 #U2 #l #m * -G -L -U1 -U2 -l -m -[ #I #G #L #l #m #J #W1 #U1 #H destruct -| #I #G #L #K #V #W #i #l #m #_ #_ #_ #_ #J #W1 #U1 #H destruct -| #a #I #G #L #V1 #V2 #T1 #T2 #l #m #_ #_ #J #W1 #U1 #H destruct -| #I #G #L #V1 #V2 #T1 #T2 #l #m #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -lemma cpy_inv_flat1: ∀I,G,L,V1,T1,U2,l,m. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ▶[l, m] U2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[l, m] V2 & - ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 & - U2 = ⓕ{I}V2.T2. -/2 width=3 by cpy_inv_flat1_aux/ qed-. - - -fact cpy_inv_refl_O2_aux: ∀G,L,T1,T2,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T2 → m = 0 → T1 = T2. -#G #L #T1 #T2 #l #m #H elim H -G -L -T1 -T2 -l -m -[ // -| #I #G #L #K #V #W #i #l #m #Hli #Hilm #_ #_ #H destruct - elim (ylt_yle_false … Hli) -Hli // -| /3 width=1 by eq_f2/ -| /3 width=1 by eq_f2/ -] -qed-. - -lemma cpy_inv_refl_O2: ∀G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 ▶[l, 0] T2 → T1 = T2. -/2 width=6 by cpy_inv_refl_O2_aux/ qed-. - -(* Basic_1: was: subst1_gen_lift_eq *) -lemma cpy_inv_lift1_eq: ∀G,T1,U1,l,m. ⬆[l, m] T1 ≡ U1 → - ∀L,U2. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → U1 = U2. -#G #T1 #U1 #l #m #HTU1 #L #U2 #HU12 elim (cpy_fwd_up … HU12 … HTU1) -HU12 -HTU1 -/2 width=4 by cpy_inv_refl_O2/ -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/substitution/cpy_cpy.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy_cpy.ma deleted file mode 100644 index 7a65a2db5..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy_cpy.ma +++ /dev/null @@ -1,123 +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/cpy_lift.ma". - -(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) - -(* Main properties **********************************************************) - -(* Basic_1: was: subst1_confluence_eq *) -theorem cpy_conf_eq: ∀G,L,T0,T1,l1,m1. ⦃G, L⦄ ⊢ T0 ▶[l1, m1] T1 → - ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶[l2, m2] T2 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L⦄ ⊢ T2 ▶[l1, m1] T. -#G #L #T0 #T1 #l1 #m1 #H elim H -G -L -T0 -T1 -l1 -m1 -[ /2 width=3 by ex2_intro/ -| #I1 #G #L #K1 #V1 #T1 #i0 #l1 #m1 #Hl1 #Hlm1 #HLK1 #HVT1 #T2 #l2 #m2 #H - elim (cpy_inv_lref1 … H) -H - [ #HX destruct /3 width=7 by cpy_subst, ex2_intro/ - | -Hl1 -Hlm1 * #I2 #K2 #V2 #_ #_ #HLK2 #HVT2 - lapply (drop_mono … HLK1 … HLK2) -HLK1 -HLK2 #H destruct - >(lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3 by ex2_intro/ - ] -| #a #I #G #L #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #X #l2 #m2 #HX - elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02) -IHV01 -HV02 #V #HV1 #HV2 - elim (IHT01 … HT02) -T0 #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT1 (L.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/ - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V2) ?) -HT2 - /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/ -| #I #G #L #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #X #l2 #m2 #HX - elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02) -V0 - elim (IHT01 … HT02) -T0 /3 width=5 by cpy_flat, ex2_intro/ -] -qed-. - -(* Basic_1: was: subst1_confluence_neq *) -theorem cpy_conf_neq: ∀G,L1,T0,T1,l1,m1. ⦃G, L1⦄ ⊢ T0 ▶[l1, m1] T1 → - ∀L2,T2,l2,m2. ⦃G, L2⦄ ⊢ T0 ▶[l2, m2] T2 → - (l1 + m1 ≤ l2 ∨ l2 + m2 ≤ l1) → - ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L1⦄ ⊢ T2 ▶[l1, m1] T. -#G #L1 #T0 #T1 #l1 #m1 #H elim H -G -L1 -T0 -T1 -l1 -m1 -[ /2 width=3 by ex2_intro/ -| #I1 #G #L1 #K1 #V1 #T1 #i0 #l1 #m1 #Hl1 #Hlm1 #HLK1 #HVT1 #L2 #T2 #l2 #m2 #H1 #H2 - elim (cpy_inv_lref1 … H1) -H1 - [ #H destruct /3 width=7 by cpy_subst, ex2_intro/ - | -HLK1 -HVT1 * #I2 #K2 #V2 #Hl2 #Hlm2 #_ #_ elim H2 -H2 #Hlml [ -Hl1 -Hlm2 | -Hl2 -Hlm1 ] - [ elim (ylt_yle_false … Hlm1) -Hlm1 /2 width=3 by yle_trans/ - | elim (ylt_yle_false … Hlm2) -Hlm2 /2 width=3 by yle_trans/ - ] - ] -| #a #I #G #L1 #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #L2 #X #l2 #m2 #HX #H - elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02 H) -IHV01 -HV02 #V #HV1 #HV2 - elim (IHT01 … HT02) -T0 - [ -H #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT1 (L2.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/ - lapply (lsuby_cpy_trans … HT2 (L1.ⓑ{I}V2) ?) -HT2 /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/ - | -HV1 -HV2 elim H -H /3 width=1 by yle_succ, or_introl, or_intror/ - ] -| #I #G #L1 #V0 #V1 #T0 #T1 #l1 #m1 #_ #_ #IHV01 #IHT01 #L2 #X #l2 #m2 #HX #H - elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02 H) -V0 - elim (IHT01 … HT02 H) -T0 -H /3 width=5 by cpy_flat, ex2_intro/ -] -qed-. - -(* Note: the constant 1 comes from cpy_subst *) -(* Basic_1: was: subst1_trans *) -theorem cpy_trans_ge: ∀G,L,T1,T0,l,m. ⦃G, L⦄ ⊢ T1 ▶[l, m] T0 → - ∀T2. ⦃G, L⦄ ⊢ T0 ▶[l, 1] T2 → 1 ≤ m → ⦃G, L⦄ ⊢ T1 ▶[l, m] T2. -#G #L #T1 #T0 #l #m #H elim H -G -L -T1 -T0 -l -m -[ #I #G #L #l #m #T2 #H #Hm - elim (cpy_inv_atom1 … H) -H - [ #H destruct // - | * #J #K #V #i #Hl2i #Hilm2 #HLK #HVT2 #H destruct - lapply (ylt_yle_trans … (l+m) … Hilm2) - /2 width=5 by cpy_subst, monotonic_yle_plus_sn/ - ] -| #I #G #L #K #V #V2 #i #l #m #Hli #Hilm #HLK #HVW #T2 #HVT2 #Hm - lapply (cpy_weak … HVT2 0 (i+1) ? ?) -HVT2 /3 width=1 by yle_plus_dx2_trans, yle_succ/ - >yplus_inj #HVT2 <(cpy_inv_lift1_eq … HVW … HVT2) -HVT2 /2 width=5 by cpy_subst/ -| #a #I #G #L #V1 #V0 #T1 #T0 #l #m #_ #_ #IHV10 #IHT10 #X #H #Hm - elim (cpy_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct - lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02 - lapply (IHT10 … HT02 Hm) -T0 /3 width=1 by cpy_bind/ -| #I #G #L #V1 #V0 #T1 #T0 #l #m #_ #_ #IHV10 #IHT10 #X #H #Hm - elim (cpy_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1 by cpy_flat/ -] -qed-. - -theorem cpy_trans_down: ∀G,L,T1,T0,l1,m1. ⦃G, L⦄ ⊢ T1 ▶[l1, m1] T0 → - ∀T2,l2,m2. ⦃G, L⦄ ⊢ T0 ▶[l2, m2] T2 → l2 + m2 ≤ l1 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[l2, m2] T & ⦃G, L⦄ ⊢ T ▶[l1, m1] T2. -#G #L #T1 #T0 #l1 #m1 #H elim H -G -L -T1 -T0 -l1 -m1 -[ /2 width=3 by ex2_intro/ -| #I #G #L #K #V #W #i1 #l1 #m1 #Hli1 #Hilm1 #HLK #HVW #T2 #l2 #m2 #HWT2 #Hlm2l1 - lapply (yle_trans … Hlm2l1 … Hli1) -Hlm2l1 #Hlm2i1 - lapply (cpy_weak … HWT2 0 (i1+1) ? ?) -HWT2 /3 width=1 by yle_succ, yle_pred_sn/ -Hlm2i1 - >yplus_inj #HWT2 <(cpy_inv_lift1_eq … HVW … HWT2) -HWT2 /3 width=9 by cpy_subst, ex2_intro/ -| #a #I #G #L #V1 #V0 #T1 #T0 #l1 #m1 #_ #_ #IHV10 #IHT10 #X #l2 #m2 #HX #lm2l1 - elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02 - elim (IHV10 … HV02) -IHV10 -HV02 // #V - elim (IHT10 … HT02) -T0 /2 width=1 by yle_succ/ #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 /3 width=6 by cpy_bind, lsuby_succ, ex2_intro/ -| #I #G #L #V1 #V0 #T1 #T0 #l1 #m1 #_ #_ #IHV10 #IHT10 #X #l2 #m2 #HX #lm2l1 - elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV10 … HV02) -V0 // - elim (IHT10 … HT02) -T0 /3 width=6 by cpy_flat, ex2_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy_lift.ma deleted file mode 100644 index 44a7f3d76..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy_lift.ma +++ /dev/null @@ -1,253 +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/drop_drop.ma". -include "basic_2/substitution/cpy.ma". - -(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) - -(* Properties on relocation *************************************************) - -(* Basic_1: was: subst1_lift_lt *) -lemma cpy_lift_le: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 → - ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K → - ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 → - lt + mt ≤ l → ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2. -#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt -[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hlmtl - lapply (ylt_yle_trans … Hlmtl … Hilmt) -Hlmtl #Hil - lapply (lift_inv_lref1_lt … H … Hil) -H #H destruct - elim (lift_trans_ge … HVW … HWU2) -W /2 width=1 by ylt_fwd_le_succ1/ - yplus_SO2 >yminus_succ2 #W #HVW #HWU2 - elim (drop_trans_le … HLK … HKV) -K /2 width=2 by ylt_fwd_le/ #X #HLK #H - elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K #Y #_ #HVY - >(lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=5 by cpy_subst/ -| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - /4 width=7 by cpy_bind, drop_skip, yle_succ/ -| #G #I #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl - 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=7 by cpy_flat/ -] -qed-. - -lemma cpy_lift_be: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 → - ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K → - ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 → - lt ≤ l → l ≤ lt + mt → ⦃G, L⦄ ⊢ U1 ▶[lt, mt + m] U2. -#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt -[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_ #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hltl #_ - elim (lift_inv_lref1 … H) -H * #Hil #H destruct - [ -Hltl - lapply (ylt_yle_trans … (lt+mt+m) … Hilmt) // -Hilmt #Hilmtm - elim (lift_trans_ge … HVW … HWU2) -W yplus_SO2 - [2: >yplus_O1 /2 width=1 by ylt_fwd_le_succ1/ ] >yminus_succ2 #W #HVW #HWU2 - elim (drop_trans_le … HLK … HKV) -K /2 width=1 by ylt_fwd_le/ #X #HLK #H - elim (drop_inv_skip2 … H) -H /2 width=1 by ylt_to_minus/ -Hil #K #Y #_ #HVY - >(lift_mono … HVY … HVW) -V #H destruct /2 width=5 by cpy_subst/ - | -Hlti - lapply (yle_trans … Hltl … Hil) -Hltl #Hlti - lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2 - lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hil - /3 width=5 by cpy_subst, drop_inv_gen, monotonic_ylt_plus_dx, yle_plus_dx1_trans/ - ] -| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hltl #Hllmt - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - /4 width=7 by cpy_bind, drop_skip, yle_succ/ -| #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hlmtl - 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=7 by cpy_flat/ -] -qed-. - -(* Basic_1: was: subst1_lift_ge *) -lemma cpy_lift_ge: ∀G,K,T1,T2,lt,mt. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 → - ∀L,U1,U2,s,l,m. ⬇[s, l, m] L ≡ K → - ⬆[l, m] T1 ≡ U1 → ⬆[l, m] T2 ≡ U2 → - l ≤ lt → ⦃G, L⦄ ⊢ U1 ▶[lt+m, mt] U2. -#G #K #T1 #T2 #lt #mt #H elim H -G -K -T1 -T2 -lt -mt -[ #I #G #K #lt #mt #L #U1 #U2 #s #l #m #_ #H1 #H2 #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #I #G #K #KV #V #W #i #lt #mt #Hlti #Hilmt #HKV #HVW #L #U1 #U2 #s #l #m #HLK #H #HWU2 #Hllt - lapply (yle_trans … Hllt … Hlti) -Hllt #Hil - lapply (lift_inv_lref1_ge … H … Hil) -H #H destruct - lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by yle_succ_dx/ >plus_plus_comm_23 #HVU2 - lapply (drop_trans_ge_comm … HLK … HKV ?) -K // -Hil - /3 width=5 by cpy_subst, drop_inv_gen, monotonic_ylt_plus_dx, monotonic_yle_plus_dx/ -| #a #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hllt - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - /4 width=6 by cpy_bind, drop_skip, yle_succ/ -| #I #G #K #V1 #V2 #T1 #T2 #lt #mt #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #l #m #HLK #H1 #H2 #Hllt - 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 by cpy_flat/ -] -qed-. - -(* Inversion lemmas on relocation *******************************************) - -(* Basic_1: was: subst1_gen_lift_lt *) -lemma cpy_inv_lift1_le: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → - ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → - lt + mt ≤ l → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, mt] T2 & ⬆[l, m] T2 ≡ U2. -#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt -[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ - | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ - ] -| #I #G #L #KV #V #W #i #lt #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hlmtl - lapply (ylt_yle_trans … Hlmtl … Hilmt) -Hlmtl #Hil - lapply (lift_inv_lref2_lt … H … Hil) -H #H destruct - elim (drop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV - elim (lift_trans_le … HUV … HVW) -V // >minus_plus yplus_SO2 >ymax_pre_sn /2 width=1 by ylt_fwd_le_succ1/ -Hil - /3 width=5 by cpy_subst, ex2_intro/ -| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl - elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2 - elim (IHU12 … HTU1) -IHU12 -HTU1 - /3 width=6 by cpy_bind, yle_succ, drop_skip, lift_bind, ex2_intro/ -| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl - elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HLK … HVW1) -W1 // - elim (IHU12 … HLK … HTU1) -U1 -HLK - /3 width=5 by cpy_flat, lift_flat, ex2_intro/ -] -qed-. - -lemma cpy_inv_lift1_be: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → - ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → - lt ≤ l → l + m ≤ lt + mt → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, mt-m] T2 & ⬆[l, m] T2 ≡ U2. -#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt -[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_ #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ - | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ - ] -| #I #G #L #KV #V #W #i #x #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hltl #Hlmlmt - elim (yle_inv_inj2 … Hlti) -Hlti #lt #Hlti #H destruct - lapply (yle_fwd_plus_yge … Hltl Hlmlmt) #Hmmt - elim (lift_inv_lref2 … H) -H * #Hil #H destruct [ -Hltl -Hilmt | -Hlti -Hlmlmt ] - [ lapply (ylt_yle_trans i l (lt+(mt-m)) ? ?) // - [ >yplus_minus_assoc_inj /2 width=1 by yle_plus1_to_minus_inj2/ ] -Hlmlmt #Hilmtm - elim (drop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV - elim (lift_trans_le … HUV … HVW) -V // - yplus_SO2 >ymax_pre_sn /2 width=1 by ylt_fwd_le_succ1/ -Hil - /4 width=5 by cpy_subst, ex2_intro, yle_inj/ - | elim (yle_inv_plus_inj2 … Hil) #Hlim #Hmi - lapply (yle_inv_inj … Hmi) -Hmi #Hmi - lapply (yle_trans … Hltl (i-m) ?) // -Hltl #Hltim - lapply (drop_conf_ge … HLK … HLKV ?) -L // #HKV - elim (lift_split … HVW l (i-m+1)) -HVW [2,3,4: /2 width=1 by yle_succ_dx, le_S_S/ ] -Hil -Hlim - #V1 #HV1 >plus_minus // yplus_minus_assoc_inj /3 width=1 by monotonic_ylt_minus_dx, yle_inj/ - ] -| #a #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hltl #Hlmlmt - elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2 - elim (IHU12 … HTU1) -U1 - /3 width=6 by cpy_bind, drop_skip, lift_bind, yle_succ, ex2_intro/ -| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hltl #Hlmlmt - elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HLK … HVW1) -W1 // - elim (IHU12 … HLK … HTU1) -U1 -HLK // - /3 width=5 by cpy_flat, lift_flat, ex2_intro/ -] -qed-. - -(* Basic_1: was: subst1_gen_lift_ge *) -lemma cpy_inv_lift1_ge: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → - ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → - l + m ≤ lt → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt-m, mt] T2 & ⬆[l, m] T2 ≡ U2. -#G #L #U1 #U2 #lt #mt #H elim H -G -L -U1 -U2 -lt -mt -[ * #i #G #L #lt #mt #K #s #l #m #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ - | elim (lift_inv_lref2 … H) -H * #Hil #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ - ] -| #I #G #L #KV #V #W #i #lt #mt #Hlti #Hilmt #HLKV #HVW #K #s #l #m #HLK #T1 #H #Hlmlt - lapply (yle_trans … Hlmlt … Hlti) #Hlmi - elim (yle_inv_plus_inj2 … Hlmlt) -Hlmlt #_ #Hmlt - elim (yle_inv_plus_inj2 … Hlmi) #Hlim #Hmi - lapply (yle_inv_inj … Hmi) -Hmi #Hmi - lapply (lift_inv_lref2_ge … H ?) -H // #H destruct - lapply (drop_conf_ge … HLK … HLKV ?) -L // #HKV - elim (lift_split … HVW l (i-m+1)) -HVW [2,3,4: /3 width=1 by yle_succ, yle_pred_sn, le_S_S/ ] -Hlmi -Hlim - #V0 #HV10 >plus_minus // yminus_succ1_inj /3 width=5 by cpy_bind, lift_bind, ex2_intro/ -| #I #G #L #W1 #W2 #U1 #U2 #lt #mt #_ #_ #IHW12 #IHU12 #K #s #l #m #HLK #X #H #Hlmtl - elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HLK … HVW1) -W1 // - elim (IHU12 … HLK … HTU1) -U1 -HLK /3 width=5 by cpy_flat, lift_flat, ex2_intro/ -] -qed-. - -(* Advanced inversion lemmas on relocation ***********************************) - -lemma cpy_inv_lift1_ge_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → - ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → - l ≤ lt → lt ≤ l + m → l + m ≤ lt + mt → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[l, lt + mt - (l + m)] T2 & ⬆[l, m] T2 ≡ U2. -#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hllt #Hltlm #Hlmlmt -elim (cpy_split_up … HU12 (l + m)) -HU12 // -Hlmlmt #U #HU1 #HU2 -lapply (cpy_weak … HU1 l m ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hllt -Hltlm #HU1 -lapply (cpy_inv_lift1_eq … HTU1 … HU1) -HU1 #HU1 destruct -elim (cpy_inv_lift1_ge … HU2 … HLK … HTU1) -U -L /2 width=3 by ex2_intro/ -qed-. - -lemma cpy_inv_lift1_be_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → - ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → - lt ≤ l → lt + mt ≤ l + m → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, l-lt] T2 & ⬆[l, m] T2 ≡ U2. -#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hlmtlm -lapply (cpy_weak … HU12 lt (l+m-lt) ? ?) -HU12 // -[ >ymax_pre_sn_comm /2 width=1 by yle_plus_dx1_trans/ ] -Hlmtlm #HU12 -elim (cpy_inv_lift1_be … HU12 … HLK … HTU1) -U1 -L /2 width=3 by ex2_intro/ -qed-. - -lemma cpy_inv_lift1_le_up: ∀G,L,U1,U2,lt,mt. ⦃G, L⦄ ⊢ U1 ▶[lt, mt] U2 → - ∀K,s,l,m. ⬇[s, l, m] L ≡ K → ∀T1. ⬆[l, m] T1 ≡ U1 → - lt ≤ l → l ≤ lt + mt → lt + mt ≤ l + m → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[lt, l - lt] T2 & ⬆[l, m] T2 ≡ U2. -#G #L #U1 #U2 #lt #mt #HU12 #K #s #l #m #HLK #T1 #HTU1 #Hltl #Hllmt #Hlmtlm -elim (cpy_split_up … HU12 l) -HU12 // #U #HU1 #HU2 -elim (cpy_inv_lift1_le … HU1 … HLK … HTU1) -U1 -[2: >ymax_pre_sn_comm // ] -Hltl #T #HT1 #HTU -lapply (cpy_weak … HU2 l m ? ?) -HU2 // -[ >ymax_pre_sn_comm // ] -Hllmt -Hlmtlm #HU2 -lapply (cpy_inv_lift1_eq … HTU … HU2) -L #H destruct /2 width=3 by ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy_nlift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy_nlift.ma deleted file mode 100644 index 3a6450724..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpy_nlift.ma +++ /dev/null @@ -1,65 +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/lift_neg.ma". -include "basic_2/substitution/lift_lift.ma". -include "basic_2/substitution/cpy.ma". - -(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) - -(* Inversion lemmas on negated relocation ***********************************) - -lemma cpy_fwd_nlift2_ge: ∀G,L,U1,U2,l,m. ⦃G, L⦄ ⊢ U1 ▶[l, m] U2 → - ∀i. l ≤ i → (∀T2. ⬆[i, 1] T2 ≡ U2 → ⊥) → - (∀T1. ⬆[i, 1] T1 ≡ U1 → ⊥) ∨ - ∃∃I,K,W,j. l ≤ yinj j & j < i & ⬇[j]L ≡ K.ⓑ{I}W & - (∀V. ⬆[⫰(i-j), 1] V ≡ W → ⊥) & (∀T1. ⬆[j, 1] T1 ≡ U1 → ⊥). -#G #L #U1 #U2 #l #m #H elim H -G -L -U1 -U2 -l -m -[ /3 width=2 by or_introl/ -| #I #G #L #K #V #W #j #l #m #Hlj #Hjlm #HLK #HVW #i #Hli #HnW - elim (ylt_split j i) #Hij - [ @or_intror @(ex5_4_intro … HLK) // -HLK - [ #X #HXV elim (lift_trans_le … HXV … HVW ?) -V // #Y #HXY - yplus_SO2 >ymax_pre_sn /2 width=2 by ylt_fwd_le_succ1/ - | -HnW /3 width=7 by lift_inv_lref2_be, ylt_inj/ - ] - | elim (lift_split … HVW i j) -HVW // - #X #_ #H elim HnW -HnW // - ] -| #a #I #G #L #W1 #W2 #U1 #U2 #l #m #_ #_ #IHW12 #IHU12 #i #Hli #H elim (nlift_inv_bind … H) -H - [ #HnW2 elim (IHW12 … HnW2) -IHW12 -HnW2 -IHU12 // - [ /4 width=9 by nlift_bind_sn, or_introl/ - | * /5 width=9 by nlift_bind_sn, ex5_4_intro, or_intror/ - ] - | #HnU2 elim (IHU12 … HnU2) -IHU12 -HnU2 -IHW12 /2 width=1 by yle_succ/ - [ /4 width=9 by nlift_bind_dx, or_introl/ - | * #J #K #W #j #Hlj elim (yle_inv_succ1 … Hlj) -Hlj #Hlj #Hj - yminus_succ #Hji #HLK #HnW - lapply (drop_inv_drop1_lt … HLK ?) /2 width=1 by ylt_O/ -HLK #HLK - (plus_minus_m_m m 1) /2 width=1 by drop_drop/ -qed. - -lemma drop_skip_lt: ∀I,L1,L2,V1,V2,s,l,m. - ⬇[s, ⫰l, m] L1 ≡ L2 → ⬆[⫰l, m] V2 ≡ V1 → 0 < l → - ⬇[s, l, m] L1. ⓑ{I} V1 ≡ L2.ⓑ{I}V2. -#I #L1 #L2 #V1 #V2 #s #l #m #HL12 #HV21 #Hl <(ylt_inv_O1 … Hl) -Hl -/2 width=1 by drop_skip/ -qed. - -lemma drop_O1_le: ∀s,m,L. m ≤ |L| → ∃K. ⬇[s, 0, m] L ≡ K. -#s #m @(nat_ind_plus … m) -m /2 width=2 by ex_intro/ -#m #IHm * -[ #H elim (le_plus_xSy_O_false … H) -| #L #I #V normalize #H elim (IHm L) -IHm /3 width=2 by drop_drop, monotonic_pred, ex_intro/ -] -qed-. - -lemma drop_O1_lt: ∀s,L,m. m < |L| → ∃∃I,K,V. ⬇[s, 0, m] L ≡ K.ⓑ{I}V. -#s #L elim L -L -[ #m #H elim (lt_zero_false … H) -| #L #I #V #IHL #m @(nat_ind_plus … m) -m /2 width=4 by drop_pair, ex1_3_intro/ - #m #_ normalize #H elim (IHL m) -IHL /3 width=4 by drop_drop, lt_plus_to_minus_r, lt_plus_to_lt_l, ex1_3_intro/ -] -qed-. - -lemma drop_O1_pair: ∀L,K,m,s. ⬇[s, yinj 0, m] L ≡ K → m ≤ |L| → ∀I,V. - ∃∃J,W. ⬇[s, yinj 0, m] L.ⓑ{I}V ≡ K.ⓑ{J}W. -#L elim L -L [| #L #Z #X #IHL ] #K #m #s #H normalize #Hm #I #V -[ elim (drop_inv_atom1 … H) -H #H <(le_n_O_to_eq … Hm) -m - #Hs destruct /2 width=3 by ex1_2_intro/ -| elim (drop_inv_O1_pair1 … H) -H * #Hm #HLK destruct /2 width=3 by ex1_2_intro/ - elim (IHL … HLK … Z X) -IHL -HLK - /3 width=3 by drop_drop_lt, le_plus_to_minus, ex1_2_intro/ -] -qed-. - -lemma drop_O1_ge: ∀L,m. |L| ≤ m → ⬇[Ⓣ, 0, m] L ≡ ⋆. -#L elim L -L [ #m #_ @drop_atom #H destruct ] -#L #I #V #IHL #m @(nat_ind_plus … m) -m [ #H elim (le_plus_xSy_O_false … H) ] -normalize /4 width=1 by drop_drop, monotonic_pred/ -qed. - -lemma drop_O1_eq: ∀L,s. ⬇[s, 0, |L|] L ≡ ⋆. -#L elim L -L /2 width=1 by drop_drop, drop_atom/ -qed. - -lemma drop_split: ∀L1,L2,l,m2,s. ⬇[s, l, m2] L1 ≡ L2 → ∀m1. m1 ≤ m2 → - ∃∃L. ⬇[s, l, m2 - m1] L1 ≡ L & ⬇[s, l, m1] L ≡ L2. -#L1 #L2 #l #m2 #s #H elim H -L1 -L2 -l -m2 -[ #l #m2 #Hs #m1 #Hm12 @(ex2_intro … (⋆)) - @drop_atom #H lapply (Hs H) -s #H destruct /2 width=1 by le_n_O_to_eq/ -| #I #L1 #V #m1 #Hm1 lapply (le_n_O_to_eq … Hm1) -Hm1 - #H destruct /2 width=3 by ex2_intro/ -| #I #L1 #L2 #V #m2 #HL12 #IHL12 #m1 @(nat_ind_plus … m1) -m1 - [ /3 width=3 by drop_drop, ex2_intro/ - | -HL12 #m1 #_ #Hm12 lapply (le_plus_to_le_r … Hm12) -Hm12 - #Hm12 elim (IHL12 … Hm12) -IHL12 >minus_plus_plus_l - #L #HL1 #HL2 elim (lt_or_ge (|L1|) (m2-m1)) #H0 - [ elim (drop_inv_O1_gt … HL1 H0) -HL1 #H1 #H2 destruct - elim (drop_inv_atom1 … HL2) -HL2 #H #_ destruct - @(ex2_intro … (⋆)) [ @drop_O1_ge normalize // ] - @drop_atom #H destruct - | elim (drop_O1_pair … HL1 H0 I V) -HL1 -H0 /3 width=5 by drop_drop, ex2_intro/ - ] - ] -| #I #L1 #L2 #V1 #V2 #l #m2 #_ #HV21 #IHL12 #m1 #Hm12 elim (IHL12 … Hm12) -IHL12 - #L #HL1 #HL2 elim (lift_split … HV21 l m1) -HV21 /3 width=5 by drop_skip, ex2_intro/ -] -qed-. - -lemma drop_FT: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → ⬇[Ⓣ, l, m] L1 ≡ L2. -#L1 #L2 #l #m #H elim H -L1 -L2 -l -m -/3 width=1 by drop_atom, drop_drop, drop_skip/ -qed. - -lemma drop_gen: ∀L1,L2,s,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → ⬇[s, l, m] L1 ≡ L2. -#L1 #L2 * /2 width=1 by drop_FT/ -qed-. - -lemma drop_T: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → ⬇[Ⓣ, l, m] L1 ≡ L2. -#L1 #L2 * /2 width=1 by drop_FT/ -qed-. - -lemma d_liftable_LTC: ∀R. d_liftable R → d_liftable (LTC … R). -#R #HR #K #T1 #T2 #H elim H -T2 -[ /3 width=10 by inj/ -| #T #T2 #_ #HT2 #IHT1 #L #s #l #m #HLK #U1 #HTU1 #U2 #HTU2 - elim (lift_total T l m) /4 width=12 by step/ -] -qed-. - -lemma d_deliftable_sn_LTC: ∀R. d_deliftable_sn R → d_deliftable_sn (LTC … R). -#R #HR #L #U1 #U2 #H elim H -U2 -[ #U2 #HU12 #K #s #l #m #HLK #T1 #HTU1 - elim (HR … HU12 … HLK … HTU1) -HR -L -U1 /3 width=3 by inj, ex2_intro/ -| #U #U2 #_ #HU2 #IHU1 #K #s #l #m #HLK #T1 #HTU1 - elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1 - elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by step, ex2_intro/ -] -qed-. - -lemma dropable_sn_TC: ∀R. dropable_sn R → dropable_sn (TC … R). -#R #HR #L1 #K1 #s #l #m #HLK1 #L2 #H elim H -L2 -[ #L2 #HL12 elim (HR … HLK1 … HL12) -HR -L1 - /3 width=3 by inj, ex2_intro/ -| #L #L2 #_ #HL2 * #K #HK1 #HLK elim (HR … HLK … HL2) -HR -L - /3 width=3 by step, ex2_intro/ -] -qed-. - -lemma dropable_dx_TC: ∀R. dropable_dx R → dropable_dx (TC … R). -#R #HR #L1 #L2 #H elim H -L2 -[ #L2 #HL12 #K2 #s #m #HLK2 elim (HR … HL12 … HLK2) -HR -L2 - /3 width=3 by inj, ex2_intro/ -| #L #L2 #_ #HL2 #IHL1 #K2 #s #m #HLK2 elim (HR … HL2 … HLK2) -HR -L2 - #K #HLK #HK2 elim (IHL1 … HLK) -L - /3 width=5 by step, ex2_intro/ -] -qed-. - -lemma d_deliftable_sn_llstar: ∀R. d_deliftable_sn R → - ∀d. d_deliftable_sn (llstar … R d). -#R #HR #d #L #U1 #U2 #H @(lstar_ind_r … d U2 H) -d -U2 -[ /2 width=3 by lstar_O, ex2_intro/ -| #d #U #U2 #_ #HU2 #IHU1 #K #s #l #m #HLK #T1 #HTU1 - elim (IHU1 … HLK … HTU1) -IHU1 -U1 #T #HTU #HT1 - elim (HR … HU2 … HLK … HTU) -HR -L -U /3 width=5 by lstar_dx, ex2_intro/ -] -qed-. - -(* Basic forward lemmas *****************************************************) - -(* Basic_1: was: drop_S *) -lemma drop_fwd_drop2: ∀L1,I2,K2,V2,s,m. ⬇[s, O, m] L1 ≡ K2. ⓑ{I2} V2 → - ⬇[s, O, m + 1] L1 ≡ K2. -#L1 elim L1 -L1 -[ #I2 #K2 #V2 #s #m #H lapply (drop_inv_atom1 … H) -H * #H destruct -| #K1 #I1 #V1 #IHL1 #I2 #K2 #V2 #s #m #H - elim (drop_inv_O1_pair1 … H) -H * #Hm #H - [ -IHL1 destruct /2 width=1 by drop_drop/ - | @drop_drop >(plus_minus_m_m m 1) /2 width=3 by/ - ] -] -qed-. - -lemma drop_fwd_length_ge: ∀L1,L2,l,m,s. ⬇[s, l, m] L1 ≡ L2 → |L1| ≤ l → |L2| = |L1|. -#L1 #L2 #l #m #s #H elim H -L1 -L2 -l -m // -[ #I #L1 #L2 #V #m #_ #_ #H elim (ylt_yle_false … H) -H normalize // -| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #IH yplus_SO2 #H - lapply (yle_inv_succ … H) -H #H normalize /3 width=1 by eq_f2/ -] -qed-. - -lemma drop_fwd_length_le_le: ∀L1,L2,l,m,s. ⬇[s, l, m] L1 ≡ L2 → l ≤ |L1| → m ≤ |L1| - l → |L2| = |L1| - m. -#L1 #L2 #l #m #s #H elim H -L1 -L2 -l -m // -[ #I #L1 #L2 #V #m #_ minus_plus_plus_l yplus_SO2 /3 width=1 by yle_inv_succ/ -| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #IHL12 yplus_SO2 #H - lapply (yle_inv_succ … H) -H #Hl1 >yminus_succ #Hml1 normalize - yminus_inj >yminus_inj yplus_SO2 >yplus_SO2 >yminus_succ - /4 width=1 by yle_inv_succ, eq_f/ -] -qed-. - -lemma drop_fwd_length: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L1| = |L2| + m. -#L1 #L2 #l #m #H elim H -L1 -L2 -l -m // normalize /2 width=1 by/ -qed-. - -lemma drop_fwd_length_minus2: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L2| = |L1| - m. -#L1 #L2 #l #m #H lapply (drop_fwd_length … H) -H /2 width=1 by plus_minus, le_n/ -qed-. - -lemma drop_fwd_length_minus4: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → m = |L1| - |L2|. -#L1 #L2 #l #m #H lapply (drop_fwd_length … H) -H // -qed-. - -lemma drop_fwd_length_le2: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → m ≤ |L1|. -#L1 #L2 #l #m #H lapply (drop_fwd_length … H) -H // -qed-. - -lemma drop_fwd_length_le4: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L2| ≤ |L1|. -#L1 #L2 #l #m #H lapply (drop_fwd_length … H) -H // -qed-. - -lemma drop_fwd_length_lt2: ∀L1,I2,K2,V2,l,m. - ⬇[Ⓕ, l, m] L1 ≡ K2. ⓑ{I2} V2 → m < |L1|. -#L1 #I2 #K2 #V2 #l #m #H -lapply (drop_fwd_length … H) normalize in ⊢ (%→?); -I2 -V2 // -qed-. - -lemma drop_fwd_length_lt4: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → 0 < m → |L2| < |L1|. -#L1 #L2 #l #m #H lapply (drop_fwd_length … H) -H /2 width=1 by lt_minus_to_plus_r/ -qed-. - -lemma drop_fwd_length_eq1: ∀L1,L2,K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 → - |L1| = |L2| → |K1| = |K2|. -#L1 #L2 #K1 #K2 #l #m #HLK1 #HLK2 #HL12 -lapply (drop_fwd_length … HLK1) -HLK1 -lapply (drop_fwd_length … HLK2) -HLK2 -/2 width=2 by injective_plus_r/ -qed-. - -lemma drop_fwd_length_eq2: ∀L1,L2,K1,K2,l,m. ⬇[Ⓕ, l, m] L1 ≡ K1 → ⬇[Ⓕ, l, m] L2 ≡ K2 → - |K1| = |K2| → |L1| = |L2|. -#L1 #L2 #K1 #K2 #l #m #HLK1 #HLK2 #HL12 -lapply (drop_fwd_length … HLK1) -HLK1 -lapply (drop_fwd_length … HLK2) -HLK2 // -qed-. - -lemma drop_fwd_lw: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → ♯{L2} ≤ ♯{L1}. -#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m // normalize -[ /2 width=3 by transitive_le/ -| #I #L1 #L2 #V1 #V2 #l #m #_ #HV21 #IHL12 - >(lift_fwd_tw … HV21) -HV21 /2 width=1 by monotonic_le_plus_l/ -] -qed-. - -lemma drop_fwd_lw_lt: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → 0 < m → ♯{L2} < ♯{L1}. -#L1 #L2 #l #m #H elim H -L1 -L2 -l -m -[ #l #m #H >H -H // -| #I #L #V #H elim (lt_refl_false … H) -| #I #L1 #L2 #V #m #HL12 #_ #_ - lapply (drop_fwd_lw … HL12) -HL12 #HL12 - @(le_to_lt_to_lt … HL12) -HL12 // -| #I #L1 #L2 #V1 #V2 #l #m #_ #HV21 #IHL12 #H normalize in ⊢ (?%%); -I - >(lift_fwd_tw … HV21) -V2 /3 by lt_minus_to_plus/ -] -qed-. - -lemma drop_fwd_rfw: ∀I,L,K,V,i. ⬇[i] L ≡ K.ⓑ{I}V → ∀T. ♯{K, V} < ♯{L, T}. -#I #L #K #V #i #HLK lapply (drop_fwd_lw … HLK) -HLK -normalize in ⊢ (%→?→?%%); /3 width=3 by le_to_lt_to_lt/ -qed-. - -(* Advanced inversion lemmas ************************************************) - -fact drop_inv_O2_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → m = 0 → L1 = L2. -#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m -[ // -| // -| #I #L1 #L2 #V #m #_ #_ >commutative_plus normalize #H destruct -| #I #L1 #L2 #V1 #V2 #l #m #_ #HV21 #IHL12 #H - >(IHL12 H) -L1 >(lift_inv_O2_aux … HV21 … H) -V2 -l -m // -] -qed-. - -(* Basic_1: was: drop_gen_refl *) -lemma drop_inv_O2: ∀L1,L2,s,l. ⬇[s, l, 0] L1 ≡ L2 → L1 = L2. -/2 width=5 by drop_inv_O2_aux/ qed-. - -lemma drop_inv_length_eq: ∀L1,L2,l,m. ⬇[Ⓕ, l, m] L1 ≡ L2 → |L1| = |L2| → m = 0. -#L1 #L2 #l #m #H #HL12 lapply (drop_fwd_length_minus4 … H) // -qed-. - -lemma drop_inv_refl: ∀L,l,m. ⬇[Ⓕ, l, m] L ≡ L → m = 0. -/2 width=5 by drop_inv_length_eq/ qed-. - -fact drop_inv_FT_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → - ∀I,K,V. L2 = K.ⓑ{I}V → s = Ⓣ → l = 0 → - ⬇[Ⓕ, l, m] L1 ≡ K.ⓑ{I}V. -#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m -[ #l #m #_ #J #K #W #H destruct -| #I #L #V #J #K #W #H destruct // -| #I #L1 #L2 #V #m #_ #IHL12 #J #K #W #H1 #H2 destruct - /3 width=1 by drop_drop/ -| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #_ #J #K #W #_ #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma drop_inv_FT: ∀I,L,K,V,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] L ≡ K.ⓑ{I}V. -/2 width=5 by drop_inv_FT_aux/ qed. - -lemma drop_inv_gen: ∀I,L,K,V,s,m. ⬇[s, 0, m] L ≡ K.ⓑ{I}V → ⬇[m] L ≡ K.ⓑ{I}V. -#I #L #K #V * /2 width=1 by drop_inv_FT/ -qed-. - -lemma drop_inv_T: ∀I,L,K,V,s,m. ⬇[Ⓣ, 0, m] L ≡ K.ⓑ{I}V → ⬇[s, 0, m] L ≡ K.ⓑ{I}V. -#I #L #K #V * /2 width=1 by drop_inv_FT/ -qed-. - -(* Basic_1: removed theorems 50: - drop_ctail drop_skip_flat - cimp_flat_sx cimp_flat_dx cimp_bind cimp_getl_conf - drop_clear drop_clear_O drop_clear_S - clear_gen_sort clear_gen_bind clear_gen_flat clear_gen_flat_r - clear_gen_all clear_clear clear_mono clear_trans clear_ctail clear_cle - getl_ctail_clen getl_gen_tail clear_getl_trans getl_clear_trans - getl_clear_bind getl_clear_conf getl_dec getl_drop getl_drop_conf_lt - getl_drop_conf_ge getl_conf_ge_drop getl_drop_conf_rev - drop_getl_trans_lt drop_getl_trans_le drop_getl_trans_ge - getl_drop_trans getl_flt getl_gen_all getl_gen_sort getl_gen_O - getl_gen_S getl_gen_2 getl_gen_flat getl_gen_bind getl_conf_le - getl_trans getl_refl getl_head getl_flat getl_ctail getl_mono -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_append.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_append.ma deleted file mode 100644 index 2ea503338..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_append.ma +++ /dev/null @@ -1,61 +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/grammar/lenv_append.ma". -include "basic_2/substitution/drop.ma". - -(* DROPPING *****************************************************************) - -(* Properties on append for local environments ******************************) - -fact drop_O1_append_sn_le_aux: ∀L1,L2,s,l,m. ⬇[s, l, m] L1 ≡ L2 → - l = yinj 0 → m ≤ |L1| → - ∀L. ⬇[s, yinj 0, m] L @@ L1 ≡ L @@ L2. -#L1 #L2 #s #l #m #H elim H -L1 -L2 -l -m // -[ #l #m #_ #_ #H <(le_n_O_to_eq … H) -H // -| normalize /4 width=1 by drop_drop, monotonic_pred/ -| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #_ #H elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma drop_O1_append_sn_le: ∀L1,L2,s,m. ⬇[s, yinj 0, m] L1 ≡ L2 → m ≤ |L1| → - ∀L. ⬇[s, yinj 0, m] L @@ L1 ≡ L @@ L2. -/2 width=3 by drop_O1_append_sn_le_aux/ qed. - -(* Inversion lemmas on append for local environments ************************) - -lemma drop_O1_inv_append1_ge: ∀K,L1,L2,s,m. ⬇[s, yinj 0, m] L1 @@ L2 ≡ K → - |L2| ≤ m → ⬇[s, yinj 0, m - |L2|] L1 ≡ K. -#K #L1 #L2 elim L2 -L2 normalize // -#L2 #I #V #IHL2 #s #m #H #H1m -elim (drop_inv_O1_pair1 … H) -H * #H2m #HL12 destruct -[ lapply (le_n_O_to_eq … H1m) -H1m -IHL2 minus_minus_comm /3 width=1 by monotonic_pred/ -] -qed-. - -lemma drop_O1_inv_append1_le: ∀K,L1,L2,s,m. ⬇[s, yinj 0, m] L1 @@ L2 ≡ K → m ≤ |L2| → - ∀K2. ⬇[s, yinj 0, m] L2 ≡ K2 → K = L1 @@ K2. -#K #L1 #L2 elim L2 -L2 normalize -[ #s #m #H1 #H2 #K2 #H3 lapply (le_n_O_to_eq … H2) -H2 - #H2 elim (drop_inv_atom1 … H3) -H3 #H3 #_ destruct - >(drop_inv_O2 … H1) -H1 // -| #L2 #I #V #IHL2 #s #m @(nat_ind_plus … m) -m [ -IHL2 ] - [ #H1 #_ #K2 #H2 - lapply (drop_inv_O2 … H1) -H1 #H1 - lapply (drop_inv_O2 … H2) -H2 #H2 destruct // - | /4 width=7 by drop_inv_drop1, le_plus_to_le_r/ - ] -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_drop.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_drop.ma deleted file mode 100644 index ea8cadb13..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_drop.ma +++ /dev/null @@ -1,218 +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/lift_lift.ma". -include "basic_2/substitution/drop.ma". - -(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************) - -(* Main properties **********************************************************) - -(* Basic_1: was: drop_mono *) -theorem drop_mono: ∀L,L1,s1,l,m. ⬇[s1, l, m] L ≡ L1 → - ∀L2,s2. ⬇[s2, l, m] L ≡ L2 → L1 = L2. -#L #L1 #s1 #l #m #H elim H -L -L1 -l -m -[ #l #m #Hm #L2 #s2 #H elim (drop_inv_atom1 … H) -H // -| #I #K #V #L2 #s2 #HL12 <(drop_inv_O2 … HL12) -L2 // -| #I #L #K #V #m #_ #IHLK #L2 #s2 #H - lapply (drop_inv_drop1 … H) -H /2 width=2 by/ -| #I #L #K1 #T #V1 #l #m #_ #HVT1 #IHLK1 #X #s2 #H - elim (drop_inv_skip1 … H) -H // >ypred_succ #K2 #V2 #HLK2 #HVT2 #H destruct - >(lift_inj … HVT1 … HVT2) -HVT1 -HVT2 - >(IHLK1 … HLK2) -IHLK1 -HLK2 // -] -qed-. - -(* Basic_1: was: drop_conf_ge *) -theorem drop_conf_ge: ∀L,L1,s1,l1,m1. ⬇[s1, l1, m1] L ≡ L1 → - ∀L2,s2,m2. ⬇[s2, 0, m2] L ≡ L2 → l1 + m1 ≤ m2 → - ⬇[s2, 0, m2 - m1] L1 ≡ L2. -#L #L1 #s1 #l1 #m1 #H elim H -L -L1 -l1 -m1 // -[ #l #m #_ #L2 #s2 #m2 #H #_ elim (drop_inv_atom1 … H) -H - #H #Hm destruct - @drop_atom #H >Hm // (**) (* explicit constructor *) -| #I #L #K #V #m #_ #IHLK #L2 #s2 #m2 #H >yplus_O1 minus_minus_comm @IHLK // -| #I #L #K #V1 #V2 #l #m #_ #_ #IHLK #L2 #s2 #m2 #H #Hlmm2 - lapply (yle_plus1_to_minus_inj2 … Hlmm2) #Hlm2m - lapply (ylt_yle_trans 0 … Hlm2m ?) // -Hlm2m #Hm2m - >yplus_succ1 in Hlmm2; #Hlmm2 - elim (yle_inv_succ1 … Hlmm2) -Hlmm2 #Hlmm2 #Hm2 - lapply (drop_inv_drop1_lt … H ?) -H /2 width=1 by ylt_inv_inj/ -Hm2 #HL2 - @drop_drop_lt /2 width=1 by ylt_inv_inj/ >minus_minus_comm - (Hm2 ?) in Hl1; // -Hm2 #Hl1 <(le_n_O_to_eq … Hl1) -l1 - /4 width=3 by drop_atom, ex2_intro/ -| normalize #I #L #V #L2 #m2 #HL2 #_ #Hm2 - lapply (le_n_O_to_eq … Hm2) -Hm2 #H destruct - lapply (drop_inv_O2 … HL2) -HL2 #H destruct /2 width=3 by drop_pair, ex2_intro/ -| normalize #I #L0 #K0 #V1 #m1 #HLK0 #IHLK0 #L2 #m2 #H #_ #Hm21 - lapply (drop_inv_O1_pair1 … H) -H * * #Hm2 #HL20 - [ -IHLK0 -Hm21 destruct plus_plus_comm_23 #_ #_ #IHLK0 #L2 #m2 #H #Hl1m2 #Hm2lm1 - elim (le_inv_plus_l … Hl1m2) #_ #Hm2 - yplus_minus_assoc_comm_inj /2 width=1 by yle_inj/ - >yplus_SO2 /3 width=3 by drop_drop_lt, ex2_intro/ - ] -] -qed-. - -(* Note: with "s2", the conclusion parameter is "s1 ∨ s2" *) -(* Basic_1: was: drop_trans_ge *) -theorem drop_trans_ge: ∀L1,L,s1,l1,m1. ⬇[s1, l1, m1] L1 ≡ L → - ∀L2,m2. ⬇[m2] L ≡ L2 → l1 ≤ m2 → ⬇[s1, 0, m1 + m2] L1 ≡ L2. -#L1 #L #s1 #l1 #m1 #H elim H -L1 -L -l1 -m1 -[ #l1 #m1 #Hm1 #L2 #m2 #H #_ elim (drop_inv_atom1 … H) -H - #H #Hm2 destruct /4 width=1 by drop_atom, eq_f2/ -| /2 width=1 by drop_gen/ -| /3 width=1 by drop_drop/ -| #I #L1 #L2 #V1 #V2 #l #m #_ #_ #IHL12 #L #m2 #H #Hlm2 - elim (yle_inv_succ1 … Hlm2) -Hlm2 #Hlm2 #Hm2 - lapply (ylt_O … Hm2) -Hm2 #Hm2 - lapply (lt_to_le_to_lt … (m + m2) Hm2 ?) // #Hmm2 - lapply (drop_inv_drop1_lt … H ?) -H // #HL2 - @drop_drop_lt // >le_plus_minus yplus_minus_assoc_comm_inj /3 width=3 by drop_drop_lt, yle_inj, ex2_intro/ - ] -] -qed-. - -(* Advanced properties ******************************************************) - -lemma d_liftable_llstar: ∀R. d_liftable R → ∀d. d_liftable (llstar … R d). -#R #HR #d #K #T1 #T2 #H @(lstar_ind_r … d T2 H) -d -T2 -[ #L #s #l #m #_ #U1 #HTU1 #U2 #HTU2 -HR -K - >(lift_mono … HTU2 … HTU1) -T1 -U2 -l -m // -| #d #T #T2 #_ #HT2 #IHT1 #L #s #l #m #HLK #U1 #HTU1 #U2 #HTU2 - elim (lift_total T l m) /3 width=12 by lstar_dx/ -] -qed-. - -(* Basic_1: was: drop_conf_lt *) -lemma drop_conf_lt: ∀L,L1,s1,l1,m1. ⬇[s1, l1, m1] L ≡ L1 → - ∀I,K2,V2,s2,m2. ⬇[s2, 0, m2] L ≡ K2.ⓑ{I}V2 → - m2 < l1 → let l ≝ ⫰(l1 - m2) in - ∃∃K1,V1. ⬇[s2, 0, m2] L1 ≡ K1.ⓑ{I}V1 & - ⬇[s1, l, m1] K2 ≡ K1 & ⬆[l, m1] V1 ≡ V2. -#L #L1 #s1 #l1 #m1 #H1 #I #K2 #V2 #s2 #m2 #H2 #Hm2l1 -elim (drop_conf_le … H1 … H2) -L /2 width=2 by ylt_fwd_le/ #K #HL1K #HK2 -elim (drop_inv_skip1 … HK2) -HK2 /2 width=1 by ylt_to_minus/ -#K1 #V1 #HK21 #HV12 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -(* Note: apparently this was missing in basic_1 *) -lemma drop_trans_lt: ∀L1,L,s1,l1,m1. ⬇[s1, l1, m1] L1 ≡ L → - ∀I,L2,V2,s2,m2. ⬇[s2, 0, m2] L ≡ L2.ⓑ{I}V2 → - m2 < l1 → let l ≝ l1 - m2 - 1 in - ∃∃L0,V0. ⬇[s2, 0, m2] L1 ≡ L0.ⓑ{I}V0 & - ⬇[s1, l, m1] L0 ≡ L2 & ⬆[l, m1] V2 ≡ V0. -#L1 #L #s1 #l1 #m1 #HL1 #I #L2 #V2 #s2 #m2 #HL2 #Hl21 -elim (drop_trans_le … HL1 … HL2) -L /2 width=1 by ylt_fwd_le/ #L0 #HL10 #HL02 -elim (drop_inv_skip2 … HL02) -HL02 /2 width=1 by ylt_to_minus/ -#L #V1 #HL2 #HV21 #H destruct /2 width=5 by ex3_2_intro/ -qed-. - -lemma drop_trans_ge_comm: ∀L1,L,L2,s1,l1,m1,m2. - ⬇[s1, l1, m1] L1 ≡ L → ⬇[m2] L ≡ L2 → l1 ≤ m2 → - ⬇[s1, 0, m2 + m1] L1 ≡ L2. -#L1 #L #L2 #s1 #l1 #m1 #m2 ->commutative_plus /2 width=5 by drop_trans_ge/ -qed. - -lemma drop_conf_div: ∀I1,L,K,V1,m1. ⬇[m1] L ≡ K.ⓑ{I1}V1 → - ∀I2,V2,m2. ⬇[m2] L ≡ K.ⓑ{I2}V2 → - ∧∧ m1 = m2 & I1 = I2 & V1 = V2. -#I1 #L #K #V1 #m1 #HLK1 #I2 #V2 #m2 #HLK2 -elim (le_or_ge m1 m2) #Hm -[ lapply (drop_conf_ge … HLK1 … HLK2 ?) -| lapply (drop_conf_ge … HLK2 … HLK1 ?) -] -HLK1 -HLK2 /2 width=1 by yle_inj/ #HK -lapply (drop_fwd_length_minus2 … HK) #H -elim (discr_minus_x_xy … H) -H -[1,3: normalize H in HK; #HK -lapply (drop_inv_O2 … HK) -HK #H destruct -lapply (inv_eq_minus_O … H) -H /3 width=1 by le_to_le_to_eq, and3_intro/ -qed-. - -(* Advanced forward lemmas **************************************************) - -lemma drop_fwd_be: ∀L,K,s,l,m,i. ⬇[s, l, m] L ≡ K → |K| ≤ i → yinj i < l → |L| ≤ i. -#L #K #s #l #m #i #HLK #HK #Hl elim (lt_or_ge i (|L|)) // -#HL elim (drop_O1_lt (Ⓕ) … HL) #I #K0 #V #HLK0 -HL -elim (drop_conf_lt … HLK … HLK0) // -HLK -HLK0 -Hl -#K1 #V1 #HK1 #_ #_ lapply (drop_fwd_length_lt2 … HK1) -I -K1 -V1 -#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_lreq.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_lreq.ma deleted file mode 100644 index 0165a8255..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/drop_lreq.ma +++ /dev/null @@ -1,92 +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/grammar/lreq_lreq.ma". -include "basic_2/substitution/drop.ma". - -(* BASIC SLICING FOR LOCAL ENVIRONMENTS *************************************) - -definition dedropable_sn: predicate (relation lenv) ≝ - λR. ∀L1,K1,s,l,m. ⬇[s, l, m] L1 ≡ K1 → ∀K2. R K1 K2 → - ∃∃L2. R L1 L2 & ⬇[s, l, m] L2 ≡ K2 & L1 ⩬[l, m] L2. - -(* Properties on equivalence ************************************************) - -lemma lreq_drop_trans_be: ∀L1,L2,l,m. L1 ⩬[l, m] L2 → - ∀I,K2,W,s,i. ⬇[s, 0, i] L2 ≡ K2.ⓑ{I}W → - l ≤ i → i < l + m → - ∃∃K1. K1 ⩬[0, ⫰(l+m-i)] K2 & ⬇[s, 0, i] L1 ≡ K1.ⓑ{I}W. -#L1 #L2 #l #m #H elim H -L1 -L2 -l -m -[ #l #m #J #K2 #W #s #i #H - elim (drop_inv_atom1 … H) -H #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J #K2 #W #s #i #_ #_ #H - elim (ylt_yle_false … H) // -| #I #L1 #L2 #V #m #HL12 #IHL12 #J #K2 #W #s #i #H #_ >yplus_O1 - elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ] - [ #_ destruct >ypred_succ - /2 width=3 by drop_pair, ex2_intro/ - | lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/ - #H yminus_succ yplus_succ1 #H lapply (ylt_inv_succ … H) -H - #Hilm lapply (drop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/ - #HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 yminus_SO2 - /4 width=3 by ylt_O, drop_drop_lt, ex2_intro/ -] -qed-. - -lemma lreq_drop_conf_be: ∀L1,L2,l,m. L1 ⩬[l, m] L2 → - ∀I,K1,W,s,i. ⬇[s, 0, i] L1 ≡ K1.ⓑ{I}W → - l ≤ i → i < l + m → - ∃∃K2. K1 ⩬[0, ⫰(l+m-i)] K2 & ⬇[s, 0, i] L2 ≡ K2.ⓑ{I}W. -#L1 #L2 #l #m #HL12 #I #K1 #W #s #i #HLK1 #Hli #Hilm -elim (lreq_drop_trans_be … (lreq_sym … HL12) … HLK1) // -L1 -Hli -Hilm -/3 width=3 by lreq_sym, ex2_intro/ -qed-. - -lemma drop_O1_ex: ∀K2,i,L1. |L1| = |K2| + i → - ∃∃L2. L1 ⩬[0, i] L2 & ⬇[i] L2 ≡ K2. -#K2 #i @(nat_ind_plus … i) -i -[ /3 width=3 by lreq_O2, ex2_intro/ -| #i #IHi #Y #Hi elim (drop_O1_lt (Ⓕ) Y 0) // - #I #L1 #V #H lapply (drop_inv_O2 … H) -H #H destruct - normalize in Hi; elim (IHi L1) -IHi - /3 width=5 by drop_drop, lreq_pair, injective_plus_l, ex2_intro/ -] -qed-. - -lemma dedropable_sn_TC: ∀R. dedropable_sn R → dedropable_sn (TC … R). -#R #HR #L1 #K1 #s #l #m #HLK1 #K2 #H elim H -K2 -[ #K2 #HK12 elim (HR … HLK1 … HK12) -HR -K1 - /3 width=4 by inj, ex3_intro/ -| #K #K2 #_ #HK2 * #L #H1L1 #HLK #H2L1 elim (HR … HLK … HK2) -HR -K - /3 width=6 by lreq_trans, step, ex3_intro/ -] -qed-. - -(* Inversion lemmas on equivalence ******************************************) - -lemma drop_O1_inj: ∀i,L1,L2,K. ⬇[i] L1 ≡ K → ⬇[i] L2 ≡ K → L1 ⩬[i, ∞] L2. -#i @(nat_ind_plus … i) -i -[ #L1 #L2 #K #H <(drop_inv_O2 … H) -K #H <(drop_inv_O2 … H) -L1 // -| #i #IHi * [2: #L1 #I1 #V1 ] * [2,4: #L2 #I2 #V2 ] #K #HLK1 #HLK2 // - lapply (drop_fwd_length … HLK1) - <(drop_fwd_length … HLK2) [ /4 width=5 by drop_inv_drop1, lreq_succ/ ] - normalize (plus_minus_m_m m 1) /2 width=3 by fqu_drop/ -qed. - -lemma fqu_lref_S_lt: ∀I,G,L,V,i. 0 < i → ⦃G, L.ⓑ{I}V, #i⦄ ⊐ ⦃G, L, #(i-1)⦄. -/4 width=3 by fqu_drop, drop_drop, lift_lref_ge_minus, yle_inj/ -qed. - -(* Basic forward lemmas *****************************************************) - -lemma fqu_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} < ♯{G1, L1, T1}. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // -#G #L #K #T #U #m #HLK #HTU -lapply (drop_fwd_lw_lt … HLK ?) -HLK // #HKL -lapply (lift_fwd_tw … HTU) -m #H -normalize in ⊢ (?%%); /2 width=1 by lt_minus_to_plus/ -qed-. - -fact fqu_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → - ∀i. T1 = #i → |L2| < |L1|. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 -[1: normalize // -|3: #a -|5: /2 width=4 by drop_fwd_length_lt4/ -] #I #G #L #V #T #j #H destruct -qed-. - -lemma fqu_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐ ⦃G2, L2, T2⦄ → |L2| < |L1|. -/2 width=7 by fqu_fwd_length_lref1_aux/ -qed-. - -(* Basic inversion lemmas ***************************************************) - -fact fqu_inv_eq_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → - G1 = G2 → |L1| = |L2| → T1 = T2 → ⊥. -#G1 #G2 #L1 #L2 #T1 #T2 * -G1 -G2 -L1 -L2 -T1 -T2 normalize -/2 width=4 by discr_tpair_xy_y, discr_tpair_xy_x, plus_xSy_x_false/ -#G #L #K #T #U #m #HLK #_ #_ #H #_ -G -T -U >(drop_fwd_length … HLK) in H; -L -/2 width=4 by plus_xySz_x_false/ -qed-. - -lemma fqu_inv_eq: ∀G,L1,L2,T. ⦃G, L1, T⦄ ⊐ ⦃G, L2, T⦄ → |L1| = |L2| → ⊥. -#G #L1 #L2 #T #H #H0 @(fqu_inv_eq_aux … H … H0) // (**) (* full auto fails *) -qed-. - -(* Advanced eliminators *****************************************************) - -lemma fqu_wf_ind: ∀R:relation3 …. ( - ∀G1,L1,T1. (∀G2,L2,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → R G2 L2 T2) → - R G1 L1 T1 - ) → ∀G1,L1,T1. R G1 L1 T1. -#R #HR @(f3_ind … fw) #x #IHx #G1 #L1 #T1 #H destruct /4 width=1 by fqu_fwd_fw/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/fquq.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/fquq.ma deleted file mode 100644 index 914af3702..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/fquq.ma +++ /dev/null @@ -1,64 +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/suptermopt_6.ma". -include "basic_2/substitution/fqu.ma". - -(* OPTIONAL SUPCLOSURE ******************************************************) - -(* activate genv *) -inductive fquq: tri_relation genv lenv term ≝ -| fquq_lref_O : ∀I,G,L,V. fquq G (L.ⓑ{I}V) (#0) G L V -| fquq_pair_sn: ∀I,G,L,V,T. fquq G L (②{I}V.T) G L V -| fquq_bind_dx: ∀a,I,G,L,V,T. fquq G L (ⓑ{a,I}V.T) G (L.ⓑ{I}V) T -| fquq_flat_dx: ∀I,G, L,V,T. fquq G L (ⓕ{I}V.T) G L T -| fquq_drop : ∀G,L,K,T,U,m. - ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → fquq G L U G K T -. - -interpretation - "optional structural successor (closure)" - 'SupTermOpt G1 L1 T1 G2 L2 T2 = (fquq G1 L1 T1 G2 L2 T2). - -(* Basic properties *********************************************************) - -lemma fquq_refl: tri_reflexive … fquq. -/2 width=3 by fquq_drop/ qed. - -lemma fqu_fquq: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 /2 width=3 by fquq_drop/ -qed. - -(* Basic forward lemmas *****************************************************) - -lemma fquq_fwd_fw: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ♯{G2, L2, T2} ≤ ♯{G1, L1, T1}. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 /2 width=1 by lt_to_le/ -#G1 #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 -lapply (drop_fwd_lw … HLK1) -HLK1 -lapply (lift_fwd_tw … HTU1) -HTU1 -/2 width=1 by le_plus, le_n/ -qed-. - -fact fquq_fwd_length_lref1_aux: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ∀i. T1 = #i → |L2| ≤ |L1|. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 // -[ #a #I #G #L #V #T #j #H destruct -| #G1 #L1 #K1 #T1 #U1 #m #HLK1 #HTU1 #i #H destruct - /2 width=3 by drop_fwd_length_le4/ -] -qed-. - -lemma fquq_fwd_length_lref1: ∀G1,G2,L1,L2,T2,i. ⦃G1, L1, #i⦄ ⊐⸮ ⦃G2, L2, T2⦄ → |L2| ≤ |L1|. -/2 width=7 by fquq_fwd_length_lref1_aux/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/fquq_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/fquq_alt.ma deleted file mode 100644 index 08a755b55..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/fquq_alt.ma +++ /dev/null @@ -1,59 +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/suptermoptalt_6.ma". -include "basic_2/substitution/fquq.ma". - -(* OPTIONAL SUPCLOSURE ******************************************************) - -(* alternative definition of fquq *) -definition fquqa: tri_relation genv lenv term ≝ tri_RC … fqu. - -interpretation - "optional structural successor (closure) alternative" - 'SupTermOptAlt G1 L1 T1 G2 L2 T2 = (fquqa G1 L1 T1 G2 L2 T2). - -(* Basic properties *********************************************************) - -lemma fquqa_refl: tri_reflexive … fquqa. -// qed. - -lemma fquqa_drop: ∀G,L,K,T,U,m. - ⬇[m] L ≡ K → ⬆[0, m] T ≡ U → ⦃G, L, U⦄ ⊐⊐⸮ ⦃G, K, T⦄. -#G #L #K #T #U #m #HLK #HTU elim (eq_or_gt m) -/3 width=5 by fqu_drop_lt, or_introl/ #H destruct ->(drop_inv_O2 … HLK) -L >(lift_inv_O2 … HTU) -T // -qed. - -(* Main properties **********************************************************) - -theorem fquq_fquqa: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⊐⸮ ⦃G2, L2, T2⦄. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -G1 -G2 -L1 -L2 -T1 -T2 -/2 width=3 by fquqa_drop, fqu_lref_O, fqu_pair_sn, fqu_bind_dx, fqu_flat_dx, or_introl/ -qed. - -(* Main inversion properties ************************************************) - -theorem fquqa_inv_fquq: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⊐⸮ ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄. -#G1 #G2 #L1 #L2 #T1 #T2 #H elim H -H /2 width=1 by fqu_fquq/ -* #H1 #H2 #H3 destruct // -qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma fquq_inv_gen: ∀G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊐⸮ ⦃G2, L2, T2⦄ → - ⦃G1, L1, T1⦄ ⊐ ⦃G2, L2, T2⦄ ∨ (∧∧ G1 = G2 & L1 = L2 & T1 = T2). -#G1 #G2 #L1 #L2 #T1 #T2 #H elim (fquq_fquqa … H) -H [| * ] -/2 width=1 by or_introl/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/gget.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/gget.ma deleted file mode 100644 index 1ea9ce7ff..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/gget.ma +++ /dev/null @@ -1,81 +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/rdrop_3.ma". -include "basic_2/grammar/genv.ma". - -(* GLOBAL ENVIRONMENT READING ***********************************************) - -inductive gget (m:nat): relation genv ≝ -| gget_gt: ∀G. |G| ≤ m → gget m G (⋆) -| gget_eq: ∀G. |G| = m + 1 → gget m G G -| gget_lt: ∀I,G1,G2,V. m < |G1| → gget m G1 G2 → gget m (G1. ⓑ{I} V) G2 -. - -interpretation "global reading" - 'RDrop m G1 G2 = (gget m G1 G2). - -(* basic inversion lemmas ***************************************************) - -lemma gget_inv_gt: ∀G1,G2,m. ⬇[m] G1 ≡ G2 → |G1| ≤ m → G2 = ⋆. -#G1 #G2 #m * -G1 -G2 // -[ #G #H >H -H >commutative_plus #H (**) (* lemma needed here *) - lapply (le_plus_to_le_r … 0 H) -H #H - lapply (le_n_O_to_eq … H) -H #H destruct -| #I #G1 #G2 #V #H1 #_ #H2 - lapply (le_to_lt_to_lt … H2 H1) -H2 -H1 normalize in ⊢ (? % ? → ?); >commutative_plus #H - lapply (lt_plus_to_lt_l … 0 H) -H #H - elim (lt_zero_false … H) -] -qed-. - -lemma gget_inv_eq: ∀G1,G2,m. ⬇[m] G1 ≡ G2 → |G1| = m + 1 → G1 = G2. -#G1 #G2 #m * -G1 -G2 // -[ #G #H1 #H2 >H2 in H1; -H2 >commutative_plus #H (**) (* lemma needed here *) - lapply (le_plus_to_le_r … 0 H) -H #H - lapply (le_n_O_to_eq … H) -H #H destruct -| #I #G1 #G2 #V #H1 #_ normalize #H2 - <(injective_plus_l … H2) in H1; -H2 #H - elim (lt_refl_false … H) -] -qed-. - -fact gget_inv_lt_aux: ∀I,G,G1,G2,V,m. ⬇[m] G ≡ G2 → G = G1. ⓑ{I} V → - m < |G1| → ⬇[m] G1 ≡ G2. -#I #G #G1 #G2 #V #m * -G -G2 -[ #G #H1 #H destruct #H2 - lapply (le_to_lt_to_lt … H1 H2) -H1 -H2 normalize in ⊢ (? % ? → ?); >commutative_plus #H - lapply (lt_plus_to_lt_l … 0 H) -H #H - elim (lt_zero_false … H) -| #G #H1 #H2 destruct >(injective_plus_l … H1) -H1 #H - elim (lt_refl_false … H) -| #J #G #G2 #W #_ #HG2 #H destruct // -] -qed-. - -lemma gget_inv_lt: ∀I,G1,G2,V,m. - ⬇[m] G1. ⓑ{I} V ≡ G2 → m < |G1| → ⬇[m] G1 ≡ G2. -/2 width=5 by gget_inv_lt_aux/ qed-. - -(* Basic properties *********************************************************) - -lemma gget_total: ∀m,G1. ∃G2. ⬇[m] G1 ≡ G2. -#m #G1 elim G1 -G1 /3 width=2 by gget_gt, ex_intro/ -#I #V #G1 * #G2 #HG12 -elim (lt_or_eq_or_gt m (|G1|)) #Hm -[ /3 width=2 by gget_lt, ex_intro/ -| destruct /3 width=2 by gget_eq, ex_intro/ -| @ex_intro [2: @gget_gt normalize /2 width=1 by/ | skip ] (**) (* explicit constructor *) -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/gget_gget.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/gget_gget.ma deleted file mode 100644 index 7bd39f09b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/gget_gget.ma +++ /dev/null @@ -1,40 +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/gget.ma". - -(* GLOBAL ENVIRONMENT READING ***********************************************) - -(* Main properties **********************************************************) - -theorem gget_mono: ∀G,G1,m. ⬇[m] G ≡ G1 → ∀G2. ⬇[m] G ≡ G2 → G1 = G2. -#G #G1 #m #H elim H -G -G1 -[ #G #Hm #G2 #H - >(gget_inv_gt … H Hm) -H -Hm // -| #G #Hm #G2 #H - >(gget_inv_eq … H Hm) -H -Hm // -| #I #G #G1 #V #Hm #_ #IHG1 #G2 #H - lapply (gget_inv_lt … H Hm) -H -Hm /2 width=1 by/ -] -qed-. - -lemma gget_dec: ∀G1,G2,m. Decidable (⬇[m] G1 ≡ G2). -#G1 #G2 #m -elim (gget_total m G1) #G #HG1 -elim (eq_genv_dec G G2) #HG2 -[ destruct /2 width=1 by or_introl/ -| @or_intror #HG12 - lapply (gget_mono … HG1 … HG12) -HG1 -HG12 /2 width=1 by/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lift.ma deleted file mode 100644 index 0d5f1dfe4..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift.ma +++ /dev/null @@ -1,408 +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 "ground_2/ynat/ynat_plus.ma". -include "basic_2/notation/relations/rlift_4.ma". -include "basic_2/grammar/term_weight.ma". -include "basic_2/grammar/term_simple.ma". - -(* BASIC TERM RELOCATION ****************************************************) - -(* Basic_1: includes: - lift_sort lift_lref_lt lift_lref_ge lift_bind lift_flat -*) -inductive lift: relation4 ynat nat term term ≝ -| lift_sort : ∀k,l,m. lift l m (⋆k) (⋆k) -| lift_lref_lt: ∀i,l,m. yinj i < l → lift l m (#i) (#i) -| lift_lref_ge: ∀i,l,m. l ≤ yinj i → lift l m (#i) (#(i + m)) -| lift_gref : ∀p,l,m. lift l m (§p) (§p) -| lift_bind : ∀a,I,V1,V2,T1,T2,l,m. - lift l m V1 V2 → lift (⫯l) m T1 T2 → - lift l m (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) -| lift_flat : ∀I,V1,V2,T1,T2,l,m. - lift l m V1 V2 → lift l m T1 T2 → - lift l m (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) -. - -interpretation "relocation" 'RLift l m T1 T2 = (lift l m T1 T2). - -(* Basic inversion lemmas ***************************************************) - -fact lift_inv_O2_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → m = 0 → T1 = T2. -#l #m #T1 #T2 #H elim H -l -m -T1 -T2 /3 width=1 by eq_f2/ -qed-. - -lemma lift_inv_O2: ∀l,T1,T2. ⬆[l, 0] T1 ≡ T2 → T1 = T2. -/2 width=4 by lift_inv_O2_aux/ qed-. - -fact lift_inv_sort1_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → ∀k. T1 = ⋆k → T2 = ⋆k. -#l #m #T1 #T2 * -l -m -T1 -T2 // -[ #i #l #m #_ #k #H destruct -| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct -] -qed-. - -lemma lift_inv_sort1: ∀l,m,T2,k. ⬆[l, m] ⋆k ≡ T2 → T2 = ⋆k. -/2 width=5 by lift_inv_sort1_aux/ qed-. - -fact lift_inv_lref1_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → ∀i. T1 = #i → - (i < l ∧ T2 = #i) ∨ (l ≤ i ∧ T2 = #(i + m)). -#l #m #T1 #T2 * -l -m -T1 -T2 -[ #k #l #m #i #H destruct -| #j #l #m #Hj #i #Hi destruct /3 width=1 by or_introl,conj/ -| #j #l #m #Hj #i #Hi destruct /3 width=1 by or_intror,conj/ -| #p #l #m #i #H destruct -| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #i #H destruct -| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #i #H destruct -] -qed-. - -lemma lift_inv_lref1: ∀l,m,T2,i. ⬆[l, m] #i ≡ T2 → - (i < l ∧ T2 = #i) ∨ (l ≤ i ∧ T2 = #(i + m)). -/2 width=3 by lift_inv_lref1_aux/ qed-. - -lemma lift_inv_lref1_lt: ∀l,m,T2,i. ⬆[l, m] #i ≡ T2 → i < l → T2 = #i. -#l #m #T2 #i #H elim (lift_inv_lref1 … H) -H * // -#Hli #_ #Hil elim (ylt_yle_false … Hli) -Hli // -qed-. - -lemma lift_inv_lref1_ge: ∀l,m,T2,i. ⬆[l, m] #i ≡ T2 → l ≤ i → T2 = #(i + m). -#l #m #T2 #i #H elim (lift_inv_lref1 … H) -H * // -#Hil #_ #Hli elim (ylt_yle_false … Hli) -Hli // -qed-. - -fact lift_inv_gref1_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → ∀p. T1 = §p → T2 = §p. -#l #m #T1 #T2 * -l -m -T1 -T2 // -[ #i #l #m #_ #k #H destruct -| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct -] -qed-. - -lemma lift_inv_gref1: ∀l,m,T2,p. ⬆[l, m] §p ≡ T2 → T2 = §p. -/2 width=5 by lift_inv_gref1_aux/ qed-. - -fact lift_inv_bind1_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → - ∀a,I,V1,U1. T1 = ⓑ{a,I}V1.U1 → - ∃∃V2,U2. ⬆[l, m] V1 ≡ V2 & ⬆[⫯l, m] U1 ≡ U2 & - T2 = ⓑ{a,I}V2.U2. -#l #m #T1 #T2 * -l -m -T1 -T2 -[ #k #l #m #a #I #V1 #U1 #H destruct -| #i #l #m #_ #a #I #V1 #U1 #H destruct -| #i #l #m #_ #a #I #V1 #U1 #H destruct -| #p #l #m #a #I #V1 #U1 #H destruct -| #b #J #W1 #W2 #T1 #T2 #l #m #HW #HT #a #I #V1 #U1 #H destruct /2 width=5 by ex3_2_intro/ -| #J #W1 #W2 #T1 #T2 #l #m #_ #HT #a #I #V1 #U1 #H destruct -] -qed-. - -lemma lift_inv_bind1: ∀l,m,T2,a,I,V1,U1. ⬆[l, m] ⓑ{a,I}V1.U1 ≡ T2 → - ∃∃V2,U2. ⬆[l, m] V1 ≡ V2 & ⬆[⫯l, m] U1 ≡ U2 & - T2 = ⓑ{a,I}V2.U2. -/2 width=3 by lift_inv_bind1_aux/ qed-. - -fact lift_inv_flat1_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → - ∀I,V1,U1. T1 = ⓕ{I}V1.U1 → - ∃∃V2,U2. ⬆[l, m] V1 ≡ V2 & ⬆[l, m] U1 ≡ U2 & - T2 = ⓕ{I}V2.U2. -#l #m #T1 #T2 * -l -m -T1 -T2 -[ #k #l #m #I #V1 #U1 #H destruct -| #i #l #m #_ #I #V1 #U1 #H destruct -| #i #l #m #_ #I #V1 #U1 #H destruct -| #p #l #m #I #V1 #U1 #H destruct -| #a #J #W1 #W2 #T1 #T2 #l #m #_ #_ #I #V1 #U1 #H destruct -| #J #W1 #W2 #T1 #T2 #l #m #HW #HT #I #V1 #U1 #H destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -lemma lift_inv_flat1: ∀l,m,T2,I,V1,U1. ⬆[l, m] ⓕ{I}V1.U1 ≡ T2 → - ∃∃V2,U2. ⬆[l, m] V1 ≡ V2 & ⬆[l, m] U1 ≡ U2 & - T2 = ⓕ{I}V2.U2. -/2 width=3 by lift_inv_flat1_aux/ qed-. - -fact lift_inv_sort2_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → ∀k. T2 = ⋆k → T1 = ⋆k. -#l #m #T1 #T2 * -l -m -T1 -T2 // -[ #i #l #m #_ #k #H destruct -| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct -| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #k #H destruct -] -qed-. - -(* Basic_1: was: lift_gen_sort *) -lemma lift_inv_sort2: ∀l,m,T1,k. ⬆[l, m] T1 ≡ ⋆k → T1 = ⋆k. -/2 width=5 by lift_inv_sort2_aux/ qed-. - -fact lift_inv_lref2_aux: ∀l,m,T1,T2. ⬆[l, m] T1 ≡ T2 → ∀i. T2 = #i → - (i < l ∧ T1 = #i) ∨ (l + m ≤ i ∧ T1 = #(i - m)). -#l #m #T1 #T2 * -l -m -T1 -T2 -[ #k #l #m #i #H destruct -| #j #l #m #Hj #i #Hi destruct /3 width=1 by or_introl, conj/ -| #j #l #m #Hj #i #Hi destruct (plus_minus_m_m i m) in ⊢ (? ? ? ? %); -elim (yle_inv_plus_inj2 … H) -H #Hlim #H -lapply (yle_inv_inj … H) -H /2 width=1 by lift_lref_ge/ -qed. - -lemma lift_lref_ge_minus_eq: ∀l,m,i,j. l + yinj m ≤ yinj i → j = i - m → ⬆[l, m] #j ≡ #i. -/2 width=1 by lift_lref_ge_minus/ qed-. - -(* Basic_1: was: lift_r *) -lemma lift_refl: ∀T,l. ⬆[l, 0] T ≡ T. -#T elim T -T -[ * #i // #l elim (ylt_split i l) /2 width=1 by lift_lref_lt, lift_lref_ge/ -| * /2 width=1 by lift_bind, lift_flat/ -] -qed. - -(* Basic_2b: first lemma *) -lemma lift_Y1: ∀T,m. ⬆[∞, m] T ≡ T. -#T elim T -T * /2 width=1 by lift_lref_lt, lift_bind, lift_flat/ -qed. - -lemma lift_total: ∀T1,l,m. ∃T2. ⬆[l, m] T1 ≡ T2. -#T1 elim T1 -T1 -[ * #i /2 width=2 by lift_sort, lift_gref, ex_intro/ - #l #m elim (ylt_split i l) /3 width=2 by lift_lref_lt, lift_lref_ge, ex_intro/ -| * [ #a ] #I #V1 #T1 #IHV1 #IHT1 #l #m - elim (IHV1 l m) -IHV1 #V2 #HV12 - [ elim (IHT1 (⫯l) m) -IHT1 /3 width=2 by lift_bind, ex_intro/ - | elim (IHT1 l m) -IHT1 /3 width=2 by lift_flat, ex_intro/ - ] -] -qed. - -(* Basic_1: was: lift_free (right to left) *) -lemma lift_split: ∀l1,m2,T1,T2. ⬆[l1, m2] T1 ≡ T2 → - ∀l2,m1. l1 ≤ l2 → l2 ≤ l1 + yinj m1 → m1 ≤ m2 → - ∃∃T. ⬆[l1, m1] T1 ≡ T & ⬆[l2, m2 - m1] T ≡ T2. -#l1 #m2 #T1 #T2 #H elim H -l1 -m2 -T1 -T2 -[ /3 width=3 by lift_sort, ex2_intro/ -| #i #l1 #m2 #Hil1 #l2 #m1 #Hl12 #_ #_ - lapply (ylt_yle_trans … Hl12 Hil1) -Hl12 #Hil2 /4 width=3 by lift_lref_lt, ex2_intro/ -| #i #l1 #m2 #Hil1 #l2 #m1 #_ #Hl21 #Hm12 - lapply (yle_trans … Hl21 (i+m1) ?) /2 width=1 by monotonic_yle_plus_dx/ -Hl21 #Hl21 - >(plus_minus_m_m m2 m1 ?) /3 width=3 by lift_lref_ge, ex2_intro/ -| /3 width=3 by lift_gref, ex2_intro/ -| #a #I #V1 #V2 #T1 #T2 #l1 #m2 #_ #_ #IHV #IHT #l2 #m1 #Hl12 #Hl21 #Hm12 - elim (IHV … Hl12 Hl21 Hm12) -IHV #V0 #HV0a #HV0b - elim (IHT (⫯l2) … ? ? Hm12) /3 width=5 by lift_bind, yle_succ, ex2_intro/ -| #I #V1 #V2 #T1 #T2 #l1 #m2 #_ #_ #IHV #IHT #l2 #m1 #Hl12 #Hl21 #Hm12 - elim (IHV … Hl12 Hl21 Hm12) -IHV #V0 #HV0a #HV0b - elim (IHT l2 … ? ? Hm12) /3 width=5 by lift_flat, ex2_intro/ -] -qed. - -(* Basic_1: was only: dnf_dec2 dnf_dec *) -lemma is_lift_dec: ∀T2,l,m. Decidable (∃T1. ⬆[l, m] T1 ≡ T2). -#T1 elim T1 -T1 -[ * [1,3: /3 width=2 by lift_sort, lift_gref, ex_intro, or_introl/ ] #i #l #m - elim (ylt_split i l) #Hli - [ /4 width=3 by lift_lref_lt, ex_intro, or_introl/ - | elim (ylt_split i (l + m)) #Hilm - [ @or_intror * #T1 #H elim (lift_inv_lref2_be … H Hli Hilm) - | -Hli /4 width=2 by lift_lref_ge_minus, ex_intro, or_introl/ - ] - ] -| * [ #a ] #I #V2 #T2 #IHV2 #IHT2 #l #m - [ elim (IHV2 l m) -IHV2 - [ * #V1 #HV12 elim (IHT2 (⫯l) m) -IHT2 - [ * #T1 #HT12 @or_introl /3 width=2 by lift_bind, ex_intro/ - | -V1 #HT2 @or_intror * #X #H - elim (lift_inv_bind2 … H) -H /3 width=2 by ex_intro/ - ] - | -IHT2 #HV2 @or_intror * #X #H - elim (lift_inv_bind2 … H) -H /3 width=2 by ex_intro/ - ] - | elim (IHV2 l m) -IHV2 - [ * #V1 #HV12 elim (IHT2 l m) -IHT2 - [ * #T1 #HT12 /4 width=2 by lift_flat, ex_intro, or_introl/ - | -V1 #HT2 @or_intror * #X #H - elim (lift_inv_flat2 … H) -H /3 width=2 by ex_intro/ - ] - | -IHT2 #HV2 @or_intror * #X #H - elim (lift_inv_flat2 … H) -H /3 width=2 by ex_intro/ - ] - ] -] -qed. - -(* Basic_1: removed theorems 7: - lift_head lift_gen_head - lift_weight_map lift_weight lift_weight_add lift_weight_add_O - lift_tlt_dx -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_lift.ma deleted file mode 100644 index 70ec68826..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_lift.ma +++ /dev/null @@ -1,226 +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/lift.ma". - -(* BASIC TERM RELOCATION ****************************************************) - -(* Main properties ***********************************************************) - -(* Basic_1: was: lift_inj *) -theorem lift_inj: ∀l,m,T1,U. ⬆[l, m] T1 ≡ U → ∀T2. ⬆[l, m] T2 ≡ U → T1 = T2. -#l #m #T1 #U #H elim H -l -m -T1 -U -[ #k #l #m #X #HX - lapply (lift_inv_sort2 … HX) -HX // -| #i #l #m #Hil #X #HX - lapply (lift_inv_lref2_lt … HX ?) -HX // -| #i #l #m #Hli #X #HX - lapply (lift_inv_lref2_ge … HX ?) -HX /2 width=1 by monotonic_yle_plus_dx/ -| #p #l #m #X #HX - lapply (lift_inv_gref2 … HX) -HX // -| #a #I #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_bind2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1 by eq_f2/ -| #I #V1 #V2 #T1 #T2 #l #m #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #HX destruct /3 width=1 by eq_f2/ -] -qed-. - -(* Basic_1: was: lift_gen_lift *) -theorem lift_div_le: ∀l1,m1,T1,T. ⬆[l1, m1] T1 ≡ T → - ∀l2,m2,T2. ⬆[l2 + m1, m2] T2 ≡ T → - l1 ≤ l2 → - ∃∃T0. ⬆[l1, m1] T0 ≡ T2 & ⬆[l2, m2] T0 ≡ T1. -#l1 #m1 #T1 #T #H elim H -l1 -m1 -T1 -T -[ #k #l1 #m1 #l2 #m2 #T2 #Hk #Hl12 - lapply (lift_inv_sort2 … Hk) -Hk #Hk destruct /3 width=3 by lift_sort, ex2_intro/ -| #i #l1 #m1 #Hil1 #l2 #m2 #T2 #Hi #Hl12 - lapply (ylt_yle_trans … Hl12 Hil1) -Hl12 #Hil2 - lapply (lift_inv_lref2_lt … Hi ?) -Hi /3 width=3 by lift_lref_lt, ylt_plus_dx1_trans, ex2_intro/ -| #i #l1 #m1 #Hil1 #l2 #m2 #T2 #Hi #Hl12 - elim (lift_inv_lref2 … Hi) -Hi * yplus_comm_23 in Hil2; #H lapply ( yle_inv_monotonic_plus_dx … H) -H #H - elim (yle_inv_plus_inj2 … H) -H >yminus_inj #Hl2im2 #H - lapply (yle_inv_inj … H) -H #Hm2i - lapply (yle_trans … Hl12 … Hl2im2) -Hl12 #Hl1im2 - >le_plus_minus_comm // >(plus_minus_m_m i m2) in ⊢ (? ? ? %); - /3 width=3 by lift_lref_ge, ex2_intro/ - ] -| #p #l1 #m1 #l2 #m2 #T2 #Hk #Hl12 - lapply (lift_inv_gref2 … Hk) -Hk #Hk destruct /3 width=3 by lift_gref, ex2_intro/ -| #a #I #W1 #W #U1 #U #l1 #m1 #_ #_ #IHW #IHU #l2 #m2 #T2 #H #Hl12 - lapply (lift_inv_bind2 … H) -H * #W2 #U2 #HW2 #HU2 #H destruct - elim (IHW … HW2) // -IHW -HW2 #W0 #HW2 #HW1 - (lift_inv_sort2 … H) -H /2 width=3 by lift_sort, ex2_intro/ -| #i #l1 #m1 #Hil1 #m #m2 #T2 #H #Hm1 #Hm1m2 - >(lift_inv_lref2_lt … H) -H /3 width=3 by ylt_plus_dx1_trans, lift_lref_lt, ex2_intro/ -| #i #l1 #m1 #Hil1 #m #m2 #T2 #H #Hm1 #Hm1m2 - elim (ylt_split (i+m1) (l1+m+m2)) #H0 - [ elim (lift_inv_lref2_be … H) -H /3 width=2 by monotonic_yle_plus, yle_inj/ - | >(lift_inv_lref2_ge … H ?) -H // - lapply (yle_plus2_to_minus_inj2 … H0) #Hl1m21i - elim (yle_inv_plus_inj2 … H0) -H0 #Hl1m12 #Hm2im1 - @ex2_intro [2: /2 width=1 by lift_lref_ge_minus/ | skip ] - @lift_lref_ge_minus_eq - [ yplus_minus_assoc_inj /2 width=1 by yle_inj/ - | /2 width=1 by minus_le_minus_minus_comm/ - ] - ] -| #p #l1 #m1 #m #m2 #T2 #H >(lift_inv_gref2 … H) -H /2 width=3 by lift_gref, ex2_intro/ -| #a #I #V1 #V #T1 #T #l1 #m1 #_ #_ #IHV1 #IHT1 #m #m2 #X #H #Hm1 #Hm1m2 - elim (lift_inv_bind2 … H) -H #V2 #T2 #HV2 (lift_inv_sort1 … HT2) -HT2 // -| #i #l1 #m1 #Hil1 #l2 #m2 #T2 #HT2 #Hl12 #_ - lapply (ylt_yle_trans … Hl12 Hil1) -Hl12 #Hil2 - lapply (lift_inv_lref1_lt … HT2 Hil2) /2 width=1 by lift_lref_lt/ -| #i #l1 #m1 #Hil1 #l2 #m2 #T2 #HT2 #_ #Hl21 - lapply (lift_inv_lref1_ge … HT2 ?) -HT2 - [ @(yle_trans … Hl21) -Hl21 /2 width=1 by monotonic_yle_plus_dx/ - | -Hl21 /2 width=1 by lift_lref_ge/ - ] -| #p #l1 #m1 #l2 #m2 #T2 #HT2 #_ #_ - >(lift_inv_gref1 … HT2) -HT2 // -| #a #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hl12 #Hl21 - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10 - lapply (IHT12 … HT20 ? ?) /2 width=1 by lift_bind, yle_succ/ (**) (* full auto a bit slow *) -| #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hl12 #Hl21 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - lapply (IHV12 … HV20 ? ?) // -IHV12 -HV20 #HV10 - lapply (IHT12 … HT20 ? ?) /2 width=1 by lift_flat/ (**) (* full auto a bit slow *) -] -qed. - -(* Basic_1: was: lift_d (right to left) *) -theorem lift_trans_le: ∀l1,m1,T1,T. ⬆[l1, m1] T1 ≡ T → - ∀l2,m2,T2. ⬆[l2, m2] T ≡ T2 → l2 ≤ l1 → - ∃∃T0. ⬆[l2, m2] T1 ≡ T0 & ⬆[l1 + m2, m1] T0 ≡ T2. -#l1 #m1 #T1 #T #H elim H -l1 -m1 -T1 -T -[ #k #l1 #m1 #l2 #m2 #X #HX #_ - >(lift_inv_sort1 … HX) -HX /2 width=3 by lift_sort, ex2_intro/ -| #i #l1 #m1 #Hil1 #l2 #m2 #X #HX #_ - lapply (ylt_yle_trans … (l1+m2) ? Hil1) // #Him2 - elim (lift_inv_lref1 … HX) -HX * #Hil2 #HX destruct - /4 width=3 by monotonic_ylt_plus_dx, monotonic_yle_plus_dx, lift_lref_ge_minus, lift_lref_lt, ex2_intro/ -| #i #l1 #m1 #Hil1 #l2 #m2 #X #HX #Hl21 - lapply (yle_trans … Hl21 … Hil1) -Hl21 #Hil2 - lapply (lift_inv_lref1_ge … HX ?) -HX /2 width=3 by yle_plus_dx1_trans/ #HX destruct - >plus_plus_comm_23 /4 width=3 by monotonic_yle_plus_dx, lift_lref_ge_minus, lift_lref_ge, ex2_intro/ -| #p #l1 #m1 #l2 #m2 #X #HX #_ - >(lift_inv_gref1 … HX) -HX /2 width=3 by lift_gref, ex2_intro/ -| #a #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hl21 - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - elim (IHV12 … HV20) -IHV12 -HV20 // - elim (IHT12 … HT20) -IHT12 -HT20 /3 width=5 by lift_bind, yle_succ, ex2_intro/ -| #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hl21 - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - elim (IHV12 … HV20) -IHV12 -HV20 // - elim (IHT12 … HT20) -IHT12 -HT20 /3 width=5 by lift_flat, ex2_intro/ -] -qed. - -(* Basic_1: was: lift_d (left to right) *) -theorem lift_trans_ge: ∀l1,m1,T1,T. ⬆[l1, m1] T1 ≡ T → - ∀l2,m2,T2. ⬆[l2, m2] T ≡ T2 → l1 + m1 ≤ l2 → - ∃∃T0. ⬆[l2 - m1, m2] T1 ≡ T0 & ⬆[l1, m1] T0 ≡ T2. -#l1 #m1 #T1 #T #H elim H -l1 -m1 -T1 -T -[ #k #l1 #m1 #l2 #m2 #X #HX #_ - >(lift_inv_sort1 … HX) -HX /2 width=3 by lift_sort, ex2_intro/ -| #i #l1 #m1 #Hil1 #l2 #m2 #X #HX #Hlml - lapply (ylt_yle_trans … (l1+m1) ? Hil1) // #Hil1m - lapply (ylt_yle_trans … (l2-m1) ? Hil1) /2 width=1 by yle_plus1_to_minus_inj2/ #Hil2m - lapply (ylt_yle_trans … Hlml Hil1m) -Hil1m -Hlml #Hil2 - lapply (lift_inv_lref1_lt … HX ?) -HX // #HX destruct /3 width=3 by lift_lref_lt, ex2_intro/ -| #i #l1 #m1 #Hil1 #l2 #m2 #X #HX #_ - elim (lift_inv_lref1 … HX) -HX * (lift_inv_gref1 … HX) -HX /2 width=3 by lift_gref, ex2_intro/ -| #a #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hlml - elim (yle_inv_plus_inj2 … Hlml) #Hlm #Hml - elim (lift_inv_bind1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - elim (IHV12 … HV20) -IHV12 -HV20 // - elim (IHT12 … HT20) -IHT12 -HT20 /2 width=1 by yle_succ/ -Hlml - #T >yminus_succ1_inj /3 width=5 by lift_bind, ex2_intro/ -| #I #V1 #V2 #T1 #T2 #l1 #m1 #_ #_ #IHV12 #IHT12 #l2 #m2 #X #HX #Hlml - elim (lift_inv_flat1 … HX) -HX #V0 #T0 #HV20 #HT20 #HX destruct - elim (IHV12 … HV20) -IHV12 -HV20 // - elim (IHT12 … HT20) -IHT12 -HT20 /3 width=5 by lift_flat, ex2_intro/ -] -qed. - -(* Advanced properties ******************************************************) - -lemma lift_conf_O1: ∀T,T1,l1,m1. ⬆[l1, m1] T ≡ T1 → ∀T2,m2. ⬆[0, m2] T ≡ T2 → - ∃∃T0. ⬆[0, m2] T1 ≡ T0 & ⬆[l1 + m2, m1] T2 ≡ T0. -#T #T1 #l1 #m1 #HT1 #T2 #m2 #HT2 -elim (lift_total T1 0 m2) #T0 #HT10 -elim (lift_trans_le … HT1 … HT10) -HT1 // #X #HTX #HT20 -lapply (lift_mono … HTX … HT2) -T #H destruct /2 width=3 by ex2_intro/ -qed. - -lemma lift_conf_be: ∀T,T1,l,m1. ⬆[l, m1] T ≡ T1 → ∀T2,m2. ⬆[l, m2] T ≡ T2 → - m1 ≤ m2 → ⬆[l + yinj m1, m2 - m1] T1 ≡ T2. -#T #T1 #l #m1 #HT1 #T2 #m2 #HT2 #Hm12 -elim (lift_split … HT2 (l+m1) m1) -HT2 // #X #H ->(lift_mono … H … HT1) -T // -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_lift_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_lift_vector.ma deleted file mode 100644 index 577682957..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_lift_vector.ma +++ /dev/null @@ -1,30 +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/lift_lift.ma". -include "basic_2/substitution/lift_vector.ma". - -(* BASIC TERM VECTOR RELOCATION *********************************************) - -(* Main properties ***********************************************************) - -theorem liftv_mono: ∀Ts,U1s,l,m. ⬆[l,m] Ts ≡ U1s → - ∀U2s:list term. ⬆[l,m] Ts ≡ U2s → U1s = U2s. -#Ts #U1s #l #m #H elim H -Ts -U1s -[ #U2s #H >(liftv_inv_nil1 … H) -H // -| #Ts #U1s #T #U1 #HTU1 #_ #IHTU1s #X #H destruct - elim (liftv_inv_cons1 … H) -H #U2 #U2s #HTU2 #HTU2s #H destruct - >(lift_mono … HTU1 … HTU2) -T /3 width=1 by eq_f/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_neg.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_neg.ma deleted file mode 100644 index 6c797befb..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_neg.ma +++ /dev/null @@ -1,68 +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/lift.ma". - -(* BASIC TERM RELOCATION ****************************************************) - -(* Properties on negated basic relocation ***********************************) - -lemma nlift_lref_be_SO: ∀X,i. ⬆[yinj i, 1] X ≡ #i → ⊥. -/3 width=7 by lift_inv_lref2_be, ylt_inj/ qed-. - -lemma nlift_bind_sn: ∀W,l,m. (∀V. ⬆[l, m] V ≡ W → ⊥) → - ∀a,I,U. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥). -#W #l #m #HW #a #I #U #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/ -qed-. - -lemma nlift_bind_dx: ∀U,l,m. (∀T. ⬆[⫯l, m] T ≡ U → ⊥) → - ∀a,I,W. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥). -#U #l #m #HU #a #I #W #X #H elim (lift_inv_bind2 … H) -H /2 width=2 by/ -qed-. - -lemma nlift_flat_sn: ∀W,l,m. (∀V. ⬆[l, m] V ≡ W → ⊥) → - ∀I,U. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥). -#W #l #m #HW #I #U #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/ -qed-. - -lemma nlift_flat_dx: ∀U,l,m. (∀T. ⬆[l, m] T ≡ U → ⊥) → - ∀I,W. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥). -#U #l #m #HU #I #W #X #H elim (lift_inv_flat2 … H) -H /2 width=2 by/ -qed-. - -(* Inversion lemmas on negated basic relocation *****************************) - -lemma nlift_inv_lref_be_SO: ∀i,j. (∀X. ⬆[i, 1] X ≡ #j → ⊥) → yinj j = i. -* [2: #j #H elim H -H // ] -#i #j elim (lt_or_eq_or_gt i j) // #Hij #H -[ elim (H (#(j-1))) -H /3 width=1 by lift_lref_ge_minus, yle_inj/ -| elim (H (#j)) -H /3 width=1 by lift_lref_lt, ylt_inj/ -] -qed-. - -lemma nlift_inv_bind: ∀a,I,W,U,l,m. (∀X. ⬆[l, m] X ≡ ⓑ{a,I}W.U → ⊥) → - (∀V. ⬆[l, m] V ≡ W → ⊥) ∨ (∀T. ⬆[⫯l, m] T ≡ U → ⊥). -#a #I #W #U #l #m #H elim (is_lift_dec W l m) -[ * /4 width=2 by lift_bind, or_intror/ -| /4 width=2 by ex_intro, or_introl/ -] -qed-. - -lemma nlift_inv_flat: ∀I,W,U,l,m. (∀X. ⬆[l, m] X ≡ ⓕ{I}W.U → ⊥) → - (∀V. ⬆[l, m] V ≡ W → ⊥) ∨ (∀T. ⬆[l, m] T ≡ U → ⊥). -#I #W #U #l #m #H elim (is_lift_dec W l m) -[ * /4 width=2 by lift_flat, or_intror/ -| /4 width=2 by ex_intro, or_introl/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_vector.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_vector.ma deleted file mode 100644 index 6876229e4..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lift_vector.ma +++ /dev/null @@ -1,62 +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/grammar/term_vector.ma". -include "basic_2/substitution/lift.ma". - -(* BASIC TERM VECTOR RELOCATION *********************************************) - -inductive liftv (l) (m): relation (list term) ≝ -| liftv_nil : liftv l m (◊) (◊) -| liftv_cons: ∀T1s,T2s,T1,T2. - ⬆[l, m] T1 ≡ T2 → liftv l m T1s T2s → - liftv l m (T1 @ T1s) (T2 @ T2s) -. - -interpretation "relocation (vector)" 'RLift l m T1s T2s = (liftv l m T1s T2s). - -(* Basic inversion lemmas ***************************************************) - -fact liftv_inv_nil1_aux: ∀T1s,T2s,l,m. ⬆[l, m] T1s ≡ T2s → T1s = ◊ → T2s = ◊. -#T1s #T2s #l #m * -T1s -T2s // -#T1s #T2s #T1 #T2 #_ #_ #H destruct -qed-. - -lemma liftv_inv_nil1: ∀T2s,l,m. ⬆[l, m] ◊ ≡ T2s → T2s = ◊. -/2 width=5 by liftv_inv_nil1_aux/ qed-. - -fact liftv_inv_cons1_aux: ∀T1s,T2s,l,m. ⬆[l, m] T1s ≡ T2s → - ∀U1,U1s. T1s = U1 @ U1s → - ∃∃U2,U2s. ⬆[l, m] U1 ≡ U2 & ⬆[l, m] U1s ≡ U2s & - T2s = U2 @ U2s. -#T1s #T2s #l #m * -T1s -T2s -[ #U1 #U1s #H destruct -| #T1s #T2s #T1 #T2 #HT12 #HT12s #U1 #U1s #H destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -lemma liftv_inv_cons1: ∀U1,U1s,T2s,l,m. ⬆[l, m] U1 @ U1s ≡ T2s → - ∃∃U2,U2s. ⬆[l, m] U1 ≡ U2 & ⬆[l, m] U1s ≡ U2s & - T2s = U2 @ U2s. -/2 width=3 by liftv_inv_cons1_aux/ qed-. - -(* Basic properties *********************************************************) - -lemma liftv_total: ∀l,m. ∀T1s:list term. ∃T2s. ⬆[l, m] T1s ≡ T2s. -#l #m #T1s elim T1s -T1s -[ /2 width=2 by liftv_nil, ex_intro/ -| #T1 #T1s * #T2s #HT12s - elim (lift_total T1 l m) /3 width=2 by liftv_cons, ex_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn.ma index 977fd887b..82664b828 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn.ma @@ -85,5 +85,5 @@ qed-. (* Basic forward lemmas *****************************************************) lemma lpx_sn_fwd_length: ∀R,L1,L2. lpx_sn R L1 L2 → |L1| = |L2|. -#R #L1 #L2 #H elim H -L1 -L2 normalize // +#R #L1 #L2 #H elim H -L1 -L2 // qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_alt.ma index e6d268bcd..5099ffa1a 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_alt.ma @@ -46,7 +46,7 @@ lemma lpx_sn_alt_inv_pair1: ∀R,I,L2,K1,V1. lpx_sn_alt R (K1.ⓑ{I}V1) L2 → elim (IH I1 I2 K1 K2 V1 V2 0) // #H #HV12 destruct @(ex3_2_intro … K2 V2) // -HV12 @conj // -HK12 -#J1 #J2 #L1 #L2 #W1 #W2 #i #HKL1 #HKL2 elim (IH J1 J2 L1 L2 W1 W2 (i+1)) -IH +#J1 #J2 #L1 #L2 #W1 #W2 #i #HKL1 #HKL2 elim (IH J1 J2 L1 L2 W1 W2 (⫯i)) -IH /2 width=1 by drop_drop, conj/ qed-. @@ -63,7 +63,7 @@ lemma lpx_sn_alt_inv_pair2: ∀R,I,L1,K2,V2. lpx_sn_alt R L1 (K2.ⓑ{I}V2) → elim (IH I1 I2 K1 K2 V1 V2 0) // #H #HV12 destruct @(ex3_2_intro … K1 V1) // -HV12 @conj // -HK12 -#J1 #J2 #L1 #L2 #W1 #W2 #i #HKL1 #HKL2 elim (IH J1 J2 L1 L2 W1 W2 (i+1)) -IH +#J1 #J2 #L1 #L2 #W1 #W2 #i #HKL1 #HKL2 elim (IH J1 J2 L1 L2 W1 W2 (⫯i)) -IH /2 width=1 by drop_drop, conj/ qed-. @@ -79,13 +79,15 @@ lemma lpx_sn_alt_pair: ∀R,I,L1,L2,V1,V2. lpx_sn_alt R L1 L2 → R L1 V1 V2 → lpx_sn_alt R (L1.ⓑ{I}V1) (L2.ⓑ{I}V2). #R #I #L1 #L2 #V1 #V2 #H #HV12 elim H -H -#HL12 #IH @conj normalize // -#I1 #I2 #K1 #K2 #W1 #W2 #i @(nat_ind_plus … i) -i +#HL12 #IH @conj // +#I1 #I2 #K1 #K2 #W1 #W2 #i @(ynat_ind … i) -i [ #HLK1 #HLK2 lapply (drop_inv_O2 … HLK1) -HLK1 #H destruct lapply (drop_inv_O2 … HLK2) -HLK2 #H destruct /2 width=1 by conj/ | -HL12 -HV12 /3 width=6 by drop_inv_drop1/ +| #H lapply (drop_fwd_Y2 … H) -H + #H elim (ylt_yle_false … H) -H // ] qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_drop.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_drop.ma index 6eb2dd8b2..612f16a37 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_drop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_drop.ma @@ -71,14 +71,12 @@ lemma lpx_sn_liftable_dedropable: ∀R. (∀L. reflexive ? (R L)) → #K2 #V2 #HK12 #HV12 #H destruct lapply (lpx_sn_fwd_length … HK12) #H @(ex3_intro … (K2.ⓑ{I}V2)) (**) (* explicit constructor *) - /3 width=1 by lpx_sn_pair, monotonic_le_plus_l/ - @lreq_O2 normalize // + /3 width=1 by lpx_sn_pair, lreq_O2/ | #I #L1 #K1 #V1 #m #_ #IHLK1 #K2 #HK12 elim (IHLK1 … HK12) -K1 /3 width=5 by drop_drop, lreq_pair, lpx_sn_pair, ex3_intro/ | #I #L1 #K1 #V1 #W1 #l #m #HLK1 #HWV1 #IHLK1 #X #H elim (lpx_sn_inv_pair1 … H) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (lift_total W2 l m) #V2 #HWV2 - lapply (H2R … HW12 … HLK1 … HWV1 … HWV2) -W1 + elim (H2R … HW12 … HLK1 … HWV1) -W1 elim (IHLK1 … HK12) -K1 /3 width=6 by drop_skip, lreq_succ, lpx_sn_pair, ex3_intro/ ] diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_lpx_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_lpx_sn.ma index fe4b2b845..031c12c96 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_lpx_sn.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lpx_sn_lpx_sn.ma @@ -35,14 +35,14 @@ qed-. theorem lpx_sn_conf: ∀R1,R2. lpx_sn_confluent R1 R2 → confluent2 … (lpx_sn R1) (lpx_sn R2). -#R1 #R2 #HR12 #L0 @(f_ind … length … L0) -L0 #x #IH * +#R1 #R2 #HR12 #L0 @(ynat_f_ind … length … L0) -L0 #x #IH * [ #_ #X1 #H1 #X2 #H2 -x >(lpx_sn_inv_atom1 … H1) -X1 >(lpx_sn_inv_atom1 … H2) -X2 /2 width=3 by lpx_sn_atom, ex2_intro/ | #L0 #I #V0 #Hx #X1 #H1 #X2 #H2 destruct elim (lpx_sn_inv_pair1 … H1) -H1 #L1 #V1 #HL01 #HV01 #H destruct elim (lpx_sn_inv_pair1 … H2) -H2 #L2 #V2 #HL02 #HV02 #H destruct - elim (IH … HL01 … HL02) -IH normalize // #L #HL1 #HL2 + elim (IH … HL01 … HL02) -IH /2 width=2 by ylt_succ2_refl/ #L #HL1 #HL2 elim (HR12 … HV01 … HV02 … HL01 … HL02) -L0 -V0 /3 width=5 by lpx_sn_pair, ex2_intro/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsuby.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lsuby.ma deleted file mode 100644 index 505a0132e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsuby.ma +++ /dev/null @@ -1,237 +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 "ground_2/ynat/ynat_plus.ma". -include "basic_2/notation/relations/lrsubeq_4.ma". -include "basic_2/substitution/drop.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR EXTENDED SUBSTITUTION *******************) - -inductive lsuby: relation4 ynat ynat lenv lenv ≝ -| lsuby_atom: ∀L,l,m. lsuby l m L (⋆) -| lsuby_zero: ∀I1,I2,L1,L2,V1,V2. - lsuby 0 0 L1 L2 → lsuby 0 0 (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2) -| lsuby_pair: ∀I1,I2,L1,L2,V,m. lsuby 0 m L1 L2 → - lsuby 0 (⫯m) (L1.ⓑ{I1}V) (L2.ⓑ{I2}V) -| lsuby_succ: ∀I1,I2,L1,L2,V1,V2,l,m. - lsuby l m L1 L2 → lsuby (⫯l) m (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2) -. - -interpretation - "local environment refinement (extended substitution)" - 'LRSubEq L1 l m L2 = (lsuby l m L1 L2). - -(* Basic properties *********************************************************) - -lemma lsuby_pair_lt: ∀I1,I2,L1,L2,V,m. L1 ⊆[0, ⫰m] L2 → 0 < m → - L1.ⓑ{I1}V ⊆[0, m] L2.ⓑ{I2}V. -#I1 #I2 #L1 #L2 #V #m #HL12 #Hm <(ylt_inv_O1 … Hm) /2 width=1 by lsuby_pair/ -qed. - -lemma lsuby_succ_lt: ∀I1,I2,L1,L2,V1,V2,l,m. L1 ⊆[⫰l, m] L2 → 0 < l → - L1.ⓑ{I1}V1 ⊆[l, m] L2. ⓑ{I2}V2. -#I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #Hl <(ylt_inv_O1 … Hl) /2 width=1 by lsuby_succ/ -qed. - -lemma lsuby_pair_O_Y: ∀L1,L2. L1 ⊆[0, ∞] L2 → - ∀I1,I2,V. L1.ⓑ{I1}V ⊆[0,∞] L2.ⓑ{I2}V. -#L1 #L2 #HL12 #I1 #I2 #V lapply (lsuby_pair I1 I2 … V … HL12) -HL12 // -qed. - -lemma lsuby_refl: ∀L,l,m. L ⊆[l, m] L. -#L elim L -L // -#L #I #V #IHL #l elim (ynat_cases … l) [| * #x ] -#Hl destruct /2 width=1 by lsuby_succ/ -#m elim (ynat_cases … m) [| * #x ] -#Hm destruct /2 width=1 by lsuby_zero, lsuby_pair/ -qed. - -lemma lsuby_O2: ∀L2,L1,l. |L2| ≤ |L1| → L1 ⊆[l, yinj 0] L2. -#L2 elim L2 -L2 // #L2 #I2 #V2 #IHL2 * normalize -[ #l #H elim (le_plus_xSy_O_false … H) -| #L1 #I1 #V1 #l #H lapply (le_plus_to_le_r … H) -H #HL12 - elim (ynat_cases l) /3 width=1 by lsuby_zero/ - * /3 width=1 by lsuby_succ/ -] -qed. - -lemma lsuby_sym: ∀l,m,L1,L2. L1 ⊆[l, m] L2 → |L1| = |L2| → L2 ⊆[l, m] L1. -#l #m #L1 #L2 #H elim H -l -m -L1 -L2 -[ #L1 #l #m #H >(length_inv_zero_dx … H) -L1 // -| /2 width=1 by lsuby_O2/ -| #I1 #I2 #L1 #L2 #V #m #_ #IHL12 #H lapply (injective_plus_l … H) - /3 width=1 by lsuby_pair/ -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL12 #H lapply (injective_plus_l … H) - /3 width=1 by lsuby_succ/ -] -qed-. - -(* Basic inversion lemmas ***************************************************) - -fact lsuby_inv_atom1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → L1 = ⋆ → L2 = ⋆. -#L1 #L2 #l #m * -L1 -L2 -l -m // -[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #H destruct -| #I1 #I2 #L1 #L2 #V #m #_ #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #H destruct -] -qed-. - -lemma lsuby_inv_atom1: ∀L2,l,m. ⋆ ⊆[l, m] L2 → L2 = ⋆. -/2 width=5 by lsuby_inv_atom1_aux/ qed-. - -fact lsuby_inv_zero1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → - ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → l = 0 → m = 0 → - L2 = ⋆ ∨ - ∃∃J2,K2,W2. K1 ⊆[0, 0] K2 & L2 = K2.ⓑ{J2}W2. -#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/ -[ #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J1 #K1 #W1 #H #_ #_ destruct - /3 width=5 by ex2_3_intro, or_intror/ -| #I1 #I2 #L1 #L2 #V #m #_ #J1 #K1 #W1 #_ #_ #H - elim (ysucc_inv_O_dx … H) -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J1 #K1 #W1 #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma lsuby_inv_zero1: ∀I1,K1,L2,V1. K1.ⓑ{I1}V1 ⊆[0, 0] L2 → - L2 = ⋆ ∨ - ∃∃I2,K2,V2. K1 ⊆[0, 0] K2 & L2 = K2.ⓑ{I2}V2. -/2 width=9 by lsuby_inv_zero1_aux/ qed-. - -fact lsuby_inv_pair1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → - ∀J1,K1,W. L1 = K1.ⓑ{J1}W → l = 0 → 0 < m → - L2 = ⋆ ∨ - ∃∃J2,K2. K1 ⊆[0, ⫰m] K2 & L2 = K2.ⓑ{J2}W. -#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/ -[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W #_ #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #m #HL12 #J1 #K1 #W #H #_ #_ destruct - /3 width=4 by ex2_2_intro, or_intror/ -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J1 #K1 #W #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma lsuby_inv_pair1: ∀I1,K1,L2,V,m. K1.ⓑ{I1}V ⊆[0, m] L2 → 0 < m → - L2 = ⋆ ∨ - ∃∃I2,K2. K1 ⊆[0, ⫰m] K2 & L2 = K2.ⓑ{I2}V. -/2 width=6 by lsuby_inv_pair1_aux/ qed-. - -fact lsuby_inv_succ1_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → - ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < l → - L2 = ⋆ ∨ - ∃∃J2,K2,W2. K1 ⊆[⫰l, m] K2 & L2 = K2.ⓑ{J2}W2. -#L1 #L2 #l #m * -L1 -L2 -l -m /2 width=1 by or_introl/ -[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W1 #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #m #_ #J1 #K1 #W1 #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #J1 #K1 #W1 #H #_ destruct - /3 width=5 by ex2_3_intro, or_intror/ -] -qed-. - -lemma lsuby_inv_succ1: ∀I1,K1,L2,V1,l,m. K1.ⓑ{I1}V1 ⊆[l, m] L2 → 0 < l → - L2 = ⋆ ∨ - ∃∃I2,K2,V2. K1 ⊆[⫰l, m] K2 & L2 = K2.ⓑ{I2}V2. -/2 width=5 by lsuby_inv_succ1_aux/ qed-. - -fact lsuby_inv_zero2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → - ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → l = 0 → m = 0 → - ∃∃J1,K1,W1. K1 ⊆[0, 0] K2 & L1 = K1.ⓑ{J1}W1. -#L1 #L2 #l #m * -L1 -L2 -l -m -[ #L1 #l #m #J2 #K2 #W1 #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J2 #K2 #W2 #H #_ #_ destruct - /2 width=5 by ex2_3_intro/ -| #I1 #I2 #L1 #L2 #V #m #_ #J2 #K2 #W2 #_ #_ #H - elim (ysucc_inv_O_dx … H) -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J2 #K2 #W2 #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma lsuby_inv_zero2: ∀I2,K2,L1,V2. L1 ⊆[0, 0] K2.ⓑ{I2}V2 → - ∃∃I1,K1,V1. K1 ⊆[0, 0] K2 & L1 = K1.ⓑ{I1}V1. -/2 width=9 by lsuby_inv_zero2_aux/ qed-. - -fact lsuby_inv_pair2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → - ∀J2,K2,W. L2 = K2.ⓑ{J2}W → l = 0 → 0 < m → - ∃∃J1,K1. K1 ⊆[0, ⫰m] K2 & L1 = K1.ⓑ{J1}W. -#L1 #L2 #l #m * -L1 -L2 -l -m -[ #L1 #l #m #J2 #K2 #W #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W #_ #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #m #HL12 #J2 #K2 #W #H #_ #_ destruct - /2 width=4 by ex2_2_intro/ -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #J2 #K2 #W #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma lsuby_inv_pair2: ∀I2,K2,L1,V,m. L1 ⊆[0, m] K2.ⓑ{I2}V → 0 < m → - ∃∃I1,K1. K1 ⊆[0, ⫰m] K2 & L1 = K1.ⓑ{I1}V. -/2 width=6 by lsuby_inv_pair2_aux/ qed-. - -fact lsuby_inv_succ2_aux: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → - ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → 0 < l → - ∃∃J1,K1,W1. K1 ⊆[⫰l, m] K2 & L1 = K1.ⓑ{J1}W1. -#L1 #L2 #l #m * -L1 -L2 -l -m -[ #L1 #l #m #J2 #K2 #W2 #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W2 #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #m #_ #J2 #K1 #W2 #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #HL12 #J2 #K2 #W2 #H #_ destruct - /2 width=5 by ex2_3_intro/ -] -qed-. - -lemma lsuby_inv_succ2: ∀I2,K2,L1,V2,l,m. L1 ⊆[l, m] K2.ⓑ{I2}V2 → 0 < l → - ∃∃I1,K1,V1. K1 ⊆[⫰l, m] K2 & L1 = K1.ⓑ{I1}V1. -/2 width=5 by lsuby_inv_succ2_aux/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lsuby_fwd_length: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → |L2| ≤ |L1|. -#L1 #L2 #l #m #H elim H -L1 -L2 -l -m normalize /2 width=1 by le_S_S/ -qed-. - -(* Properties on basic slicing **********************************************) - -lemma lsuby_drop_trans_be: ∀L1,L2,l,m. L1 ⊆[l, m] L2 → - ∀I2,K2,W,s,i. ⬇[s, 0, i] L2 ≡ K2.ⓑ{I2}W → - l ≤ i → i < l + m → - ∃∃I1,K1. K1 ⊆[0, ⫰(l+m-i)] K2 & ⬇[s, 0, i] L1 ≡ K1.ⓑ{I1}W. -#L1 #L2 #l #m #H elim H -L1 -L2 -l -m -[ #L1 #l #m #J2 #K2 #W #s #i #H - elim (drop_inv_atom1 … H) -H #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J2 #K2 #W #s #i #_ #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #m #HL12 #IHL12 #J2 #K2 #W #s #i #H #_ >yplus_O1 - elim (drop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ] - [ #_ destruct -I2 >ypred_succ - /2 width=4 by drop_pair, ex2_2_intro/ - | lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/ - #H yminus_succ yplus_succ1 #H lapply (ylt_inv_succ … H) -H - #Hilm lapply (drop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/ - #HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 yminus_SO2 - /4 width=4 by ylt_O, drop_drop_lt, ex2_2_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsuby_lsuby.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lsuby_lsuby.ma deleted file mode 100644 index ff914eca7..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lsuby_lsuby.ma +++ /dev/null @@ -1,32 +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/lsuby.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR EXTENDED SUBSTITUTION *******************) - -(* Main properties **********************************************************) - -theorem lsuby_trans: ∀l,m. Transitive … (lsuby l m). -#l #m #L1 #L2 #H elim H -L1 -L2 -l -m -[ #L1 #l #m #X #H lapply (lsuby_inv_atom1 … H) -H - #H destruct // -| #I1 #I2 #L1 #L #V1 #V #_ #IHL1 #X #H elim (lsuby_inv_zero1 … H) -H // - * #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lsuby_zero/ -| #I1 #I2 #L1 #L2 #V #m #_ #IHL1 #X #H elim (lsuby_inv_pair1 … H) -H // - * #I2 #L2 #HL2 #H destruct /3 width=1 by lsuby_pair/ -| #I1 #I2 #L1 #L2 #V1 #V2 #l #m #_ #IHL1 #X #H elim (lsuby_inv_succ1 … H) -H // - * #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lsuby_succ/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_after.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_after.ma index 4bc0233a2..78dd6fb9a 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_after.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_after.ma @@ -110,9 +110,21 @@ qed-. lemma after_inv_inh3: ∀cs1,cs2,tl,b. cs1 ⊚ cs2 ≡ b @ tl → (∃∃tl1,tl2. cs1 = Ⓣ @ tl1 & cs2 = b @ tl2 & tl1 ⊚ tl2 ≡ tl) ∨ - ∃∃tl1. cs1 = Ⓕ @ tl1 & b = Ⓕ & tl1 ⊚ cs2 ≡ tl. + ∃∃tl1. cs1 = Ⓕ @ tl1 & b = Ⓕ & tl1 ⊚ cs2 ≡ tl. /2 width=3 by after_inv_inh3_aux/ qed-. +lemma after_inv_true3: ∀cs1,cs2,tl. cs1 ⊚ cs2 ≡ Ⓣ @ tl → + ∃∃tl1,tl2. cs1 = Ⓣ @ tl1 & cs2 = Ⓣ @ tl2 & tl1 ⊚ tl2 ≡ tl. +#cs1 #cs2 #tl #H elim (after_inv_inh3 … H) -H // * +#tl1 #_ #H destruct +qed-. + +lemma after_inv_false3: ∀cs1,cs2,tl. cs1 ⊚ cs2 ≡ Ⓕ @ tl → + (∃∃tl1,tl2. cs1 = Ⓣ @ tl1 & cs2 = Ⓕ @ tl2 & tl1 ⊚ tl2 ≡ tl) ∨ + ∃∃tl1. cs1 = Ⓕ @ tl1 & tl1 ⊚ cs2 ≡ tl. +#cs1 #cs2 #tl #H elim (after_inv_inh3 … H) -H /2 width=1 by or_introl/ * /3 width=3 by ex2_intro, or_intror/ +qed-. + lemma after_inv_length: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → ∧∧ ∥cs1∥ = |cs2| & |cs| = |cs1| & ∥cs∥ = ∥cs2∥. #cs1 #cs2 #cs #H elim H -cs1 -cs2 -cs /2 width=1 by and3_intro/ diff --git a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma index 0aa9a4c03..d87487c89 100644 --- a/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma +++ b/matita/matita/contribs/lambdadelta/ground_2/relocation/trace_isid.ma @@ -58,24 +58,43 @@ qed-. (* Properties on composition ************************************************) -lemma isid_after_sn: ∀cs1,cs2. cs1 ⊚ cs2 ≡ cs2 → 𝐈⦃cs1⦄ . +lemma isid_after_sn: ∀cs2. ∃∃cs1. 𝐈⦃cs1⦄ & cs1 ⊚ cs2 ≡ cs2. +#cs2 elim cs2 -cs2 /2 width=3 by after_empty, ex2_intro/ +#b #cs2 * /3 width=3 by isid_true, after_true, ex2_intro/ +qed-. + +lemma isid_after_dx: ∀cs1. ∃∃cs2. 𝐈⦃cs2⦄ & cs1 ⊚ cs2 ≡ cs1. +#cs1 elim cs1 -cs1 /2 width=3 by after_empty, ex2_intro/ +* #cs1 * /3 width=3 by isid_true, after_true, after_false, ex2_intro/ +qed-. + +lemma after_isid_sn: ∀cs1,cs2. cs1 ⊚ cs2 ≡ cs2 → 𝐈⦃cs1⦄ . #cs1 #cs2 #H elim (after_inv_length … H) -H // qed. -lemma isid_after_dx: ∀cs1,cs2. cs1 ⊚ cs2 ≡ cs1 → 𝐈⦃cs2⦄ . +lemma after_isid_dx: ∀cs1,cs2. cs1 ⊚ cs2 ≡ cs1 → 𝐈⦃cs2⦄ . #cs1 #cs2 #H elim (after_inv_length … H) -H // qed. (* Inversion lemmas on composition ******************************************) -lemma isid_inv_after_sn: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → 𝐈⦃cs1⦄ → cs = cs2. +lemma after_isid_inv_sn: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → 𝐈⦃cs1⦄ → cs = cs2. #cs1 #cs2 #cs #H elim H -cs1 -cs2 -cs // #cs1 #cs2 #cs #_ [ #b ] #IH #H [ >IH -IH // | elim (isid_inv_false … H) ] qed-. -lemma isid_inv_after_dx: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → 𝐈⦃cs2⦄ → cs = cs1. +lemma after_isid_inv_dx: ∀cs1,cs2,cs. cs1 ⊚ cs2 ≡ cs → 𝐈⦃cs2⦄ → cs = cs1. #cs1 #cs2 #cs #H elim H -cs1 -cs2 -cs // #cs1 #cs2 #cs #_ [ #b ] #IH #H [ elim (isid_inv_cons … H) -H #H >IH -IH // | >IH -IH // ] qed-. + +lemma after_inv_isid3: ∀t1,t2,t. t1 ⊚ t2 ≡ t → 𝐈⦃t⦄ → 𝐈⦃t1⦄ ∧ 𝐈⦃t2⦄. +#t1 #t2 #t #H elim H -t1 -t2 -t +[ /2 width=1 by conj/ +| #t1 #t2 #t #_ #b #IHt #H elim (isid_inv_cons … H) -H + #Ht #H elim (IHt Ht) -t /2 width=1 by isid_true, conj/ +| #t1 #t2 #t #_ #_ #H elim (isid_inv_false … H) +] +qed-.