From: Ferruccio Guidi Date: Tue, 16 Apr 2013 19:45:41 +0000 (+0000) Subject: - we commit just the components before "reducibility" X-Git-Tag: make_still_working~1190 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=713673ecf863cb6187291f016ed4490b12a03ac0;p=helm.git - we commit just the components before "reducibility" - reaxiomatized substitution and reduction now commute with respect to subclosure - delift and thin removed for now --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/aaa_ltpss_dx.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/aaa_ltpss_dx.etc new file mode 100644 index 000000000..2f2d07360 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/aaa_ltpss_dx.etc @@ -0,0 +1,79 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/tpss_tpss.ma". +include "basic_2/unfold/ltpss_dx_ldrop.ma". +include "basic_2/static/aaa_lift.ma". + +(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) + +(* Properties about dx parallel unfold **************************************) + +(* Note: lemma 500 *) +lemma aaa_ltpss_dx_tpss_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → + ∀L2,d,e. L1 ▶* [d, e] L2 → + ∀T2. L2 ⊢ T1 ▶* [d, e] T2 → L2 ⊢ T2 ⁝ A. +#L1 #T1 #A #H elim H -L1 -T1 -A +[ #L1 #k #L2 #d #e #_ #T2 #H + >(tpss_inv_sort1 … H) -H // +| #I #L1 #K1 #V1 #B #i #HLK1 #_ #IHV1 #L2 #d #e #HL12 #T2 #H + elim (tpss_inv_lref1 … H) -H + [ #H destruct + elim (lt_or_ge i d) #Hdi + [ elim (ltpss_dx_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2 + elim (ltpss_dx_inv_tpss11 … H ?) -H /2 width=1/ -Hdi #K2 #V2 #HK12 #HV12 #H destruct + /3 width=8 by aaa_lref/ (**) (* too slow without trace *) + | elim (lt_or_ge i (d + e)) #Hide + [ elim (ltpss_dx_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2 + elim (ltpss_dx_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K2 #V2 #HK12 #HV12 #H destruct + /3 width=8 by aaa_lref/ (**) (* too slow without trace *) + | -Hdi + lapply (ltpss_dx_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide + /3 width=8 by aaa_lref/ (**) (* too slow without trace *) + ] + ] + | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #HVW2 #HWT2 + elim (ltpss_dx_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0 + elim (ltpss_dx_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #HK12 #HV12 #H destruct + lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 #H destruct + lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2 + lapply (tpss_trans_eq … HV12 HVW2) -V2 /3 width=7/ + ] +| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H + elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=4/ +| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H + elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=4/ +| #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H + elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /3 width=4/ +| #L1 #V1 #T1 #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H + elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /3 width=4/ +] +qed. + +lemma aaa_ltpss_dx_tps_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → + ∀L2,d,e. L1 ▶* [d, e] L2 → + ∀T2. L2 ⊢ T1 ▶ [d, e] T2 → L2 ⊢ T2 ⁝ A. +/3 width=7/ qed. + +lemma aaa_ltpss_dx_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → + ∀L2,d,e. L1 ▶* [d, e] L2 → L2 ⊢ T ⁝ A. +/2 width=7/ qed. + +lemma aaa_tpss_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → + ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T2 ⁝ A. +/2 width=7/ qed. + +lemma aaa_tps_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → + ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 → L ⊢ T2 ⁝ A. +/2 width=7/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/aaa_ltpss_sn.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/aaa_ltpss_sn.etc new file mode 100644 index 000000000..4f2a44827 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/aaa_ltpss_sn.etc @@ -0,0 +1,37 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/ltpss_sn_alt.ma". +include "basic_2/static/aaa_ltpss_dx.ma". + +(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) + +(* Properties about sn parallel unfold **************************************) + +lemma aaa_ltpss_sn_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → + ∀L2,d,e. L1 ⊢ ▶* [d, e] L2 → L2 ⊢ T ⁝ A. +#L1 #T #A #HT #L2 #d #e #HL12 +lapply (ltpss_sn_ltpssa … HL12) -HL12 #HL12 +@(TC_Conf3 … (λL,A. L ⊢ T ⁝ A) … HT ? HL12) /2 width=5/ +qed. + +lemma aaa_ltpss_sn_tpss_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → + ∀L2,d,e. L1 ⊢ ▶* [d, e] L2 → + ∀T2. L2 ⊢ T1 ▶* [d, e] T2 → L2 ⊢ T2 ⁝ A. +/3 width=5/ qed. + +lemma aaa_ltpss_sn_tps_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → + ∀L2,d,e. L1 ⊢ ▶* [d, e] L2 → + ∀T2. L2 ⊢ T1 ▶ [d, e] T2 → L2 ⊢ T2 ⁝ A. +/3 width=5/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr.etc new file mode 100644 index 000000000..d82299c38 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr.etc @@ -0,0 +1,111 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/tpss.ma". +include "basic_2/reducibility/tpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Basic_1: includes: pr2_delta1 *) +definition cpr: lenv → relation term ≝ + λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T ▶* [0, |L|] T2. + +interpretation + "context-sensitive parallel reduction (term)" + 'PRed L T1 T2 = (cpr L T1 T2). + +(* Basic properties *********************************************************) + +lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ➡ T2. +/3 width=5/ qed-. + +(* Basic_1: was by definition: pr2_free *) +lemma cpr_tpr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2. +/2 width=3/ qed. + +lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ➡ T2. +/3 width=5/ qed. + +lemma cpr_refl: ∀L,T. L ⊢ T ➡ T. +/2 width=1/ qed. + +(* Note: new property *) +(* Basic_1: was only: pr2_thin_dx *) +lemma cpr_flat: ∀I,L,V1,V2,T1,T2. + L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2. +#I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/ +qed. + +lemma cpr_cast: ∀L,V,T1,T2. + L ⊢ T1 ➡ T2 → L ⊢ ⓝV. T1 ➡ T2. +#L #V #T1 #T2 * /3 width=3/ +qed. + +(* Note: it does not hold replacing |L1| with |L2| *) +(* Basic_1: was only: pr2_change *) +lemma cpr_lsubr_trans: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 → + ∀L2. L2 ⊑ [0, |L1|] L1 → L2 ⊢ T1 ➡ T2. +#L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 +lapply (tpss_lsubr_trans … HT2 … HL12) -HT2 -HL12 /3 width=4/ +qed. + +(* Basic inversion lemmas ***************************************************) + +(* Basic_1: was: pr2_gen_csort *) +lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → T1 ➡ T2. +#T1 #T2 * #T #HT normalize #HT2 +<(tpss_inv_refl_O2 … HT2) -HT2 // +qed-. + +(* Basic_1: was: pr2_gen_sort *) +lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k. +#L #T2 #k * #X #H +>(tpr_inv_atom1 … H) -H #H +>(tpss_inv_sort1 … H) -H // +qed-. + +(* Basic_1: was: pr2_gen_cast *) +lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓝV1. T1 ➡ U2 → ( + ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 & + U2 = ⓝV2. T2 + ) ∨ L ⊢ T1 ➡ U2. +#L #V1 #T1 #U2 * #X #H #HU2 +elim (tpr_inv_cast1 … H) -H /3 width=3/ +* #V #T #HV1 #HT1 #H destruct +elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/ +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma cpr_fwd_bind1_minus: ∀I,L,V1,T1,T. L ⊢ -ⓑ{I}V1.T1 ➡ T → ∀b. + ∃∃V2,T2. L ⊢ ⓑ{b,I}V1.T1 ➡ ⓑ{b,I}V2.T2 & + T = -ⓑ{I}V2.T2. +#I #L #V1 #T1 #T * #X #H1 #H2 #b +elim (tpr_fwd_bind1_minus … H1 b) -H1 #V0 #T0 #HT10 #H destruct +elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV02 #HT02 #H destruct /4 width=5/ +qed-. + +lemma cpr_fwd_shift1: ∀L,L1,T1,T. L ⊢ L1 @@ T1 ➡ T → + ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. +#L #L1 #T1 #T * #X #H1 #H2 +elim (tpr_fwd_shift1 … H1) -H1 #L0 #T0 #HL10 #H destruct +elim (tpss_fwd_shift1 … H2) -H2 #L2 #T2 #HL02 #H destruct /2 width=4/ +qed-. + +(* Basic_1: removed theorems 6: + pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans + pr2_gen_ctail pr2_ctail + Basic_1: removed local theorems 3: + pr2_free_free pr2_free_delta pr2_delta_delta +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_lift.etc new file mode 100644 index 000000000..0a7afddab --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/cpr_lift.etc @@ -0,0 +1,180 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/tpss_lift.ma". +include "basic_2/reducibility/tpr_lift.ma". +include "basic_2/reducibility/cpr.ma". + +(* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************) + +(* Advanced properties ******************************************************) + +lemma cpr_cdelta: ∀L,K,V1,W1,W2,i. + ⇩[0, i] L ≡ K. ⓓV1 → K ⊢ V1 ▶* [0, |L| - i - 1] W1 → + ⇧[0, i + 1] W1 ≡ W2 → L ⊢ #i ➡ W2. +#L #K #V1 #W1 #W2 #i #HLK #HVW1 #HW12 +lapply (ldrop_fwd_ldrop2_length … HLK) #Hi +@ex2_intro [2: // | skip | @tpss_subst /width=6/ ] (**) (* /3 width=6/ is too slow *) +qed. + +lemma cpr_abst: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀V,T1,T2. L.ⓛV ⊢ T1 ➡ T2 → + ∀a,I. L ⊢ ⓑ{a,I}V1. T1 ➡ ⓑ{a,I}V2. T2. +#L #V1 #V2 * #V0 #HV10 #HV02 #V #T1 #T2 * #T0 #HT10 #HT02 #a #I +lapply (tpss_inv_S2 … HT02 L V ?) -HT02 // #HT02 +lapply (tpss_lsubr_trans … HT02 (L.ⓑ{I}V2) ?) -HT02 /2 width=1/ #HT02 +@(ex2_intro … (ⓑ{a,I}V0.T0)) /2 width=1/ (* explicit constructors *) +qed. + +lemma cpr_beta: ∀a,L,V1,V2,W,T1,T2. + L ⊢ V1 ➡ V2 → L.ⓛW ⊢ T1 ➡ T2 → L ⊢ ⓐV1.ⓛ{a}W.T1 ➡ ⓓ{a}V2.T2. +#a #L #V1 #V2 #W #T1 #T2 * #V #HV1 #HV2 * #T #HT1 #HT2 +lapply (tpss_inv_S2 … HT2 L W ?) -HT2 // #HT2 +lapply (tpss_lsubr_trans … HT2 (L.ⓓV2) ?) -HT2 /2 width=1/ #HT2 +@(ex2_intro … (ⓓ{a}V.T)) /2 width=1/ (**) (* explicit constructor, /3/ is too slow *) +qed. + +lemma cpr_beta_dx: ∀a,L,V1,V2,W,T1,T2. + V1 ➡ V2 → L.ⓛW ⊢ T1 ➡ T2 → L ⊢ ⓐV1.ⓛ{a}W.T1 ➡ ⓓ{a}V2.T2. +/3 width=1/ qed. + +(* Advanced inversion lemmas ************************************************) + +(* Basic_1: was: pr2_gen_lref *) +lemma cpr_inv_lref1: ∀L,T2,i. L ⊢ #i ➡ T2 → + T2 = #i ∨ + ∃∃K,V1,T1. ⇩[0, i] L ≡ K. ⓓV1 & + K ⊢ V1 ▶* [0, |L| - i - 1] T1 & + ⇧[0, i + 1] T1 ≡ T2 & + i < |L|. +#L #T2 #i * #X #H +>(tpr_inv_atom1 … H) -H #H +elim (tpss_inv_lref1 … H) -H /2 width=1/ +* /3 width=6/ +qed-. + +(* Basic_1: was pr2_gen_abbr *) +lemma cpr_inv_abbr1: ∀a,L,V1,T1,U2. L ⊢ ⓓ{a}V1. T1 ➡ U2 → + (∃∃V,V2,T2. V1 ➡ V & L ⊢ V ▶* [O, |L|] V2 & + L. ⓓV ⊢ T1 ➡ T2 & + U2 = ⓓ{a}V2. T2 + ) ∨ + ∃∃T2. L.ⓓV1 ⊢ T1 ➡ T2 & ⇧[0,1] U2 ≡ T2 & a = true. +#a #L #V1 #T1 #Y * #X #H1 #H2 +elim (tpr_inv_abbr1 … H1) -H1 * +[ #V #T #T0 #HV1 #HT1 #HT0 #H destruct + elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT02 #H destruct + lapply (tps_lsubr_trans … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0 + lapply (tps_weak_full … HT0) -HT0 #HT0 + lapply (tpss_lsubr_trans … HT02 (L. ⓓV) ?) -HT02 /2 width=1/ #HT02 + lapply (tpss_weak_full … HT02) -HT02 #HT02 + lapply (tpss_strap2 … HT0 HT02) -T0 /4 width=7/ +| #T2 #HT12 #HXT2 #H destruct + elim (lift_total Y 0 1) #Z #HYZ + lapply (tpss_lift_ge … H2 (L.ⓓV1) … HXT2 … HYZ) -X // /2 width=1/ #H + lapply (cpr_intro … HT12 … H) -T2 /3 width=3/ +] +qed-. + +(* Basic_1: was: pr2_gen_abst *) +lemma cpr_inv_abst1: ∀a,L,V1,T1,U2. L ⊢ ⓛ{a}V1. T1 ➡ U2 → ∀I,W. + ∃∃V2,T2. L ⊢ V1 ➡ V2 & L. ⓑ{I} W ⊢ T1 ➡ T2 & U2 = ⓛ{a}V2. T2. +#a #L #V1 #T1 #Y * #X #H1 #H2 #I #W +elim (tpr_inv_abst1 … H1) -H1 #V #T #HV1 #HT1 #H destruct +elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct +lapply (tpss_lsubr_trans … HT2 (L. ⓑ{I} W) ?) -HT2 /2 width=1/ /4 width=5/ +qed-. + +(* Basic_1: was pr2_gen_appl *) +lemma cpr_inv_appl1: ∀L,V1,U0,U2. L ⊢ ⓐV1. U0 ➡ U2 → + ∨∨ ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ U0 ➡ T2 & + U2 = ⓐV2. T2 + | ∃∃a,V2,W,T1,T2. L ⊢ V1 ➡ V2 & L. ⓓV2 ⊢ T1 ➡ T2 & + U0 = ⓛ{a}W. T1 & + U2 = ⓓ{a}V2. T2 + | ∃∃a,V2,V,W1,W2,T1,T2. L ⊢ V1 ➡ V2 & L ⊢ W1 ➡ W2 & L. ⓓW2 ⊢ T1 ➡ T2 & + ⇧[0,1] V2 ≡ V & + U0 = ⓓ{a}W1. T1 & + U2 = ⓓ{a}W2. ⓐV. T2. +#L #V1 #U0 #Y * #X #H1 #H2 +elim (tpr_inv_appl1 … H1) -H1 * +[ #V #U #HV1 #HU0 #H destruct + elim (tpss_inv_flat1 … H2) -H2 #V2 #U2 #HV2 #HU2 #H destruct /4 width=5/ +| #a #V #W #T0 #T #HV1 #HT0 #H #H1 destruct + elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct + lapply (tpss_weak … HT2 0 (|L|+1) ? ?) -HT2 // /4 width=9/ +| #a #V0 #V #W #W0 #T #T0 #HV10 #HW0 #HT0 #HV0 #H #H1 destruct + elim (tpss_inv_bind1 … H2) -H2 #W2 #X #HW02 #HX #HY destruct + elim (tpss_inv_flat1 … HX) -HX #V2 #T2 #HV2 #HT2 #H destruct + elim (tpss_inv_lift1_ge … HV2 … HV0 ?) -V // [3: /2 width=1/ |2: skip ] #V commutative_plus normalize #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #HV12 #IHL12 #He destruct + >(IHL12 ?) -IHL12 // >(tpss_inv_refl_O2 … HV12) // +] +qed. + +lemma ltpss_dx_inv_refl_O2: ∀d,L1,L2. L1 ▶* [d, 0] L2 → L1 = L2. +/2 width=4/ qed-. + +fact ltpss_dx_inv_atom1_aux: ∀d,e,L1,L2. + L1 ▶* [d, e] L2 → L1 = ⋆ → L2 = ⋆. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ // +| #L #I #V #H destruct +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct +] +qed. + +lemma ltpss_dx_inv_atom1: ∀d,e,L2. ⋆ ▶* [d, e] L2 → L2 = ⋆. +/2 width=5/ qed-. + +fact ltpss_dx_inv_tpss21_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → d = 0 → 0 < e → + ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → + ∃∃K2,V2. K1 ▶* [0, e - 1] K2 & + K2 ⊢ V1 ▶* [0, e - 1] V2 & + L2 = K2. ⓑ{I} V2. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #_ #K1 #I #V1 #H destruct +| #L1 #I #V #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K1 #J #W1 #H destruct /2 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma ltpss_dx_inv_tpss21: ∀e,K1,I,V1,L2. K1. ⓑ{I} V1 ▶* [0, e] L2 → 0 < e → + ∃∃K2,V2. K1 ▶* [0, e - 1] K2 & + K2 ⊢ V1 ▶* [0, e - 1] V2 & + L2 = K2. ⓑ{I} V2. +/2 width=5/ qed-. + +fact ltpss_dx_inv_tpss11_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → 0 < d → + ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → + ∃∃K2,V2. K1 ▶* [d - 1, e] K2 & + K2 ⊢ V1 ▶* [d - 1, e] V2 & + L2 = K2. ⓑ{I} V2. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #I #K1 #V1 #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K1 #W1 #H destruct /2 width=5/ +] +qed. + +lemma ltpss_dx_inv_tpss11: ∀d,e,I,K1,V1,L2. K1. ⓑ{I} V1 ▶* [d, e] L2 → 0 < d → + ∃∃K2,V2. K1 ▶* [d - 1, e] K2 & + K2 ⊢ V1 ▶* [d - 1, e] V2 & + L2 = K2. ⓑ{I} V2. +/2 width=3/ qed-. + +fact ltpss_dx_inv_atom2_aux: ∀d,e,L1,L2. + L1 ▶* [d, e] L2 → L2 = ⋆ → L1 = ⋆. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ // +| #L #I #V #H destruct +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct +] +qed. + +lemma ltpss_dx_inv_atom2: ∀d,e,L1. L1 ▶* [d, e] ⋆ → L1 = ⋆. +/2 width=5/ qed-. + +fact ltpss_dx_inv_tpss22_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → d = 0 → 0 < e → + ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 ▶* [0, e - 1] K2 & + K2 ⊢ V1 ▶* [0, e - 1] V2 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #_ #K1 #I #V1 #H destruct +| #L1 #I #V #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K2 #J #W2 #H destruct /2 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma ltpss_dx_inv_tpss22: ∀e,L1,K2,I,V2. L1 ▶* [0, e] K2. ⓑ{I} V2 → 0 < e → + ∃∃K1,V1. K1 ▶* [0, e - 1] K2 & + K2 ⊢ V1 ▶* [0, e - 1] V2 & + L1 = K1. ⓑ{I} V1. +/2 width=5/ qed-. + +fact ltpss_dx_inv_tpss12_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → 0 < d → + ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 ▶* [d - 1, e] K2 & + K2 ⊢ V1 ▶* [d - 1, e] V2 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #I #K2 #V2 #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K2 #W2 #H destruct /2 width=5/ +] +qed. + +lemma ltpss_dx_inv_tpss12: ∀L1,K2,I,V2,d,e. L1 ▶* [d, e] K2. ⓑ{I} V2 → 0 < d → + ∃∃K1,V1. K1 ▶* [d - 1, e] K2 & + K2 ⊢ V1 ▶* [d - 1, e] V2 & + L1 = K1. ⓑ{I} V1. +/2 width=3/ qed-. + +(* Basic properties *********************************************************) + +lemma ltpss_dx_tps2: ∀L1,L2,I,V1,V2,e. + L1 ▶* [0, e] L2 → L2 ⊢ V1 ▶ [0, e] V2 → + L1. ⓑ{I} V1 ▶* [0, e + 1] L2. ⓑ{I} V2. +/3 width=1/ qed. + +lemma ltpss_dx_tps1: ∀L1,L2,I,V1,V2,d,e. + L1 ▶* [d, e] L2 → L2 ⊢ V1 ▶ [d, e] V2 → + L1. ⓑ{I} V1 ▶* [d + 1, e] L2. ⓑ{I} V2. +/3 width=1/ qed. + +lemma ltpss_dx_tpss2_lt: ∀L1,L2,I,V1,V2,e. + L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶* [0, e - 1] V2 → + 0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He +>(plus_minus_m_m e 1) /2 width=1/ +qed. + +lemma ltpss_dx_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶* [d - 1, e] V2 → + 0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd +>(plus_minus_m_m d 1) /2 width=1/ +qed. + +lemma ltpss_dx_tps2_lt: ∀L1,L2,I,V1,V2,e. + L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶ [0, e - 1] V2 → + 0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2. +/3 width=1/ qed. + +lemma ltpss_dx_tps1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶ [d - 1, e] V2 → + 0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2. +/3 width=1/ qed. + +(* Basic_1: was by definition: csubst1_refl *) +lemma ltpss_dx_refl: ∀L,d,e. L ▶* [d, e] L. +#L elim L -L // +#L #I #V #IHL * /2 width=1/ * /2 width=1/ +qed. + +lemma ltpss_dx_weak: ∀L1,L2,d1,e1. L1 ▶* [d1, e1] L2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → L1 ▶* [d2, e2] L2. +#L1 #L2 #d1 #e1 #H elim H -L1 -L2 -d1 -e1 // +[ #L1 #L2 #I #V1 #V2 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd2 #Hde2 + lapply (le_n_O_to_eq … Hd2) #H destruct normalize in Hde2; + lapply (lt_to_le_to_lt 0 … Hde2) // #He2 + lapply (le_plus_to_minus_r … Hde2) -Hde2 /3 width=5/ +| #L1 #L2 #I #V1 #V2 #d1 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd21 #Hde12 + >plus_plus_comm_23 in Hde12; #Hde12 + elim (le_to_or_lt_eq 0 d2 ?) // #H destruct + [ lapply (le_plus_to_minus_r … Hde12) -Hde12 plus_plus_comm_23 + /4 width=5 by ltpss_dx_tpss2, tpss_append, tpss_weak, monotonic_le_plus_r/ (**) (* too slow without trace *) +| #K1 #K2 #I #V1 #V2 #d #x #_ #HV12 #IHK12 normalize (ldrop_inv_atom1 … H) -H // +| // +| normalize #K0 #K1 #I #V0 #V1 #e1 #_ #_ #IHK01 #L2 #e2 #H #He12 + elim (le_inv_plus_l … He12) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ +| #K0 #K1 #I #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ +] +qed. + +lemma ltpss_dx_ldrop_trans_ge: ∀L1,L0,d1,e1. L1 ▶* [d1, e1] L0 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → + d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. +#L1 #L0 #d1 #e1 #H elim H -L1 -L0 -d1 -e1 +[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H // +| // +| normalize #K1 #K0 #I #V1 #V0 #e1 #_ #_ #IHK10 #L2 #e2 #H #He12 + elim (le_inv_plus_l … He12) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK10 … HK0L2 ?) -K0 /2 width=1/ +| #K0 #K1 #I #V1 #V0 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK10 … HK0L2 ?) -IHK10 -HK0L2 /2 width=1/ +] +qed. + +lemma ltpss_dx_ldrop_conf_be: ∀L0,L1,d1,e1. L0 ▶* [d1, e1] L1 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → + ∃∃L. L2 ▶* [0, d1 + e1 - e2] L & ⇩[0, e2] L1 ≡ L. +#L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 +[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H /2 width=3/ +| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 + lapply (le_n_O_to_eq … He2) -He2 #H destruct + lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ +| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #_ #He21 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK01 -He21 destruct plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 #He2de1 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + (ldrop_inv_atom1 … H) -H /2 width=3/ +| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 + lapply (le_n_O_to_eq … He2) -He2 #H destruct + lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ +| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #_ #He21 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK10 -He21 destruct plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 #He2de1 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + (ldrop_inv_atom1 … H) -H /2 width=3/ +| /2 width=3/ +| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #_ #L2 #e2 #H #He2 + lapply (le_n_O_to_eq … He2) -He2 #He2 destruct + lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ +| #K0 #K1 #I #V0 #V1 #d1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #He2d1 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK01 -He2d1 destruct (ldrop_inv_atom1 … H) -H /2 width=3/ +| /2 width=3/ +| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #_ #L2 #e2 #H #He2 + lapply (le_n_O_to_eq … He2) -He2 #He2 destruct + lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ +| #K1 #K0 #I #V1 #V0 #d1 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #He2d1 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK10 -He2d1 destruct (ltpss_dx_inv_atom1 … H) -H /2 width=3/ +| /2 width=3/ +| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #Hd + elim (IHLK1 … HK12 Hd) -K1 -Hd /3 width=5/ +| #L1 #K1 #I #V1 #W1 #d1 #e1 #_ #HWV1 #IHLK1 #X #d2 #e2 #H #Hd12 + elim (le_inv_plus_l … Hd12) -Hd12 #Hd12 #Hd2 + elim (ltpss_dx_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (IHLK1 … HK12 … Hd12) -IHLK1 -HK12 (ltpss_dx_inv_atom1 … H) -H /2 width=3/ +| #K1 #I #V1 #K2 #d2 #e2 #HK12 #H #_ + lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/ +| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #H1 #H2 + elim (IHLK1 … HK12 H1 H2) -K1 -H2 + lapply (le_n_O_to_eq … H1) -H1 #H destruct /3 width=5/ +| #L1 #K1 #I #V1 #W1 #d1 #e1 #_ #HWV1 #IHLK1 #X #d2 #e2 #H #Hd21 #Hd12 + elim (eq_or_gt d2) #Hd2 [ -Hd21 elim (eq_or_gt e2) #He2 ] destruct + [ lapply (le_n_O_to_eq … Hd12) -Hd12 plus_minus_commutative // #L2 #HL12 #HLK2 + elim (lift_total W2 d1 e1) #V2 #HWV2 + lapply (tpss_lift_be … HW12 … HLK2 HWV1 … HWV2) -W1 // /2 width=1/ + >plus_minus // >commutative_plus /4 width=5/ + | elim (ltpss_dx_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (IHLK1 … HK12 …) -IHLK1 [2: >plus_minus // ] /2 width=1/ #L2 #HL12 #HLK2 + elim (lift_total W2 d1 e1) #V2 #HWV2 + lapply (tpss_lift_be … HW12 … HLK2 HWV1 … HWV2) -W1 [ >plus_minus // ] /2 width=1/ + >commutative_plus /3 width=5/ + ] +] +qed-. + +lemma ldrop_ltpss_dx_trans_ge: ∀L1,K1,d1,e1. ⇩[d1, e1] L1 ≡ K1 → + ∀K2,d2,e2. K1 ▶* [d2, e2] K2 → d2 + e2 ≤ d1 → + ∃∃L2. L1 ▶* [d2, e2] L2 & ⇩[d1, e1] L2 ≡ K2. +#L1 #K1 #d1 #e1 #H elim H -L1 -K1 -d1 -e1 +[ #d1 #e1 #K2 #d2 #e2 #H #_ + >(ltpss_dx_inv_atom1 … H) -H /2 width=3/ +| #K1 #I #V1 #K2 #d2 #e2 #HK12 #H + elim (plus_le_0 … H) -H #H1 #H2 destruct /2 width=3/ +| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #H + elim (IHLK1 … HK12 H) -K1 + elim (plus_le_0 … H) -H #H1 #H2 destruct #L2 #HL12 + >(ltpss_dx_inv_refl_O2 … HL12) -L1 /3 width=5/ +| #L1 #K1 #I #V1 #W1 #d1 #e1 #HLK1 #HWV1 #IHLK1 #X #d2 #e2 #H #Hd21 + elim (eq_or_gt d2) #Hd2 [ elim (eq_or_gt e2) #He2 ] destruct + [ -IHLK1 -Hd21 <(ltpss_dx_inv_refl_O2 … H) -X /3 width=5/ + | elim (ltpss_dx_inv_tpss21 … H He2) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (IHLK1 … HK12 …) -IHLK1 /2 width=1/ #L2 #HL12 #HLK2 + elim (lift_total W2 d1 e1) #V2 #HWV2 + lapply (tpss_lift_le … HW12 … HLK2 HWV1 … HWV2) -W1 /2 width=1/ /3 width=5/ + | elim (ltpss_dx_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (IHLK1 … HK12 …) -IHLK1 [2: >plus_minus // /2 width=1/ ] #L2 #HL12 #HLK2 + elim (lift_total W2 d1 e1) #V2 #HWV2 + lapply (tpss_lift_le … HW12 … HLK2 HWV1 … HWV2) -W1 [ >plus_minus // /2 width=1/ ] /3 width=5/ + ] +] +qed-. + +(* Properties on supclosure *************************************************) + +lemma fsup_tps_trans_full: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ▶[0,|L2|] U2 → + ∃∃L,U1. L1 ▶*[0,|L|] L & L ⊢ T1 ▶[0,|L|] U1 & ⦃L, U1⦄ ⊃ ⦃L2, T2⦄. +#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [1,2,3,4,5: /3 width=5/ ] +#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 +elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 +elim (lift_total T d e) #U #HTU +elim (le_or_ge d (|K|)) #Hd +[ elim (ldrop_ltpss_dx_trans_be … HLK1 … HK1 … Hd) // -HLK1 -HK1 #L2 #HL12 #HL2K + lapply (tps_lift_be … HT1 … HL2K … HTU1 HTU … Hd) // -HT1 -HTU1 #HU1 +| elim (ldrop_ltpss_dx_trans_ge … HLK1 … HK1 Hd) -HLK1 -HK1 #L2 #HL12 #HL2K + lapply (tps_lift_le … HT1 … HL2K … HTU1 HTU Hd) -HT1 -HTU1 #HU1 +] +lapply (ltpss_dx_weak_full … HL12) -HL12 #HL12 +lapply (tps_weak_full … HU1) -HU1 #HU1 +@(ex3_2_intro … L2 U) // /2 width=7/ (**) (* explicit constructor: auto /3 width=14/ too slow *) +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_ltpss_dx.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_ltpss_dx.etc new file mode 100644 index 000000000..74b898d69 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_ltpss_dx.etc @@ -0,0 +1,106 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/tpss_tpss.ma". +include "basic_2/unfold/ltpss_dx_tpss.ma". + +(* DX PARTIAL UNFOLD ON LOCAL ENVIRONMENTS **********************************) + +(* Advanced properties ******************************************************) + +lemma ltpss_dx_tpss_conf: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶* [d2, e2] U2 → + ∀L1,d1,e1. L0 ▶* [d1, e1] L1 → + ∃∃T. L1 ⊢ T2 ▶* [d2, e2] T & + L1 ⊢ U2 ▶* [d1, e1] T. +#L0 #T2 #U2 #d2 #e2 #H #L1 #d1 #e1 #HL01 @(tpss_ind … H) -U2 /2 width=3/ +#U #U2 #_ #HU2 * #X2 #HTX2 #HUX2 +elim (ltpss_dx_tps_conf … HU2 … HL01) -L0 #X1 #HUX1 #HU2X1 +elim (tpss_strip_eq … HUX2 … HUX1) -U #X #HX2 #HX1 +lapply (tpss_trans_eq … HU2X1 … HX1) -X1 /3 width=3/ +qed. + +lemma ltpss_dx_tpss_trans_down: ∀L0,L1,T2,U2,d1,e1,d2,e2. d2 + e2 ≤ d1 → + L1 ▶* [d1, e1] L0 → L0 ⊢ T2 ▶* [d2, e2] U2 → + ∃∃T. L1 ⊢ T2 ▶* [d2, e2] T & L0 ⊢ T ▶* [d1, e1] U2. +#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde2d1 #HL10 #H @(tpss_ind … H) -U2 +[ /2 width=3/ +| #U #U2 #_ #HU2 * #T #HT2 #HTU + elim (tpss_strap1_down … HTU … HU2 ?) -U // #U #HTU #HU2 + elim (ltpss_dx_tps_trans … HTU … HL10) -HTU -HL10 #X #HTX #HXU + lapply (tpss_trans_eq … HXU HU2) -U /3 width=3/ +] +qed. + +lemma ltpss_dx_tpss_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 ▶* [d, e] U2 → + ∀L0. L0 ▶* [d, e] L1 → L0 ⊢ T2 ▶* [d, e] U2. +#L1 #T2 @(f2_ind … fw … L1 T2) -L1 -T2 #n #IH #L1 * +[ #I #Hn #W2 #d #e #H #L0 #HL01 destruct + elim (tpss_inv_atom1 … H) -H // * + #K1 #V1 #V2 #i #Hdi #Hide #HLK1 #HV12 #HVW2 #H destruct + lapply (ldrop_fwd_lw … HLK1) #H1 normalize in H1; + elim (ltpss_dx_ldrop_trans_be … HL01 … HLK1 ? ?) -HL01 -HLK1 // /2 width=2/ #X #H #HLK0 + elim (ltpss_dx_inv_tpss22 … H ?) -H /2 width=1/ #K0 #V0 #HK01 #HV01 #H destruct + lapply (tpss_fwd_tw … HV01) #H2 + lapply (transitive_le (♯{K1} + ♯{V0}) … H1) -H1 /2 width=1/ -H2 #H + lapply (tpss_trans_eq … HV01 HV12) -V1 #HV02 + lapply (IH … HV02 … HK01) -IH -HV02 -HK01 + [ normalize /2 width=1/ | /2 width=6/ ] +| * [ #a ] #I #V2 #T2 #Hn #X #d #e #H #L0 #HL0 destruct + [ elim (tpss_inv_bind1 … H) -H #W2 #U2 #HVW2 #HTU2 #H destruct + lapply (tpss_lsubr_trans … HTU2 (L1. ⓑ{I} V2) ?) -HTU2 /2 width=1/ #HTU2 + lapply (IH … HVW2 … HL0) -HVW2 [ /2 width=2/ ] #HVW2 + lapply (IH … HTU2 (L0. ⓑ{I} V2) ?) -IH -HTU2 // /2 width=2/ -HL0 #HTU2 + lapply (tpss_lsubr_trans … HTU2 (L0. ⓑ{I} W2) ?) -HTU2 /2 width=1/ + | elim (tpss_inv_flat1 … H) -H #W2 #U2 #HVW2 #HTU2 #H destruct + lapply (IH … HVW2 … HL0) -HVW2 // + lapply (IH … HTU2 … HL0) -IH -HTU2 // -HL0 /2 width=1/ +] +qed. + +lemma ltpss_dx_tps_trans_eq: ∀L0,L1,T2,U2,d,e. L0 ▶* [d, e] L1 → + L1 ⊢ T2 ▶ [d, e] U2 → L0 ⊢ T2 ▶* [d, e] U2. +/3 width=3/ qed. + +(* Main properties **********************************************************) + +theorem ltpss_dx_conf: ∀L0,L1,d1,e1. L0 ▶* [d1, e1] L1 → + ∀L2,d2,e2. L0 ▶* [d2, e2] L2 → + ∃∃L. L1 ▶* [d2, e2] L & L2 ▶* [d1, e1] L. +#L0 @(f_ind … lw … L0) -L0 #n #IH * +[ #_ #L1 #d1 #e1 #H1 #L2 #d2 #e2 #H2 -n + >(ltpss_dx_inv_atom1 … H1) -L1 + >(ltpss_dx_inv_atom1 … H2) -L2 /2 width=3/ +| #K0 #I0 #V0 #Hn #L1 #d1 #e1 #H1 #L2 #d2 #e2 #H2 destruct + elim (eq_or_gt d1) #Hd1 [ elim (eq_or_gt e1) #He1 ] destruct + [ lapply (ltpss_dx_inv_refl_O2 … H1) -H1 #H1 + | elim (ltpss_dx_inv_tpss21 … H1 … He1) -H1 #K1 #V1 #HK01 #HV01 #H1 + | elim (ltpss_dx_inv_tpss11 … H1 … Hd1) -H1 #K1 #V1 #HK01 #HV01 #H1 + ] destruct + elim (eq_or_gt d2) #Hd2 [1,3,5: elim (eq_or_gt e2) #He2 ] destruct + [1,3,5: lapply (ltpss_dx_inv_refl_O2 … H2) -H2 #H2 + |2,4,6: elim (ltpss_dx_inv_tpss21 … H2 … He2) -H2 #K2 #V2 #HK02 #HV02 #H2 + |7,8,9: elim (ltpss_dx_inv_tpss11 … H2 … Hd2) -H2 #K2 #V2 #HK02 #HV02 #H2 + ] destruct + [1: -IH /2 width=3/ + |2,3,4,7: -IH /3 width=5/ + |5,6,8,9: + elim (IH … HK01 … HK02) // -K0 #K #HK1 #HK2 + elim (ltpss_dx_tpss_conf … HV01 … HK1) -HV01 #W1 #HW1 #HVW1 + elim (ltpss_dx_tpss_conf … HV02 … HK2) -HV02 #W2 #HW2 #HVW2 + elim (tpss_conf_eq … HW1 … HW2) -V0 #V #HW1 #HW2 + lapply (tpss_trans_eq … HVW1 HW1) -W1 + lapply (tpss_trans_eq … HVW2 HW2) -W2 /3 width=5/ + ] +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_tps.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_tps.etc new file mode 100644 index 000000000..b10306677 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_tps.etc @@ -0,0 +1,47 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/ltpss_dx_ldrop.ma". + +(* DX PARALLEL UNFOLD ON LOCAL ENVIRONMENTS *********************************) + +(* Properties concerning partial substitution on terms **********************) + +lemma ltpss_dx_tps_conf_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → + ∀L1,d1,e1. L0 ▶* [d1, e1] L1 → d1 + e1 ≤ d2 → + L1 ⊢ T2 ▶ [d2, e2] U2. +#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 +[ // +| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL01 #Hde1d2 + lapply (transitive_le … Hde1d2 Hdi2) -Hde1d2 #Hde1i2 + lapply (ltpss_dx_ldrop_conf_ge … HL01 … HLK0 ?) -L0 // /2 width=4/ +| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 #Hde1d2 + @tps_bind [ /2 width=4/ | @IHTU2 -IHTU2 -IHVW2 [3: /2 by ltpss_dx_tpss1/ |1,2: skip | /2 width=1/ ] ] (**) (* explicit constructor *) +| /3 width=4/ +] +qed. + +lemma ltpss_dx_tps_trans_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → + ∀L1,d1,e1. L1 ▶* [d1, e1] L0 → d1 + e1 ≤ d2 → + L1 ⊢ T2 ▶ [d2, e2] U2. +#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 +[ // +| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 #Hde1d2 + lapply (transitive_le … Hde1d2 Hdi2) -Hde1d2 #Hde1i2 + lapply (ltpss_dx_ldrop_trans_ge … HL10 … HLK0 ?) -L0 // /2 width=4/ +| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 #Hde1d2 + @tps_bind [ /2 width=4/ | @IHTU2 -IHTU2 -IHVW2 [3: /2 by ltpss_dx_tpss1/ |1,2: skip | /2 width=1/ ] ] (**) (* explicit constructor *) +| /3 width=4/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_tpss.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_tpss.etc new file mode 100644 index 000000000..9c7875927 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_dx_tpss.etc @@ -0,0 +1,101 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/ltpss_dx_tps.ma". + +(* DX PARALLEL UNFOLD ON LOCAL ENVIRONMENTS *********************************) + +(* Properties concerning partial unfold on terms ****************************) + +lemma ltpss_dx_tpss_conf_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶* [d2, e2] U2 → + ∀L1,d1,e1. L0 ▶* [d1, e1] L1 → d1 + e1 ≤ d2 → + L1 ⊢ T2 ▶* [d2, e2] U2. +#L0 #T2 #U2 #d2 #e2 #H #L1 #d1 #e1 #HL01 #Hde1d2 @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU +lapply (ltpss_dx_tps_conf_ge … HU2 … HL01 ?) -L0 // -Hde1d2 /2 width=3/ +qed. + +(* Basic_1: was: subst1_subst1_back *) +lemma ltpss_dx_tps_conf: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → + ∀L1,d1,e1. L0 ▶* [d1, e1] L1 → + ∃∃T. L1 ⊢ T2 ▶ [d2, e2] T & + L1 ⊢ U2 ▶* [d1, e1] T. +#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 +[ /2 width=3/ +| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL01 + elim (lt_or_ge i2 d1) #Hi2d1 + [ elim (ltpss_dx_ldrop_conf_le … HL01 … HLK0 ?) -L0 /2 width=2/ #X #H #HLK1 + elim (ltpss_dx_inv_tpss11 … H ?) -H /2 width=1/ #K1 #V1 #_ #HV01 #H destruct + lapply (ldrop_fwd_ldrop2 … HLK1) #H + elim (lift_total V1 0 (i2 + 1)) #W1 #HVW1 + lapply (tpss_lift_ge … HV01 … H HVW0 … HVW1) -V0 -H // >minus_plus minus_plus >commutative_plus /2 width=1/ + | lapply (ltpss_dx_ldrop_conf_ge … HL01 … HLK0 ?) -L0 // /3 width=4/ + ] + ] +| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 + elim (IHVW2 … HL01) -IHVW2 #V #HV2 #HVW2 + elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL01 /3 width=5/ +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 + elim (IHVW2 … HL01) -IHVW2 + elim (IHTU2 … HL01) -IHTU2 -HL01 /3 width=5/ +] +qed. + +lemma ltpss_dx_tpss_trans_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶* [d2, e2] U2 → + ∀L1,d1,e1. L1 ▶* [d1, e1] L0 → d1 + e1 ≤ d2 → + L1 ⊢ T2 ▶* [d2, e2] U2. +#L0 #T2 #U2 #d2 #e2 #H #L1 #d1 #e1 #HL01 #Hde1d2 @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU +lapply (ltpss_dx_tps_trans_ge … HU2 … HL01 ?) -L0 // -Hde1d2 /2 width=3/ +qed. + +(* Basic_1: was: subst1_subst1 *) +lemma ltpss_dx_tps_trans: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → + ∀L1,d1,e1. L1 ▶* [d1, e1] L0 → + ∃∃T. L1 ⊢ T2 ▶ [d2, e2] T & + L0 ⊢ T ▶* [d1, e1] U2. +#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 +[ /2 width=3/ +| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 + elim (lt_or_ge i2 d1) #Hi2d1 + [ elim (ltpss_dx_ldrop_trans_le … HL10 … HLK0 ?) -HL10 /2 width=2/ #X #H #HLK1 + elim (ltpss_dx_inv_tpss12 … H ?) -H /2 width=1/ #K1 #V1 #_ #HV01 #H destruct + lapply (ldrop_fwd_ldrop2 … HLK0) -HLK0 #H + elim (lift_total V1 0 (i2 + 1)) #W1 #HVW1 + lapply (tpss_lift_ge … HV01 … H HVW1 … HVW0) -V0 -H // >minus_plus minus_plus >commutative_plus /2 width=1/ + | lapply (ltpss_dx_ldrop_trans_ge … HL10 … HLK0 ?) -HL10 -HLK0 // /3 width=4/ + ] + ] +| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 + elim (IHVW2 … HL10) -IHVW2 #V #HV2 #HVW2 + elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL10 /3 width=5/ +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 + elim (IHVW2 … HL10) -IHVW2 + elim (IHTU2 … HL10) -IHTU2 -HL10 /3 width=5/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_sn.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_sn.etc new file mode 100644 index 000000000..577c1506d --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_sn.etc @@ -0,0 +1,259 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "hvbox( T1 break ⊢ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'PSubstStarSn $T1 $d $e $T2 }. + +include "basic_2/unfold/tpss.ma". + +(* SN PARALLEL UNFOLD ON LOCAL ENVIRONMENTS *********************************) + +inductive ltpss_sn: nat → nat → relation lenv ≝ +| ltpss_sn_atom : ∀d,e. ltpss_sn d e (⋆) (⋆) +| ltpss_sn_pair : ∀L,I,V. ltpss_sn 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V) +| ltpss_sn_tpss2: ∀L1,L2,I,V1,V2,e. + ltpss_sn 0 e L1 L2 → L1 ⊢ V1 ▶* [0, e] V2 → + ltpss_sn 0 (e + 1) (L1. ⓑ{I} V1) (L2. ⓑ{I} V2) +| ltpss_sn_tpss1: ∀L1,L2,I,V1,V2,d,e. + ltpss_sn d e L1 L2 → L1 ⊢ V1 ▶* [d, e] V2 → + ltpss_sn (d + 1) e (L1. ⓑ{I} V1) (L2. ⓑ{I} V2) +. + +interpretation "parallel unfold (local environment, sn variant)" + 'PSubstStarSn L1 d e L2 = (ltpss_sn d e L1 L2). + +(* Basic inversion lemmas ***************************************************) + +fact ltpss_sn_inv_refl_O2_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → e = 0 → L1 = L2. +#d #e #L1 #L2 #H elim H -d -e -L1 -L2 // +[ #L1 #L2 #I #V1 #V2 #e #_ #_ #_ >commutative_plus normalize #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #HV12 #IHL12 #He destruct + >(IHL12 ?) -IHL12 // >(tpss_inv_refl_O2 … HV12) // +] +qed. + +lemma ltpss_sn_inv_refl_O2: ∀d,L1,L2. L1 ⊢ ▶* [d, 0] L2 → L1 = L2. +/2 width=4/ qed-. + +fact ltpss_sn_inv_atom1_aux: ∀d,e,L1,L2. + L1 ⊢ ▶* [d, e] L2 → L1 = ⋆ → L2 = ⋆. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ // +| #L #I #V #H destruct +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct +] +qed. + +lemma ltpss_sn_inv_atom1: ∀d,e,L2. ⋆ ⊢ ▶* [d, e] L2 → L2 = ⋆. +/2 width=5/ qed-. + +fact ltpss_sn_inv_tpss21_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → d = 0 → 0 < e → + ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → + ∃∃K2,V2. K1 ⊢ ▶* [0, e - 1] K2 & + K1 ⊢ V1 ▶* [0, e - 1] V2 & + L2 = K2. ⓑ{I} V2. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #_ #K1 #I #V1 #H destruct +| #L1 #I #V #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K1 #J #W1 #H destruct /2 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma ltpss_sn_inv_tpss21: ∀e,K1,I,V1,L2. K1. ⓑ{I} V1 ⊢ ▶* [0, e] L2 → 0 < e → + ∃∃K2,V2. K1 ⊢ ▶* [0, e - 1] K2 & + K1 ⊢ V1 ▶* [0, e - 1] V2 & + L2 = K2. ⓑ{I} V2. +/2 width=5/ qed-. + +fact ltpss_sn_inv_tpss11_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → 0 < d → + ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → + ∃∃K2,V2. K1 ⊢ ▶* [d - 1, e] K2 & + K1 ⊢ V1 ▶* [d - 1, e] V2 & + L2 = K2. ⓑ{I} V2. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #I #K1 #V1 #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K1 #W1 #H destruct /2 width=5/ +] +qed. + +lemma ltpss_sn_inv_tpss11: ∀d,e,I,K1,V1,L2. K1. ⓑ{I} V1 ⊢ ▶* [d, e] L2 → 0 < d → + ∃∃K2,V2. K1 ⊢ ▶* [d - 1, e] K2 & + K1 ⊢ V1 ▶* [d - 1, e] V2 & + L2 = K2. ⓑ{I} V2. +/2 width=3/ qed-. + +fact ltpss_sn_inv_atom2_aux: ∀d,e,L1,L2. + L1 ⊢ ▶* [d, e] L2 → L2 = ⋆ → L1 = ⋆. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ // +| #L #I #V #H destruct +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct +] +qed. + +lemma ltpss_sn_inv_atom2: ∀d,e,L1. L1 ⊢ ▶* [d, e] ⋆ → L1 = ⋆. +/2 width=5/ qed-. + +fact ltpss_sn_inv_tpss22_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → d = 0 → 0 < e → + ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 ⊢ ▶* [0, e - 1] K2 & + K1 ⊢ V1 ▶* [0, e - 1] V2 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #_ #K1 #I #V1 #H destruct +| #L1 #I #V #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K2 #J #W2 #H destruct /2 width=5/ +| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct +] +qed. + +lemma ltpss_sn_inv_tpss22: ∀e,L1,K2,I,V2. L1 ⊢ ▶* [0, e] K2. ⓑ{I} V2 → 0 < e → + ∃∃K1,V1. K1 ⊢ ▶* [0, e - 1] K2 & + K1 ⊢ V1 ▶* [0, e - 1] V2 & + L1 = K1. ⓑ{I} V1. +/2 width=5/ qed-. + +fact ltpss_sn_inv_tpss12_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → 0 < d → + ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → + ∃∃K1,V1. K1 ⊢ ▶* [d - 1, e] K2 & + K1 ⊢ V1 ▶* [d - 1, e] V2 & + L1 = K1. ⓑ{I} V1. +#d #e #L1 #L2 * -d -e -L1 -L2 +[ #d #e #_ #I #K2 #V2 #H destruct +| #L #I #V #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) +| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K2 #W2 #H destruct /2 width=5/ +] +qed. + +lemma ltpss_sn_inv_tpss12: ∀L1,K2,I,V2,d,e. L1 ⊢ ▶* [d, e] K2. ⓑ{I} V2 → 0 < d → + ∃∃K1,V1. K1 ⊢ ▶* [d - 1, e] K2 & + K1 ⊢ V1 ▶* [d - 1, e] V2 & + L1 = K1. ⓑ{I} V1. +/2 width=3/ qed-. + +(* Basic properties *********************************************************) + +lemma ltpss_sn_tps2: ∀L1,L2,I,V1,V2,e. + L1 ⊢ ▶* [0, e] L2 → L1 ⊢ V1 ▶ [0, e] V2 → + L1. ⓑ{I} V1 ⊢ ▶* [0, e + 1] L2. ⓑ{I} V2. +/3 width=1/ qed. + +lemma ltpss_sn_tps1: ∀L1,L2,I,V1,V2,d,e. + L1 ⊢ ▶* [d, e] L2 → L1 ⊢ V1 ▶ [d, e] V2 → + L1. ⓑ{I} V1 ⊢ ▶* [d + 1, e] L2. ⓑ{I} V2. +/3 width=1/ qed. + +lemma ltpss_sn_tpss2_lt: ∀L1,L2,I,V1,V2,e. + L1 ⊢ ▶* [0, e - 1] L2 → L1 ⊢ V1 ▶* [0, e - 1] V2 → + 0 < e → L1. ⓑ{I} V1 ⊢ ▶* [0, e] L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He +>(plus_minus_m_m e 1) /2 width=1/ +qed. + +lemma ltpss_sn_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 ⊢ ▶* [d - 1, e] L2 → L1 ⊢ V1 ▶* [d - 1, e] V2 → + 0 < d → L1. ⓑ{I} V1 ⊢ ▶* [d, e] L2. ⓑ{I} V2. +#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd +>(plus_minus_m_m d 1) /2 width=1/ +qed. + +lemma ltpss_sn_tps2_lt: ∀L1,L2,I,V1,V2,e. + L1 ⊢ ▶* [0, e - 1] L2 → L1 ⊢ V1 ▶ [0, e - 1] V2 → + 0 < e → L1. ⓑ{I} V1 ⊢ ▶* [0, e] L2. ⓑ{I} V2. +/3 width=1/ qed. + +lemma ltpss_sn_tps1_lt: ∀L1,L2,I,V1,V2,d,e. + L1 ⊢ ▶* [d - 1, e] L2 → L1 ⊢ V1 ▶ [d - 1, e] V2 → + 0 < d → L1. ⓑ{I} V1 ⊢ ▶* [d, e] L2. ⓑ{I} V2. +/3 width=1/ qed. + +lemma ltpss_sn_refl: ∀L,d,e. L ⊢ ▶* [d, e] L. +#L elim L -L // +#L #I #V #IHL * /2 width=1/ * /2 width=1/ +qed. + +lemma ltpss_sn_weak: ∀L1,L2,d1,e1. L1 ⊢ ▶* [d1, e1] L2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → L1 ⊢ ▶* [d2, e2] L2. +#L1 #L2 #d1 #e1 #H elim H -L1 -L2 -d1 -e1 // +[ #L1 #L2 #I #V1 #V2 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd2 #Hde2 + lapply (le_n_O_to_eq … Hd2) #H destruct normalize in Hde2; + lapply (lt_to_le_to_lt 0 … Hde2) // #He2 + lapply (le_plus_to_minus_r … Hde2) -Hde2 /3 width=5/ +| #L1 #L2 #I #V1 #V2 #d1 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd21 #Hde12 + >plus_plus_comm_23 in Hde12; #Hde12 + elim (le_to_or_lt_eq 0 d2 ?) // #H destruct + [ lapply (le_plus_to_minus_r … Hde12) -Hde12 plus_plus_comm_23 + /4 width=5 by ltpss_sn_tpss2, tpss_append, tpss_weak, monotonic_le_plus_r/ (**) (* too slow without trace *) +| #K1 #K2 #I #V1 #V2 #d #x #_ #HV12 #IHK12 normalize (ldrop_inv_atom1 … H) -H // +| // +| normalize #K0 #K1 #I #V0 #V1 #e1 #_ #_ #IHK01 #L2 #e2 #H #He12 + elim (le_inv_plus_l … He12) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ +| #K0 #K1 #I #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ +] +qed. + +lemma ltpss_sn_ldrop_trans_ge: ∀L1,L0,d1,e1. L1 ⊢ ▶* [d1, e1] L0 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → + d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. +#L1 #L0 #d1 #e1 #H elim H -L1 -L0 -d1 -e1 +[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H // +| // +| normalize #K1 #K0 #I #V1 #V0 #e1 #_ #_ #IHK10 #L2 #e2 #H #He12 + elim (le_inv_plus_l … He12) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK10 … HK0L2 ?) -K0 /2 width=1/ +| #K0 #K1 #I #V1 #V0 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 + lapply (IHK10 … HK0L2 ?) -IHK10 -HK0L2 /2 width=1/ +] +qed. + +lemma ltpss_sn_ldrop_conf_be: ∀L0,L1,d1,e1. L0 ⊢ ▶* [d1, e1] L1 → + ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → + ∃∃L. L2 ⊢ ▶* [0, d1 + e1 - e2] L & ⇩[0, e2] L1 ≡ L. +#L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 +[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H /2 width=3/ +| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 + lapply (le_n_O_to_eq … He2) -He2 #H destruct + lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ +| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #_ #He21 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK01 -He21 destruct plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 #He2de1 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + (ldrop_inv_atom1 … H) -H /2 width=3/ +| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 + lapply (le_n_O_to_eq … He2) -He2 #H destruct + lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ +| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #_ #He21 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK10 -He21 destruct plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 #He2de1 + elim (le_inv_plus_l … Hd1e2) #_ #He2 + (ldrop_inv_atom1 … H) -H /2 width=3/ +| /2 width=3/ +| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #_ #L2 #e2 #H #He2 + lapply (le_n_O_to_eq … He2) -He2 #He2 destruct + lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ +| #K0 #K1 #I #V0 #V1 #d1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #He2d1 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK01 -He2d1 destruct (ldrop_inv_atom1 … H) -H /2 width=3/ +| /2 width=3/ +| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #_ #L2 #e2 #H #He2 + lapply (le_n_O_to_eq … He2) -He2 #He2 destruct + lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ +| #K1 #K0 #I #V1 #V0 #d1 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #He2d1 + lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 + [ -IHK10 -He2d1 destruct (ltpss_sn_inv_atom1 … H) -H /2 width=3/ +| /2 width=3/ +| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #Hd + elim (IHLK1 … HK12 Hd) -K1 -Hd /3 width=5/ +| #L1 #K1 #I #V1 #W1 #d1 #e1 #HLK1 #HWV1 #IHLK1 #X #d2 #e2 #H #Hd12 + elim (le_inv_plus_l … Hd12) -Hd12 #Hd12 #Hd2 + elim (ltpss_sn_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (IHLK1 … HK12 … Hd12) -IHLK1 -HK12 (ltpss_sn_inv_atom1 … H) -H /2 width=3/ +| #K1 #I #V1 #K2 #d2 #e2 #HK12 #H #_ + lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/ +| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #H1 #H2 + elim (IHLK1 … HK12 H1 H2) -K1 -H2 + lapply (le_n_O_to_eq … H1) -H1 #H destruct /3 width=5/ +| #L1 #K1 #I #V1 #W1 #d1 #e1 #HLK1 #HWV1 #IHLK1 #X #d2 #e2 #H #Hd21 #Hd12 + elim (eq_or_gt d2) #Hd2 [ -Hd21 elim (eq_or_gt e2) #He2 ] destruct + [ lapply (le_n_O_to_eq … Hd12) -Hd12 plus_minus_commutative // #L2 #HL12 #HLK2 + elim (lift_total W2 d1 e1) #V2 #HWV2 + lapply (tpss_lift_be … HW12 … HLK1 HWV1 … HWV2) -HLK1 -W1 // /2 width=1/ + >plus_minus // >commutative_plus /4 width=5/ + | elim (ltpss_sn_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (IHLK1 … HK12 …) -IHLK1 [2: >plus_minus // ] /2 width=1/ #L2 #HL12 #HLK2 + elim (lift_total W2 d1 e1) #V2 #HWV2 + lapply (tpss_lift_be … HW12 … HLK1 HWV1 … HWV2) -HLK1 -W1 [ >plus_minus // ] /2 width=1/ + >commutative_plus /3 width=5/ + ] +] +qed-. + +lemma ldrop_ltpss_sn_trans_ge: ∀L1,K1,d1,e1. ⇩[d1, e1] L1 ≡ K1 → + ∀K2,d2,e2. K1 ⊢ ▶* [d2, e2] K2 → d2 + e2 ≤ d1 → + ∃∃L2. L1 ⊢ ▶* [d2, e2] L2 & ⇩[d1, e1] L2 ≡ K2. +#L1 #K1 #d1 #e1 #H elim H -L1 -K1 -d1 -e1 +[ #d1 #e1 #K2 #d2 #e2 #H #_ + >(ltpss_sn_inv_atom1 … H) -H /2 width=3/ +| #K1 #I #V1 #K2 #d2 #e2 #HK12 #H + elim (plus_le_0 … H) -H #H1 #H2 destruct /2 width=3/ +| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #H + elim (IHLK1 … HK12 H) -K1 + elim (plus_le_0 … H) -H #H1 #H2 destruct #L2 #HL12 + >(ltpss_sn_inv_refl_O2 … HL12) -L1 /3 width=5/ +| #L1 #K1 #I #V1 #W1 #d1 #e1 #HLK1 #HWV1 #IHLK1 #X #d2 #e2 #H #Hd21 + elim (eq_or_gt d2) #Hd2 [ elim (eq_or_gt e2) #He2 ] destruct + [ -IHLK1 -Hd21 <(ltpss_sn_inv_refl_O2 … H) -X /3 width=5/ + | elim (ltpss_sn_inv_tpss21 … H He2) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (IHLK1 … HK12 …) -IHLK1 /2 width=1/ #L2 #HL12 #HLK2 + elim (lift_total W2 d1 e1) #V2 #HWV2 + lapply (tpss_lift_le … HW12 … HLK1 HWV1 … HWV2) -HLK1 -W1 /2 width=1/ /3 width=5/ + | elim (ltpss_sn_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct + elim (IHLK1 … HK12 …) -IHLK1 [2: >plus_minus // /2 width=1/ ] #L2 #HL12 #HLK2 + elim (lift_total W2 d1 e1) #V2 #HWV2 + lapply (tpss_lift_le … HW12 … HLK1 HWV1 … HWV2) -HLK1 -W1 [ >plus_minus // /2 width=1/ ] /3 width=5/ + ] +] +qed-. + +(* Properties on supclosure *************************************************) + +lemma fsup_tpss_trans_full: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ▶*[0,|L2|] U2 → + ∃∃L,U1. L1 ⊢ ▶*[0,|L1|] L & L ⊢ T1 ▶*[0,|L|] U1 & ⦃L, U1⦄ ⊃ ⦃L2, T2⦄. +#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [1,2,3,4,5: /3 width=5/ ] +#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 +elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 +elim (lift_total T d e) #U #HTU +lapply (ltpss_sn_fwd_length … HK1) #H >H in HK1; -H #HK1 +elim (le_or_ge d (|K|)) #Hd +[ elim (ldrop_ltpss_sn_trans_be … HLK1 … HK1 … Hd) // -HLK1 -HK1 #L2 #HL12 #HL2K + lapply (tpss_lift_be … HT1 … Hd HL2K HTU1 … HTU) // -HT1 -HTU1 #HU1 +| elim (ldrop_ltpss_sn_trans_ge … HLK1 … HK1 Hd) -HLK1 -HK1 #L2 #HL12 #HL2K + lapply (tpss_lift_le … HT1 … Hd HL2K HTU1 … HTU) -HT1 -HTU1 #HU1 +] +lapply (ltpss_sn_weak_full … HL12) -HL12 #HL12 +lapply (tpss_weak_full … HU1) -HU1 #HU1 +@(ex3_2_intro … L2 U) // /2 width=7/ (**) (* explicit constructor: auto /3 width=14/ too slow *) +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_sn_ltpss_sn.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_sn_ltpss_sn.etc new file mode 100644 index 000000000..67c287e2b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ltpss_sn_ltpss_sn.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/unfold/tpss_tpss.ma". +include "basic_2/unfold/ltpss_sn_tpss.ma". + +(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) + +(* Advanced properties ******************************************************) + +lemma ltpss_sn_tpss_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 ▶* [d, e] U2 → + ∀L0. L0 ⊢ ▶* [d, e] L1 → L0 ⊢ T2 ▶* [d, e] U2. +#L1 #T2 @(f2_ind … fw … L1 T2) -L1 -T2 #n #IH #L1 * +[ #I #Hn #W2 #d #e #H #L0 #HL01 destruct + elim (tpss_inv_atom1 … H) -H // * + #K1 #V1 #V2 #i #Hdi #Hide #HLK1 #HV12 #HVW2 #H destruct + lapply (ldrop_fwd_lw … HLK1) #H1 normalize in H1; + elim (ltpss_sn_ldrop_trans_be … HL01 … HLK1 ? ?) -HL01 -HLK1 // /2 width=2/ #X #H #HLK0 + elim (ltpss_sn_inv_tpss22 … H ?) -H /2 width=1/ #K0 #V0 #HK01 #HV01 #H destruct + lapply (IH … HV12 … HK01) -IH -HV12 -HK01 [ normalize /2 width=1/ ] #HV12 + lapply (tpss_trans_eq … HV01 HV12) -V1 /2 width=6/ +| * [ #a ] #I #V2 #T2 #Hn #X #d #e #H #L0 #HL0 destruct + [ elim (tpss_inv_bind1 … H) -H #W2 #U2 #HVW2 #HTU2 #H destruct + lapply (tpss_lsubr_trans … HTU2 (L1. ⓑ{I} V2) ?) -HTU2 /2 width=1/ #HTU2 + lapply (IH … HVW2 … HL0) -HVW2 [ /2 width=2/ ] #HVW2 + lapply (IH … HTU2 (L0. ⓑ{I} V2) ?) -IH -HTU2 // /2 width=2/ -HL0 #HTU2 + lapply (tpss_lsubr_trans … HTU2 (L0. ⓑ{I} W2) ?) -HTU2 /2 width=1/ + | elim (tpss_inv_flat1 … H) -H #W2 #U2 #HVW2 #HTU2 #H destruct + lapply (IH … HVW2 … HL0) -HVW2 // + lapply (IH … HTU2 … HL0) -IH -HTU2 // -HL0 /2 width=1/ +] +qed. + +lemma ltpss_sn_tps_trans_eq: ∀L0,L1,T2,U2,d,e. L0 ⊢ ▶* [d, e] L1 → + L1 ⊢ T2 ▶ [d, e] U2 → L0 ⊢ T2 ▶* [d, e] U2. +/3 width=3/ qed. + +(* Main properties **********************************************************) + +theorem ltpss_sn_trans_eq: ∀L1,L,d,e. L1 ⊢ ▶* [d, e] L → + ∀L2. L ⊢ ▶* [d, e] L2 → L1 ⊢ ▶* [d, e] L2. +#L1 #L #d #e #H elim H -L1 -L -d -e // +[ #L1 #L #I #V1 #V #e #HL1 #HV1 #IHL1 #X #H + elim (ltpss_sn_inv_tpss21 … H ?) -H // shift_append_assoc #H + elim (tps_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + elim (IH … HT12) -IH -HT12 #L2 #T #HL12 #HT1 #H destruct + append_length minus_plus minus_plus >commutative_plus /2 width=1/ + | lapply (ltpss_sn_ldrop_conf_ge … HL01 … HLK0 ?) -HL01 -HLK0 // /3 width=4/ + ] + ] +| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 + elim (IHVW2 … HL01) -IHVW2 #V #HV2 #HVW2 + elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL01 #T #HT2 #H + lapply (tpss_lsubr_trans … H (L0.ⓑ{I}V) ?) -H /2 width=1/ /3 width=5/ +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 + elim (IHVW2 … HL01) -IHVW2 + elim (IHTU2 … HL01) -IHTU2 -HL01 /3 width=5/ +] +qed. + +lemma ltpss_sn_tpss_trans_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶* [d2, e2] U2 → + ∀L1,d1,e1. L1 ⊢ ▶* [d1, e1] L0 → d1 + e1 ≤ d2 → + L1 ⊢ T2 ▶* [d2, e2] U2. +#L0 #T2 #U2 #d2 #e2 #H #L1 #d1 #e1 #HL01 #Hde1d2 @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU +lapply (ltpss_sn_tps_trans_ge … HU2 … HL01 ?) -L0 // -Hde1d2 /2 width=3/ +qed. + +lemma ltpss_sn_tps_trans: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → + ∀L1,d1,e1. L1 ⊢ ▶* [d1, e1] L0 → + ∃∃T. L1 ⊢ T2 ▶ [d2, e2] T & + L1 ⊢ T ▶* [d1, e1] U2. +#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 +[ /2 width=3/ +| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 + elim (lt_or_ge i2 d1) #Hi2d1 + [ elim (ltpss_sn_ldrop_trans_le … HL10 … HLK0 ?) -L0 /2 width=2/ #X #H #HLK1 + elim (ltpss_sn_inv_tpss12 … H ?) -H /2 width=1/ #K1 #V1 #_ #HV01 #H destruct + lapply (ldrop_fwd_ldrop2 … HLK1) #H + elim (lift_total V1 0 (i2 + 1)) #W1 #HVW1 + lapply (tpss_lift_ge … HV01 … H HVW1 … HVW0) -V0 -H // >minus_plus minus_plus >commutative_plus /2 width=1/ + | lapply (ltpss_sn_ldrop_trans_ge … HL10 … HLK0 ?) -HL10 -HLK0 // /3 width=4/ + ] + ] +| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 + elim (IHVW2 … HL10) -IHVW2 #V #HV2 #HVW2 + elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL10 #T #HT2 #H + lapply (tpss_lsubr_trans … H (L1.ⓑ{I}W2) ?) -H /2 width=1/ /3 width=5/ +| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 + elim (IHVW2 … HL10) -IHVW2 + elim (IHTU2 … HL10) -IHTU2 -HL10 /3 width=5/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ssta_ltpss_dx.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ssta_ltpss_dx.etc new file mode 100644 index 000000000..72eff7850 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/ssta_ltpss_dx.etc @@ -0,0 +1,125 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/tpss_tpss.ma". +include "basic_2/unfold/ltpss_dx_tpss.ma". +include "basic_2/static/ssta_lift.ma". + +(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************) + +(* Properties about dx parallel unfold **************************************) + +(* Note: apparently this was missing in basic_1 *) +lemma ssta_ltpss_dx_tpss_conf: ∀h,g,L1,T1,U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ → + ∀L2,d,e. L1 ▶* [d, e] L2 → + ∀T2. L2 ⊢ T1 ▶* [d, e] T2 → + ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄ & + L2 ⊢ U1 ▶* [d, e] U2. +#h #g #L1 #T1 #U #l #H elim H -L1 -T1 -U -l +[ #L1 #k1 #l1 #Hkl1 #L2 #d #e #_ #T2 #H + >(tpss_inv_sort1 … H) -H /3 width=3/ +| #L1 #K1 #V1 #W1 #U1 #i #l #HLK1 #HVW1 #HWU1 #IHVW1 #L2 #d #e #HL12 #T2 #H + elim (tpss_inv_lref1 … H) -H [ | -HVW1 ] + [ #H destruct + elim (lt_or_ge i d) #Hdi [ -HVW1 | ] + [ elim (ltpss_dx_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2 + elim (ltpss_dx_inv_tpss11 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct + elim (IHVW1 … HK12 … HV12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12 + lapply (ldrop_fwd_ldrop2 … HLK2) #H + elim (lift_total W2 0 (i+1)) #U2 #HWU2 + lapply (tpss_lift_ge … HW12 … H … HWU1 … HWU2) // -HW12 -H -HWU1 + >minus_plus minus_plus #H + lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus minus_plus #H + lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus minus_plus minus_plus #H + lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus shift_append_assoc normalize #H + elim (tpr_inv_bind1 … H) -H * + [ #V0 #T0 #X0 #_ #HT10 #H0 #H destruct + elim (IH … HT10) -IH -T1 #L #T #HL1 #H destruct + elim (tps_fwd_shift1 … H0) -T #L2 #T2 #HL2 #H destruct + >append_length >HL1 >HL2 -L1 -L + @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] // /2 width=3/ (**) (* explicit constructor *) + | #T #_ #_ #H destruct + ] +] +qed-. + +(* Basic_1: removed theorems 3: + pr0_subst0_back pr0_subst0_fwd pr0_subst0 +*) +(* Basic_1: removed local theorems: 1: pr0_delta_tau *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpr_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpr_lift.etc new file mode 100644 index 000000000..8d6b0363a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpr_lift.etc @@ -0,0 +1,121 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/tps_lift.ma". +include "basic_2/reducibility/tpr.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) + +(* Relocation properties ****************************************************) + +(* Basic_1: was: pr0_lift *) +lemma tpr_lift: t_liftable tpr. +#T1 #T2 #H elim H -T1 -T2 +[ * #i #U1 #d #e #HU1 #U2 #HU2 + lapply (lift_mono … HU1 … HU2) -HU1 #H destruct + [ lapply (lift_inv_sort1 … HU2) -HU2 #H destruct // + | lapply (lift_inv_lref1 … HU2) * * #Hid #H destruct // + | lapply (lift_inv_gref1 … HU2) -HU2 #H destruct // + ] +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #X1 #d #e #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct + elim (lift_inv_flat1 … HX2) -HX2 #W2 #U2 #HVW2 #HTU2 #HX2 destruct /3 width=4/ +| #a #V1 #V2 #W #T1 #T2 #_ #_ #IHV12 #IHT12 #X1 #d #e #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct + elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct + elim (lift_inv_bind1 … HX2) -HX2 #V3 #T3 #HV23 #HT23 #HX2 destruct /3 width=4/ +| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV12 #IHT1 #X1 #d #e #HX1 #X2 #HX2 + elim (lift_inv_bind1 … HX1) -HX1 #W1 #U1 #HVW1 #HTU1 #HX1 destruct + elim (lift_inv_bind1 … HX2) -HX2 #W2 #U0 #HVW2 #HTU0 #HX2 destruct + elim (lift_total T (d + 1) e) #U #HTU + @tpr_delta + [4: @(tps_lift_le … HT2 … HTU HTU0 ?) /2 width=1/ |1: skip |2: /2 width=4/ |3: /2 width=4/ ] (**) (*/3. is too slow *) +| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #X1 #d #e #HX1 #X2 #HX2 + elim (lift_inv_flat1 … HX1) -HX1 #V0 #X #HV10 #HX #HX1 destruct + elim (lift_inv_bind1 … HX) -HX #W0 #T0 #HW0 #HT10 #HX destruct + elim (lift_inv_bind1 … HX2) -HX2 #W3 #X #HW23 #HX #HX2 destruct + elim (lift_inv_flat1 … HX) -HX #V3 #T3 #HV3 #HT23 #HX destruct + elim (lift_trans_ge … HV2 … HV3 ?) -V // /3 width=4/ +| #V #T1 #T #T2 #_ #HT2 #IHT1 #X #d #e #H #U2 #HTU2 + elim (lift_inv_bind1 … H) -H #V3 #T3 #_ #HT13 #H destruct -V + elim (lift_conf_O1 … HTU2 … HT2) -T2 /3 width=4/ +| #V #T1 #T2 #_ #IHT12 #X #d #e #HX #T #HT2 + elim (lift_inv_flat1 … HX) -HX #V0 #T0 #_ #HT0 #HX destruct /3 width=4/ +] +qed. + +(* Basic_1: was: pr0_gen_lift *) +lemma tpr_inv_lift1: t_deliftable_sn tpr. +#T1 #T2 #H elim H -T1 -T2 +[ * #i #X #d #e #HX + [ lapply (lift_inv_sort2 … HX) -HX #H destruct /2 width=3/ + | lapply (lift_inv_lref2 … HX) -HX * * #Hid #H destruct /3 width=3/ + | lapply (lift_inv_gref2 … HX) -HX #H destruct /2 width=3/ + ] +| #I #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #X #d #e #HX + elim (lift_inv_flat2 … HX) -HX #V0 #T0 #HV01 #HT01 #HX destruct + elim (IHV12 … HV01) -V1 + elim (IHT12 … HT01) -T1 /3 width=5/ +| #a #V1 #V2 #W1 #T1 #T2 #_ #_ #IHV12 #IHT12 #X #d #e #HX + elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct + elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct + elim (IHV12 … HV01) -V1 + elim (IHT12 … HT01) -T1 /3 width=5/ +| #a #I #V1 #V2 #T1 #T #T2 #_ #_ #HT2 #IHV12 #IHT1 #X #d #e #HX + elim (lift_inv_bind2 … HX) -HX #W1 #U1 #HWV1 #HUT1 #HX destruct + elim (IHV12 … HWV1) -V1 #W2 #HWV2 #HW12 + elim (IHT1 … HUT1) -T1 #U #HUT #HU1 + elim (tps_inv_lift1_le … HT2 … HUT ?) -T // [3: /2 width=5/ |2: skip ] #U2 #HU2 #HUT2 + @ex2_intro [2: /2 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /3 width=5/ is slow *) +| #a #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #HV2 #_ #_ #IHV12 #IHW12 #IHT12 #X #d #e #HX + elim (lift_inv_flat2 … HX) -HX #V0 #Y #HV01 #HY #HX destruct + elim (lift_inv_bind2 … HY) -HY #W0 #T0 #HW01 #HT01 #HY destruct + elim (IHV12 … HV01) -V1 #V3 #HV32 #HV03 + elim (IHW12 … HW01) -W1 #W3 #HW32 #HW03 + elim (IHT12 … HT01) -T1 #T3 #HT32 #HT03 + elim (lift_trans_le … HV32 … HV2 ?) -V2 // #V2 #HV32 #HV2 + @ex2_intro [2: /3 width=2/ |1: skip |3: /2 width=3/ ] (**) (* /4 width=5/ is slow *) +| #V #T1 #T #T2 #_ #HT2 #IHT1 #X #d #e #HX + elim (lift_inv_bind2 … HX) -HX #V3 #T3 #_ #HT31 #H destruct + elim (IHT1 … HT31) -T1 #T1 #HT1 #HT31 + elim (lift_div_le … HT2 … HT1 ?) -T // /3 width=5/ +| #V #T1 #T2 #_ #IHT12 #X #d #e #HX + elim (lift_inv_flat2 … HX) -HX #V0 #T0 #_ #HT01 #H destruct + elim (IHT12 … HT01) -T1 /3 width=3/ +] +qed-. + +(* Advanced inversion lemmas ************************************************) + +(* Basic_1: was pr0_gen_abst *) +lemma tpr_inv_abst1: ∀a,V1,T1,U2. ⓛ{a}V1. T1 ➡ U2 → + ∃∃V2,T2. V1 ➡ V2 & T1 ➡ T2 & U2 = ⓛ{a}V2. T2. +#a #V1 #T1 #U2 #H elim (tpr_inv_bind1 … H) -H * +[ #V2 #T #T2 #HV12 #HT1 #HT2 + lapply (tps_inv_refl_SO2 … HT2 ???) -HT2 // /2 width=5/ +| #T2 #_ #_ #_ #H destruct +] +qed-. + +(* Advanced forward lemmas **************************************************) + +lemma tpr_fwd_abst1: ∀a,V1,T1,U2. ⓛ{a}V1.T1 ➡ U2 → ∀b,I,W. + ∃∃V2,T2. ⓑ{b,I}W.T1 ➡ ⓑ{b,I}W.T2 & + U2 = ⓛ{a}V2.T2. +#a #V1 #T1 #U2 #H #b #I #W elim (tpr_inv_bind1 … H) -H * +[ #V2 #T #T2 #HV12 #HT1 #HT2 + lapply (tps_inv_refl_SO2 … HT2 ???) -HT2 // /3 width=4/ +| #T2 #_ #_ #_ #H destruct +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpr_tpr.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpr_tpr.etc new file mode 100644 index 000000000..a733345bf --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpr_tpr.etc @@ -0,0 +1,261 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reducibility/ltpr_tpss.ma". + +(* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************) + +(* Confluence lemmas ********************************************************) + +fact tpr_conf_atom_atom: ∀I. ∃∃X. ⓪{I} ➡ X & ⓪{I} ➡ X. +/2 width=3/ qed. + +fact tpr_conf_flat_flat: + ∀I,V0,V1,T0,T1,V2,T2. ( + ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → + ∃∃T0. ⓕ{I} V1. T1 ➡ T0 & ⓕ{I} V2. T2 ➡ T0. +#I #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02 +elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /3 width=5/ +qed. + +fact tpr_conf_flat_beta: + ∀a,V0,V1,T1,V2,W0,U0,T2. ( + ∀X0:term. ♯{X0} < ♯{V0} + (♯{W0} + ♯{U0} + 1) + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → + U0 ➡ T2 → ⓛ{a}W0. U0 ➡ T1 → + ∃∃X. ⓐV1. T1 ➡ X & ⓓ{a}V2. T2 ➡ X. +#a #V0 #V1 #T1 #V2 #W0 #U0 #T2 #IH #HV01 #HV02 #HT02 #H +elim (tpr_inv_abst1 … H) -H #W1 #U1 #HW01 #HU01 #H destruct +elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (IH … HT02 … HU01) -HT02 -HU01 -IH /2 width=1/ /3 width=5/ +qed. + +(* Basic-1: was: + pr0_cong_upsilon_refl pr0_cong_upsilon_zeta + pr0_cong_upsilon_cong pr0_cong_upsilon_delta +*) +fact tpr_conf_flat_theta: + ∀a,V0,V1,T1,V2,V,W0,W2,U0,U2. ( + ∀X0:term. ♯{X0} < ♯{V0} + (♯{W0} + ♯{U0} + 1) + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → ⇧[O,1] V2 ≡ V → + W0 ➡ W2 → U0 ➡ U2 → ⓓ{a}W0. U0 ➡ T1 → + ∃∃X. ⓐV1. T1 ➡ X & ⓓ{a}W2. ⓐV. U2 ➡ X. +#a #V0 #V1 #T1 #V2 #V #W0 #W2 #U0 #U2 #IH #HV01 #HV02 #HV2 #HW02 #HU02 #H +elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #VV #HVV1 #HVV2 +elim (lift_total VV 0 1) #VVV #HVV +lapply (tpr_lift … HVV2 … HV2 … HVV) #HVVV +elim (tpr_inv_abbr1 … H) -H * +(* case 1: delta *) +[ -HV2 -HVV2 #WW2 #UU2 #UU #HWW2 #HUU02 #HUU2 #H destruct + elim (IH … HW02 … HWW2) -HW02 -HWW2 /2 width=1/ #W #HW02 #HWW2 + elim (IH … HU02 … HUU02) -HU02 -HUU02 -IH /2 width=1/ #U #HU2 #HUUU2 + elim (tpr_tps_conf_bind … HWW2 HUUU2 … HUU2) -UU2 #UUU #HUUU2 #HUUU1 + @ex2_intro + [2: @tpr_theta [6: @HVV |7: @HWW2 |8: @HUUU2 |1,2,3,4: skip | // ] + |1:skip + |3: @tpr_delta [3: @tpr_flat |1: skip ] /2 width=5/ + ] (**) (* /5 width=14/ is too slow *) +(* case 3: zeta *) +| -HV2 -HW02 -HVV2 #U1 #HU01 #HTU1 + elim (IH … HU01 … HU02) -HU01 -HU02 -IH // -U0 #U #HU1 #HU2 + elim (tpr_inv_lift1 … HU1 … HTU1) -U1 #UU #HUU #HT1UU #H destruct + @(ex2_intro … (ⓐVV.UU)) /2 width=1/ /3 width=5/ (**) (* /4 width=9/ is too slow *) +] +qed. + +fact tpr_conf_flat_cast: + ∀X2,V0,V1,T0,T1. ( + ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → T0 ➡ T1 → T0 ➡ X2 → + ∃∃X. ⓝV1. T1 ➡ X & X2 ➡ X. +#X2 #V0 #V1 #T0 #T1 #IH #_ #HT01 #HT02 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /3 width=3/ +qed. + +fact tpr_conf_beta_beta: + ∀a. ∀W0:term. ∀V0,V1,T0,T1,V2,T2. ( + ∀X0:term. ♯{X0} < ♯{V0} + (♯{W0} + ♯{T0} + 1) + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → + ∃∃X. ⓓ{a}V1. T1 ➡X & ⓓ{a}V2. T2 ➡ X. +#a #W0 #V0 #V1 #T0 #T1 #V2 #T2 #IH #HV01 #HV02 #HT01 #HT02 +elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ +elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ /3 width=5/ +qed. + +(* Basic_1: was: pr0_cong_delta pr0_delta_delta *) +fact tpr_conf_delta_delta: + ∀a,I1,V0,V1,T0,T1,TT1,V2,T2,TT2. ( + ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → T0 ➡ T1 → T0 ➡ T2 → + ⋆. ⓑ{I1} V1 ⊢ T1 ▶ [O, 1] TT1 → + ⋆. ⓑ{I1} V2 ⊢ T2 ▶ [O, 1] TT2 → + ∃∃X. ⓑ{a,I1} V1. TT1 ➡ X & ⓑ{a,I1} V2. TT2 ➡ X. +#a #I1 #V0 #V1 #T0 #T1 #TT1 #V2 #T2 #TT2 #IH #HV01 #HV02 #HT01 #HT02 #HTT1 #HTT2 +elim (IH … HV01 … HV02) -HV01 -HV02 // #V #HV1 #HV2 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH // #T #HT1 #HT2 +elim (tpr_tps_conf_bind … HV1 HT1 … HTT1) -T1 #U1 #TTU1 #HTU1 +elim (tpr_tps_conf_bind … HV2 HT2 … HTT2) -T2 #U2 #TTU2 #HTU2 +elim (tps_conf_eq … HTU1 … HTU2) -T #U #HU1 #HU2 +@ex2_intro [2,3: @tpr_delta |1: skip ] /width=10/ (**) (* /3 width=10/ is too slow *) +qed. + +fact tpr_conf_delta_zeta: + ∀X2,V0,V1,T0,T1,TT1,T2. ( + ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → T0 ➡ T1 → ⋆. ⓓV1 ⊢ T1 ▶ [O,1] TT1 → + T0 ➡ T2 → ⇧[O, 1] X2 ≡ T2 → + ∃∃X. +ⓓV1. TT1 ➡ X & X2 ➡ X. +#X2 #V0 #V1 #T0 #T1 #TT1 #T2 #IH #_ #HT01 #HTT1 #HT02 #HXT2 +elim (IH … HT01 … HT02) -IH -HT01 -HT02 // -V0 -T0 #T #HT1 #HT2 +elim (tpr_tps_conf_bind ? ? V1 … HT1 HTT1) -T1 // #TT #HTT1 #HTT +elim (tpr_inv_lift1 … HT2 … HXT2) -T2 #X #HXT #HX2 +lapply (tps_inv_lift1_eq … HTT … HXT) -HTT #H destruct /3 width=3/ +qed. + +(* Basic_1: was: pr0_upsilon_upsilon *) +fact tpr_conf_theta_theta: + ∀a,VV1,V0,V1,W0,W1,T0,T1,V2,VV2,W2,T2. ( + ∀X0:term. ♯{X0} < ♯{V0} + (♯{W0} + ♯{T0} + 1) + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + V0 ➡ V1 → V0 ➡ V2 → W0 ➡ W1 → W0 ➡ W2 → T0 ➡ T1 → T0 ➡ T2 → + ⇧[O, 1] V1 ≡ VV1 → ⇧[O, 1] V2 ≡ VV2 → + ∃∃X. ⓓ{a}W1. ⓐVV1. T1 ➡ X & ⓓ{a}W2. ⓐVV2. T2 ➡ X. +#a #VV1 #V0 #V1 #W0 #W1 #T0 #T1 #V2 #VV2 #W2 #T2 #IH #HV01 #HV02 #HW01 #HW02 #HT01 #HT02 #HVV1 #HVV2 +elim (IH … HV01 … HV02) -HV01 -HV02 /2 width=1/ #V #HV1 #HV2 +elim (IH … HW01 … HW02) -HW01 -HW02 /2 width=1/ #W #HW1 #HW2 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH /2 width=1/ #T #HT1 #HT2 +elim (lift_total V 0 1) #VV #HVV +lapply (tpr_lift … HV1 … HVV1 … HVV) -V1 #HVV1 +lapply (tpr_lift … HV2 … HVV2 … HVV) -V2 -HVV #HVV2 +@ex2_intro [2,3: @tpr_bind |1:skip ] /2 width=5/ (**) (* /4 width=5/ is too slow *) +qed. + +fact tpr_conf_zeta_zeta: + ∀V0:term. ∀X2,TT0,T0,T1,TT2. ( + ∀X0:term. ♯{X0} < ♯{V0} + ♯{TT0} + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + TT0 ➡ T0 → ⇧[O, 1] T1 ≡ T0 → + TT0 ➡ TT2 → ⇧[O, 1] X2 ≡ TT2 → + ∃∃X. T1 ➡ X & X2 ➡ X. +#V0 #X2 #TT0 #T0 #T1 #TT2 #IH #HTT0 #HT10 #HTT02 #HXTT2 +elim (IH … HTT0 … HTT02) -IH -HTT0 -HTT02 // -V0 -TT0 #T #HT0 #HTT2 +elim (tpr_inv_lift1 … HT0 … HT10) -T0 #T0 #HT0 #HT10 +elim (tpr_inv_lift1 … HTT2 … HXTT2) -TT2 #TT2 #HTT2 #HXTT2 +lapply (lift_inj … HTT2 … HT0) -HTT2 #H destruct /2 width=3/ +qed. + +fact tpr_conf_tau_tau: + ∀V0,T0:term. ∀X2,T1. ( + ∀X0:term. ♯{X0} < ♯{V0} + ♯{T0} + 1 → + ∀X1,X2. X0 ➡ X1 → X0 ➡ X2 → + ∃∃X. X1 ➡ X & X2 ➡ X + ) → + T0 ➡ T1 → T0 ➡ X2 → + ∃∃X. T1 ➡ X & X2 ➡ X. +#V0 #T0 #X2 #T1 #IH #HT01 #HT02 +elim (IH … HT01 … HT02) -HT01 -HT02 -IH // /2 width=3/ +qed. + +(* Confluence ***************************************************************) + +(* Basic_1: was: pr0_confluence *) +theorem tpr_conf: ∀T0:term. ∀T1,T2. T0 ➡ T1 → T0 ➡ T2 → + ∃∃T. T1 ➡ T & T2 ➡ T. +#T0 @(f_ind … tw … T0) -T0 #n #IH * +[ #I #_ #X1 #X2 #H1 #H2 -n + >(tpr_inv_atom1 … H1) -X1 + >(tpr_inv_atom1 … H2) -X2 +(* case 1: atom, atom *) + // +| * [ #a ] #I #V0 #T0 #Hn #X1 #X2 #H1 #H2 + [ elim (tpr_inv_bind1 … H1) -H1 * + [ #V1 #T1 #U1 #HV01 #HT01 #HTU1 #H1 + | #T1 #HT01 #HXT1 #H11 #H12 + ] + elim (tpr_inv_bind1 … H2) -H2 * + [1,3: #V2 #T2 #U2 #HV02 #HT02 #HTU2 #H2 + |2,4: #T2 #HT02 #HXT2 #H21 #H22 + ] destruct +(* case 2: delta, delta *) + [ /3 width=11 by tpr_conf_delta_delta/ (**) (* /3 width=11/ is too slow *) +(* case 3: zeta, delta (repeated) *) + | @ex2_commute /3 width=10 by tpr_conf_delta_zeta/ +(* case 4: delta, zeta *) + | /3 width=10 by tpr_conf_delta_zeta/ (**) (* /3 width=10/ is too slow *) +(* case 5: zeta, zeta *) + | /3 width=9 by tpr_conf_zeta_zeta/ (**) (* /3 width=9/ is too slow *) + ] + | elim (tpr_inv_flat1 … H1) -H1 * + [ #V1 #T1 #HV01 #HT01 #H1 + | #b1 #V1 #W1 #U1 #T1 #HV01 #HUT1 #H11 #H12 #H13 + | #b1 #V1 #Y1 #W1 #Z1 #U1 #T1 #HV01 #HWZ1 #HUT1 #HVY1 #H11 #H12 #H13 + | #HX1 #H1 + ] + elim (tpr_inv_flat1 … H2) -H2 * + [1,5,9,13: #V2 #T2 #HV02 #HT02 #H2 + |2,6,10,14: #b2 #V2 #W2 #U2 #T2 #HV02 #HUT2 #H21 #H22 #H23 + |3,7,11,15: #b2 #V2 #Y2 #W2 #Z2 #U2 #T2 #HV02 #HWZ2 #HUT2 #HVY2 #H21 #H22 #H23 + |4,8,12,16: #HX2 #H2 + ] destruct +(* case 6: flat, flat *) + [ /3 width=7 by tpr_conf_flat_flat/ (**) (* /3 width=7/ is too slow *) +(* case 7: beta, flat (repeated) *) + | @ex2_commute /3 width=8 by tpr_conf_flat_beta/ +(* case 8: theta, flat (repeated) *) + | @ex2_commute /3 width=11 by tpr_conf_flat_theta/ +(* case 9: tau, flat (repeated) *) + | @ex2_commute /3 width=6 by tpr_conf_flat_cast/ +(* case 10: flat, beta *) + | /3 width=8 by tpr_conf_flat_beta/ (**) (* /3 width=8/ is too slow *) +(* case 11: beta, beta *) + | /3 width=8 by tpr_conf_beta_beta/ (**) (* /3 width=8/ is too slow *) +(* case 12: flat, theta *) + | /3 width=11 by tpr_conf_flat_theta/ (**) (* /3 width=11/ is too slow *) +(* case 13: theta, theta *) + | /3 width=14 by tpr_conf_theta_theta/ (**) (* /3 width=14/ is too slow *) +(* case 14: flat, tau *) + | /3 width=6 by tpr_conf_flat_cast/ (**) (* /3 width=6/ is too slow *) +(* case 15: tau, tau *) + | /3 width=5 by tpr_conf_tau_tau/ + ] + ] +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tps.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tps.etc new file mode 100644 index 000000000..7d89243d5 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tps.etc @@ -0,0 +1,288 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +notation "hvbox( L ⊢ break term 46 T1 break ▶ [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'PSubst $L $T1 $d $e $T2 }. + +include "basic_2/substitution/ldrop_append.ma". + +(* PARALLEL SUBSTITUTION ON TERMS *******************************************) + +inductive tps: nat → nat → lenv → relation term ≝ +| tps_atom : ∀L,I,d,e. tps d e L (⓪{I}) (⓪{I}) +| tps_subst: ∀L,K,V,W,i,d,e. d ≤ i → i < d + e → + ⇩[0, i] L ≡ K. ⓓV → ⇧[0, i + 1] V ≡ W → tps d e L (#i) W +| tps_bind : ∀L,a,I,V1,V2,T1,T2,d,e. + tps d e L V1 V2 → tps (d + 1) e (L. ⓑ{I} V2) T1 T2 → + tps d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) +| tps_flat : ∀L,I,V1,V2,T1,T2,d,e. + tps d e L V1 V2 → tps d e L T1 T2 → + tps d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) +. + +interpretation "parallel substritution (term)" + 'PSubst L T1 d e T2 = (tps d e L T1 T2). + +(* Basic properties *********************************************************) + +lemma tps_lsubr_trans: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶ [d, e] T2 → + ∀L2. L2 ⊑ [d, e] L1 → L2 ⊢ T1 ▶ [d, e] T2. +#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e +[ // +| #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12 + elim (ldrop_lsubr_ldrop2_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /2 width=4/ +| /4 width=1/ +| /3 width=1/ +] +qed. + +lemma tps_refl: ∀T,L,d,e. L ⊢ T ▶ [d, e] T. +#T elim T -T // +#I elim I -I /2 width=1/ +qed. + +(* Basic_1: was: subst1_ex *) +lemma tps_full: ∀K,V,T1,L,d. ⇩[0, d] L ≡ (K. ⓓV) → + ∃∃T2,T. L ⊢ T1 ▶ [d, 1] T2 & ⇧[d, 1] T ≡ T2. +#K #V #T1 elim T1 -T1 +[ * #i #L #d #HLK /2 width=4/ + elim (lt_or_eq_or_gt i d) #Hid /3 width=4/ + destruct + elim (lift_total V 0 (i+1)) #W #HVW + elim (lift_split … HVW i i ? ? ?) // /3 width=4/ +| * [ #a ] #I #W1 #U1 #IHW1 #IHU1 #L #d #HLK + elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 + [ elim (IHU1 (L. ⓑ{I} W2) (d+1) ?) -IHU1 /2 width=1/ -HLK /3 width=9/ + | elim (IHU1 … HLK) -IHU1 -HLK /3 width=8/ + ] +] +qed. + +lemma tps_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 ▶ [d1, e1] T2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → + L ⊢ T1 ▶ [d2, e2] T2. +#L #T1 #T2 #d1 #e1 #H elim H -L -T1 -T2 -d1 -e1 +[ // +| #L #K #V #W #i #d1 #e1 #Hid1 #Hide1 #HLK #HVW #d2 #e2 #Hd12 #Hde12 + lapply (transitive_le … Hd12 … Hid1) -Hd12 -Hid1 #Hid2 + lapply (lt_to_le_to_lt … Hide1 … Hde12) -Hide1 /2 width=4/ +| /4 width=3/ +| /4 width=1/ +] +qed. + +lemma tps_weak_top: ∀L,T1,T2,d,e. + L ⊢ T1 ▶ [d, e] T2 → L ⊢ T1 ▶ [d, |L| - d] T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +[ // +| #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW + lapply (ldrop_fwd_ldrop2_length … HLK) #Hi + lapply (le_to_lt_to_lt … Hdi Hi) /3 width=4/ +| normalize /2 width=1/ +| /2 width=1/ +] +qed. + +lemma tps_weak_full: ∀L,T1,T2,d,e. + L ⊢ T1 ▶ [d, e] T2 → L ⊢ T1 ▶ [0, |L|] T2. +#L #T1 #T2 #d #e #HT12 +lapply (tps_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 +lapply (tps_weak_top … HT12) // +qed. + +lemma tps_split_up: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ∀i. d ≤ i → i ≤ d + e → + ∃∃T. L ⊢ T1 ▶ [d, i - d] T & L ⊢ T ▶ [i, d + e - i] T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +[ /2 width=3/ +| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde + elim (lt_or_ge i j) + [ -Hide -Hjde + >(plus_minus_m_m j d) in ⊢ (% → ?); // -Hdj /3 width=4/ + | -Hdi -Hdj #Hij + lapply (plus_minus_m_m … Hjde) -Hjde /3 width=8/ + ] +| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 + elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ + -Hdi -Hide >arith_c1x #T #HT1 #HT2 + lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // + -Hdi -Hide /3 width=5/ +] +qed. + +lemma tps_split_down: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → + ∀i. d ≤ i → i ≤ d + e → + ∃∃T. L ⊢ T1 ▶ [i, d + e - i] T & + L ⊢ T ▶ [d, i - d] T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +[ /2 width=3/ +| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde + elim (lt_or_ge i j) + [ -Hide -Hjde >(plus_minus_m_m j d) in ⊢ (% → ?); // -Hdj /3 width=8/ + | -Hdi -Hdj + >(plus_minus_m_m (d+e) j) in Hide; // -Hjde /3 width=4/ + ] +| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 + elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ + -Hdi -Hide >arith_c1x #T #HT1 #HT2 + lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide + elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // + -Hdi -Hide /3 width=5/ +] +qed. + +lemma tps_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶ [d, e] T2 → + ∀L. L @@ K ⊢ T1 ▶ [d, e] T2. +#K #T1 #T2 #d #e #H elim H -K -T1 -T2 -d -e // /2 width=1/ +#K #K0 #V #W #i #d #e #Hdi #Hide #HK0 #HVW #L +lapply (ldrop_fwd_ldrop2_length … HK0) #H +@(tps_subst … (L@@K0) … HVW) // (**) (* /3/ does not work *) +@(ldrop_O1_append_sn_le … HK0) /2 width=2/ +qed. + +(* Basic inversion lemmas ***************************************************) + +fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ∀I. T1 = ⓪{I} → + T2 = ⓪{I} ∨ + ∃∃K,V,i. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K. ⓓV & + ⇧[O, i + 1] V ≡ T2 & + I = LRef i. +#L #T1 #T2 #d #e * -L -T1 -T2 -d -e +[ #L #I #d #e #J #H destruct /2 width=1/ +| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #I #H destruct /3 width=8/ +| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +] +qed. + +lemma tps_inv_atom1: ∀L,T2,I,d,e. L ⊢ ⓪{I} ▶ [d, e] T2 → + T2 = ⓪{I} ∨ + ∃∃K,V,i. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K. ⓓV & + ⇧[O, i + 1] V ≡ T2 & + I = LRef i. +/2 width=3/ qed-. + + +(* Basic_1: was: subst1_gen_sort *) +lemma tps_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k ▶ [d, e] T2 → T2 = ⋆k. +#L #T2 #k #d #e #H +elim (tps_inv_atom1 … H) -H // +* #K #V #i #_ #_ #_ #_ #H destruct +qed-. + +(* Basic_1: was: subst1_gen_lref *) +lemma tps_inv_lref1: ∀L,T2,i,d,e. L ⊢ #i ▶ [d, e] T2 → + T2 = #i ∨ + ∃∃K,V. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K. ⓓV & + ⇧[O, i + 1] V ≡ T2. +#L #T2 #i #d #e #H +elim (tps_inv_atom1 … H) -H /2 width=1/ +* #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=4/ +qed-. + +lemma tps_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p ▶ [d, e] T2 → T2 = §p. +#L #T2 #p #d #e #H +elim (tps_inv_atom1 … H) -H // +* #K #V #i #_ #_ #_ #_ #H destruct +qed-. + +fact tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 → + ∀a,I,V1,T1. U1 = ⓑ{a,I} V1. T1 → + ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & + L. ⓑ{I} V2 ⊢ T1 ▶ [d + 1, e] T2 & + U2 = ⓑ{a,I} V2. T2. +#d #e #L #U1 #U2 * -d -e -L -U1 -U2 +[ #L #k #d #e #a #I #V1 #T1 #H destruct +| #L #K #V #W #i #d #e #_ #_ #_ #_ #a #I #V1 #T1 #H destruct +| #L #b #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #a #I #V #T #H destruct /2 width=5/ +| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #a #I #V #T #H destruct +] +qed. + +lemma tps_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. L ⊢ ⓑ{a,I} V1. T1 ▶ [d, e] U2 → + ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & + L. ⓑ{I} V2 ⊢ T1 ▶ [d + 1, e] T2 & + U2 = ⓑ{a,I} V2. T2. +/2 width=3/ qed-. + +fact tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 → + ∀I,V1,T1. U1 = ⓕ{I} V1. T1 → + ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & L ⊢ T1 ▶ [d, e] T2 & + U2 = ⓕ{I} V2. T2. +#d #e #L #U1 #U2 * -d -e -L -U1 -U2 +[ #L #k #d #e #I #V1 #T1 #H destruct +| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct +| #L #a #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct +| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ +] +qed. + +lemma tps_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 ▶ [d, e] U2 → + ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & L ⊢ T1 ▶ [d, e] T2 & + U2 = ⓕ{I} V2. T2. +/2 width=3/ qed-. + +fact tps_inv_refl_O2_aux: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → e = 0 → T1 = T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e +[ // +| #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct + lapply (le_to_lt_to_lt … Hdi … Hide) -Hdi -Hide shift_append_assoc normalize #H + elim (tps_inv_bind1 … H) -H + #V0 #T0 #_ #HT10 #H destruct + elim (IH … HT10) -IH -HT10 #L2 #T2 #HL12 #H destruct + >append_length >HL12 -HL12 + @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] // /2 width=3/ (**) (* explicit constructor *) +] +qed-. + +(* Basic_1: removed theorems 25: + subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt + subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans + subst0_lift_lt subst0_lift_ge subst0_lift_ge_S subst0_lift_ge_s + subst0_subst0 subst0_subst0_back subst0_weight_le subst0_weight_lt + subst0_confluence_neq subst0_confluence_eq subst0_tlt_head + subst0_confluence_lift subst0_tlt + subst1_head subst1_gen_head subst1_lift_S subst1_confluence_lift +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tps_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tps_lift.etc new file mode 100644 index 000000000..1d65d8d63 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tps_lift.etc @@ -0,0 +1,294 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/ldrop_ldrop.ma". +include "basic_2/substitution/tps.ma". + +(* PARTIAL SUBSTITUTION ON TERMS ********************************************) + +(* Advanced inversion lemmas ************************************************) + +fact tps_inv_S2_aux: ∀L,T1,T2,d,e1. L ⊢ T1 ▶ [d, e1] T2 → ∀e2. e1 = e2 + 1 → + ∀K,V. ⇩[0, d] L ≡ K. ⓛV → L ⊢ T1 ▶ [d + 1, e2] T2. +#L #T1 #T2 #d #e1 #H elim H -L -T1 -T2 -d -e1 +[ // +| #L #K0 #V0 #W #i #d #e1 #Hdi #Hide1 #HLK0 #HV0 #e2 #He12 #K #V #HLK destruct + elim (lt_or_ge i (d+1)) #HiSd + [ -Hide1 -HV0 + lapply (le_to_le_to_eq … Hdi ?) /2 width=1/ #H destruct + lapply (ldrop_mono … HLK0 … HLK) #H destruct + | -V -Hdi /2 width=4/ + ] +| /4 width=3/ +| /3 width=3/ +] +qed. + +lemma tps_inv_S2: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e + 1] T2 → + ∀K,V. ⇩[0, d] L ≡ K. ⓛV → L ⊢ T1 ▶ [d + 1, e] T2. +/2 width=3/ qed-. + +lemma tps_inv_refl_SO2: ∀L,T1,T2,d. L ⊢ T1 ▶ [d, 1] T2 → + ∀K,V. ⇩[0, d] L ≡ K. ⓛV → T1 = T2. +#L #T1 #T2 #d #HT12 #K #V #HLK +lapply (tps_inv_S2 … T1 T2 … 0 … HLK) -K // -HT12 #HT12 +lapply (tps_inv_refl_O2 … HT12) -HT12 // +qed-. + +(* Relocation properties ****************************************************) + +(* Basic_1: was: subst1_lift_lt *) +lemma tps_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → + ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + dt + et ≤ d → + L ⊢ U1 ▶ [dt, et] U2. +#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et +[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdetd + lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid + lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct + elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=4/ +| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + @tps_bind [ /2 width=6/ | @IHT12 /2 width=6/ ] (**) (* /3 width=6/ is too slow, arith3 needed to avoid crash *) +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ +] +qed. + +lemma tps_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → + ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + dt ≤ d → d ≤ dt + et → + L ⊢ U1 ▶ [dt, et + e] U2. +#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et +[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdtd #_ + elim (lift_inv_lref1 … H) -H * #Hid #H destruct + [ -Hdtd + lapply (lt_to_le_to_lt … (dt+et+e) Hidet ?) // -Hidet #Hidete + elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -V #H destruct /2 width=4/ + | -Hdti + lapply (transitive_le … Hdtd Hid) -Hdtd #Hdti + lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ + ] +| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdtd #Hddet + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + @tps_bind [ /2 width=6/ | @IHT12 [3,4: // | skip |5,6: /2 width=1/ | /2 width=1/ ] + ] (**) (* /3 width=6/ is too slow, simplification like tps_lift_le is too slow *) +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd + elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ +] +qed. + +(* Basic_1: was: subst1_lift_ge *) +lemma tps_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → + ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + d ≤ dt → + L ⊢ U1 ▶ [dt + e, et] U2. +#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et +[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hddt + lapply (transitive_le … Hddt … Hdti) -Hddt #Hid + lapply (lift_inv_lref1_ge … H … Hid) -H #H destruct + lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ +| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt + elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct + @tps_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *) +| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt + elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 + elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=5/ +] +qed. + +(* Basic_1: was: subst1_gen_lift_lt *) +lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt + et ≤ d → + ∃∃T2. K ⊢ T1 ▶ [dt, et] T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et +[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ + ] +| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdetd + lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid + lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct + elim (ldrop_conf_lt … HLK … HLKV ?) -L // #L #U #HKL #_ #HUV + elim (lift_trans_le … HUV … HVW ?) -V // >minus_plus minus_plus plus_minus // commutative_plus >plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) + ] +| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet + elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct + elim (IHV12 … HLK … HWV1 ? ?) -V1 // #W2 #HW12 #HWV2 + elim (IHU12 … HTU1 ? ?) -U1 [5: @ldrop_skip // |2: skip |3: >plus_plus_comm_23 >(plus_plus_comm_23 dt) /2 width=1/ |4: /2 width=1/ ] (**) (* 29s without the rewrites *) + /3 width=5/ +| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet + elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct + elim (IHV12 … HLK … HWV1 ? ?) -V1 // + elim (IHU12 … HLK … HTU1 ? ?) -U1 -HLK // /3 width=5/ +] +qed. + +(* Basic_1: was: subst1_gen_lift_ge *) +lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d + e ≤ dt → + ∃∃T2. K ⊢ T1 ▶ [dt - e, et] T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et +[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ + ] +| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdedt + lapply (transitive_le … Hdedt … Hdti) #Hdei + elim (le_inv_plus_l … Hdedt) -Hdedt #_ #Hedt + elim (le_inv_plus_l … Hdei) #Hdie #Hei + lapply (lift_inv_lref2_ge … H … Hdei) -H #H destruct + lapply (ldrop_conf_ge … HLK … HLKV ?) -L // #HKV + elim (lift_split … HVW d (i - e + 1) ? ? ?) -HVW [4: // |3: /2 width=1/ |2: /3 width=1/ ] -Hdei -Hdie + #V0 #HV10 >plus_minus // plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) +| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd + elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct + elim (le_inv_plus_l … Hdetd) #_ #Hedt + elim (IHV12 … HLK … HWV1 ?) -V1 // #W2 #HW12 #HWV2 + elim (IHU12 … HTU1 ?) -U1 [4: @ldrop_skip // |2: skip |3: /2 width=1/ ] + IHV12 // >IHT12 // +| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX + elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct + >IHV12 // >IHT12 // +] +qed. +(* + Theorem subst0_gen_lift_rev_ge: (t1,v,u2,i,h,d:?) + (subst0 i v t1 (lift h d u2)) -> + (le (plus d h) i) -> + (EX u1 | (subst0 (minus i h) v u1 u2) & + t1 = (lift h d u1) + ). + + + Theorem subst0_gen_lift_rev_lelt: (t1,v,u2,i,h,d:?) + (subst0 i v t1 (lift h d u2)) -> + (le d i) -> (lt i (plus d h)) -> + (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)). +*) +lemma tps_inv_lift1_ge_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → + ∃∃T2. K ⊢ T1 ▶ [d, dt + et - (d + e)] T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet +elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 +lapply (tps_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 +lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct +elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -U -L // commutative_plus /2 width=1/ ] -Hdtd #T #HT1 #HTU +lapply (tps_weak … HU2 d e ? ?) -HU2 // [ >commutative_plus (lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3/ + ] +| #L #a #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V1) ?) -HT02 /2 width=1/ #HT02 + elim (IHV01 … HV02) -V0 #V #HV1 #HV2 + elim (IHT01 … HT02) -T0 #T #HT1 #HT2 + lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubr_trans … HT2 (L. ⓑ{I} V) ?) -HT2 /3 width=5/ +| #L #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02) -V0 + elim (IHT01 … HT02) -T0 /3 width=5/ +] +qed. + +(* Basic_1: was: subst1_confluence_neq *) +theorem tps_conf_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 ▶ [d1, e1] T1 → + ∀L2,T2,d2,e2. L2 ⊢ T0 ▶ [d2, e2] T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. L2 ⊢ T1 ▶ [d2, e2] T & L1 ⊢ T2 ▶ [d1, e1] T. +#L1 #T0 #T1 #d1 #e1 #H elim H -L1 -T0 -T1 -d1 -e1 +[ /2 width=3/ +| #L1 #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #L2 #T2 #d2 #e2 #H1 #H2 + elim (tps_inv_lref1 … H1) -H1 + [ #H destruct /3 width=6/ + | -HLK1 -HVT1 * #K2 #V2 #Hd2 #Hde2 #_ #_ elim H2 -H2 #Hded + [ -Hd1 -Hde2 + lapply (transitive_le … Hded Hd2) -Hded -Hd2 #H + lapply (lt_to_le_to_lt … Hde1 H) -Hde1 -H #H + elim (lt_refl_false … H) + | -Hd2 -Hde1 + lapply (transitive_le … Hded Hd1) -Hded -Hd1 #H + lapply (lt_to_le_to_lt … Hde2 H) -Hde2 -H #H + elim (lt_refl_false … H) + ] + ] +| #L1 #a #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -V0 #V #HV1 #HV2 + elim (IHT01 … HT02 ?) -T0 + [ -H #T #HT1 #HT2 + lapply (tps_lsubr_trans … HT1 (L2. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubr_trans … HT2 (L1. ⓑ{I} V) ?) -HT2 /2 width=1/ /3 width=5/ + | -HV1 -HV2 >plus_plus_comm_23 >plus_plus_comm_23 in ⊢ (? ? %); elim H -H #H + [ @or_introl | @or_intror ] /2 by monotonic_le_plus_l/ (**) (* /3 / is too slow *) + ] +| #L1 #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -V0 + elim (IHT01 … HT02 H) -T0 -H /3 width=5/ +] +qed. + +(* Note: the constant 1 comes from tps_subst *) +(* Basic_1: was: subst1_trans *) +theorem tps_trans_ge: ∀L,T1,T0,d,e. L ⊢ T1 ▶ [d, e] T0 → + ∀T2. L ⊢ T0 ▶ [d, 1] T2 → 1 ≤ e → + L ⊢ T1 ▶ [d, e] T2. +#L #T1 #T0 #d #e #H elim H -L -T1 -T0 -d -e +[ #L #I #d #e #T2 #H #He + elim (tps_inv_atom1 … H) -H + [ #H destruct // + | * #K #V #i #Hd2i #Hide2 #HLK #HVT2 #H destruct + lapply (lt_to_le_to_lt … (d + e) Hide2 ?) /2 width=4/ + ] +| #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He + lapply (tps_weak … HVT2 0 (i +1) ? ?) -HVT2 /2 width=1/ #HVT2 + <(tps_inv_lift1_eq … HVT2 … HVW) -HVT2 /2 width=4/ +| #L #a #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He + elim (tps_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct + lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 + lapply (IHT10 … HT02 He) -T0 #HT12 + lapply (tps_lsubr_trans … HT12 (L. ⓑ{I} V2) ?) -HT12 /2 width=1/ /3 width=1/ +| #L #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He + elim (tps_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1/ +] +qed. + +theorem tps_trans_down: ∀L,T1,T0,d1,e1. L ⊢ T1 ▶ [d1, e1] T0 → + ∀T2,d2,e2. L ⊢ T0 ▶ [d2, e2] T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 ▶ [d2, e2] T & L ⊢ T ▶ [d1, e1] T2. +#L #T1 #T0 #d1 #e1 #H elim H -L -T1 -T0 -d1 -e1 +[ /2 width=3/ +| #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1 + lapply (transitive_le … Hde2d1 Hdi1) -Hde2d1 #Hde2i1 + lapply (tps_weak … HWT2 0 (i1 + 1) ? ?) -HWT2 normalize /2 width=1/ -Hde2i1 #HWT2 + <(tps_inv_lift1_eq … HWT2 … HVW) -HWT2 /3 width=8/ +| #L #a #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 + elim (IHV10 … HV02 ?) -IHV10 -HV02 // #V + elim (IHT10 … HT02 ?) -T0 /2 width=1/ #T #HT1 #HT2 + lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ + lapply (tps_lsubr_trans … HT2 (L. ⓑ{I} V2) ?) -HT2 /2 width=1/ /3 width=6/ +| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV10 … HV02 ?) -V0 // + elim (IHT10 … HT02 ?) -T0 // /3 width=6/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss.etc new file mode 100644 index 000000000..433718803 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss.etc @@ -0,0 +1,188 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "hvbox( L ⊢ break term 46 T1 break ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'PSubstStar $L $T1 $d $e $T2 }. + +include "basic_2/substitution/tps.ma". + +(* PARTIAL UNFOLD ON TERMS **************************************************) + +definition tpss: nat → nat → lenv → relation term ≝ + λd,e,L. TC … (tps d e L). + +interpretation "partial unfold (term)" + 'PSubstStar L T1 d e T2 = (tpss d e L T1 T2). + +(* Basic eliminators ********************************************************) + +lemma tpss_ind: ∀d,e,L,T1. ∀R:predicate term. R T1 → + (∀T,T2. L ⊢ T1 ▶* [d, e] T → L ⊢ T ▶ [d, e] T2 → R T → R T2) → + ∀T2. L ⊢ T1 ▶* [d, e] T2 → R T2. +#d #e #L #T1 #R #HT1 #IHT1 #T2 #HT12 +@(TC_star_ind … HT1 IHT1 … HT12) // +qed-. + +lemma tpss_ind_dx: ∀d,e,L,T2. ∀R:predicate term. R T2 → + (∀T1,T. L ⊢ T1 ▶ [d, e] T → L ⊢ T ▶* [d, e] T2 → R T → R T1) → + ∀T1. L ⊢ T1 ▶* [d, e] T2 → R T1. +#d #e #L #T2 #R #HT2 #IHT2 #T1 #HT12 +@(TC_star_ind_dx … HT2 IHT2 … HT12) // +qed-. + +(* Basic properties *********************************************************) + +lemma tps_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. +/2 width=1/ qed. + +lemma tpss_strap1: ∀L,T1,T,T2,d,e. + L ⊢ T1 ▶* [d, e] T → L ⊢ T ▶ [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. +/2 width=3/ qed. + +lemma tpss_strap2: ∀L,T1,T,T2,d,e. + L ⊢ T1 ▶ [d, e] T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. +/2 width=3/ qed. + +lemma tpss_lsubr_trans: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶* [d, e] T2 → + ∀L2. L2 ⊑ [d, e] L1 → L2 ⊢ T1 ▶* [d, e] T2. +/3 width=3/ qed. + +lemma tpss_refl: ∀d,e,L,T. L ⊢ T ▶* [d, e] T. +/2 width=1/ qed. + +lemma tpss_bind: ∀L,V1,V2,d,e. L ⊢ V1 ▶* [d, e] V2 → + ∀a,I,T1,T2. L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 → + L ⊢ ⓑ{a,I} V1. T1 ▶* [d, e] ⓑ{a,I} V2. T2. +#L #V1 #V2 #d #e #HV12 elim HV12 -V2 +[ #V2 #HV12 #a #I #T1 #T2 #HT12 elim HT12 -T2 + [ /3 width=5/ + | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) + ] +| #V #V2 #_ #HV12 #IHV #a #I #T1 #T2 #HT12 + lapply (tpss_lsubr_trans … HT12 (L. ⓑ{I} V) ?) -HT12 /2 width=1/ #HT12 + lapply (IHV a … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) +] +qed. + +lemma tpss_flat: ∀L,I,V1,V2,T1,T2,d,e. + L ⊢ V1 ▶* [d, e] V2 → L ⊢ T1 ▶* [d, e] T2 → + L ⊢ ⓕ{I} V1. T1 ▶* [d, e] ⓕ{I} V2. T2. +#L #I #V1 #V2 #T1 #T2 #d #e #HV12 elim HV12 -V2 +[ #V2 #HV12 #HT12 elim HT12 -T2 + [ /3 width=1/ + | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) + ] +| #V #V2 #_ #HV12 #IHV #HT12 + lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) +] +qed. + +lemma tpss_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 ▶* [d1, e1] T2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → + L ⊢ T1 ▶* [d2, e2] T2. +#L #T1 #T2 #d1 #e1 #H #d1 #d2 #Hd21 #Hde12 @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT12 #IHT + lapply (tps_weak … HT12 … Hd21 Hde12) -HT12 -Hd21 -Hde12 /2 width=3/ +] +qed. + +lemma tpss_weak_top: ∀L,T1,T2,d,e. + L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ▶* [d, |L| - d] T2. +#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT12 #IHT + lapply (tps_weak_top … HT12) -HT12 /2 width=3/ +] +qed. + +lemma tpss_weak_full: ∀L,T1,T2,d,e. + L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ▶* [0, |L|] T2. +#L #T1 #T2 #d #e #HT12 +lapply (tpss_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 +lapply (tpss_weak_top … HT12) // +qed. + +lemma tpss_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶* [d, e] T2 → + ∀L. L @@ K ⊢ T1 ▶* [d, e] T2. +#K #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // /3 width=3/ +qed. + +(* Basic inversion lemmas ***************************************************) + +(* Note: this can be derived from tpss_inv_atom1 *) +lemma tpss_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k ▶* [d, e] T2 → T2 = ⋆k. +#L #T2 #k #d #e #H @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT2 #IHT destruct + >(tps_inv_sort1 … HT2) -HT2 // +] +qed-. + +(* Note: this can be derived from tpss_inv_atom1 *) +lemma tpss_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p ▶* [d, e] T2 → T2 = §p. +#L #T2 #p #d #e #H @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT2 #IHT destruct + >(tps_inv_gref1 … HT2) -HT2 // +] +qed-. + +lemma tpss_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. L ⊢ ⓑ{a,I} V1. T1 ▶* [d, e] U2 → + ∃∃V2,T2. L ⊢ V1 ▶* [d, e] V2 & + L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 & + U2 = ⓑ{a,I} V2. T2. +#d #e #L #a #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 +[ /2 width=5/ +| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct + elim (tps_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H + lapply (tpss_lsubr_trans … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ +] +qed-. + +lemma tpss_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 ▶* [d, e] U2 → + ∃∃V2,T2. L ⊢ V1 ▶* [d, e] V2 & L ⊢ T1 ▶* [d, e] T2 & + U2 = ⓕ{I} V2. T2. +#d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 +[ /2 width=5/ +| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct + elim (tps_inv_flat1 … HU2) -HU2 /3 width=5/ +] +qed-. + +lemma tpss_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 ▶* [d, 0] T2 → T1 = T2. +#L #T1 #T2 #d #H @(tpss_ind … H) -T2 +[ // +| #T #T2 #_ #HT2 #IHT <(tps_inv_refl_O2 … HT2) -HT2 // +] +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma tpss_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → ♯{T1} ≤ ♯{T2}. +#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT1 +lapply (tps_fwd_tw … HT2) -HT2 #HT2 +@(transitive_le … IHT1) // +qed-. + +lemma tpss_fwd_shift1: ∀L,L1,T1,T,d,e. L ⊢ L1 @@ T1 ▶*[d, e] T → + ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. +#L #L1 #T1 #T #d #e #H @(tpss_ind … H) -T +[ /2 width=4/ +| #T #X #_ #H0 * #L0 #T0 #HL10 #H destruct + elim (tps_fwd_shift1 … H0) -H0 #L2 #T2 #HL02 #H destruct /2 width=4/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_alt.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_alt.etc new file mode 100644 index 000000000..f7670eaa8 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_alt.etc @@ -0,0 +1,105 @@ +(**************************************************************************) +(* ___ *) +(* ||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 "hvbox( L ⊢ break term 46 T1 break ▶ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'PSubstStarAlt $L $T1 $d $e $T2 }. + +include "basic_2/unfold/tpss_lift.ma". + +(* PARALLEL UNFOLD ON TERMS *************************************************) + +(* alternative definition of tpss *) +inductive tpssa: nat → nat → lenv → relation term ≝ +| tpssa_atom : ∀L,I,d,e. tpssa d e L (⓪{I}) (⓪{I}) +| tpssa_subst: ∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e → + ⇩[0, i] L ≡ K. ⓓV1 → tpssa 0 (d + e - i - 1) K V1 V2 → + ⇧[0, i + 1] V2 ≡ W2 → tpssa d e L (#i) W2 +| tpssa_bind : ∀L,a,I,V1,V2,T1,T2,d,e. + tpssa d e L V1 V2 → tpssa (d + 1) e (L. ⓑ{I} V2) T1 T2 → + tpssa d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) +| tpssa_flat : ∀L,I,V1,V2,T1,T2,d,e. + tpssa d e L V1 V2 → tpssa d e L T1 T2 → + tpssa d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) +. + +interpretation "parallel unfold (term) alternative" + 'PSubstStarAlt L T1 d e T2 = (tpssa d e L T1 T2). + +(* Basic properties *********************************************************) + +lemma tpssa_lsubr_trans: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶▶* [d, e] T2 → + ∀L2. L2 ⊑ [d, e] L1 → L2 ⊢ T1 ▶▶* [d, e] T2. +#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e +[ // +| #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 + elim (ldrop_lsubr_ldrop2_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /3 width=6/ +| /4 width=1/ +| /3 width=1/ +] +qed. + +lemma tpssa_refl: ∀T,L,d,e. L ⊢ T ▶▶* [d, e] T. +#T elim T -T // +#I elim I -I /2 width=1/ +qed. + +lemma tpssa_tps_trans: ∀L,T1,T,d,e. L ⊢ T1 ▶▶* [d, e] T → + ∀T2. L ⊢ T ▶ [d, e] T2 → L ⊢ T1 ▶▶* [d, e] T2. +#L #T1 #T #d #e #H elim H -L -T1 -T -d -e +[ #L #I #d #e #X #H + elim (tps_inv_atom1 … H) -H // * /2 width=6/ +| #L #K #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK #_ #HVW2 #IHV12 #T2 #H + lapply (ldrop_fwd_ldrop2 … HLK) #H0LK + lapply (tps_weak … H 0 (d+e) ? ?) -H // #H + elim (tps_inv_lift1_be … H … H0LK … HVW2 ? ?) -H -H0LK -HVW2 // /3 width=6/ +| #L #a #I #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H + elim (tps_inv_bind1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct + lapply (tps_lsubr_trans … HT2 (L.ⓑ{I}V) ?) -HT2 /2 width=1/ #HT2 + lapply (IHV1 … HV2) -IHV1 -HV2 #HV12 + lapply (IHT1 … HT2) -IHT1 -HT2 #HT12 + lapply (tpssa_lsubr_trans … HT12 (L.ⓑ{I}V2) ?) -HT12 /2 width=1/ +| #L #I #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H + elim (tps_inv_flat1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct /3 width=1/ +] +qed. + +lemma tpss_tpssa: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ▶▶* [d, e] T2. +#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // /2 width=3/ +qed. + +(* Basic inversion lemmas ***************************************************) + +lemma tpssa_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶▶* [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e // /2 width=6/ +qed-. + +lemma tpss_ind_alt: ∀R:nat→nat→lenv→relation term. + (∀L,I,d,e. R d e L (⓪{I}) (⓪{I})) → + (∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e → + ⇩[O, i] L ≡ K.ⓓV1 → K ⊢ V1 ▶* [O, d + e - i - 1] V2 → + ⇧[O, i + 1] V2 ≡ W2 → R O (d+e-i-1) K V1 V2 → R d e L (#i) W2 + ) → + (∀L,a,I,V1,V2,T1,T2,d,e. L ⊢ V1 ▶* [d, e] V2 → + L.ⓑ{I}V2 ⊢ T1 ▶* [d + 1, e] T2 → R d e L V1 V2 → + R (d+1) e (L.ⓑ{I}V2) T1 T2 → R d e L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) + ) → + (∀L,I,V1,V2,T1,T2,d,e. L ⊢ V1 ▶* [d, e] V2 → + L ⊢ T1 ▶* [d, e] T2 → R d e L V1 V2 → + R d e L T1 T2 → R d e L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) + ) → + ∀d,e,L,T1,T2. L ⊢ T1 ▶* [d, e] T2 → R d e L T1 T2. +#R #H1 #H2 #H3 #H4 #d #e #L #T1 #T2 #H elim (tpss_tpssa … H) -L -T1 -T2 -d -e +// /3 width=1 by tpssa_tpss/ /3 width=7 by tpssa_tpss/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_lift.etc new file mode 100644 index 000000000..a68f86e32 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_lift.etc @@ -0,0 +1,196 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/tps_lift.ma". +include "basic_2/unfold/tpss.ma". + +(* PARTIAL UNFOLD ON TERMS **************************************************) + +(* Advanced properties ******************************************************) + +lemma tpss_subst: ∀L,K,V,U1,i,d,e. + d ≤ i → i < d + e → + ⇩[0, i] L ≡ K. ⓓV → K ⊢ V ▶* [0, d + e - i - 1] U1 → + ∀U2. ⇧[0, i + 1] U1 ≡ U2 → L ⊢ #i ▶* [d, e] U2. +#L #K #V #U1 #i #d #e #Hdi #Hide #HLK #H @(tpss_ind … H) -U1 +[ /3 width=4/ +| #U #U1 #_ #HU1 #IHU #U2 #HU12 + elim (lift_total U 0 (i+1)) #U0 #HU0 + lapply (IHU … HU0) -IHU #H + lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK + lapply (tps_lift_ge … HU1 … HLK HU0 HU12 ?) -HU1 -HLK -HU0 -HU12 // normalize #HU02 + lapply (tps_weak … HU02 d e ? ?) -HU02 [ >minus_plus >commutative_plus /2 width=1/ | /2 width=1/ | /2 width=3/ ] +] +qed. + +(* Advanced inverion lemmas *************************************************) + +lemma tpss_inv_atom1: ∀L,T2,I,d,e. L ⊢ ⓪{I} ▶* [d, e] T2 → + T2 = ⓪{I} ∨ + ∃∃K,V1,V2,i. d ≤ i & i < d + e & + ⇩[O, i] L ≡ K. ⓓV1 & + K ⊢ V1 ▶* [0, d + e - i - 1] V2 & + ⇧[O, i + 1] V2 ≡ T2 & + I = LRef i. +#L #T2 #I #d #e #H @(tpss_ind … H) -T2 +[ /2 width=1/ +| #T #T2 #_ #HT2 * + [ #H destruct + elim (tps_inv_atom1 … HT2) -HT2 [ /2 width=1/ | * /3 width=10/ ] + | * #K #V1 #V #i #Hdi #Hide #HLK #HV1 #HVT #HI + lapply (ldrop_fwd_ldrop2 … HLK) #H + elim (tps_inv_lift1_ge_up … HT2 … H … HVT ? ? ?) normalize -HT2 -H -HVT [2,3,4: /2 width=1/ ] #V2 (lift_mono … HTU1 … H) -H // +| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 + elim (lift_total T d e) #U #HTU + lapply (IHT … HTU) -IHT #HU1 + lapply (tps_lift_le … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ +] +qed. + +lemma tpss_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 ▶* [dt, et] T2 → + ∀L,U1,d,e. dt ≤ d → d ≤ dt + et → + ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → + ∀U2. ⇧[d, e] T2 ≡ U2 → L ⊢ U1 ▶* [dt, et + e] U2. +#K #T1 #T2 #dt #et #H #L #U1 #d #e #Hdtd #Hddet #HLK #HTU1 @(tpss_ind … H) -T2 +[ #U2 #H >(lift_mono … HTU1 … H) -H // +| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 + elim (lift_total T d e) #U #HTU + lapply (IHT … HTU) -IHT #HU1 + lapply (tps_lift_be … HT2 … HLK HTU HTU2 ? ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ +] +qed. + +lemma tpss_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 ▶* [dt, et] T2 → + ∀L,U1,d,e. d ≤ dt → ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → + L ⊢ U1 ▶* [dt + e, et] U2. +#K #T1 #T2 #dt #et #H #L #U1 #d #e #Hddt #HLK #HTU1 @(tpss_ind … H) -T2 +[ #U2 #H >(lift_mono … HTU1 … H) -H // +| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 + elim (lift_total T d e) #U #HTU + lapply (IHT … HTU) -IHT #HU1 + lapply (tps_lift_ge … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ +] +qed. + +lemma tpss_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt + et ≤ d → + ∃∃T2. K ⊢ T1 ▶* [dt, et] T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdetd @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_le … HU2 … HLK … HTU ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. + +lemma tpss_inv_lift1_be: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → d + e ≤ dt + et → + ∃∃T2. K ⊢ T1 ▶* [dt, et - e] T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hdedet @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_be … HU2 … HLK … HTU ? ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. + +lemma tpss_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d + e ≤ dt → + ∃∃T2. K ⊢ T1 ▶* [dt - e, et] T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdedt @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_ge … HU2 … HLK … HTU ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. + +lemma tpss_inv_lift1_eq: ∀L,U1,U2,d,e. + L ⊢ U1 ▶* [d, e] U2 → ∀T1. ⇧[d, e] T1 ≡ U1 → U1 = U2. +#L #U1 #U2 #d #e #H #T1 #HTU1 @(tpss_ind … H) -U2 // +#U #U2 #_ #HU2 #IHU destruct +<(tps_inv_lift1_eq … HU2 … HTU1) -HU2 -HTU1 // +qed. + +lemma tpss_inv_lift1_ge_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → + ∃∃T2. K ⊢ T1 ▶* [d, dt + et - (d + e)] T2 & + ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_ge_up … HU2 … HLK … HTU ? ? ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. + +lemma tpss_inv_lift1_be_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → dt + et ≤ d + e → + ∃∃T2. K ⊢ T1 ▶* [dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_be_up … HU2 … HLK … HTU ? ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. + +lemma tpss_inv_lift1_le_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → d ≤ dt + et → dt + et ≤ d + e → + ∃∃T2. K ⊢ T1 ▶* [dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hddet #Hdetde @(tpss_ind … H) -U2 +[ /2 width=3/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (tps_inv_lift1_le_up … HU2 … HLK … HTU ? ? ?) -HU2 -HLK -HTU // /3 width=3/ +] +qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_tpss.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_tpss.etc new file mode 100644 index 000000000..d124bb32c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/cpr/tpss_tpss.etc @@ -0,0 +1,96 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/tps_tps.ma". +include "basic_2/unfold/tpss_lift.ma". + +(* PARTIAL UNFOLD ON TERMS **************************************************) + +(* Advanced inversion lemmas ************************************************) + +lemma tpss_inv_SO2: ∀L,T1,T2,d. L ⊢ T1 ▶* [d, 1] T2 → L ⊢ T1 ▶ [d, 1] T2. +#L #T1 #T2 #d #H @(tpss_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT1 +lapply (tps_trans_ge … IHT1 … HT2 ?) // +qed-. + +(* Advanced properties ******************************************************) + +lemma tpss_strip_eq: ∀L,T0,T1,d1,e1. L ⊢ T0 ▶* [d1, e1] T1 → + ∀T2,d2,e2. L ⊢ T0 ▶ [d2, e2] T2 → + ∃∃T. L ⊢ T1 ▶ [d2, e2] T & L ⊢ T2 ▶* [d1, e1] T. +/3 width=3/ qed. + +lemma tpss_strip_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 ▶* [d1, e1] T1 → + ∀L2,T2,d2,e2. L2 ⊢ T0 ▶ [d2, e2] T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. L2 ⊢ T1 ▶ [d2, e2] T & L1 ⊢ T2 ▶* [d1, e1] T. +/3 width=3/ qed. + +lemma tpss_strap1_down: ∀L,T1,T0,d1,e1. L ⊢ T1 ▶* [d1, e1] T0 → + ∀T2,d2,e2. L ⊢ T0 ▶ [d2, e2] T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 ▶ [d2, e2] T & L ⊢ T ▶* [d1, e1] T2. +/3 width=3/ qed. + +lemma tpss_strap2_down: ∀L,T1,T0,d1,e1. L ⊢ T1 ▶ [d1, e1] T0 → + ∀T2,d2,e2. L ⊢ T0 ▶* [d2, e2] T2 → d2 + e2 ≤ d1 → + ∃∃T. L ⊢ T1 ▶* [d2, e2] T & L ⊢ T ▶ [d1, e1] T2. +/3 width=3/ qed. + +lemma tpss_split_up: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → + ∀i. d ≤ i → i ≤ d + e → + ∃∃T. L ⊢ T1 ▶* [d, i - d] T & L ⊢ T ▶* [i, d + e - i] T2. +#L #T1 #T2 #d #e #H #i #Hdi #Hide @(tpss_ind … H) -T2 +[ /2 width=3/ +| #T #T2 #_ #HT12 * #T3 #HT13 #HT3 + elim (tps_split_up … HT12 … Hdi Hide) -HT12 -Hide #T0 #HT0 #HT02 + elim (tpss_strap1_down … HT3 … HT0 ?) -T [2: >commutative_plus /2 width=1/ ] + /3 width=7 by ex2_intro, step/ (**) (* just /3 width=7/ is too slow *) +] +qed. + +lemma tpss_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → + ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → + ∃∃T2. K ⊢ T1 ▶* [d, dt + et - (d + e)] T2 & + ⇧[d, e] T2 ≡ U2. +#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet +elim (tpss_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 +lapply (tpss_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 +lapply (tpss_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct +elim (tpss_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 -HLK -HTU1 // (tpss_inv_sort1 … HU) -HU #HU2 +>(lift_inv_sort2 … HU2) -HU2 // +qed-. + +lemma delift_inv_gref1: ∀L,U2,d,e,p. L ⊢ ▼*[d, e] §p ≡ U2 → U2 = §p. +#L #U #d #e #p * #U #HU +>(tpss_inv_gref1 … HU) -HU #HU2 +>(lift_inv_gref2 … HU2) -HU2 // +qed-. + +lemma delift_inv_bind1: ∀a,I,L,V1,T1,U2,d,e. L ⊢ ▼*[d, e] ⓑ{a,I} V1. T1 ≡ U2 → + ∃∃V2,T2. L ⊢ ▼*[d, e] V1 ≡ V2 & + L. ⓑ{I} V2 ⊢ ▼*[d+1, e] T1 ≡ T2 & + U2 = ⓑ{a,I} V2. T2. +#a #I #L #V1 #T1 #U2 #d #e * #U #HU #HU2 +elim (tpss_inv_bind1 … HU) -HU #V #T #HV1 #HT1 #X destruct +elim (lift_inv_bind2 … HU2) -HU2 #V2 #T2 #HV2 #HT2 +lapply (tpss_lsubr_trans … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ +qed-. + +lemma delift_inv_flat1: ∀I,L,V1,T1,U2,d,e. L ⊢ ▼*[d, e] ⓕ{I} V1. T1 ≡ U2 → + ∃∃V2,T2. L ⊢ ▼*[d, e] V1 ≡ V2 & + L ⊢ ▼*[d, e] T1 ≡ T2 & + U2 = ⓕ{I} V2. T2. +#I #L #V1 #T1 #U2 #d #e * #U #HU #HU2 +elim (tpss_inv_flat1 … HU) -HU #V #T #HV1 #HT1 #X destruct +elim (lift_inv_flat2 … HU2) -HU2 /3 width=5/ +qed-. + +lemma delift_inv_refl_O2: ∀L,T1,T2,d. L ⊢ ▼*[d, 0] T1 ≡ T2 → T1 = T2. +#L #T1 #T2 #d * #T #HT1 +>(tpss_inv_refl_O2 … HT1) -HT1 #HT2 +>(lift_inv_refl_O2 … HT2) -HT2 // +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma delift_fwd_tw: ∀L,T1,T2,d,e. L ⊢ ▼*[d, e] T1 ≡ T2 → ♯{T1} ≤ ♯{T2}. +#L #T1 #T2 #d #e * #T #HT1 #HT2 +>(lift_fwd_tw … HT2) -T2 /2 width=4 by tpss_fwd_tw/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_alt.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_alt.etc new file mode 100644 index 000000000..b9322cd1e --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_alt.etc @@ -0,0 +1,100 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/delift_lift.ma". + +(* INVERSE BASIC TERM RELOCATION *******************************************) + +(* alternative definition of inverse basic term relocation *) +inductive delifta: nat → nat → lenv → relation term ≝ +| delifta_sort : ∀L,d,e,k. delifta d e L (⋆k) (⋆k) +| delifta_lref_lt: ∀L,d,e,i. i < d → delifta d e L (#i) (#i) +| delifta_lref_be: ∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e → + ⇩[0, i] L ≡ K. ⓓV1 → delifta 0 (d + e - i - 1) K V1 V2 → + ⇧[0, d] V2 ≡ W2 → delifta d e L (#i) W2 +| delifta_lref_ge: ∀L,d,e,i. d + e ≤ i → delifta d e L (#i) (#(i - e)) +| delifta_gref : ∀L,d,e,p. delifta d e L (§p) (§p) +| delifta_bind : ∀L,a,I,V1,V2,T1,T2,d,e. + delifta d e L V1 V2 → delifta (d + 1) e (L. ⓑ{I} V2) T1 T2 → + delifta d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) +| delifta_flat : ∀L,I,V1,V2,T1,T2,d,e. + delifta d e L V1 V2 → delifta d e L T1 T2 → + delifta d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) +. + +interpretation "inverse basic relocation (term) alternative" + 'TSubstAlt L T1 d e T2 = (delifta d e L T1 T2). + +(* Basic properties *********************************************************) + +lemma delifta_lsubr_trans: ∀L1,T1,T2,d,e. L1 ⊢ ▼▼*[d, e] T1 ≡ T2 → + ∀L2. L2 ⊑ [d, e] L1 → L2 ⊢ ▼▼*[d, e] T1 ≡ T2. +#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e // /2 width=1/ +[ #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 + elim (ldrop_lsubr_ldrop2_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /3 width=6/ +| /4 width=1/ +| /3 width=1/ +] +qed. + +lemma delift_delifta: ∀L,T1,T2,d,e. L ⊢ ▼*[d, e] T1 ≡ T2 → L ⊢ ▼▼*[d, e] T1 ≡ T2. +#L #T1 @(f2_ind … fw … L T1) -L -T1 #n #IH #L * +[ * #i #Hn #T2 #d #e #H destruct + [ >(delift_inv_sort1 … H) -H // + | elim (delift_inv_lref1 … H) -H * /2 width=1/ + #K #V1 #V2 #Hdi #Hide #HLK #HV12 #HVT2 + lapply (ldrop_pair2_fwd_fw … HLK) #H + lapply (IH … HV12) // -H /2 width=6/ + | >(delift_inv_gref1 … H) -H // + ] +| * [ #a ] #I #V1 #T1 #Hn #X #d #e #H + [ elim (delift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + lapply (delift_lsubr_trans … HT12 (L.ⓑ{I}V1) ?) -HT12 /2 width=1/ #HT12 + lapply (IH … HV12) -HV12 // #HV12 + lapply (IH … HT12) -IH -HT12 /2 width=1/ #HT12 + lapply (delifta_lsubr_trans … HT12 (L.ⓑ{I}V2) ?) -HT12 /2 width=1/ + | elim (delift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + lapply (IH … HV12) -HV12 // + lapply (IH … HT12) -IH -HT12 // /2 width=1/ + ] +] +qed. + +(* Basic inversion lemmas ***************************************************) + +lemma delifta_delift: ∀L,T1,T2,d,e. L ⊢ ▼▼*[d, e] T1 ≡ T2 → L ⊢ ▼*[d, e] T1 ≡ T2. +#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e // /2 width=1/ /2 width=6/ +qed-. + +lemma delift_ind_alt: ∀R:ℕ→ℕ→lenv→relation term. + (∀L,d,e,k. R d e L (⋆k) (⋆k)) → + (∀L,d,e,i. i < d → R d e L (#i) (#i)) → + (∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e → + ⇩[O, i] L ≡ K.ⓓV1 → K ⊢ ▼*[O, d + e - i - 1] V1 ≡ V2 → + ⇧[O, d] V2 ≡ W2 → R O (d+e-i-1) K V1 V2 → R d e L (#i) W2 + ) → + (∀L,d,e,i. d + e ≤ i → R d e L (#i) (#(i - e))) → + (∀L,d,e,p. R d e L (§p) (§p)) → + (∀L,a,I,V1,V2,T1,T2,d,e. L ⊢ ▼*[d, e] V1 ≡ V2 → + L.ⓑ{I}V2 ⊢ ▼*[d + 1, e] T1 ≡ T2 → R d e L V1 V2 → + R (d+1) e (L.ⓑ{I}V2) T1 T2 → R d e L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) + ) → + (∀L,I,V1,V2,T1,T2,d,e. L ⊢ ▼*[d, e] V1 ≡ V2 → + L⊢ ▼*[d, e] T1 ≡ T2 → R d e L V1 V2 → + R d e L T1 T2 → R d e L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) + ) → + ∀d,e,L,T1,T2. L ⊢ ▼*[d, e] T1 ≡ T2 → R d e L T1 T2. +#R #H1 #H2 #H3 #H4 #H5 #H6 #H7 #d #e #L #T1 #T2 #H elim (delift_delifta … H) -L -T1 -T2 -d -e +// /2 width=1 by delifta_delift/ /3 width=1 by delifta_delift/ /3 width=7 by delifta_delift/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_delift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_delift.etc new file mode 100644 index 000000000..a5c563565 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_delift.etc @@ -0,0 +1,29 @@ +(**************************************************************************) +(* ___ *) +(* ||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/unfold/tpss_tpss.ma". +include "basic_2/unfold/delift.ma". + +(* INVERSE BASIC TERM RELOCATION *******************************************) + +(* Main properties **********************************************************) + +theorem delift_mono: ∀L,T,T1,T2,d,e. + L ⊢ ▼*[d, e] T ≡ T1 → L ⊢ ▼*[d, e] T ≡ T2 → T1 = T2. +#L #T #T1 #T2 #d #e * #U1 #H1TU1 #H2TU1 * #U2 #H1TU2 #H2TU2 +elim (tpss_conf_eq … H1TU1 … H1TU2) -T #U #HU1 #HU2 +lapply (tpss_inv_lift1_eq … HU1 … H2TU1) -HU1 #H destruct +lapply (tpss_inv_lift1_eq … HU2 … H2TU2) -HU2 #H destruct +lapply (lift_inj … H2TU1 … H2TU2) // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_lift.etc new file mode 100644 index 000000000..8abaeb9f9 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/delift/delift_lift.etc @@ -0,0 +1,164 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/ldrop_lbotr.ma". +include "basic_2/unfold/tpss_lift.ma". +include "basic_2/unfold/delift.ma". + +(* INVERSE BASIC TERM RELOCATION *******************************************) + +(* Advanced properties ******************************************************) + +lemma delift_lref_be: ∀L,K,V1,V2,U2,i,d,e. d ≤ i → i < d + e → + ⇩[0, i] L ≡ K. ⓓV1 → K ⊢ ▼*[0, d + e - i - 1] V1 ≡ V2 → + ⇧[0, d] V2 ≡ U2 → L ⊢ ▼*[d, e] #i ≡ U2. +#L #K #V1 #V2 #U2 #i #d #e #Hdi #Hide #HLK * #V #HV1 #HV2 #HVU2 +elim (lift_total V 0 (i+1)) #U #HVU +lapply (lift_trans_be … HV2 … HVU ? ?) -HV2 // >minus_plus commutative_plus in ⊢ (??%??→?); H -H /2 width=1/ ] -Hde -H #V2 #V12 (**) (* H erased two times *) + elim (lift_total V2 0 d) /3 width=7/ +| #a #I #V1 #T1 #H #d #e #Hde #HL destruct + elim (IH … V1 … Hde HL) // #V2 #HV12 + elim (IH (L.ⓑ{I}V1) T1 … (d+1) e ??) -IH // [2,3: /2 width=1/ ] -Hde -HL #T2 #HT12 + lapply (delift_lsubr_trans … HT12 (L.ⓑ{I}V2) ?) -HT12 /2 width=1/ /3 width=4/ +| #I #V1 #T1 #H #d #e #Hde #HL destruct + elim (IH … V1 … Hde HL) // #V2 #HV12 + elim (IH … T1 … Hde HL) -IH -Hde -HL // /3 width=2/ +] +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma delift_inv_lref1_lt: ∀L,U2,i,d,e. L ⊢ ▼*[d, e] #i ≡ U2 → i < d → U2 = #i. +#L #U2 #i #d #e * #U #HU #HU2 #Hid +elim (tpss_inv_lref1 … HU) -HU +[ #H destruct >(lift_inv_lref2_lt … HU2) // +| * #K #V1 #V2 #Hdi + lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi + elim (lt_refl_false … Hi) +] +qed-. + +lemma delift_inv_lref1_be: ∀L,U2,d,e,i. L ⊢ ▼*[d, e] #i ≡ U2 → + d ≤ i → i < d + e → + ∃∃K,V1,V2. ⇩[0, i] L ≡ K. ⓓV1 & + K ⊢ ▼*[0, d + e - i - 1] V1 ≡ V2 & + ⇧[0, d] V2 ≡ U2. +#L #U2 #d #e #i * #U #HU #HU2 #Hdi #Hide +elim (tpss_inv_lref1 … HU) -HU +[ #H destruct elim (lift_inv_lref2_be … HU2 ? ?) // +| * #K #V1 #V #_ #_ #HLK #HV1 #HVU + elim (lift_div_be … HVU … HU2 ? ?) -U // /2 width=1/ /3 width=6/ +] +qed-. + +lemma delift_inv_lref1_ge: ∀L,U2,i,d,e. L ⊢ ▼*[d, e] #i ≡ U2 → + d + e ≤ i → U2 = #(i - e). +#L #U2 #i #d #e * #U #HU #HU2 #Hdei +elim (tpss_inv_lref1 … HU) -HU +[ #H destruct >(lift_inv_lref2_ge … HU2) // +| * #K #V1 #V2 #_ #Hide + lapply (lt_to_le_to_lt … Hide Hdei) -Hide -Hdei #Hi + elim (lt_refl_false … Hi) +] +qed-. + +lemma delift_inv_lref1: ∀L,U2,i,d,e. L ⊢ ▼*[d, e] #i ≡ U2 → + ∨∨ (i < d ∧ U2 = #i) + | (∃∃K,V1,V2. d ≤ i & i < d + e & + ⇩[0, i] L ≡ K. ⓓV1 & + K ⊢ ▼*[0, d + e - i - 1] V1 ≡ V2 & + ⇧[0, d] V2 ≡ U2 + ) + | (d + e ≤ i ∧ U2 = #(i - e)). +#L #U2 #i #d #e #H +elim (lt_or_ge i d) #Hdi +[ elim (delift_inv_lref1_lt … H Hdi) -H /3 width=1/ +| elim (lt_or_ge i (d+e)) #Hide + [ elim (delift_inv_lref1_be … H Hdi Hide) -H /3 width=6/ + | elim (delift_inv_lref1_ge … H Hide) -H /3 width=1/ + ] +] +qed-. + +(* Properties on basic term relocation **************************************) + +lemma delift_lift_le: ∀K,T1,T2,dt,et. K ⊢ ▼*[dt, et] T1 ≡ T2 → + ∀L,U1,d,e. dt + et ≤ d → ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d - et, e] T2 ≡ U2 → + L ⊢ ▼*[dt, et] U1 ≡ U2. +#K #T1 #T2 #dt #et * #T #HT1 #HT2 #L #U1 #d #e #Hdetd #HLK #HTU1 #U2 #HTU2 +elim (lift_total T d e) #U #HTU +lapply (tpss_lift_le … HT1 … HLK HTU1 … HTU) -T1 -HLK // #HU1 +elim (lift_trans_ge … HT2 … HTU ?) -T // -Hdetd #T #HT2 #HTU +>(lift_mono … HTU2 … HT2) -T2 /2 width=3/ +qed. + +lemma delift_lift_be: ∀K,T1,T2,dt,et. K ⊢ ▼*[dt, et] T1 ≡ T2 → + ∀L,U1,d,e. dt ≤ d → d ≤ dt + et → + ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → + L ⊢ ▼*[dt, et + e] U1 ≡ T2. +#K #T1 #T2 #dt #et * #T #HT1 #HT2 #L #U1 #d #e #Hdtd #Hddet #HLK #HTU1 +elim (lift_total T d e) #U #HTU +lapply (tpss_lift_be … HT1 … HLK HTU1 … HTU) -T1 -HLK // #HU1 +lapply (lift_trans_be … HT2 … HTU ? ?) -T // -Hdtd -Hddet /2 width=3/ +qed. + +lemma delift_lift_ge: ∀K,T1,T2,dt,et. K ⊢ ▼*[dt, et] T1 ≡ T2 → + ∀L,U1,d,e. d ≤ dt → ⇩[d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → + L ⊢ ▼*[dt + e, et] U1 ≡ U2. +#K #T1 #T2 #dt #et * #T #HT1 #HT2 #L #U1 #d #e #Hddt #HLK #HTU1 #U2 #HTU2 +elim (lift_total T d e) #U #HTU +lapply (tpss_lift_ge … HT1 … HLK HTU1 … HTU) -T1 -HLK // #HU1 +elim (lift_trans_le … HT2 … HTU ?) -T // -Hddt #T #HT2 #HTU +>(lift_mono … HTU2 … HT2) -T2 /2 width=3/ +qed. + +lemma delift_inv_lift1_eq: ∀L,U1,T2,d,e. L ⊢ ▼*[d, e] U1 ≡ T2 → + ∀K. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → T1 = T2. +#L #U1 #T2 #d #e * #U2 #HU12 #HTU2 #K #HLK #T1 #HTU1 +lapply (tpss_inv_lift1_eq … HU12 … HTU1) -L -K #H destruct +lapply (lift_inj … HTU1 … HTU2) -U2 // +qed-. + +lemma delift_lift_div_be: ∀L,T1,T,d,e,i. L ⊢ ▼*[i, d + e - i] T1 ≡ T → + ∀T2. ⇧[d, i - d] T2 ≡ T → d ≤ i → i ≤ d + e → + L ⊢ ▼*[d, e] T1 ≡ T2. +#L #T1 #T #d #e #i * #T0 #HT10 #HT0 #T2 #HT2 #Hdi #Hide +lapply (tpss_weak … HT10 d e ? ?) -HT10 // [ >commutative_plus /2 width=1/ ] #HT10 +lapply (lift_trans_be … HT2 … HT0 ? ?) -T // +>commutative_plus >commutative_plus in ⊢ (? ? (? % ?) ? ? → ?); +(lpx_sn_inv_atom1 … H1) -X1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation.ma b/matita/matita/contribs/lambdadelta/basic_2/notation.ma index 5e730aec2..689d6edc2 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/notation.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/notation.ma @@ -306,17 +306,17 @@ notation "hvbox( L ⊢ break 𝐇𝐍 ⦃ term 46 T ⦄ )" non associative with precedence 45 for @{ 'HdNormal $L $T }. -notation "hvbox( T1 ➡ break term 46 T2 )" - non associative with precedence 45 - for @{ 'PRed $T1 $T2 }. - notation "hvbox( L ⊢ break term 46 T1 ➡ break term 46 T2 )" non associative with precedence 45 for @{ 'PRed $L $T1 $T2 }. -notation "hvbox( ⦃ term 46 L1 ⦄ ➡ break ⦃ term 46 L2 ⦄ )" +notation "hvbox( L1 ⊢ ➡ break term 46 L2 )" + non associative with precedence 45 + for @{ 'PRedSn $L1 $L2 }. + +notation "hvbox( L1 ⊢ ➡ ➡ break term 46 L2 )" non associative with precedence 45 - for @{ 'FocalizedPRed $L1 $L2 }. + for @{ 'PRedSnAlt $L1 $L2 }. notation "hvbox( ⦃ term 46 L1, break term 46 T1 ⦄ ➡ break ⦃ term 46 L2 , break term 46 T2 ⦄ )" non associative with precedence 45 @@ -326,10 +326,6 @@ notation "hvbox( L ⊢ break ⦃ term 46 L1, break term 46 T1 ⦄ ➡ break ⦃ non associative with precedence 45 for @{ 'FocalizedPRed $L $L1 $T1 $L2 $T2 }. -notation "hvbox( ⦃ term 46 L1 ⦄ ➡ ➡ break ⦃ term 46 L2 ⦄ )" - non associative with precedence 45 - for @{ 'FocalizedPRedAlt $L1 $L2 }. - (* Computation **************************************************************) notation "hvbox( T1 ➡ * break term 46 T2 )" diff --git a/matita/matita/contribs/lambdadelta/basic_2/restricted/cpqs.ma b/matita/matita/contribs/lambdadelta/basic_2/restricted/cpqs.ma index 9d5983292..4992de0b0 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/restricted/cpqs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/restricted/cpqs.ma @@ -48,9 +48,8 @@ lemma cpqs_lsubr_trans: ∀L1,T1,T2. L1 ⊢ T1 ➤* T2 → elim (ldrop_lsubr_ldrop2_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // -Hi (cpss_inv_sort1 … H) -H // +| #I #L1 #K1 #V1 #B #i #HLK1 #_ #IHV1 #X #H #L2 #HL12 + elim (cpss_inv_lref1 … H) -H + [ #H destruct + elim (lpss_ldrop_conf … HLK1 … HL12) -L1 #X #H #HLK2 + elim (lpss_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct /3 width=6/ + | * #Y #Z #V2 #H #HV12 #HV2 + lapply (ldrop_mono … H … HLK1) -H #H destruct + elim (lpss_ldrop_conf … HLK1 … HL12) -L1 #Z #H #HLK2 + elim (lpss_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct + lapply (ldrop_fwd_ldrop2 … HLK2) -V0 /3 width=7/ + ] +| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12 + elim (cpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=2/ +| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12 + elim (cpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=1/ +| #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12 + elim (cpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /3 width=3/ +| #L1 #V1 #T1 #A #_ #_ #IHV1 #IHT1 #X #H #L2 #HL12 + elim (cpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /3 width=1/ +] +qed-. + +lemma aaa_cpss_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → ∀T2. L ⊢ T1 ▶* T2 → L ⊢ T2 ⁝ A. +/2 width=5 by aaa_cpss_lpss_conf/ qed-. + +lemma aaa_lpss_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → ∀L2. L1 ⊢ ▶* L2 → L2 ⊢ T ⁝ A. +/2 width=5 by aaa_cpss_lpss_conf/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ltpss_dx.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ltpss_dx.ma deleted file mode 100644 index 2f2d07360..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ltpss_dx.ma +++ /dev/null @@ -1,79 +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/unfold/tpss_tpss.ma". -include "basic_2/unfold/ltpss_dx_ldrop.ma". -include "basic_2/static/aaa_lift.ma". - -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) - -(* Properties about dx parallel unfold **************************************) - -(* Note: lemma 500 *) -lemma aaa_ltpss_dx_tpss_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → - ∀L2,d,e. L1 ▶* [d, e] L2 → - ∀T2. L2 ⊢ T1 ▶* [d, e] T2 → L2 ⊢ T2 ⁝ A. -#L1 #T1 #A #H elim H -L1 -T1 -A -[ #L1 #k #L2 #d #e #_ #T2 #H - >(tpss_inv_sort1 … H) -H // -| #I #L1 #K1 #V1 #B #i #HLK1 #_ #IHV1 #L2 #d #e #HL12 #T2 #H - elim (tpss_inv_lref1 … H) -H - [ #H destruct - elim (lt_or_ge i d) #Hdi - [ elim (ltpss_dx_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2 - elim (ltpss_dx_inv_tpss11 … H ?) -H /2 width=1/ -Hdi #K2 #V2 #HK12 #HV12 #H destruct - /3 width=8 by aaa_lref/ (**) (* too slow without trace *) - | elim (lt_or_ge i (d + e)) #Hide - [ elim (ltpss_dx_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HLK2 - elim (ltpss_dx_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K2 #V2 #HK12 #HV12 #H destruct - /3 width=8 by aaa_lref/ (**) (* too slow without trace *) - | -Hdi - lapply (ltpss_dx_ldrop_conf_ge … HL12 … HLK1 ?) -L1 // -Hide - /3 width=8 by aaa_lref/ (**) (* too slow without trace *) - ] - ] - | * #K2 #V2 #W2 #Hdi #Hide #HLK2 #HVW2 #HWT2 - elim (ltpss_dx_ldrop_conf_be … HL12 … HLK1 ? ?) -L1 // /2 width=2/ #X #H #HL2K0 - elim (ltpss_dx_inv_tpss21 … H ?) -H /2 width=1/ -Hdi -Hide #K0 #V0 #HK12 #HV12 #H destruct - lapply (ldrop_mono … HL2K0 … HLK2) -HL2K0 #H destruct - lapply (ldrop_fwd_ldrop2 … HLK2) -HLK2 #HLK2 - lapply (tpss_trans_eq … HV12 HVW2) -V2 /3 width=7/ - ] -| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H - elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=4/ -| #a #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H - elim (tpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /4 width=4/ -| #L1 #V1 #T1 #B #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H - elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /3 width=4/ -| #L1 #V1 #T1 #A #_ #_ #IHV1 #IHT1 #L2 #d #e #HL12 #X #H - elim (tpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct /3 width=4/ -] -qed. - -lemma aaa_ltpss_dx_tps_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → - ∀L2,d,e. L1 ▶* [d, e] L2 → - ∀T2. L2 ⊢ T1 ▶ [d, e] T2 → L2 ⊢ T2 ⁝ A. -/3 width=7/ qed. - -lemma aaa_ltpss_dx_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → - ∀L2,d,e. L1 ▶* [d, e] L2 → L2 ⊢ T ⁝ A. -/2 width=7/ qed. - -lemma aaa_tpss_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → - ∀T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T2 ⁝ A. -/2 width=7/ qed. - -lemma aaa_tps_conf: ∀L,T1,A. L ⊢ T1 ⁝ A → - ∀T2,d,e. L ⊢ T1 ▶ [d, e] T2 → L ⊢ T2 ⁝ A. -/2 width=7/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ltpss_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ltpss_sn.ma deleted file mode 100644 index 4f2a44827..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/aaa_ltpss_sn.ma +++ /dev/null @@ -1,37 +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/unfold/ltpss_sn_alt.ma". -include "basic_2/static/aaa_ltpss_dx.ma". - -(* ATONIC ARITY ASSIGNMENT ON TERMS *****************************************) - -(* Properties about sn parallel unfold **************************************) - -lemma aaa_ltpss_sn_conf: ∀L1,T,A. L1 ⊢ T ⁝ A → - ∀L2,d,e. L1 ⊢ ▶* [d, e] L2 → L2 ⊢ T ⁝ A. -#L1 #T #A #HT #L2 #d #e #HL12 -lapply (ltpss_sn_ltpssa … HL12) -HL12 #HL12 -@(TC_Conf3 … (λL,A. L ⊢ T ⁝ A) … HT ? HL12) /2 width=5/ -qed. - -lemma aaa_ltpss_sn_tpss_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → - ∀L2,d,e. L1 ⊢ ▶* [d, e] L2 → - ∀T2. L2 ⊢ T1 ▶* [d, e] T2 → L2 ⊢ T2 ⁝ A. -/3 width=5/ qed. - -lemma aaa_ltpss_sn_tps_conf: ∀L1,T1,A. L1 ⊢ T1 ⁝ A → - ∀L2,d,e. L1 ⊢ ▶* [d, e] L2 → - ∀T2. L2 ⊢ T1 ▶ [d, e] T2 → L2 ⊢ T2 ⁝ A. -/3 width=5/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lpss.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lpss.ma new file mode 100644 index 000000000..0aacb2b63 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_lpss.ma @@ -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/unfold/lpss_ldrop.ma". +include "basic_2/static/ssta_lift.ma". + +(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************) + +(* Properties about sn parallel unfold **************************************) + +(* Note: apparently this was missing in basic_1 *) +lemma ssta_tpss_lpss_conf: ∀h,g,L1,T1,U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ → + ∀T2. L1 ⊢ T1 ▶* T2 → ∀L2. L1 ⊢ ▶* L2 → + ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄ & L1 ⊢ U1 ▶* U2. +#h #g #L1 #T1 #U1 #l #H elim H -L1 -T1 -U1 -l +[ #L1 #k1 #l1 #Hkl1 #X #H + >(cpss_inv_sort1 … H) -H /3 width=3/ +| #L1 #K1 #V1 #W1 #U1 #i #l #HLK1 #_ #HWU1 #IHVW1 #X #H #L2 #HL12 + elim (cpss_inv_lref1 … H) -H + [ #H destruct + elim (lpss_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2 + elim (lpss_inv_pair1 … H) -H #K2 #V2 #HK12 #HV12 #H destruct + lapply (ldrop_fwd_ldrop2 … HLK1) -HLK1 #HLK1 + elim (IHVW1 … HV12 … HK12) -IHVW1 -HV12 -HK12 #W2 #HVW2 #HW12 + elim (lift_total W2 0 (i+1)) #U2 #HWU2 + lapply (cpss_lift … HW12 … HLK1 … HWU1 … HWU2) -HW12 -HLK1 -HWU1 /3 width=6/ + | * #Y #Z #V2 #H #HV12 #HV2 + lapply (ldrop_mono … H … HLK1) -H #H destruct + elim (lpss_ldrop_conf … HLK1 … HL12) -HL12 #Z #H #HLK2 + elim (lpss_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct + lapply (ldrop_fwd_ldrop2 … HLK2) -V0 #HLK2 + lapply (ldrop_fwd_ldrop2 … HLK1) -HLK1 #HLK1 + elim (IHVW1 … HV12 … HK12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12 + elim (lift_total W2 0 (i+1)) #U2 #HWU2 + lapply (ssta_lift … HVW2 … HLK2 … HV2 … HWU2) -HVW2 -HLK2 -HV2 + lapply (cpss_lift … HW12 … HLK1 … HWU1 … HWU2) -HW12 -HLK1 -HWU1 -HWU2 /3 width=3/ + ] +| #L1 #K1 #W1 #V1 #U1 #i #l #HLK1 #_ #HWU1 #IHWV1 #X #H #L2 #HL12 + elim (cpss_inv_lref1 … H) -H [ | -IHWV1 -HWU1 -HL12 ] + [ #H destruct + elim (lpss_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2 + elim (lpss_inv_pair1 … H) -H #K2 #W2 #HK12 #HW12 #H destruct + lapply (ldrop_fwd_ldrop2 … HLK1) -HLK1 #HLK1 + elim (IHWV1 … HW12 … HK12) -IHWV1 -HK12 #V2 #HWV2 #_ + elim (lift_total W2 0 (i+1)) #U2 #HWU2 + lapply (cpss_lift … HW12 … HLK1 … HWU1 … HWU2) -HW12 -HLK1 -HWU1 /3 width=6/ + | * #K2 #V2 #W2 #HLK2 #_ #_ + lapply (ldrop_mono … HLK2 … HLK1) -HLK1 -HLK2 #H destruct + ] +| #a #I #L1 #V1 #T1 #U1 #l #_ #IHTU1 #X #H #L2 #HL12 + elim (cpss_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + elim (IHTU1 … HT12 (L2.ⓑ{I}V2)) -IHTU1 -HT12 /2 width=1/ -HL12 /3 width=5/ +| #L1 #V1 #T1 #U1 #l #_ #IHTU1 #X #H #L2 #HL12 + elim (cpss_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct + elim (IHTU1 … HT12 … HL12) -IHTU1 -HT12 -HL12 /3 width=5/ +| #L1 #W1 #T1 #U1 #l #_ #IHTU1 #X #H #L2 #HL12 + elim (cpss_inv_flat1 … H) -H #W2 #T2 #HW12 #HT12 #H destruct + elim (IHTU1 … HT12 … HL12) -IHTU1 -HT12 -HL12 /3 width=3/ +] +qed-. + +lemma ssta_cpss_conf: ∀h,g,L,T1,U1,l. ⦃h, L⦄ ⊢ T1 •[g] ⦃l, U1⦄ → + ∀T2. L ⊢ T1 ▶* T2 → + ∃∃U2. ⦃h, L⦄ ⊢ T2 •[g] ⦃l, U2⦄ & L ⊢ U1 ▶* U2. +/2 width=3 by ssta_tpss_lpss_conf/ qed-. + +lemma ssta_lpss_conf: ∀h,g,L1,T,U1,l. ⦃h, L1⦄ ⊢ T •[g] ⦃l, U1⦄ → + ∀L2. L1 ⊢ ▶* L2 → + ∃∃U2. ⦃h, L2⦄ ⊢ T •[g] ⦃l, U2⦄ & L1 ⊢ U1 ▶* U2. +/2 width=3 by ssta_tpss_lpss_conf/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_ltpss_dx.ma b/matita/matita/contribs/lambdadelta/basic_2/static/ssta_ltpss_dx.ma deleted file mode 100644 index 72eff7850..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/static/ssta_ltpss_dx.ma +++ /dev/null @@ -1,125 +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/unfold/tpss_tpss.ma". -include "basic_2/unfold/ltpss_dx_tpss.ma". -include "basic_2/static/ssta_lift.ma". - -(* STRATIFIED STATIC TYPE ASSIGNMENT ON TERMS *******************************) - -(* Properties about dx parallel unfold **************************************) - -(* Note: apparently this was missing in basic_1 *) -lemma ssta_ltpss_dx_tpss_conf: ∀h,g,L1,T1,U1,l. ⦃h, L1⦄ ⊢ T1 •[g] ⦃l, U1⦄ → - ∀L2,d,e. L1 ▶* [d, e] L2 → - ∀T2. L2 ⊢ T1 ▶* [d, e] T2 → - ∃∃U2. ⦃h, L2⦄ ⊢ T2 •[g] ⦃l, U2⦄ & - L2 ⊢ U1 ▶* [d, e] U2. -#h #g #L1 #T1 #U #l #H elim H -L1 -T1 -U -l -[ #L1 #k1 #l1 #Hkl1 #L2 #d #e #_ #T2 #H - >(tpss_inv_sort1 … H) -H /3 width=3/ -| #L1 #K1 #V1 #W1 #U1 #i #l #HLK1 #HVW1 #HWU1 #IHVW1 #L2 #d #e #HL12 #T2 #H - elim (tpss_inv_lref1 … H) -H [ | -HVW1 ] - [ #H destruct - elim (lt_or_ge i d) #Hdi [ -HVW1 | ] - [ elim (ltpss_dx_ldrop_conf_le … HL12 … HLK1 ?) -L1 /2 width=2/ #X #H #HLK2 - elim (ltpss_dx_inv_tpss11 … H ?) -H /2 width=1/ #K2 #V2 #HK12 #HV12 #H destruct - elim (IHVW1 … HK12 … HV12) -IHVW1 -HK12 -HV12 #W2 #HVW2 #HW12 - lapply (ldrop_fwd_ldrop2 … HLK2) #H - elim (lift_total W2 0 (i+1)) #U2 #HWU2 - lapply (tpss_lift_ge … HW12 … H … HWU1 … HWU2) // -HW12 -H -HWU1 - >minus_plus minus_plus #H - lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus minus_plus #H - lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus minus_plus minus_plus #H - lapply (tpss_weak … H d e ? ?) [1,2: normalize [ >commutative_plus (plus_minus_m_m j d) in ⊢ (% → ?); // -Hdj /3 width=4/ - | -Hdi -Hdj #Hij - lapply (plus_minus_m_m … Hjde) -Hjde /3 width=8/ - ] -| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 - elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ - -Hdi -Hide >arith_c1x #T #HT1 #HT2 - lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // - -Hdi -Hide /3 width=5/ -] -qed. - -lemma tps_split_down: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → - ∀i. d ≤ i → i ≤ d + e → - ∃∃T. L ⊢ T1 ▶ [i, d + e - i] T & - L ⊢ T ▶ [d, i - d] T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e -[ /2 width=3/ -| #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hdj #Hjde - elim (lt_or_ge i j) - [ -Hide -Hjde >(plus_minus_m_m j d) in ⊢ (% → ?); // -Hdj /3 width=8/ - | -Hdi -Hdj - >(plus_minus_m_m (d+e) j) in Hide; // -Hjde /3 width=4/ - ] -| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // #V #HV1 #HV2 - elim (IHT12 (i + 1) ? ?) -IHT12 /2 width=1/ - -Hdi -Hide >arith_c1x #T #HT1 #HT2 - lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /3 width=5/ -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hdi #Hide - elim (IHV12 i ? ?) -IHV12 // elim (IHT12 i ? ?) -IHT12 // - -Hdi -Hide /3 width=5/ -] -qed. - -lemma tps_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶ [d, e] T2 → - ∀L. L @@ K ⊢ T1 ▶ [d, e] T2. -#K #T1 #T2 #d #e #H elim H -K -T1 -T2 -d -e // /2 width=1/ -#K #K0 #V #W #i #d #e #Hdi #Hide #HK0 #HVW #L -lapply (ldrop_fwd_ldrop2_length … HK0) #H -@(tps_subst … (L@@K0) … HVW) // (**) (* /3/ does not work *) -@(ldrop_O1_append_sn_le … HK0) /2 width=2/ -qed. - -(* Basic inversion lemmas ***************************************************) - -fact tps_inv_atom1_aux: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → ∀I. T1 = ⓪{I} → - T2 = ⓪{I} ∨ - ∃∃K,V,i. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2 & - I = LRef i. -#L #T1 #T2 #d #e * -L -T1 -T2 -d -e -[ #L #I #d #e #J #H destruct /2 width=1/ -| #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #I #H destruct /3 width=8/ -| #L #a #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct -] -qed. - -lemma tps_inv_atom1: ∀L,T2,I,d,e. L ⊢ ⓪{I} ▶ [d, e] T2 → - T2 = ⓪{I} ∨ - ∃∃K,V,i. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2 & - I = LRef i. -/2 width=3/ qed-. - - -(* Basic_1: was: subst1_gen_sort *) -lemma tps_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k ▶ [d, e] T2 → T2 = ⋆k. -#L #T2 #k #d #e #H -elim (tps_inv_atom1 … H) -H // -* #K #V #i #_ #_ #_ #_ #H destruct -qed-. - -(* Basic_1: was: subst1_gen_lref *) -lemma tps_inv_lref1: ∀L,T2,i,d,e. L ⊢ #i ▶ [d, e] T2 → - T2 = #i ∨ - ∃∃K,V. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV & - ⇧[O, i + 1] V ≡ T2. -#L #T2 #i #d #e #H -elim (tps_inv_atom1 … H) -H /2 width=1/ -* #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=4/ -qed-. - -lemma tps_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p ▶ [d, e] T2 → T2 = §p. -#L #T2 #p #d #e #H -elim (tps_inv_atom1 … H) -H // -* #K #V #i #_ #_ #_ #_ #H destruct -qed-. - -fact tps_inv_bind1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 → - ∀a,I,V1,T1. U1 = ⓑ{a,I} V1. T1 → - ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & - L. ⓑ{I} V2 ⊢ T1 ▶ [d + 1, e] T2 & - U2 = ⓑ{a,I} V2. T2. -#d #e #L #U1 #U2 * -d -e -L -U1 -U2 -[ #L #k #d #e #a #I #V1 #T1 #H destruct -| #L #K #V #W #i #d #e #_ #_ #_ #_ #a #I #V1 #T1 #H destruct -| #L #b #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #a #I #V #T #H destruct /2 width=5/ -| #L #J #V1 #V2 #T1 #T2 #d #e #_ #_ #a #I #V #T #H destruct -] -qed. - -lemma tps_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. L ⊢ ⓑ{a,I} V1. T1 ▶ [d, e] U2 → - ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & - L. ⓑ{I} V2 ⊢ T1 ▶ [d + 1, e] T2 & - U2 = ⓑ{a,I} V2. T2. -/2 width=3/ qed-. - -fact tps_inv_flat1_aux: ∀d,e,L,U1,U2. L ⊢ U1 ▶ [d, e] U2 → - ∀I,V1,T1. U1 = ⓕ{I} V1. T1 → - ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & L ⊢ T1 ▶ [d, e] T2 & - U2 = ⓕ{I} V2. T2. -#d #e #L #U1 #U2 * -d -e -L -U1 -U2 -[ #L #k #d #e #I #V1 #T1 #H destruct -| #L #K #V #W #i #d #e #_ #_ #_ #_ #I #V1 #T1 #H destruct -| #L #a #J #V1 #V2 #T1 #T2 #d #e #_ #_ #I #V #T #H destruct -| #L #J #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #I #V #T #H destruct /2 width=5/ -] -qed. - -lemma tps_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 ▶ [d, e] U2 → - ∃∃V2,T2. L ⊢ V1 ▶ [d, e] V2 & L ⊢ T1 ▶ [d, e] T2 & - U2 = ⓕ{I} V2. T2. -/2 width=3/ qed-. - -fact tps_inv_refl_O2_aux: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → e = 0 → T1 = T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e -[ // -| #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct - lapply (le_to_lt_to_lt … Hdi … Hide) -Hdi -Hide shift_append_assoc normalize #H - elim (tps_inv_bind1 … H) -H - #V0 #T0 #_ #HT10 #H destruct - elim (IH … HT10) -IH -HT10 #L2 #T2 #HL12 #H destruct - >append_length >HL12 -HL12 - @(ex2_2_intro … (⋆.ⓑ{I}V0@@L2) T2) [ >append_length ] // /2 width=3/ (**) (* explicit constructor *) -] -qed-. - -(* Basic_1: removed theorems 25: - subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt - subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans - subst0_lift_lt subst0_lift_ge subst0_lift_ge_S subst0_lift_ge_s - subst0_subst0 subst0_subst0_back subst0_weight_le subst0_weight_lt - subst0_confluence_neq subst0_confluence_eq subst0_tlt_head - subst0_confluence_lift subst0_tlt - subst1_head subst1_gen_head subst1_lift_S subst1_confluence_lift -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/tps_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/tps_lift.ma deleted file mode 100644 index 1d65d8d63..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/tps_lift.ma +++ /dev/null @@ -1,294 +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/ldrop_ldrop.ma". -include "basic_2/substitution/tps.ma". - -(* PARTIAL SUBSTITUTION ON TERMS ********************************************) - -(* Advanced inversion lemmas ************************************************) - -fact tps_inv_S2_aux: ∀L,T1,T2,d,e1. L ⊢ T1 ▶ [d, e1] T2 → ∀e2. e1 = e2 + 1 → - ∀K,V. ⇩[0, d] L ≡ K. ⓛV → L ⊢ T1 ▶ [d + 1, e2] T2. -#L #T1 #T2 #d #e1 #H elim H -L -T1 -T2 -d -e1 -[ // -| #L #K0 #V0 #W #i #d #e1 #Hdi #Hide1 #HLK0 #HV0 #e2 #He12 #K #V #HLK destruct - elim (lt_or_ge i (d+1)) #HiSd - [ -Hide1 -HV0 - lapply (le_to_le_to_eq … Hdi ?) /2 width=1/ #H destruct - lapply (ldrop_mono … HLK0 … HLK) #H destruct - | -V -Hdi /2 width=4/ - ] -| /4 width=3/ -| /3 width=3/ -] -qed. - -lemma tps_inv_S2: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e + 1] T2 → - ∀K,V. ⇩[0, d] L ≡ K. ⓛV → L ⊢ T1 ▶ [d + 1, e] T2. -/2 width=3/ qed-. - -lemma tps_inv_refl_SO2: ∀L,T1,T2,d. L ⊢ T1 ▶ [d, 1] T2 → - ∀K,V. ⇩[0, d] L ≡ K. ⓛV → T1 = T2. -#L #T1 #T2 #d #HT12 #K #V #HLK -lapply (tps_inv_S2 … T1 T2 … 0 … HLK) -K // -HT12 #HT12 -lapply (tps_inv_refl_O2 … HT12) -HT12 // -qed-. - -(* Relocation properties ****************************************************) - -(* Basic_1: was: subst1_lift_lt *) -lemma tps_lift_le: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → - ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - dt + et ≤ d → - L ⊢ U1 ▶ [dt, et] U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdetd - lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid - lapply (lift_inv_lref1_lt … H … Hid) -H #H destruct - elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -Y -HVW #H destruct /2 width=4/ -| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=6/ | @IHT12 /2 width=6/ ] (**) (* /3 width=6/ is too slow, arith3 needed to avoid crash *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd - elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ -] -qed. - -lemma tps_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → - ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - dt ≤ d → d ≤ dt + et → - L ⊢ U1 ▶ [dt, et + e] U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hdtd #_ - elim (lift_inv_lref1 … H) -H * #Hid #H destruct - [ -Hdtd - lapply (lt_to_le_to_lt … (dt+et+e) Hidet ?) // -Hidet #Hidete - elim (lift_trans_ge … HVW … HWU2 ?) -W // (lift_mono … HVY … HVW) -V #H destruct /2 width=4/ - | -Hdti - lapply (transitive_le … Hdtd Hid) -Hdtd #Hdti - lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 - lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ - ] -| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdtd #Hddet - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=6/ | @IHT12 [3,4: // | skip |5,6: /2 width=1/ | /2 width=1/ ] - ] (**) (* /3 width=6/ is too slow, simplification like tps_lift_le is too slow *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hdetd - elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=6/ -] -qed. - -(* Basic_1: was: subst1_lift_ge *) -lemma tps_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 ▶ [dt, et] T2 → - ∀L,U1,U2,d,e. ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - d ≤ dt → - L ⊢ U1 ▶ [dt + e, et] U2. -#K #T1 #T2 #dt #et #H elim H -K -T1 -T2 -dt -et -[ #K #I #dt #et #L #U1 #U2 #d #e #_ #H1 #H2 #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #d #e #HLK #H #HWU2 #Hddt - lapply (transitive_le … Hddt … Hdti) -Hddt #Hid - lapply (lift_inv_lref1_ge … H … Hid) -H #H destruct - lapply (lift_trans_be … HVW … HWU2 ? ?) -W // /2 width=1/ >plus_plus_comm_23 #HVU2 - lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid /3 width=4/ -| #K #a #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt - elim (lift_inv_bind1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_bind1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct - @tps_bind [ /2 width=5/ | /3 width=5/ ] (**) (* explicit constructor *) -| #K #I #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #d #e #HLK #H1 #H2 #Hddt - elim (lift_inv_flat1 … H1) -H1 #VV1 #TT1 #HVV1 #HTT1 #H1 - elim (lift_inv_flat1 … H2) -H2 #VV2 #TT2 #HVV2 #HTT2 #H2 destruct /3 width=5/ -] -qed. - -(* Basic_1: was: subst1_gen_lift_lt *) -lemma tps_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt + et ≤ d → - ∃∃T2. K ⊢ T1 ▶ [dt, et] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et -[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ - ] -| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdetd - lapply (lt_to_le_to_lt … Hidet … Hdetd) -Hdetd #Hid - lapply (lift_inv_lref2_lt … H … Hid) -H #H destruct - elim (ldrop_conf_lt … HLK … HLKV ?) -L // #L #U #HKL #_ #HUV - elim (lift_trans_le … HUV … HVW ?) -V // >minus_plus minus_plus plus_minus // commutative_plus >plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) - ] -| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet - elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct - elim (IHV12 … HLK … HWV1 ? ?) -V1 // #W2 #HW12 #HWV2 - elim (IHU12 … HTU1 ? ?) -U1 [5: @ldrop_skip // |2: skip |3: >plus_plus_comm_23 >(plus_plus_comm_23 dt) /2 width=1/ |4: /2 width=1/ ] (**) (* 29s without the rewrites *) - /3 width=5/ -| #L #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdtd #Hdedet - elim (lift_inv_flat2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct - elim (IHV12 … HLK … HWV1 ? ?) -V1 // - elim (IHU12 … HLK … HTU1 ? ?) -U1 -HLK // /3 width=5/ -] -qed. - -(* Basic_1: was: subst1_gen_lift_ge *) -lemma tps_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d + e ≤ dt → - ∃∃T2. K ⊢ T1 ▶ [dt - e, et] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H elim H -L -U1 -U2 -dt -et -[ #L * #i #dt #et #K #d #e #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3/ - ] -| #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #d #e #HLK #T1 #H #Hdedt - lapply (transitive_le … Hdedt … Hdti) #Hdei - elim (le_inv_plus_l … Hdedt) -Hdedt #_ #Hedt - elim (le_inv_plus_l … Hdei) #Hdie #Hei - lapply (lift_inv_lref2_ge … H … Hdei) -H #H destruct - lapply (ldrop_conf_ge … HLK … HLKV ?) -L // #HKV - elim (lift_split … HVW d (i - e + 1) ? ? ?) -HVW [4: // |3: /2 width=1/ |2: /3 width=1/ ] -Hdei -Hdie - #V0 #HV10 >plus_minus // plus_minus // /2 width=1/ ] ] (**) (* explicit constructor, uses monotonic_lt_minus_l *) -| #L #a #I #V1 #V2 #U1 #U2 #dt #et #_ #_ #IHV12 #IHU12 #K #d #e #HLK #X #H #Hdetd - elim (lift_inv_bind2 … H) -H #W1 #T1 #HWV1 #HTU1 #H destruct - elim (le_inv_plus_l … Hdetd) #_ #Hedt - elim (IHV12 … HLK … HWV1 ?) -V1 // #W2 #HW12 #HWV2 - elim (IHU12 … HTU1 ?) -U1 [4: @ldrop_skip // |2: skip |3: /2 width=1/ ] - IHV12 // >IHT12 // -| #L #I #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #X #HX - elim (lift_inv_flat2 … HX) -HX #V #T #HV1 #HT1 #H destruct - >IHV12 // >IHT12 // -] -qed. -(* - Theorem subst0_gen_lift_rev_ge: (t1,v,u2,i,h,d:?) - (subst0 i v t1 (lift h d u2)) -> - (le (plus d h) i) -> - (EX u1 | (subst0 (minus i h) v u1 u2) & - t1 = (lift h d u1) - ). - - - Theorem subst0_gen_lift_rev_lelt: (t1,v,u2,i,h,d:?) - (subst0 i v t1 (lift h d u2)) -> - (le d i) -> (lt i (plus d h)) -> - (EX u1 | t1 = (lift (minus (plus d h) (S i)) (S i) u1)). -*) -lemma tps_inv_lift1_ge_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶ [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 ▶ [d, dt + et - (d + e)] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet -elim (tps_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 -lapply (tps_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 -lapply (tps_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -elim (tps_inv_lift1_ge … HU2 … HLK … HTU1 ?) -U -L // commutative_plus /2 width=1/ ] -Hdtd #T #HT1 #HTU -lapply (tps_weak … HU2 d e ? ?) -HU2 // [ >commutative_plus (lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3/ - ] -| #L #a #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX - elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V1) ?) -HT02 /2 width=1/ #HT02 - elim (IHV01 … HV02) -V0 #V #HV1 #HV2 - elim (IHT01 … HT02) -T0 #T #HT1 #HT2 - lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ - lapply (tps_lsubr_trans … HT2 (L. ⓑ{I} V) ?) -HT2 /3 width=5/ -| #L #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX - elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02) -V0 - elim (IHT01 … HT02) -T0 /3 width=5/ -] -qed. - -(* Basic_1: was: subst1_confluence_neq *) -theorem tps_conf_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 ▶ [d1, e1] T1 → - ∀L2,T2,d2,e2. L2 ⊢ T0 ▶ [d2, e2] T2 → - (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → - ∃∃T. L2 ⊢ T1 ▶ [d2, e2] T & L1 ⊢ T2 ▶ [d1, e1] T. -#L1 #T0 #T1 #d1 #e1 #H elim H -L1 -T0 -T1 -d1 -e1 -[ /2 width=3/ -| #L1 #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #L2 #T2 #d2 #e2 #H1 #H2 - elim (tps_inv_lref1 … H1) -H1 - [ #H destruct /3 width=6/ - | -HLK1 -HVT1 * #K2 #V2 #Hd2 #Hde2 #_ #_ elim H2 -H2 #Hded - [ -Hd1 -Hde2 - lapply (transitive_le … Hded Hd2) -Hded -Hd2 #H - lapply (lt_to_le_to_lt … Hde1 H) -Hde1 -H #H - elim (lt_refl_false … H) - | -Hd2 -Hde1 - lapply (transitive_le … Hded Hd1) -Hded -Hd1 #H - lapply (lt_to_le_to_lt … Hde2 H) -Hde2 -H #H - elim (lt_refl_false … H) - ] - ] -| #L1 #a #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H - elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02 H) -V0 #V #HV1 #HV2 - elim (IHT01 … HT02 ?) -T0 - [ -H #T #HT1 #HT2 - lapply (tps_lsubr_trans … HT1 (L2. ⓑ{I} V) ?) -HT1 /2 width=1/ - lapply (tps_lsubr_trans … HT2 (L1. ⓑ{I} V) ?) -HT2 /2 width=1/ /3 width=5/ - | -HV1 -HV2 >plus_plus_comm_23 >plus_plus_comm_23 in ⊢ (? ? %); elim H -H #H - [ @or_introl | @or_intror ] /2 by monotonic_le_plus_l/ (**) (* /3 / is too slow *) - ] -| #L1 #I #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H - elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02 H) -V0 - elim (IHT01 … HT02 H) -T0 -H /3 width=5/ -] -qed. - -(* Note: the constant 1 comes from tps_subst *) -(* Basic_1: was: subst1_trans *) -theorem tps_trans_ge: ∀L,T1,T0,d,e. L ⊢ T1 ▶ [d, e] T0 → - ∀T2. L ⊢ T0 ▶ [d, 1] T2 → 1 ≤ e → - L ⊢ T1 ▶ [d, e] T2. -#L #T1 #T0 #d #e #H elim H -L -T1 -T0 -d -e -[ #L #I #d #e #T2 #H #He - elim (tps_inv_atom1 … H) -H - [ #H destruct // - | * #K #V #i #Hd2i #Hide2 #HLK #HVT2 #H destruct - lapply (lt_to_le_to_lt … (d + e) Hide2 ?) /2 width=4/ - ] -| #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He - lapply (tps_weak … HVT2 0 (i +1) ? ?) -HVT2 /2 width=1/ #HVT2 - <(tps_inv_lift1_eq … HVT2 … HVW) -HVT2 /2 width=4/ -| #L #a #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He - elim (tps_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct - lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 - lapply (IHT10 … HT02 He) -T0 #HT12 - lapply (tps_lsubr_trans … HT12 (L. ⓑ{I} V2) ?) -HT12 /2 width=1/ /3 width=1/ -| #L #I #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He - elim (tps_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1/ -] -qed. - -theorem tps_trans_down: ∀L,T1,T0,d1,e1. L ⊢ T1 ▶ [d1, e1] T0 → - ∀T2,d2,e2. L ⊢ T0 ▶ [d2, e2] T2 → d2 + e2 ≤ d1 → - ∃∃T. L ⊢ T1 ▶ [d2, e2] T & L ⊢ T ▶ [d1, e1] T2. -#L #T1 #T0 #d1 #e1 #H elim H -L -T1 -T0 -d1 -e1 -[ /2 width=3/ -| #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1 - lapply (transitive_le … Hde2d1 Hdi1) -Hde2d1 #Hde2i1 - lapply (tps_weak … HWT2 0 (i1 + 1) ? ?) -HWT2 normalize /2 width=1/ -Hde2i1 #HWT2 - <(tps_inv_lift1_eq … HWT2 … HVW) -HWT2 /3 width=8/ -| #L #a #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 - elim (tps_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - lapply (tps_lsubr_trans … HT02 (L. ⓑ{I} V0) ?) -HT02 /2 width=1/ #HT02 - elim (IHV10 … HV02 ?) -IHV10 -HV02 // #V - elim (IHT10 … HT02 ?) -T0 /2 width=1/ #T #HT1 #HT2 - lapply (tps_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ - lapply (tps_lsubr_trans … HT2 (L. ⓑ{I} V2) ?) -HT2 /2 width=1/ /3 width=6/ -| #L #I #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 - elim (tps_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV10 … HV02 ?) -V0 // - elim (IHT10 … HT02 ?) -T0 // /3 width=6/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/delift.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/delift.ma deleted file mode 100644 index 3cd3241bc..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/delift.ma +++ /dev/null @@ -1,108 +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/unfold/tpss.ma". - -(* INVERSE BASIC TERM RELOCATION *******************************************) - -definition delift: nat → nat → lenv → relation term ≝ - λd,e,L,T1,T2. ∃∃T. L ⊢ T1 ▶* [d, e] T & ⇧[d, e] T2 ≡ T. - -interpretation "inverse basic relocation (term)" - 'TSubst L T1 d e T2 = (delift d e L T1 T2). - -(* Basic properties *********************************************************) - -lemma lift_delift: ∀T1,T2,d,e. ⇧[d, e] T1 ≡ T2 → - ∀L. L ⊢ ▼*[d, e] T2 ≡ T1. -/2 width=3/ qed. - -lemma delift_refl_O2: ∀L,T,d. L ⊢ ▼*[d, 0] T ≡ T. -/2 width=3/ qed. - -lemma delift_lsubr_trans: ∀L1,T1,T2,d,e. L1 ⊢ ▼*[d, e] T1 ≡ T2 → - ∀L2. L2 ⊑ [d, e] L1 → L2 ⊢ ▼*[d, e] T1 ≡ T2. -#L1 #T1 #T2 #d #e * /3 width=3/ -qed. - -lemma delift_sort: ∀L,d,e,k. L ⊢ ▼*[d, e] ⋆k ≡ ⋆k. -/2 width=3/ qed. - -lemma delift_lref_lt: ∀L,d,e,i. i < d → L ⊢ ▼*[d, e] #i ≡ #i. -/3 width=3/ qed. - -lemma delift_lref_ge: ∀L,d,e,i. d + e ≤ i → L ⊢ ▼*[d, e] #i ≡ #(i - e). -/3 width=3/ qed. - -lemma delift_gref: ∀L,d,e,p. L ⊢ ▼*[d, e] §p ≡ §p. -/2 width=3/ qed. - -lemma delift_bind: ∀a,I,L,V1,V2,T1,T2,d,e. - L ⊢ ▼*[d, e] V1 ≡ V2 → L. ⓑ{I} V2 ⊢ ▼*[d+1, e] T1 ≡ T2 → - L ⊢ ▼*[d, e] ⓑ{a,I} V1. T1 ≡ ⓑ{a,I} V2. T2. -#a #I #L #V1 #V2 #T1 #T2 #d #e * #V #HV1 #HV2 * #T #HT1 #HT2 -lapply (tpss_lsubr_trans … HT1 (L. ⓑ{I} V) ?) -HT1 /2 width=1/ /3 width=5/ -qed. - -lemma delift_flat: ∀I,L,V1,V2,T1,T2,d,e. - L ⊢ ▼*[d, e] V1 ≡ V2 → L ⊢ ▼*[d, e] T1 ≡ T2 → - L ⊢ ▼*[d, e] ⓕ{I} V1. T1 ≡ ⓕ{I} V2. T2. -#I #L #V1 #V2 #T1 #T2 #d #e * #V #HV1 #HV2 * /3 width=5/ -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma delift_inv_sort1: ∀L,U2,d,e,k. L ⊢ ▼*[d, e] ⋆k ≡ U2 → U2 = ⋆k. -#L #U2 #d #e #k * #U #HU ->(tpss_inv_sort1 … HU) -HU #HU2 ->(lift_inv_sort2 … HU2) -HU2 // -qed-. - -lemma delift_inv_gref1: ∀L,U2,d,e,p. L ⊢ ▼*[d, e] §p ≡ U2 → U2 = §p. -#L #U #d #e #p * #U #HU ->(tpss_inv_gref1 … HU) -HU #HU2 ->(lift_inv_gref2 … HU2) -HU2 // -qed-. - -lemma delift_inv_bind1: ∀a,I,L,V1,T1,U2,d,e. L ⊢ ▼*[d, e] ⓑ{a,I} V1. T1 ≡ U2 → - ∃∃V2,T2. L ⊢ ▼*[d, e] V1 ≡ V2 & - L. ⓑ{I} V2 ⊢ ▼*[d+1, e] T1 ≡ T2 & - U2 = ⓑ{a,I} V2. T2. -#a #I #L #V1 #T1 #U2 #d #e * #U #HU #HU2 -elim (tpss_inv_bind1 … HU) -HU #V #T #HV1 #HT1 #X destruct -elim (lift_inv_bind2 … HU2) -HU2 #V2 #T2 #HV2 #HT2 -lapply (tpss_lsubr_trans … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ -qed-. - -lemma delift_inv_flat1: ∀I,L,V1,T1,U2,d,e. L ⊢ ▼*[d, e] ⓕ{I} V1. T1 ≡ U2 → - ∃∃V2,T2. L ⊢ ▼*[d, e] V1 ≡ V2 & - L ⊢ ▼*[d, e] T1 ≡ T2 & - U2 = ⓕ{I} V2. T2. -#I #L #V1 #T1 #U2 #d #e * #U #HU #HU2 -elim (tpss_inv_flat1 … HU) -HU #V #T #HV1 #HT1 #X destruct -elim (lift_inv_flat2 … HU2) -HU2 /3 width=5/ -qed-. - -lemma delift_inv_refl_O2: ∀L,T1,T2,d. L ⊢ ▼*[d, 0] T1 ≡ T2 → T1 = T2. -#L #T1 #T2 #d * #T #HT1 ->(tpss_inv_refl_O2 … HT1) -HT1 #HT2 ->(lift_inv_refl_O2 … HT2) -HT2 // -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma delift_fwd_tw: ∀L,T1,T2,d,e. L ⊢ ▼*[d, e] T1 ≡ T2 → ♯{T1} ≤ ♯{T2}. -#L #T1 #T2 #d #e * #T #HT1 #HT2 ->(lift_fwd_tw … HT2) -T2 /2 width=4 by tpss_fwd_tw/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_alt.ma deleted file mode 100644 index b9322cd1e..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_alt.ma +++ /dev/null @@ -1,100 +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/unfold/delift_lift.ma". - -(* INVERSE BASIC TERM RELOCATION *******************************************) - -(* alternative definition of inverse basic term relocation *) -inductive delifta: nat → nat → lenv → relation term ≝ -| delifta_sort : ∀L,d,e,k. delifta d e L (⋆k) (⋆k) -| delifta_lref_lt: ∀L,d,e,i. i < d → delifta d e L (#i) (#i) -| delifta_lref_be: ∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e → - ⇩[0, i] L ≡ K. ⓓV1 → delifta 0 (d + e - i - 1) K V1 V2 → - ⇧[0, d] V2 ≡ W2 → delifta d e L (#i) W2 -| delifta_lref_ge: ∀L,d,e,i. d + e ≤ i → delifta d e L (#i) (#(i - e)) -| delifta_gref : ∀L,d,e,p. delifta d e L (§p) (§p) -| delifta_bind : ∀L,a,I,V1,V2,T1,T2,d,e. - delifta d e L V1 V2 → delifta (d + 1) e (L. ⓑ{I} V2) T1 T2 → - delifta d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) -| delifta_flat : ∀L,I,V1,V2,T1,T2,d,e. - delifta d e L V1 V2 → delifta d e L T1 T2 → - delifta d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) -. - -interpretation "inverse basic relocation (term) alternative" - 'TSubstAlt L T1 d e T2 = (delifta d e L T1 T2). - -(* Basic properties *********************************************************) - -lemma delifta_lsubr_trans: ∀L1,T1,T2,d,e. L1 ⊢ ▼▼*[d, e] T1 ≡ T2 → - ∀L2. L2 ⊑ [d, e] L1 → L2 ⊢ ▼▼*[d, e] T1 ≡ T2. -#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e // /2 width=1/ -[ #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 - elim (ldrop_lsubr_ldrop2_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /3 width=6/ -| /4 width=1/ -| /3 width=1/ -] -qed. - -lemma delift_delifta: ∀L,T1,T2,d,e. L ⊢ ▼*[d, e] T1 ≡ T2 → L ⊢ ▼▼*[d, e] T1 ≡ T2. -#L #T1 @(f2_ind … fw … L T1) -L -T1 #n #IH #L * -[ * #i #Hn #T2 #d #e #H destruct - [ >(delift_inv_sort1 … H) -H // - | elim (delift_inv_lref1 … H) -H * /2 width=1/ - #K #V1 #V2 #Hdi #Hide #HLK #HV12 #HVT2 - lapply (ldrop_pair2_fwd_fw … HLK) #H - lapply (IH … HV12) // -H /2 width=6/ - | >(delift_inv_gref1 … H) -H // - ] -| * [ #a ] #I #V1 #T1 #Hn #X #d #e #H - [ elim (delift_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - lapply (delift_lsubr_trans … HT12 (L.ⓑ{I}V1) ?) -HT12 /2 width=1/ #HT12 - lapply (IH … HV12) -HV12 // #HV12 - lapply (IH … HT12) -IH -HT12 /2 width=1/ #HT12 - lapply (delifta_lsubr_trans … HT12 (L.ⓑ{I}V2) ?) -HT12 /2 width=1/ - | elim (delift_inv_flat1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - lapply (IH … HV12) -HV12 // - lapply (IH … HT12) -IH -HT12 // /2 width=1/ - ] -] -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma delifta_delift: ∀L,T1,T2,d,e. L ⊢ ▼▼*[d, e] T1 ≡ T2 → L ⊢ ▼*[d, e] T1 ≡ T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e // /2 width=1/ /2 width=6/ -qed-. - -lemma delift_ind_alt: ∀R:ℕ→ℕ→lenv→relation term. - (∀L,d,e,k. R d e L (⋆k) (⋆k)) → - (∀L,d,e,i. i < d → R d e L (#i) (#i)) → - (∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e → - ⇩[O, i] L ≡ K.ⓓV1 → K ⊢ ▼*[O, d + e - i - 1] V1 ≡ V2 → - ⇧[O, d] V2 ≡ W2 → R O (d+e-i-1) K V1 V2 → R d e L (#i) W2 - ) → - (∀L,d,e,i. d + e ≤ i → R d e L (#i) (#(i - e))) → - (∀L,d,e,p. R d e L (§p) (§p)) → - (∀L,a,I,V1,V2,T1,T2,d,e. L ⊢ ▼*[d, e] V1 ≡ V2 → - L.ⓑ{I}V2 ⊢ ▼*[d + 1, e] T1 ≡ T2 → R d e L V1 V2 → - R (d+1) e (L.ⓑ{I}V2) T1 T2 → R d e L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) - ) → - (∀L,I,V1,V2,T1,T2,d,e. L ⊢ ▼*[d, e] V1 ≡ V2 → - L⊢ ▼*[d, e] T1 ≡ T2 → R d e L V1 V2 → - R d e L T1 T2 → R d e L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) - ) → - ∀d,e,L,T1,T2. L ⊢ ▼*[d, e] T1 ≡ T2 → R d e L T1 T2. -#R #H1 #H2 #H3 #H4 #H5 #H6 #H7 #d #e #L #T1 #T2 #H elim (delift_delifta … H) -L -T1 -T2 -d -e -// /2 width=1 by delifta_delift/ /3 width=1 by delifta_delift/ /3 width=7 by delifta_delift/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_delift.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_delift.ma deleted file mode 100644 index a5c563565..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_delift.ma +++ /dev/null @@ -1,29 +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/unfold/tpss_tpss.ma". -include "basic_2/unfold/delift.ma". - -(* INVERSE BASIC TERM RELOCATION *******************************************) - -(* Main properties **********************************************************) - -theorem delift_mono: ∀L,T,T1,T2,d,e. - L ⊢ ▼*[d, e] T ≡ T1 → L ⊢ ▼*[d, e] T ≡ T2 → T1 = T2. -#L #T #T1 #T2 #d #e * #U1 #H1TU1 #H2TU1 * #U2 #H1TU2 #H2TU2 -elim (tpss_conf_eq … H1TU1 … H1TU2) -T #U #HU1 #HU2 -lapply (tpss_inv_lift1_eq … HU1 … H2TU1) -HU1 #H destruct -lapply (tpss_inv_lift1_eq … HU2 … H2TU2) -HU2 #H destruct -lapply (lift_inj … H2TU1 … H2TU2) // -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_lift.ma deleted file mode 100644 index 8abaeb9f9..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/delift_lift.ma +++ /dev/null @@ -1,164 +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/ldrop_lbotr.ma". -include "basic_2/unfold/tpss_lift.ma". -include "basic_2/unfold/delift.ma". - -(* INVERSE BASIC TERM RELOCATION *******************************************) - -(* Advanced properties ******************************************************) - -lemma delift_lref_be: ∀L,K,V1,V2,U2,i,d,e. d ≤ i → i < d + e → - ⇩[0, i] L ≡ K. ⓓV1 → K ⊢ ▼*[0, d + e - i - 1] V1 ≡ V2 → - ⇧[0, d] V2 ≡ U2 → L ⊢ ▼*[d, e] #i ≡ U2. -#L #K #V1 #V2 #U2 #i #d #e #Hdi #Hide #HLK * #V #HV1 #HV2 #HVU2 -elim (lift_total V 0 (i+1)) #U #HVU -lapply (lift_trans_be … HV2 … HVU ? ?) -HV2 // >minus_plus commutative_plus in ⊢ (??%??→?); H -H /2 width=1/ ] -Hde -H #V2 #V12 (**) (* H erased two times *) - elim (lift_total V2 0 d) /3 width=7/ -| #a #I #V1 #T1 #H #d #e #Hde #HL destruct - elim (IH … V1 … Hde HL) // #V2 #HV12 - elim (IH (L.ⓑ{I}V1) T1 … (d+1) e ??) -IH // [2,3: /2 width=1/ ] -Hde -HL #T2 #HT12 - lapply (delift_lsubr_trans … HT12 (L.ⓑ{I}V2) ?) -HT12 /2 width=1/ /3 width=4/ -| #I #V1 #T1 #H #d #e #Hde #HL destruct - elim (IH … V1 … Hde HL) // #V2 #HV12 - elim (IH … T1 … Hde HL) -IH -Hde -HL // /3 width=2/ -] -qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma delift_inv_lref1_lt: ∀L,U2,i,d,e. L ⊢ ▼*[d, e] #i ≡ U2 → i < d → U2 = #i. -#L #U2 #i #d #e * #U #HU #HU2 #Hid -elim (tpss_inv_lref1 … HU) -HU -[ #H destruct >(lift_inv_lref2_lt … HU2) // -| * #K #V1 #V2 #Hdi - lapply (lt_to_le_to_lt … Hid Hdi) -Hid -Hdi #Hi - elim (lt_refl_false … Hi) -] -qed-. - -lemma delift_inv_lref1_be: ∀L,U2,d,e,i. L ⊢ ▼*[d, e] #i ≡ U2 → - d ≤ i → i < d + e → - ∃∃K,V1,V2. ⇩[0, i] L ≡ K. ⓓV1 & - K ⊢ ▼*[0, d + e - i - 1] V1 ≡ V2 & - ⇧[0, d] V2 ≡ U2. -#L #U2 #d #e #i * #U #HU #HU2 #Hdi #Hide -elim (tpss_inv_lref1 … HU) -HU -[ #H destruct elim (lift_inv_lref2_be … HU2 ? ?) // -| * #K #V1 #V #_ #_ #HLK #HV1 #HVU - elim (lift_div_be … HVU … HU2 ? ?) -U // /2 width=1/ /3 width=6/ -] -qed-. - -lemma delift_inv_lref1_ge: ∀L,U2,i,d,e. L ⊢ ▼*[d, e] #i ≡ U2 → - d + e ≤ i → U2 = #(i - e). -#L #U2 #i #d #e * #U #HU #HU2 #Hdei -elim (tpss_inv_lref1 … HU) -HU -[ #H destruct >(lift_inv_lref2_ge … HU2) // -| * #K #V1 #V2 #_ #Hide - lapply (lt_to_le_to_lt … Hide Hdei) -Hide -Hdei #Hi - elim (lt_refl_false … Hi) -] -qed-. - -lemma delift_inv_lref1: ∀L,U2,i,d,e. L ⊢ ▼*[d, e] #i ≡ U2 → - ∨∨ (i < d ∧ U2 = #i) - | (∃∃K,V1,V2. d ≤ i & i < d + e & - ⇩[0, i] L ≡ K. ⓓV1 & - K ⊢ ▼*[0, d + e - i - 1] V1 ≡ V2 & - ⇧[0, d] V2 ≡ U2 - ) - | (d + e ≤ i ∧ U2 = #(i - e)). -#L #U2 #i #d #e #H -elim (lt_or_ge i d) #Hdi -[ elim (delift_inv_lref1_lt … H Hdi) -H /3 width=1/ -| elim (lt_or_ge i (d+e)) #Hide - [ elim (delift_inv_lref1_be … H Hdi Hide) -H /3 width=6/ - | elim (delift_inv_lref1_ge … H Hide) -H /3 width=1/ - ] -] -qed-. - -(* Properties on basic term relocation **************************************) - -lemma delift_lift_le: ∀K,T1,T2,dt,et. K ⊢ ▼*[dt, et] T1 ≡ T2 → - ∀L,U1,d,e. dt + et ≤ d → ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d - et, e] T2 ≡ U2 → - L ⊢ ▼*[dt, et] U1 ≡ U2. -#K #T1 #T2 #dt #et * #T #HT1 #HT2 #L #U1 #d #e #Hdetd #HLK #HTU1 #U2 #HTU2 -elim (lift_total T d e) #U #HTU -lapply (tpss_lift_le … HT1 … HLK HTU1 … HTU) -T1 -HLK // #HU1 -elim (lift_trans_ge … HT2 … HTU ?) -T // -Hdetd #T #HT2 #HTU ->(lift_mono … HTU2 … HT2) -T2 /2 width=3/ -qed. - -lemma delift_lift_be: ∀K,T1,T2,dt,et. K ⊢ ▼*[dt, et] T1 ≡ T2 → - ∀L,U1,d,e. dt ≤ d → d ≤ dt + et → - ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → - L ⊢ ▼*[dt, et + e] U1 ≡ T2. -#K #T1 #T2 #dt #et * #T #HT1 #HT2 #L #U1 #d #e #Hdtd #Hddet #HLK #HTU1 -elim (lift_total T d e) #U #HTU -lapply (tpss_lift_be … HT1 … HLK HTU1 … HTU) -T1 -HLK // #HU1 -lapply (lift_trans_be … HT2 … HTU ? ?) -T // -Hdtd -Hddet /2 width=3/ -qed. - -lemma delift_lift_ge: ∀K,T1,T2,dt,et. K ⊢ ▼*[dt, et] T1 ≡ T2 → - ∀L,U1,d,e. d ≤ dt → ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → - L ⊢ ▼*[dt + e, et] U1 ≡ U2. -#K #T1 #T2 #dt #et * #T #HT1 #HT2 #L #U1 #d #e #Hddt #HLK #HTU1 #U2 #HTU2 -elim (lift_total T d e) #U #HTU -lapply (tpss_lift_ge … HT1 … HLK HTU1 … HTU) -T1 -HLK // #HU1 -elim (lift_trans_le … HT2 … HTU ?) -T // -Hddt #T #HT2 #HTU ->(lift_mono … HTU2 … HT2) -T2 /2 width=3/ -qed. - -lemma delift_inv_lift1_eq: ∀L,U1,T2,d,e. L ⊢ ▼*[d, e] U1 ≡ T2 → - ∀K. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → T1 = T2. -#L #U1 #T2 #d #e * #U2 #HU12 #HTU2 #K #HLK #T1 #HTU1 -lapply (tpss_inv_lift1_eq … HU12 … HTU1) -L -K #H destruct -lapply (lift_inj … HTU1 … HTU2) -U2 // -qed-. - -lemma delift_lift_div_be: ∀L,T1,T,d,e,i. L ⊢ ▼*[i, d + e - i] T1 ≡ T → - ∀T2. ⇧[d, i - d] T2 ≡ T → d ≤ i → i ≤ d + e → - L ⊢ ▼*[d, e] T1 ≡ T2. -#L #T1 #T #d #e #i * #T0 #HT10 #HT0 #T2 #HT2 #Hdi #Hide -lapply (tpss_weak … HT10 d e ? ?) -HT10 // [ >commutative_plus /2 width=1/ ] #HT10 -lapply (lift_trans_be … HT2 … HT0 ? ?) -T // ->commutative_plus >commutative_plus in ⊢ (? ? (? % ?) ? ? → ?); -commutative_plus normalize #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #HV12 #IHL12 #He destruct - >(IHL12 ?) -IHL12 // >(tpss_inv_refl_O2 … HV12) // -] -qed. - -lemma ltpss_dx_inv_refl_O2: ∀d,L1,L2. L1 ▶* [d, 0] L2 → L1 = L2. -/2 width=4/ qed-. - -fact ltpss_dx_inv_atom1_aux: ∀d,e,L1,L2. - L1 ▶* [d, e] L2 → L1 = ⋆ → L2 = ⋆. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| #L #I #V #H destruct -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct -] -qed. - -lemma ltpss_dx_inv_atom1: ∀d,e,L2. ⋆ ▶* [d, e] L2 → L2 = ⋆. -/2 width=5/ qed-. - -fact ltpss_dx_inv_tpss21_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → d = 0 → 0 < e → - ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. K1 ▶* [0, e - 1] K2 & - K2 ⊢ V1 ▶* [0, e - 1] V2 & - L2 = K2. ⓑ{I} V2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #_ #K1 #I #V1 #H destruct -| #L1 #I #V #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K1 #J #W1 #H destruct /2 width=5/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma ltpss_dx_inv_tpss21: ∀e,K1,I,V1,L2. K1. ⓑ{I} V1 ▶* [0, e] L2 → 0 < e → - ∃∃K2,V2. K1 ▶* [0, e - 1] K2 & - K2 ⊢ V1 ▶* [0, e - 1] V2 & - L2 = K2. ⓑ{I} V2. -/2 width=5/ qed-. - -fact ltpss_dx_inv_tpss11_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → 0 < d → - ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. K1 ▶* [d - 1, e] K2 & - K2 ⊢ V1 ▶* [d - 1, e] V2 & - L2 = K2. ⓑ{I} V2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K1 #V1 #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K1 #W1 #H destruct /2 width=5/ -] -qed. - -lemma ltpss_dx_inv_tpss11: ∀d,e,I,K1,V1,L2. K1. ⓑ{I} V1 ▶* [d, e] L2 → 0 < d → - ∃∃K2,V2. K1 ▶* [d - 1, e] K2 & - K2 ⊢ V1 ▶* [d - 1, e] V2 & - L2 = K2. ⓑ{I} V2. -/2 width=3/ qed-. - -fact ltpss_dx_inv_atom2_aux: ∀d,e,L1,L2. - L1 ▶* [d, e] L2 → L2 = ⋆ → L1 = ⋆. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| #L #I #V #H destruct -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct -] -qed. - -lemma ltpss_dx_inv_atom2: ∀d,e,L1. L1 ▶* [d, e] ⋆ → L1 = ⋆. -/2 width=5/ qed-. - -fact ltpss_dx_inv_tpss22_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → d = 0 → 0 < e → - ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 ▶* [0, e - 1] K2 & - K2 ⊢ V1 ▶* [0, e - 1] V2 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #_ #K1 #I #V1 #H destruct -| #L1 #I #V #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K2 #J #W2 #H destruct /2 width=5/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma ltpss_dx_inv_tpss22: ∀e,L1,K2,I,V2. L1 ▶* [0, e] K2. ⓑ{I} V2 → 0 < e → - ∃∃K1,V1. K1 ▶* [0, e - 1] K2 & - K2 ⊢ V1 ▶* [0, e - 1] V2 & - L1 = K1. ⓑ{I} V1. -/2 width=5/ qed-. - -fact ltpss_dx_inv_tpss12_aux: ∀d,e,L1,L2. L1 ▶* [d, e] L2 → 0 < d → - ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 ▶* [d - 1, e] K2 & - K2 ⊢ V1 ▶* [d - 1, e] V2 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K2 #V2 #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K2 #W2 #H destruct /2 width=5/ -] -qed. - -lemma ltpss_dx_inv_tpss12: ∀L1,K2,I,V2,d,e. L1 ▶* [d, e] K2. ⓑ{I} V2 → 0 < d → - ∃∃K1,V1. K1 ▶* [d - 1, e] K2 & - K2 ⊢ V1 ▶* [d - 1, e] V2 & - L1 = K1. ⓑ{I} V1. -/2 width=3/ qed-. - -(* Basic properties *********************************************************) - -lemma ltpss_dx_tps2: ∀L1,L2,I,V1,V2,e. - L1 ▶* [0, e] L2 → L2 ⊢ V1 ▶ [0, e] V2 → - L1. ⓑ{I} V1 ▶* [0, e + 1] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_dx_tps1: ∀L1,L2,I,V1,V2,d,e. - L1 ▶* [d, e] L2 → L2 ⊢ V1 ▶ [d, e] V2 → - L1. ⓑ{I} V1 ▶* [d + 1, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_dx_tpss2_lt: ∀L1,L2,I,V1,V2,e. - L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶* [0, e - 1] V2 → - 0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He ->(plus_minus_m_m e 1) /2 width=1/ -qed. - -lemma ltpss_dx_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶* [d - 1, e] V2 → - 0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd ->(plus_minus_m_m d 1) /2 width=1/ -qed. - -lemma ltpss_dx_tps2_lt: ∀L1,L2,I,V1,V2,e. - L1 ▶* [0, e - 1] L2 → L2 ⊢ V1 ▶ [0, e - 1] V2 → - 0 < e → L1. ⓑ{I} V1 ▶* [0, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_dx_tps1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 ▶* [d - 1, e] L2 → L2 ⊢ V1 ▶ [d - 1, e] V2 → - 0 < d → L1. ⓑ{I} V1 ▶* [d, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -(* Basic_1: was by definition: csubst1_refl *) -lemma ltpss_dx_refl: ∀L,d,e. L ▶* [d, e] L. -#L elim L -L // -#L #I #V #IHL * /2 width=1/ * /2 width=1/ -qed. - -lemma ltpss_dx_weak: ∀L1,L2,d1,e1. L1 ▶* [d1, e1] L2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → L1 ▶* [d2, e2] L2. -#L1 #L2 #d1 #e1 #H elim H -L1 -L2 -d1 -e1 // -[ #L1 #L2 #I #V1 #V2 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd2 #Hde2 - lapply (le_n_O_to_eq … Hd2) #H destruct normalize in Hde2; - lapply (lt_to_le_to_lt 0 … Hde2) // #He2 - lapply (le_plus_to_minus_r … Hde2) -Hde2 /3 width=5/ -| #L1 #L2 #I #V1 #V2 #d1 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd21 #Hde12 - >plus_plus_comm_23 in Hde12; #Hde12 - elim (le_to_or_lt_eq 0 d2 ?) // #H destruct - [ lapply (le_plus_to_minus_r … Hde12) -Hde12 plus_plus_comm_23 - /4 width=5 by ltpss_dx_tpss2, tpss_append, tpss_weak, monotonic_le_plus_r/ (**) (* too slow without trace *) -| #K1 #K2 #I #V1 #V2 #d #x #_ #HV12 #IHK12 normalize (ldrop_inv_atom1 … H) -H // -| // -| normalize #K0 #K1 #I #V0 #V1 #e1 #_ #_ #IHK01 #L2 #e2 #H #He12 - elim (le_inv_plus_l … He12) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ -| #K0 #K1 #I #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ -] -qed. - -lemma ltpss_dx_ldrop_trans_ge: ∀L1,L0,d1,e1. L1 ▶* [d1, e1] L0 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → - d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. -#L1 #L0 #d1 #e1 #H elim H -L1 -L0 -d1 -e1 -[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H // -| // -| normalize #K1 #K0 #I #V1 #V0 #e1 #_ #_ #IHK10 #L2 #e2 #H #He12 - elim (le_inv_plus_l … He12) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK10 … HK0L2 ?) -K0 /2 width=1/ -| #K0 #K1 #I #V1 #V0 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK10 … HK0L2 ?) -IHK10 -HK0L2 /2 width=1/ -] -qed. - -lemma ltpss_dx_ldrop_conf_be: ∀L0,L1,d1,e1. L0 ▶* [d1, e1] L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → - ∃∃L. L2 ▶* [0, d1 + e1 - e2] L & ⇩[0, e2] L1 ≡ L. -#L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 -[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H /2 width=3/ -| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 - lapply (le_n_O_to_eq … He2) -He2 #H destruct - lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ -| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #_ #He21 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK01 -He21 destruct plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 #He2de1 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - (ldrop_inv_atom1 … H) -H /2 width=3/ -| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 - lapply (le_n_O_to_eq … He2) -He2 #H destruct - lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ -| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #_ #He21 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK10 -He21 destruct plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 #He2de1 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - (ldrop_inv_atom1 … H) -H /2 width=3/ -| /2 width=3/ -| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #_ #L2 #e2 #H #He2 - lapply (le_n_O_to_eq … He2) -He2 #He2 destruct - lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ -| #K0 #K1 #I #V0 #V1 #d1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #He2d1 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK01 -He2d1 destruct (ldrop_inv_atom1 … H) -H /2 width=3/ -| /2 width=3/ -| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #_ #L2 #e2 #H #He2 - lapply (le_n_O_to_eq … He2) -He2 #He2 destruct - lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ -| #K1 #K0 #I #V1 #V0 #d1 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #He2d1 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK10 -He2d1 destruct (ltpss_dx_inv_atom1 … H) -H /2 width=3/ -| /2 width=3/ -| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #Hd - elim (IHLK1 … HK12 Hd) -K1 -Hd /3 width=5/ -| #L1 #K1 #I #V1 #W1 #d1 #e1 #_ #HWV1 #IHLK1 #X #d2 #e2 #H #Hd12 - elim (le_inv_plus_l … Hd12) -Hd12 #Hd12 #Hd2 - elim (ltpss_dx_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (IHLK1 … HK12 … Hd12) -IHLK1 -HK12 (ltpss_dx_inv_atom1 … H) -H /2 width=3/ -| #K1 #I #V1 #K2 #d2 #e2 #HK12 #H #_ - lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/ -| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #H1 #H2 - elim (IHLK1 … HK12 H1 H2) -K1 -H2 - lapply (le_n_O_to_eq … H1) -H1 #H destruct /3 width=5/ -| #L1 #K1 #I #V1 #W1 #d1 #e1 #_ #HWV1 #IHLK1 #X #d2 #e2 #H #Hd21 #Hd12 - elim (eq_or_gt d2) #Hd2 [ -Hd21 elim (eq_or_gt e2) #He2 ] destruct - [ lapply (le_n_O_to_eq … Hd12) -Hd12 plus_minus_commutative // #L2 #HL12 #HLK2 - elim (lift_total W2 d1 e1) #V2 #HWV2 - lapply (tpss_lift_be … HW12 … HLK2 HWV1 … HWV2) -W1 // /2 width=1/ - >plus_minus // >commutative_plus /4 width=5/ - | elim (ltpss_dx_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (IHLK1 … HK12 …) -IHLK1 [2: >plus_minus // ] /2 width=1/ #L2 #HL12 #HLK2 - elim (lift_total W2 d1 e1) #V2 #HWV2 - lapply (tpss_lift_be … HW12 … HLK2 HWV1 … HWV2) -W1 [ >plus_minus // ] /2 width=1/ - >commutative_plus /3 width=5/ - ] -] -qed-. - -lemma ldrop_ltpss_dx_trans_ge: ∀L1,K1,d1,e1. ⇩[d1, e1] L1 ≡ K1 → - ∀K2,d2,e2. K1 ▶* [d2, e2] K2 → d2 + e2 ≤ d1 → - ∃∃L2. L1 ▶* [d2, e2] L2 & ⇩[d1, e1] L2 ≡ K2. -#L1 #K1 #d1 #e1 #H elim H -L1 -K1 -d1 -e1 -[ #d1 #e1 #K2 #d2 #e2 #H #_ - >(ltpss_dx_inv_atom1 … H) -H /2 width=3/ -| #K1 #I #V1 #K2 #d2 #e2 #HK12 #H - elim (plus_le_0 … H) -H #H1 #H2 destruct /2 width=3/ -| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #H - elim (IHLK1 … HK12 H) -K1 - elim (plus_le_0 … H) -H #H1 #H2 destruct #L2 #HL12 - >(ltpss_dx_inv_refl_O2 … HL12) -L1 /3 width=5/ -| #L1 #K1 #I #V1 #W1 #d1 #e1 #HLK1 #HWV1 #IHLK1 #X #d2 #e2 #H #Hd21 - elim (eq_or_gt d2) #Hd2 [ elim (eq_or_gt e2) #He2 ] destruct - [ -IHLK1 -Hd21 <(ltpss_dx_inv_refl_O2 … H) -X /3 width=5/ - | elim (ltpss_dx_inv_tpss21 … H He2) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (IHLK1 … HK12 …) -IHLK1 /2 width=1/ #L2 #HL12 #HLK2 - elim (lift_total W2 d1 e1) #V2 #HWV2 - lapply (tpss_lift_le … HW12 … HLK2 HWV1 … HWV2) -W1 /2 width=1/ /3 width=5/ - | elim (ltpss_dx_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (IHLK1 … HK12 …) -IHLK1 [2: >plus_minus // /2 width=1/ ] #L2 #HL12 #HLK2 - elim (lift_total W2 d1 e1) #V2 #HWV2 - lapply (tpss_lift_le … HW12 … HLK2 HWV1 … HWV2) -W1 [ >plus_minus // /2 width=1/ ] /3 width=5/ - ] -] -qed-. - -(* Properties on supclosure *************************************************) - -lemma fsup_tps_trans_full: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ▶[0,|L2|] U2 → - ∃∃L,U1. L1 ▶*[0,|L|] L & L ⊢ T1 ▶[0,|L|] U1 & ⦃L, U1⦄ ⊃ ⦃L2, T2⦄. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [1,2,3,4,5: /3 width=5/ ] -#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 -elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 -elim (lift_total T d e) #U #HTU -elim (le_or_ge d (|K|)) #Hd -[ elim (ldrop_ltpss_dx_trans_be … HLK1 … HK1 … Hd) // -HLK1 -HK1 #L2 #HL12 #HL2K - lapply (tps_lift_be … HT1 … HL2K … HTU1 HTU … Hd) // -HT1 -HTU1 #HU1 -| elim (ldrop_ltpss_dx_trans_ge … HLK1 … HK1 Hd) -HLK1 -HK1 #L2 #HL12 #HL2K - lapply (tps_lift_le … HT1 … HL2K … HTU1 HTU Hd) -HT1 -HTU1 #HU1 -] -lapply (ltpss_dx_weak_full … HL12) -HL12 #HL12 -lapply (tps_weak_full … HU1) -HU1 #HU1 -@(ex3_2_intro … L2 U) // /2 width=7/ (**) (* explicit constructor: auto /3 width=14/ too slow *) -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_ltpss_dx.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_ltpss_dx.ma deleted file mode 100644 index 74b898d69..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_ltpss_dx.ma +++ /dev/null @@ -1,106 +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/unfold/tpss_tpss.ma". -include "basic_2/unfold/ltpss_dx_tpss.ma". - -(* DX PARTIAL UNFOLD ON LOCAL ENVIRONMENTS **********************************) - -(* Advanced properties ******************************************************) - -lemma ltpss_dx_tpss_conf: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶* [d2, e2] U2 → - ∀L1,d1,e1. L0 ▶* [d1, e1] L1 → - ∃∃T. L1 ⊢ T2 ▶* [d2, e2] T & - L1 ⊢ U2 ▶* [d1, e1] T. -#L0 #T2 #U2 #d2 #e2 #H #L1 #d1 #e1 #HL01 @(tpss_ind … H) -U2 /2 width=3/ -#U #U2 #_ #HU2 * #X2 #HTX2 #HUX2 -elim (ltpss_dx_tps_conf … HU2 … HL01) -L0 #X1 #HUX1 #HU2X1 -elim (tpss_strip_eq … HUX2 … HUX1) -U #X #HX2 #HX1 -lapply (tpss_trans_eq … HU2X1 … HX1) -X1 /3 width=3/ -qed. - -lemma ltpss_dx_tpss_trans_down: ∀L0,L1,T2,U2,d1,e1,d2,e2. d2 + e2 ≤ d1 → - L1 ▶* [d1, e1] L0 → L0 ⊢ T2 ▶* [d2, e2] U2 → - ∃∃T. L1 ⊢ T2 ▶* [d2, e2] T & L0 ⊢ T ▶* [d1, e1] U2. -#L0 #L1 #T2 #U2 #d1 #e1 #d2 #e2 #Hde2d1 #HL10 #H @(tpss_ind … H) -U2 -[ /2 width=3/ -| #U #U2 #_ #HU2 * #T #HT2 #HTU - elim (tpss_strap1_down … HTU … HU2 ?) -U // #U #HTU #HU2 - elim (ltpss_dx_tps_trans … HTU … HL10) -HTU -HL10 #X #HTX #HXU - lapply (tpss_trans_eq … HXU HU2) -U /3 width=3/ -] -qed. - -lemma ltpss_dx_tpss_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 ▶* [d, e] U2 → - ∀L0. L0 ▶* [d, e] L1 → L0 ⊢ T2 ▶* [d, e] U2. -#L1 #T2 @(f2_ind … fw … L1 T2) -L1 -T2 #n #IH #L1 * -[ #I #Hn #W2 #d #e #H #L0 #HL01 destruct - elim (tpss_inv_atom1 … H) -H // * - #K1 #V1 #V2 #i #Hdi #Hide #HLK1 #HV12 #HVW2 #H destruct - lapply (ldrop_fwd_lw … HLK1) #H1 normalize in H1; - elim (ltpss_dx_ldrop_trans_be … HL01 … HLK1 ? ?) -HL01 -HLK1 // /2 width=2/ #X #H #HLK0 - elim (ltpss_dx_inv_tpss22 … H ?) -H /2 width=1/ #K0 #V0 #HK01 #HV01 #H destruct - lapply (tpss_fwd_tw … HV01) #H2 - lapply (transitive_le (♯{K1} + ♯{V0}) … H1) -H1 /2 width=1/ -H2 #H - lapply (tpss_trans_eq … HV01 HV12) -V1 #HV02 - lapply (IH … HV02 … HK01) -IH -HV02 -HK01 - [ normalize /2 width=1/ | /2 width=6/ ] -| * [ #a ] #I #V2 #T2 #Hn #X #d #e #H #L0 #HL0 destruct - [ elim (tpss_inv_bind1 … H) -H #W2 #U2 #HVW2 #HTU2 #H destruct - lapply (tpss_lsubr_trans … HTU2 (L1. ⓑ{I} V2) ?) -HTU2 /2 width=1/ #HTU2 - lapply (IH … HVW2 … HL0) -HVW2 [ /2 width=2/ ] #HVW2 - lapply (IH … HTU2 (L0. ⓑ{I} V2) ?) -IH -HTU2 // /2 width=2/ -HL0 #HTU2 - lapply (tpss_lsubr_trans … HTU2 (L0. ⓑ{I} W2) ?) -HTU2 /2 width=1/ - | elim (tpss_inv_flat1 … H) -H #W2 #U2 #HVW2 #HTU2 #H destruct - lapply (IH … HVW2 … HL0) -HVW2 // - lapply (IH … HTU2 … HL0) -IH -HTU2 // -HL0 /2 width=1/ -] -qed. - -lemma ltpss_dx_tps_trans_eq: ∀L0,L1,T2,U2,d,e. L0 ▶* [d, e] L1 → - L1 ⊢ T2 ▶ [d, e] U2 → L0 ⊢ T2 ▶* [d, e] U2. -/3 width=3/ qed. - -(* Main properties **********************************************************) - -theorem ltpss_dx_conf: ∀L0,L1,d1,e1. L0 ▶* [d1, e1] L1 → - ∀L2,d2,e2. L0 ▶* [d2, e2] L2 → - ∃∃L. L1 ▶* [d2, e2] L & L2 ▶* [d1, e1] L. -#L0 @(f_ind … lw … L0) -L0 #n #IH * -[ #_ #L1 #d1 #e1 #H1 #L2 #d2 #e2 #H2 -n - >(ltpss_dx_inv_atom1 … H1) -L1 - >(ltpss_dx_inv_atom1 … H2) -L2 /2 width=3/ -| #K0 #I0 #V0 #Hn #L1 #d1 #e1 #H1 #L2 #d2 #e2 #H2 destruct - elim (eq_or_gt d1) #Hd1 [ elim (eq_or_gt e1) #He1 ] destruct - [ lapply (ltpss_dx_inv_refl_O2 … H1) -H1 #H1 - | elim (ltpss_dx_inv_tpss21 … H1 … He1) -H1 #K1 #V1 #HK01 #HV01 #H1 - | elim (ltpss_dx_inv_tpss11 … H1 … Hd1) -H1 #K1 #V1 #HK01 #HV01 #H1 - ] destruct - elim (eq_or_gt d2) #Hd2 [1,3,5: elim (eq_or_gt e2) #He2 ] destruct - [1,3,5: lapply (ltpss_dx_inv_refl_O2 … H2) -H2 #H2 - |2,4,6: elim (ltpss_dx_inv_tpss21 … H2 … He2) -H2 #K2 #V2 #HK02 #HV02 #H2 - |7,8,9: elim (ltpss_dx_inv_tpss11 … H2 … Hd2) -H2 #K2 #V2 #HK02 #HV02 #H2 - ] destruct - [1: -IH /2 width=3/ - |2,3,4,7: -IH /3 width=5/ - |5,6,8,9: - elim (IH … HK01 … HK02) // -K0 #K #HK1 #HK2 - elim (ltpss_dx_tpss_conf … HV01 … HK1) -HV01 #W1 #HW1 #HVW1 - elim (ltpss_dx_tpss_conf … HV02 … HK2) -HV02 #W2 #HW2 #HVW2 - elim (tpss_conf_eq … HW1 … HW2) -V0 #V #HW1 #HW2 - lapply (tpss_trans_eq … HVW1 HW1) -W1 - lapply (tpss_trans_eq … HVW2 HW2) -W2 /3 width=5/ - ] -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_tps.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_tps.ma deleted file mode 100644 index b10306677..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_tps.ma +++ /dev/null @@ -1,47 +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/unfold/ltpss_dx_ldrop.ma". - -(* DX PARALLEL UNFOLD ON LOCAL ENVIRONMENTS *********************************) - -(* Properties concerning partial substitution on terms **********************) - -lemma ltpss_dx_tps_conf_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → - ∀L1,d1,e1. L0 ▶* [d1, e1] L1 → d1 + e1 ≤ d2 → - L1 ⊢ T2 ▶ [d2, e2] U2. -#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 -[ // -| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL01 #Hde1d2 - lapply (transitive_le … Hde1d2 Hdi2) -Hde1d2 #Hde1i2 - lapply (ltpss_dx_ldrop_conf_ge … HL01 … HLK0 ?) -L0 // /2 width=4/ -| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 #Hde1d2 - @tps_bind [ /2 width=4/ | @IHTU2 -IHTU2 -IHVW2 [3: /2 by ltpss_dx_tpss1/ |1,2: skip | /2 width=1/ ] ] (**) (* explicit constructor *) -| /3 width=4/ -] -qed. - -lemma ltpss_dx_tps_trans_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → - ∀L1,d1,e1. L1 ▶* [d1, e1] L0 → d1 + e1 ≤ d2 → - L1 ⊢ T2 ▶ [d2, e2] U2. -#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 -[ // -| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 #Hde1d2 - lapply (transitive_le … Hde1d2 Hdi2) -Hde1d2 #Hde1i2 - lapply (ltpss_dx_ldrop_trans_ge … HL10 … HLK0 ?) -L0 // /2 width=4/ -| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 #Hde1d2 - @tps_bind [ /2 width=4/ | @IHTU2 -IHTU2 -IHVW2 [3: /2 by ltpss_dx_tpss1/ |1,2: skip | /2 width=1/ ] ] (**) (* explicit constructor *) -| /3 width=4/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_tpss.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_tpss.ma deleted file mode 100644 index 9c7875927..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_dx_tpss.ma +++ /dev/null @@ -1,101 +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/unfold/ltpss_dx_tps.ma". - -(* DX PARALLEL UNFOLD ON LOCAL ENVIRONMENTS *********************************) - -(* Properties concerning partial unfold on terms ****************************) - -lemma ltpss_dx_tpss_conf_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶* [d2, e2] U2 → - ∀L1,d1,e1. L0 ▶* [d1, e1] L1 → d1 + e1 ≤ d2 → - L1 ⊢ T2 ▶* [d2, e2] U2. -#L0 #T2 #U2 #d2 #e2 #H #L1 #d1 #e1 #HL01 #Hde1d2 @(tpss_ind … H) -U2 // -#U #U2 #_ #HU2 #IHU -lapply (ltpss_dx_tps_conf_ge … HU2 … HL01 ?) -L0 // -Hde1d2 /2 width=3/ -qed. - -(* Basic_1: was: subst1_subst1_back *) -lemma ltpss_dx_tps_conf: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → - ∀L1,d1,e1. L0 ▶* [d1, e1] L1 → - ∃∃T. L1 ⊢ T2 ▶ [d2, e2] T & - L1 ⊢ U2 ▶* [d1, e1] T. -#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 -[ /2 width=3/ -| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL01 - elim (lt_or_ge i2 d1) #Hi2d1 - [ elim (ltpss_dx_ldrop_conf_le … HL01 … HLK0 ?) -L0 /2 width=2/ #X #H #HLK1 - elim (ltpss_dx_inv_tpss11 … H ?) -H /2 width=1/ #K1 #V1 #_ #HV01 #H destruct - lapply (ldrop_fwd_ldrop2 … HLK1) #H - elim (lift_total V1 0 (i2 + 1)) #W1 #HVW1 - lapply (tpss_lift_ge … HV01 … H HVW0 … HVW1) -V0 -H // >minus_plus minus_plus >commutative_plus /2 width=1/ - | lapply (ltpss_dx_ldrop_conf_ge … HL01 … HLK0 ?) -L0 // /3 width=4/ - ] - ] -| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 - elim (IHVW2 … HL01) -IHVW2 #V #HV2 #HVW2 - elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL01 /3 width=5/ -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 - elim (IHVW2 … HL01) -IHVW2 - elim (IHTU2 … HL01) -IHTU2 -HL01 /3 width=5/ -] -qed. - -lemma ltpss_dx_tpss_trans_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶* [d2, e2] U2 → - ∀L1,d1,e1. L1 ▶* [d1, e1] L0 → d1 + e1 ≤ d2 → - L1 ⊢ T2 ▶* [d2, e2] U2. -#L0 #T2 #U2 #d2 #e2 #H #L1 #d1 #e1 #HL01 #Hde1d2 @(tpss_ind … H) -U2 // -#U #U2 #_ #HU2 #IHU -lapply (ltpss_dx_tps_trans_ge … HU2 … HL01 ?) -L0 // -Hde1d2 /2 width=3/ -qed. - -(* Basic_1: was: subst1_subst1 *) -lemma ltpss_dx_tps_trans: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → - ∀L1,d1,e1. L1 ▶* [d1, e1] L0 → - ∃∃T. L1 ⊢ T2 ▶ [d2, e2] T & - L0 ⊢ T ▶* [d1, e1] U2. -#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 -[ /2 width=3/ -| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 - elim (lt_or_ge i2 d1) #Hi2d1 - [ elim (ltpss_dx_ldrop_trans_le … HL10 … HLK0 ?) -HL10 /2 width=2/ #X #H #HLK1 - elim (ltpss_dx_inv_tpss12 … H ?) -H /2 width=1/ #K1 #V1 #_ #HV01 #H destruct - lapply (ldrop_fwd_ldrop2 … HLK0) -HLK0 #H - elim (lift_total V1 0 (i2 + 1)) #W1 #HVW1 - lapply (tpss_lift_ge … HV01 … H HVW1 … HVW0) -V0 -H // >minus_plus minus_plus >commutative_plus /2 width=1/ - | lapply (ltpss_dx_ldrop_trans_ge … HL10 … HLK0 ?) -HL10 -HLK0 // /3 width=4/ - ] - ] -| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 - elim (IHVW2 … HL10) -IHVW2 #V #HV2 #HVW2 - elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL10 /3 width=5/ -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 - elim (IHVW2 … HL10) -IHVW2 - elim (IHTU2 … HL10) -IHTU2 -HL10 /3 width=5/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_sn.ma deleted file mode 100644 index 577c1506d..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_sn.ma +++ /dev/null @@ -1,259 +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 *) -(* *) -(**************************************************************************) - -notation "hvbox( T1 break ⊢ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" - non associative with precedence 45 - for @{ 'PSubstStarSn $T1 $d $e $T2 }. - -include "basic_2/unfold/tpss.ma". - -(* SN PARALLEL UNFOLD ON LOCAL ENVIRONMENTS *********************************) - -inductive ltpss_sn: nat → nat → relation lenv ≝ -| ltpss_sn_atom : ∀d,e. ltpss_sn d e (⋆) (⋆) -| ltpss_sn_pair : ∀L,I,V. ltpss_sn 0 0 (L. ⓑ{I} V) (L. ⓑ{I} V) -| ltpss_sn_tpss2: ∀L1,L2,I,V1,V2,e. - ltpss_sn 0 e L1 L2 → L1 ⊢ V1 ▶* [0, e] V2 → - ltpss_sn 0 (e + 1) (L1. ⓑ{I} V1) (L2. ⓑ{I} V2) -| ltpss_sn_tpss1: ∀L1,L2,I,V1,V2,d,e. - ltpss_sn d e L1 L2 → L1 ⊢ V1 ▶* [d, e] V2 → - ltpss_sn (d + 1) e (L1. ⓑ{I} V1) (L2. ⓑ{I} V2) -. - -interpretation "parallel unfold (local environment, sn variant)" - 'PSubstStarSn L1 d e L2 = (ltpss_sn d e L1 L2). - -(* Basic inversion lemmas ***************************************************) - -fact ltpss_sn_inv_refl_O2_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → e = 0 → L1 = L2. -#d #e #L1 #L2 #H elim H -d -e -L1 -L2 // -[ #L1 #L2 #I #V1 #V2 #e #_ #_ #_ >commutative_plus normalize #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #HV12 #IHL12 #He destruct - >(IHL12 ?) -IHL12 // >(tpss_inv_refl_O2 … HV12) // -] -qed. - -lemma ltpss_sn_inv_refl_O2: ∀d,L1,L2. L1 ⊢ ▶* [d, 0] L2 → L1 = L2. -/2 width=4/ qed-. - -fact ltpss_sn_inv_atom1_aux: ∀d,e,L1,L2. - L1 ⊢ ▶* [d, e] L2 → L1 = ⋆ → L2 = ⋆. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| #L #I #V #H destruct -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct -] -qed. - -lemma ltpss_sn_inv_atom1: ∀d,e,L2. ⋆ ⊢ ▶* [d, e] L2 → L2 = ⋆. -/2 width=5/ qed-. - -fact ltpss_sn_inv_tpss21_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → d = 0 → 0 < e → - ∀K1,I,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. K1 ⊢ ▶* [0, e - 1] K2 & - K1 ⊢ V1 ▶* [0, e - 1] V2 & - L2 = K2. ⓑ{I} V2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #_ #K1 #I #V1 #H destruct -| #L1 #I #V #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K1 #J #W1 #H destruct /2 width=5/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma ltpss_sn_inv_tpss21: ∀e,K1,I,V1,L2. K1. ⓑ{I} V1 ⊢ ▶* [0, e] L2 → 0 < e → - ∃∃K2,V2. K1 ⊢ ▶* [0, e - 1] K2 & - K1 ⊢ V1 ▶* [0, e - 1] V2 & - L2 = K2. ⓑ{I} V2. -/2 width=5/ qed-. - -fact ltpss_sn_inv_tpss11_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → 0 < d → - ∀I,K1,V1. L1 = K1. ⓑ{I} V1 → - ∃∃K2,V2. K1 ⊢ ▶* [d - 1, e] K2 & - K1 ⊢ V1 ▶* [d - 1, e] V2 & - L2 = K2. ⓑ{I} V2. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K1 #V1 #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K1 #W1 #H destruct /2 width=5/ -] -qed. - -lemma ltpss_sn_inv_tpss11: ∀d,e,I,K1,V1,L2. K1. ⓑ{I} V1 ⊢ ▶* [d, e] L2 → 0 < d → - ∃∃K2,V2. K1 ⊢ ▶* [d - 1, e] K2 & - K1 ⊢ V1 ▶* [d - 1, e] V2 & - L2 = K2. ⓑ{I} V2. -/2 width=3/ qed-. - -fact ltpss_sn_inv_atom2_aux: ∀d,e,L1,L2. - L1 ⊢ ▶* [d, e] L2 → L2 = ⋆ → L1 = ⋆. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ // -| #L #I #V #H destruct -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H destruct -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ #H destruct -] -qed. - -lemma ltpss_sn_inv_atom2: ∀d,e,L1. L1 ⊢ ▶* [d, e] ⋆ → L1 = ⋆. -/2 width=5/ qed-. - -fact ltpss_sn_inv_tpss22_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → d = 0 → 0 < e → - ∀K2,I,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 ⊢ ▶* [0, e - 1] K2 & - K1 ⊢ V1 ▶* [0, e - 1] V2 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #_ #K1 #I #V1 #H destruct -| #L1 #I #V #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #_ #_ #K2 #J #W2 #H destruct /2 width=5/ -| #L1 #L2 #I #V1 #V2 #d #e #_ #_ >commutative_plus normalize #H destruct -] -qed. - -lemma ltpss_sn_inv_tpss22: ∀e,L1,K2,I,V2. L1 ⊢ ▶* [0, e] K2. ⓑ{I} V2 → 0 < e → - ∃∃K1,V1. K1 ⊢ ▶* [0, e - 1] K2 & - K1 ⊢ V1 ▶* [0, e - 1] V2 & - L1 = K1. ⓑ{I} V1. -/2 width=5/ qed-. - -fact ltpss_sn_inv_tpss12_aux: ∀d,e,L1,L2. L1 ⊢ ▶* [d, e] L2 → 0 < d → - ∀I,K2,V2. L2 = K2. ⓑ{I} V2 → - ∃∃K1,V1. K1 ⊢ ▶* [d - 1, e] K2 & - K1 ⊢ V1 ▶* [d - 1, e] V2 & - L1 = K1. ⓑ{I} V1. -#d #e #L1 #L2 * -d -e -L1 -L2 -[ #d #e #_ #I #K2 #V2 #H destruct -| #L #I #V #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #e #_ #_ #H elim (lt_refl_false … H) -| #L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #_ #J #K2 #W2 #H destruct /2 width=5/ -] -qed. - -lemma ltpss_sn_inv_tpss12: ∀L1,K2,I,V2,d,e. L1 ⊢ ▶* [d, e] K2. ⓑ{I} V2 → 0 < d → - ∃∃K1,V1. K1 ⊢ ▶* [d - 1, e] K2 & - K1 ⊢ V1 ▶* [d - 1, e] V2 & - L1 = K1. ⓑ{I} V1. -/2 width=3/ qed-. - -(* Basic properties *********************************************************) - -lemma ltpss_sn_tps2: ∀L1,L2,I,V1,V2,e. - L1 ⊢ ▶* [0, e] L2 → L1 ⊢ V1 ▶ [0, e] V2 → - L1. ⓑ{I} V1 ⊢ ▶* [0, e + 1] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_sn_tps1: ∀L1,L2,I,V1,V2,d,e. - L1 ⊢ ▶* [d, e] L2 → L1 ⊢ V1 ▶ [d, e] V2 → - L1. ⓑ{I} V1 ⊢ ▶* [d + 1, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_sn_tpss2_lt: ∀L1,L2,I,V1,V2,e. - L1 ⊢ ▶* [0, e - 1] L2 → L1 ⊢ V1 ▶* [0, e - 1] V2 → - 0 < e → L1. ⓑ{I} V1 ⊢ ▶* [0, e] L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #e #HL12 #HV12 #He ->(plus_minus_m_m e 1) /2 width=1/ -qed. - -lemma ltpss_sn_tpss1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 ⊢ ▶* [d - 1, e] L2 → L1 ⊢ V1 ▶* [d - 1, e] V2 → - 0 < d → L1. ⓑ{I} V1 ⊢ ▶* [d, e] L2. ⓑ{I} V2. -#L1 #L2 #I #V1 #V2 #d #e #HL12 #HV12 #Hd ->(plus_minus_m_m d 1) /2 width=1/ -qed. - -lemma ltpss_sn_tps2_lt: ∀L1,L2,I,V1,V2,e. - L1 ⊢ ▶* [0, e - 1] L2 → L1 ⊢ V1 ▶ [0, e - 1] V2 → - 0 < e → L1. ⓑ{I} V1 ⊢ ▶* [0, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_sn_tps1_lt: ∀L1,L2,I,V1,V2,d,e. - L1 ⊢ ▶* [d - 1, e] L2 → L1 ⊢ V1 ▶ [d - 1, e] V2 → - 0 < d → L1. ⓑ{I} V1 ⊢ ▶* [d, e] L2. ⓑ{I} V2. -/3 width=1/ qed. - -lemma ltpss_sn_refl: ∀L,d,e. L ⊢ ▶* [d, e] L. -#L elim L -L // -#L #I #V #IHL * /2 width=1/ * /2 width=1/ -qed. - -lemma ltpss_sn_weak: ∀L1,L2,d1,e1. L1 ⊢ ▶* [d1, e1] L2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → L1 ⊢ ▶* [d2, e2] L2. -#L1 #L2 #d1 #e1 #H elim H -L1 -L2 -d1 -e1 // -[ #L1 #L2 #I #V1 #V2 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd2 #Hde2 - lapply (le_n_O_to_eq … Hd2) #H destruct normalize in Hde2; - lapply (lt_to_le_to_lt 0 … Hde2) // #He2 - lapply (le_plus_to_minus_r … Hde2) -Hde2 /3 width=5/ -| #L1 #L2 #I #V1 #V2 #d1 #e1 #_ #HV12 #IHL12 #d2 #e2 #Hd21 #Hde12 - >plus_plus_comm_23 in Hde12; #Hde12 - elim (le_to_or_lt_eq 0 d2 ?) // #H destruct - [ lapply (le_plus_to_minus_r … Hde12) -Hde12 plus_plus_comm_23 - /4 width=5 by ltpss_sn_tpss2, tpss_append, tpss_weak, monotonic_le_plus_r/ (**) (* too slow without trace *) -| #K1 #K2 #I #V1 #V2 #d #x #_ #HV12 #IHK12 normalize (ldrop_inv_atom1 … H) -H // -| // -| normalize #K0 #K1 #I #V0 #V1 #e1 #_ #_ #IHK01 #L2 #e2 #H #He12 - elim (le_inv_plus_l … He12) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ -| #K0 #K1 #I #V0 #V1 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK01 … HK0L2 ?) -K0 /2 width=1/ -] -qed. - -lemma ltpss_sn_ldrop_trans_ge: ∀L1,L0,d1,e1. L1 ⊢ ▶* [d1, e1] L0 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → - d1 + e1 ≤ e2 → ⇩[0, e2] L1 ≡ L2. -#L1 #L0 #d1 #e1 #H elim H -L1 -L0 -d1 -e1 -[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H // -| // -| normalize #K1 #K0 #I #V1 #V0 #e1 #_ #_ #IHK10 #L2 #e2 #H #He12 - elim (le_inv_plus_l … He12) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK10 … HK0L2 ?) -K0 /2 width=1/ -| #K0 #K1 #I #V1 #V0 #d1 #e1 >plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - lapply (ldrop_inv_ldrop1 … H ?) -H // #HK0L2 - lapply (IHK10 … HK0L2 ?) -IHK10 -HK0L2 /2 width=1/ -] -qed. - -lemma ltpss_sn_ldrop_conf_be: ∀L0,L1,d1,e1. L0 ⊢ ▶* [d1, e1] L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → - ∃∃L. L2 ⊢ ▶* [0, d1 + e1 - e2] L & ⇩[0, e2] L1 ≡ L. -#L0 #L1 #d1 #e1 #H elim H -L0 -L1 -d1 -e1 -[ #d1 #e1 #L2 #e2 #H >(ldrop_inv_atom1 … H) -H /2 width=3/ -| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 - lapply (le_n_O_to_eq … He2) -He2 #H destruct - lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ -| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #_ #He21 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK01 -He21 destruct plus_plus_comm_23 #_ #_ #IHK01 #L2 #e2 #H #Hd1e2 #He2de1 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - (ldrop_inv_atom1 … H) -H /2 width=3/ -| normalize #L #I #V #L2 #e2 #HL2 #_ #He2 - lapply (le_n_O_to_eq … He2) -He2 #H destruct - lapply (ldrop_inv_refl … HL2) -HL2 #H destruct /2 width=3/ -| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #_ #He21 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK10 -He21 destruct plus_plus_comm_23 #_ #_ #IHK10 #L2 #e2 #H #Hd1e2 #He2de1 - elim (le_inv_plus_l … Hd1e2) #_ #He2 - (ldrop_inv_atom1 … H) -H /2 width=3/ -| /2 width=3/ -| normalize #K0 #K1 #I #V0 #V1 #e1 #HK01 #HV01 #_ #L2 #e2 #H #He2 - lapply (le_n_O_to_eq … He2) -He2 #He2 destruct - lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ -| #K0 #K1 #I #V0 #V1 #d1 #e1 #HK01 #HV01 #IHK01 #L2 #e2 #H #He2d1 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK01 -He2d1 destruct (ldrop_inv_atom1 … H) -H /2 width=3/ -| /2 width=3/ -| normalize #K1 #K0 #I #V1 #V0 #e1 #HK10 #HV10 #_ #L2 #e2 #H #He2 - lapply (le_n_O_to_eq … He2) -He2 #He2 destruct - lapply (ldrop_inv_refl … H) -H #H destruct /3 width=3/ -| #K1 #K0 #I #V1 #V0 #d1 #e1 #HK10 #HV10 #IHK10 #L2 #e2 #H #He2d1 - lapply (ldrop_inv_O1 … H) -H * * #He2 #HK0L2 - [ -IHK10 -He2d1 destruct (ltpss_sn_inv_atom1 … H) -H /2 width=3/ -| /2 width=3/ -| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #Hd - elim (IHLK1 … HK12 Hd) -K1 -Hd /3 width=5/ -| #L1 #K1 #I #V1 #W1 #d1 #e1 #HLK1 #HWV1 #IHLK1 #X #d2 #e2 #H #Hd12 - elim (le_inv_plus_l … Hd12) -Hd12 #Hd12 #Hd2 - elim (ltpss_sn_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (IHLK1 … HK12 … Hd12) -IHLK1 -HK12 (ltpss_sn_inv_atom1 … H) -H /2 width=3/ -| #K1 #I #V1 #K2 #d2 #e2 #HK12 #H #_ - lapply (le_n_O_to_eq … H) -H #H destruct /2 width=3/ -| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #H1 #H2 - elim (IHLK1 … HK12 H1 H2) -K1 -H2 - lapply (le_n_O_to_eq … H1) -H1 #H destruct /3 width=5/ -| #L1 #K1 #I #V1 #W1 #d1 #e1 #HLK1 #HWV1 #IHLK1 #X #d2 #e2 #H #Hd21 #Hd12 - elim (eq_or_gt d2) #Hd2 [ -Hd21 elim (eq_or_gt e2) #He2 ] destruct - [ lapply (le_n_O_to_eq … Hd12) -Hd12 plus_minus_commutative // #L2 #HL12 #HLK2 - elim (lift_total W2 d1 e1) #V2 #HWV2 - lapply (tpss_lift_be … HW12 … HLK1 HWV1 … HWV2) -HLK1 -W1 // /2 width=1/ - >plus_minus // >commutative_plus /4 width=5/ - | elim (ltpss_sn_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (IHLK1 … HK12 …) -IHLK1 [2: >plus_minus // ] /2 width=1/ #L2 #HL12 #HLK2 - elim (lift_total W2 d1 e1) #V2 #HWV2 - lapply (tpss_lift_be … HW12 … HLK1 HWV1 … HWV2) -HLK1 -W1 [ >plus_minus // ] /2 width=1/ - >commutative_plus /3 width=5/ - ] -] -qed-. - -lemma ldrop_ltpss_sn_trans_ge: ∀L1,K1,d1,e1. ⇩[d1, e1] L1 ≡ K1 → - ∀K2,d2,e2. K1 ⊢ ▶* [d2, e2] K2 → d2 + e2 ≤ d1 → - ∃∃L2. L1 ⊢ ▶* [d2, e2] L2 & ⇩[d1, e1] L2 ≡ K2. -#L1 #K1 #d1 #e1 #H elim H -L1 -K1 -d1 -e1 -[ #d1 #e1 #K2 #d2 #e2 #H #_ - >(ltpss_sn_inv_atom1 … H) -H /2 width=3/ -| #K1 #I #V1 #K2 #d2 #e2 #HK12 #H - elim (plus_le_0 … H) -H #H1 #H2 destruct /2 width=3/ -| #L1 #K1 #I #V #e1 #_ #IHLK1 #K2 #d2 #e2 #HK12 #H - elim (IHLK1 … HK12 H) -K1 - elim (plus_le_0 … H) -H #H1 #H2 destruct #L2 #HL12 - >(ltpss_sn_inv_refl_O2 … HL12) -L1 /3 width=5/ -| #L1 #K1 #I #V1 #W1 #d1 #e1 #HLK1 #HWV1 #IHLK1 #X #d2 #e2 #H #Hd21 - elim (eq_or_gt d2) #Hd2 [ elim (eq_or_gt e2) #He2 ] destruct - [ -IHLK1 -Hd21 <(ltpss_sn_inv_refl_O2 … H) -X /3 width=5/ - | elim (ltpss_sn_inv_tpss21 … H He2) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (IHLK1 … HK12 …) -IHLK1 /2 width=1/ #L2 #HL12 #HLK2 - elim (lift_total W2 d1 e1) #V2 #HWV2 - lapply (tpss_lift_le … HW12 … HLK1 HWV1 … HWV2) -HLK1 -W1 /2 width=1/ /3 width=5/ - | elim (ltpss_sn_inv_tpss11 … H Hd2) -H #K2 #W2 #HK12 #HW12 #H destruct - elim (IHLK1 … HK12 …) -IHLK1 [2: >plus_minus // /2 width=1/ ] #L2 #HL12 #HLK2 - elim (lift_total W2 d1 e1) #V2 #HWV2 - lapply (tpss_lift_le … HW12 … HLK1 HWV1 … HWV2) -HLK1 -W1 [ >plus_minus // /2 width=1/ ] /3 width=5/ - ] -] -qed-. - -(* Properties on supclosure *************************************************) - -lemma fsup_tpss_trans_full: ∀L1,L2,T1,T2. ⦃L1, T1⦄ ⊃ ⦃L2, T2⦄ → ∀U2. L2 ⊢ T2 ▶*[0,|L2|] U2 → - ∃∃L,U1. L1 ⊢ ▶*[0,|L1|] L & L ⊢ T1 ▶*[0,|L|] U1 & ⦃L, U1⦄ ⊃ ⦃L2, T2⦄. -#L1 #L2 #T1 #T2 #H elim H -L1 -L2 -T1 -T2 [1,2,3,4,5: /3 width=5/ ] -#L1 #K1 #K2 #T1 #T2 #U1 #d #e #HLK1 #HTU1 #_ #IHT12 #U2 #HTU2 -elim (IHT12 … HTU2) -IHT12 -HTU2 #K #T #HK1 #HT1 #HT2 -elim (lift_total T d e) #U #HTU -lapply (ltpss_sn_fwd_length … HK1) #H >H in HK1; -H #HK1 -elim (le_or_ge d (|K|)) #Hd -[ elim (ldrop_ltpss_sn_trans_be … HLK1 … HK1 … Hd) // -HLK1 -HK1 #L2 #HL12 #HL2K - lapply (tpss_lift_be … HT1 … Hd HL2K HTU1 … HTU) // -HT1 -HTU1 #HU1 -| elim (ldrop_ltpss_sn_trans_ge … HLK1 … HK1 Hd) -HLK1 -HK1 #L2 #HL12 #HL2K - lapply (tpss_lift_le … HT1 … Hd HL2K HTU1 … HTU) -HT1 -HTU1 #HU1 -] -lapply (ltpss_sn_weak_full … HL12) -HL12 #HL12 -lapply (tpss_weak_full … HU1) -HU1 #HU1 -@(ex3_2_intro … L2 U) // /2 width=7/ (**) (* explicit constructor: auto /3 width=14/ too slow *) -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_sn_ltpss_sn.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_sn_ltpss_sn.ma deleted file mode 100644 index 67c287e2b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/ltpss_sn_ltpss_sn.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/unfold/tpss_tpss.ma". -include "basic_2/unfold/ltpss_sn_tpss.ma". - -(* PARTIAL UNFOLD ON LOCAL ENVIRONMENTS *************************************) - -(* Advanced properties ******************************************************) - -lemma ltpss_sn_tpss_trans_eq: ∀L1,T2,U2,d,e. L1 ⊢ T2 ▶* [d, e] U2 → - ∀L0. L0 ⊢ ▶* [d, e] L1 → L0 ⊢ T2 ▶* [d, e] U2. -#L1 #T2 @(f2_ind … fw … L1 T2) -L1 -T2 #n #IH #L1 * -[ #I #Hn #W2 #d #e #H #L0 #HL01 destruct - elim (tpss_inv_atom1 … H) -H // * - #K1 #V1 #V2 #i #Hdi #Hide #HLK1 #HV12 #HVW2 #H destruct - lapply (ldrop_fwd_lw … HLK1) #H1 normalize in H1; - elim (ltpss_sn_ldrop_trans_be … HL01 … HLK1 ? ?) -HL01 -HLK1 // /2 width=2/ #X #H #HLK0 - elim (ltpss_sn_inv_tpss22 … H ?) -H /2 width=1/ #K0 #V0 #HK01 #HV01 #H destruct - lapply (IH … HV12 … HK01) -IH -HV12 -HK01 [ normalize /2 width=1/ ] #HV12 - lapply (tpss_trans_eq … HV01 HV12) -V1 /2 width=6/ -| * [ #a ] #I #V2 #T2 #Hn #X #d #e #H #L0 #HL0 destruct - [ elim (tpss_inv_bind1 … H) -H #W2 #U2 #HVW2 #HTU2 #H destruct - lapply (tpss_lsubr_trans … HTU2 (L1. ⓑ{I} V2) ?) -HTU2 /2 width=1/ #HTU2 - lapply (IH … HVW2 … HL0) -HVW2 [ /2 width=2/ ] #HVW2 - lapply (IH … HTU2 (L0. ⓑ{I} V2) ?) -IH -HTU2 // /2 width=2/ -HL0 #HTU2 - lapply (tpss_lsubr_trans … HTU2 (L0. ⓑ{I} W2) ?) -HTU2 /2 width=1/ - | elim (tpss_inv_flat1 … H) -H #W2 #U2 #HVW2 #HTU2 #H destruct - lapply (IH … HVW2 … HL0) -HVW2 // - lapply (IH … HTU2 … HL0) -IH -HTU2 // -HL0 /2 width=1/ -] -qed. - -lemma ltpss_sn_tps_trans_eq: ∀L0,L1,T2,U2,d,e. L0 ⊢ ▶* [d, e] L1 → - L1 ⊢ T2 ▶ [d, e] U2 → L0 ⊢ T2 ▶* [d, e] U2. -/3 width=3/ qed. - -(* Main properties **********************************************************) - -theorem ltpss_sn_trans_eq: ∀L1,L,d,e. L1 ⊢ ▶* [d, e] L → - ∀L2. L ⊢ ▶* [d, e] L2 → L1 ⊢ ▶* [d, e] L2. -#L1 #L #d #e #H elim H -L1 -L -d -e // -[ #L1 #L #I #V1 #V #e #HL1 #HV1 #IHL1 #X #H - elim (ltpss_sn_inv_tpss21 … H ?) -H // shift_append_assoc #H - elim (tps_inv_bind1 … H) -H #V2 #T2 #HV12 #HT12 #H destruct - elim (IH … HT12) -IH -HT12 #L2 #T #HL12 #HT1 #H destruct - append_length minus_plus minus_plus >commutative_plus /2 width=1/ - | lapply (ltpss_sn_ldrop_conf_ge … HL01 … HLK0 ?) -HL01 -HLK0 // /3 width=4/ - ] - ] -| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 - elim (IHVW2 … HL01) -IHVW2 #V #HV2 #HVW2 - elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL01 #T #HT2 #H - lapply (tpss_lsubr_trans … H (L0.ⓑ{I}V) ?) -H /2 width=1/ /3 width=5/ -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL01 - elim (IHVW2 … HL01) -IHVW2 - elim (IHTU2 … HL01) -IHTU2 -HL01 /3 width=5/ -] -qed. - -lemma ltpss_sn_tpss_trans_ge: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶* [d2, e2] U2 → - ∀L1,d1,e1. L1 ⊢ ▶* [d1, e1] L0 → d1 + e1 ≤ d2 → - L1 ⊢ T2 ▶* [d2, e2] U2. -#L0 #T2 #U2 #d2 #e2 #H #L1 #d1 #e1 #HL01 #Hde1d2 @(tpss_ind … H) -U2 // -#U #U2 #_ #HU2 #IHU -lapply (ltpss_sn_tps_trans_ge … HU2 … HL01 ?) -L0 // -Hde1d2 /2 width=3/ -qed. - -lemma ltpss_sn_tps_trans: ∀L0,T2,U2,d2,e2. L0 ⊢ T2 ▶ [d2, e2] U2 → - ∀L1,d1,e1. L1 ⊢ ▶* [d1, e1] L0 → - ∃∃T. L1 ⊢ T2 ▶ [d2, e2] T & - L1 ⊢ T ▶* [d1, e1] U2. -#L0 #T2 #U2 #d2 #e2 #H elim H -L0 -T2 -U2 -d2 -e2 -[ /2 width=3/ -| #L0 #K0 #V0 #W0 #i2 #d2 #e2 #Hdi2 #Hide2 #HLK0 #HVW0 #L1 #d1 #e1 #HL10 - elim (lt_or_ge i2 d1) #Hi2d1 - [ elim (ltpss_sn_ldrop_trans_le … HL10 … HLK0 ?) -L0 /2 width=2/ #X #H #HLK1 - elim (ltpss_sn_inv_tpss12 … H ?) -H /2 width=1/ #K1 #V1 #_ #HV01 #H destruct - lapply (ldrop_fwd_ldrop2 … HLK1) #H - elim (lift_total V1 0 (i2 + 1)) #W1 #HVW1 - lapply (tpss_lift_ge … HV01 … H HVW1 … HVW0) -V0 -H // >minus_plus minus_plus >commutative_plus /2 width=1/ - | lapply (ltpss_sn_ldrop_trans_ge … HL10 … HLK0 ?) -HL10 -HLK0 // /3 width=4/ - ] - ] -| #L0 #a #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 - elim (IHVW2 … HL10) -IHVW2 #V #HV2 #HVW2 - elim (IHTU2 (L1. ⓑ{I} V) (d1 + 1) e1 ?) -IHTU2 /2 width=1/ -HL10 #T #HT2 #H - lapply (tpss_lsubr_trans … H (L1.ⓑ{I}W2) ?) -H /2 width=1/ /3 width=5/ -| #L0 #I #V2 #W2 #T2 #U2 #d2 #e2 #_ #_ #IHVW2 #IHTU2 #L1 #d1 #e1 #HL10 - elim (IHVW2 … HL10) -IHVW2 - elim (IHTU2 … HL10) -IHTU2 -HL10 /3 width=5/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/thin.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/thin.ma deleted file mode 100644 index 65fb76fe0..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/thin.ma +++ /dev/null @@ -1,37 +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/unfold/ltpss_sn.ma". - -(* BASIC LOCAL ENVIRONMENT THINNING *****************************************) - -definition thin: nat → nat → relation lenv ≝ - λd,e,L1,L2. ∃∃L. L1 ⊢ ▶* [d, e] L & ⇩[d, e] L ≡ L2. - -interpretation "basic thinning (local environment)" - 'TSubst L1 d e L2 = (thin d e L1 L2). - -(* Basic properties *********************************************************) - -lemma ldrop_thin: ∀L1,L2,d,e. ⇩[d, e] L1 ≡ L2 → ▼*[d, e] L1 ≡ L2. -/2 width=3/ qed. - -(* Basic inversion lemmas ***************************************************) - -lemma thin_inv_thin1: ∀I,K1,V1,L2,e. ▼*[0, e] K1.ⓑ{I} V1 ≡ L2 → 0 < e → - ▼*[0, e - 1] K1 ≡ L2. -#I #K1 #V1 #L2 #e * #X #HK1 #HL2 #e -elim (ltpss_sn_inv_tpss21 … HK1 ?) -HK1 // #K #V #HK1 #_ #H destruct -lapply (ldrop_inv_ldrop1 … HL2 ?) -HL2 // /2 width=3/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/thin_delift.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/thin_delift.ma deleted file mode 100644 index b5ffc5e4f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/thin_delift.ma +++ /dev/null @@ -1,102 +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/unfold/delift_tpss.ma". -include "basic_2/unfold/delift_ltpss.ma". -include "basic_2/unfold/thin.ma". - -(* BASIC DELIFT ON LOCAL ENVIRONMENTS ***************************************) - -(* Inversion lemmas on inverse basic term relocation ************************) - -lemma thin_inv_delift1: ∀I,K1,V1,L2,d,e. ▼*[d, e] K1. ⓑ{I} V1 ≡ L2 → 0 < d → - ∃∃K2,V2. ▼*[d - 1, e] K1 ≡ K2 & - K1 ⊢ ▼*[d - 1, e] V1 ≡ V2 & - L2 = K2. ⓑ{I} V2. -#I #K1 #V1 #L2 #d #e * #X #HK1 #HL2 #e -elim (ltpss_sn_inv_tpss11 … HK1 ?) -HK1 // #K #V #HK1 #HV1 #H destruct -elim (ldrop_inv_skip1 … HL2 ?) -HL2 // #K2 #V2 #HK2 #HV2 #H destruct /3 width=5/ -qed-. - -(* Properties on inverse basic term relocation ******************************) - -lemma thin_delift: ∀L1,L2,d,e. ▼*[d, e] L1 ≡ L2 → ∀V1,V2. L1 ⊢ ▼*[d, e] V1 ≡ V2 → - ∀I. ▼*[d + 1, e] L1.ⓑ{I}V1 ≡ L2.ⓑ{I}V2. -#L1 #L2 #d #e * #L #HL1 #HL2 #V1 #V2 * #V #HV1 #HV2 #I -elim (ltpss_sn_tpss_conf … HV1 … HL1) -HV1 #V0 #HV10 #HV0 -lapply (tpss_inv_lift1_eq … HV0 … HV2) -HV0 #H destruct -lapply (ltpss_sn_tpss_trans_eq … HV10 … HL1) -HV10 /3 width=5/ -qed. - -lemma thin_delift_tpss_conf_le: ∀L,U1,U2,d,e. L ⊢ U1 ▶* [d, e] U2 → - ∀T1,dd,ee. L ⊢ ▼*[dd, ee] U1 ≡ T1 → - ∀K. ▼*[dd, ee] L ≡ K → d + e ≤ dd → - ∃∃T2. K ⊢ T1 ▶* [d, e] T2 & - L ⊢ ▼*[dd, ee] U2 ≡ T2. -#L #U1 #U2 #d #e #HU12 #T1 #dd #ee #HUT1 #K * #Y #HLY #HYK #Hdedd -lapply (delift_ltpss_sn_conf_eq … HUT1 … HLY) -HUT1 #HUT1 -elim (ltpss_sn_tpss_conf … HU12 … HLY) -HU12 #U #HU1 #HU2 -elim (delift_tpss_conf_le … HU1 … HUT1 … HYK ?) -HU1 -HUT1 -HYK // -Hdedd #T #HT1 #HUT -lapply (ltpss_sn_delift_trans_eq … HLY … HUT) -HLY -HUT #HUT -lapply (tpss_delift_trans_eq … HU2 … HUT) -U /2 width=3/ -qed. - -lemma thin_delift_tps_conf_le: ∀L,U1,U2,d,e. L ⊢ U1 ▶ [d, e] U2 → - ∀T1,dd,ee. L ⊢ ▼*[dd, ee] U1 ≡ T1 → - ∀K. ▼*[dd, ee] L ≡ K → d + e ≤ dd → - ∃∃T2. K ⊢ T1 ▶* [d, e] T2 & - L ⊢ ▼*[dd, ee] U2 ≡ T2. -/3 width=3/ qed. - -lemma thin_delift_tpss_conf_le_up: ∀L,U1,U2,d,e. L ⊢ U1 ▶* [d, e] U2 → - ∀T1,dd,ee. L ⊢ ▼*[dd, ee] U1 ≡ T1 → - ∀K. ▼*[dd, ee] L ≡ K → - d ≤ dd → dd ≤ d + e → d + e ≤ dd + ee → - ∃∃T2. K ⊢ T1 ▶* [d, dd - d] T2 & - L ⊢ ▼*[dd, ee] U2 ≡ T2. -#L #U1 #U2 #d #e #HU12 #T1 #dd #ee #HUT1 #K * #Y #HLY #HYK #Hdd #Hdde #Hddee -lapply (delift_ltpss_sn_conf_eq … HUT1 … HLY) -HUT1 #HUT1 -elim (ltpss_sn_tpss_conf … HU12 … HLY) -HU12 #U #HU1 #HU2 -elim (delift_tpss_conf_le_up … HU1 … HUT1 … HYK ? ? ?) -HU1 -HUT1 -HYK // -Hdd -Hdde -Hddee #T #HT1 #HUT -lapply (ltpss_sn_delift_trans_eq … HLY … HUT) -HLY -HUT #HUT -lapply (tpss_delift_trans_eq … HU2 … HUT) -U /2 width=3/ -qed. - -lemma thin_delift_tps_conf_le_up: ∀L,U1,U2,d,e. L ⊢ U1 ▶ [d, e] U2 → - ∀T1,dd,ee. L ⊢ ▼*[dd, ee] U1 ≡ T1 → - ∀K. ▼*[dd, ee] L ≡ K → - d ≤ dd → dd ≤ d + e → d + e ≤ dd + ee → - ∃∃T2. K ⊢ T1 ▶* [d, dd - d] T2 & - L ⊢ ▼*[dd, ee] U2 ≡ T2. -/3 width=6 by thin_delift_tpss_conf_le_up, tpss_strap2/ qed. (**) (* too slow without trace *) - -lemma thin_delift_tpss_conf_be: ∀L,U1,U2,d,e. L ⊢ U1 ▶* [d, e] U2 → - ∀T1,dd,ee. L ⊢ ▼*[dd, ee] U1 ≡ T1 → - ∀K. ▼*[dd, ee] L ≡ K → d ≤ dd → dd + ee ≤ d + e → - ∃∃T2. K ⊢ T1 ▶* [d, e - ee] T2 & - L ⊢ ▼*[dd, ee] U2 ≡ T2. -#L #U1 #U2 #d #e #HU12 #T1 #dd #ee #HUT1 #K * #Y #HLY #HYK #Hdd #Hddee -lapply (delift_ltpss_sn_conf_eq … HUT1 … HLY) -HUT1 #HUT1 -elim (ltpss_sn_tpss_conf … HU12 … HLY) -HU12 #U #HU1 #HU2 -elim (delift_tpss_conf_be … HU1 … HUT1 … HYK ? ?) -HU1 -HUT1 -HYK // -Hdd -Hddee #T #HT1 #HUT -lapply (ltpss_sn_delift_trans_eq … HLY … HUT) -HLY -HUT #HUT -lapply (tpss_delift_trans_eq … HU2 … HUT) -U /2 width=3/ -qed. - -lemma thin_delift_tps_conf_be: ∀L,U1,U2,d,e. L ⊢ U1 ▶ [d, e] U2 → - ∀T1,dd,ee. L ⊢ ▼*[dd, ee] U1 ≡ T1 → - ∀K. ▼*[dd, ee] L ≡ K → d ≤ dd → dd + ee ≤ d + e → - ∃∃T2. K ⊢ T1 ▶* [d, e - ee] T2 & - L ⊢ ▼*[dd, ee] U2 ≡ T2. -/3 width=3/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/thin_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/thin_ldrop.ma deleted file mode 100644 index f26c26f01..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/thin_ldrop.ma +++ /dev/null @@ -1,58 +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/unfold/ltpss_sn_ldrop.ma". -include "basic_2/unfold/thin.ma". - -(* BASIC LOCAL ENVIRONMENT THINNING *****************************************) - -(* Properties on local environment slicing **********************************) - -lemma thin_ldrop_conf_ge: ∀L0,L1,d1,e1. ▼*[d1, e1] L0 ≡ L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → - d1 + e1 ≤ e2 → ⇩[0, e2 - e1] L1 ≡ L2. -#L0 #L1 #d1 #e1 * /3 width=8 by ltpss_sn_ldrop_conf_ge, ldrop_conf_ge/ -qed. - -lemma thin_ldrop_conf_be: ∀L0,L1,d1,e1. ▼*[d1, e1] L0 ≡ L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → d1 ≤ e2 → e2 ≤ d1 + e1 → - ∃∃L. ▼*[0, d1 + e1 - e2] L2 ≡ L & ⇩[0, d1] L1 ≡ L. -#L0 #L1 #d1 #e1 * #L #HL0 #HL1 #L2 #e2 #HL02 #Hd1e2 #He2de1 -elim (ltpss_sn_ldrop_conf_be … HL0 … HL02 ? ?) -L0 // #L0 #HL20 #HL0 -elim (ldrop_conf_be … HL1 … HL0 ? ?) -L // -Hd1e2 -He2de1 /3 width=3/ -qed. - -lemma thin_ldrop_conf_le: ∀L0,L1,d1,e1. ▼*[d1, e1] L0 ≡ L1 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → e2 ≤ d1 → - ∃∃L. ▼*[d1 - e2, e1] L2 ≡ L & ⇩[0, e2] L1 ≡ L. -#L0 #L1 #d1 #e1 * #L #HL0 #HL1 #L2 #e2 #HL02 #He2d1 -elim (ltpss_sn_ldrop_conf_le … HL0 … HL02 ?) -L0 // #L0 #HL20 #HL0 -elim (ldrop_conf_le … HL1 … HL0 ?) -L // -He2d1 /3 width=3/ -qed. - -lemma thin_ldrop_trans_ge: ∀L1,L0,d1,e1. ▼*[d1, e1] L1 ≡ L0 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → - d1 ≤ e2 → ⇩[0, e1 + e2] L1 ≡ L2. -#L1 #L0 #d1 #e1 * #L #HL1 #HL0 #L2 #e2 #HL02 #Hd1e2 -lapply (ldrop_trans_ge … HL0 … HL02 ?) -L0 // #HL2 -lapply (ltpss_sn_ldrop_trans_ge … HL1 … HL2 ?) -L // /2 width=1/ -qed. - -lemma thin_ldrop_trans_le: ∀L1,L0,d1,e1. ▼*[d1, e1] L1 ≡ L0 → - ∀L2,e2. ⇩[0, e2] L0 ≡ L2 → e2 ≤ d1 → - ∃∃L. ▼*[d1 - e2, e1] L ≡ L2 & ⇩[0, e2] L1 ≡ L. -#L1 #L0 #d1 #e1 * #L #HL1 #HL0 #L2 #e2 #HL02 #He2d1 -elim (ldrop_trans_le … HL0 … HL02 He2d1) -L0 #L0 #HL0 #HL02 -elim (ltpss_sn_ldrop_trans_le … HL1 … HL0 He2d1) -L -He2d1 /3 width=3/ -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss.ma deleted file mode 100644 index 433718803..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss.ma +++ /dev/null @@ -1,188 +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 *) -(* *) -(**************************************************************************) - -notation "hvbox( L ⊢ break term 46 T1 break ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" - non associative with precedence 45 - for @{ 'PSubstStar $L $T1 $d $e $T2 }. - -include "basic_2/substitution/tps.ma". - -(* PARTIAL UNFOLD ON TERMS **************************************************) - -definition tpss: nat → nat → lenv → relation term ≝ - λd,e,L. TC … (tps d e L). - -interpretation "partial unfold (term)" - 'PSubstStar L T1 d e T2 = (tpss d e L T1 T2). - -(* Basic eliminators ********************************************************) - -lemma tpss_ind: ∀d,e,L,T1. ∀R:predicate term. R T1 → - (∀T,T2. L ⊢ T1 ▶* [d, e] T → L ⊢ T ▶ [d, e] T2 → R T → R T2) → - ∀T2. L ⊢ T1 ▶* [d, e] T2 → R T2. -#d #e #L #T1 #R #HT1 #IHT1 #T2 #HT12 -@(TC_star_ind … HT1 IHT1 … HT12) // -qed-. - -lemma tpss_ind_dx: ∀d,e,L,T2. ∀R:predicate term. R T2 → - (∀T1,T. L ⊢ T1 ▶ [d, e] T → L ⊢ T ▶* [d, e] T2 → R T → R T1) → - ∀T1. L ⊢ T1 ▶* [d, e] T2 → R T1. -#d #e #L #T2 #R #HT2 #IHT2 #T1 #HT12 -@(TC_star_ind_dx … HT2 IHT2 … HT12) // -qed-. - -(* Basic properties *********************************************************) - -lemma tps_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶ [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. -/2 width=1/ qed. - -lemma tpss_strap1: ∀L,T1,T,T2,d,e. - L ⊢ T1 ▶* [d, e] T → L ⊢ T ▶ [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. -/2 width=3/ qed. - -lemma tpss_strap2: ∀L,T1,T,T2,d,e. - L ⊢ T1 ▶ [d, e] T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. -/2 width=3/ qed. - -lemma tpss_lsubr_trans: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶* [d, e] T2 → - ∀L2. L2 ⊑ [d, e] L1 → L2 ⊢ T1 ▶* [d, e] T2. -/3 width=3/ qed. - -lemma tpss_refl: ∀d,e,L,T. L ⊢ T ▶* [d, e] T. -/2 width=1/ qed. - -lemma tpss_bind: ∀L,V1,V2,d,e. L ⊢ V1 ▶* [d, e] V2 → - ∀a,I,T1,T2. L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 → - L ⊢ ⓑ{a,I} V1. T1 ▶* [d, e] ⓑ{a,I} V2. T2. -#L #V1 #V2 #d #e #HV12 elim HV12 -V2 -[ #V2 #HV12 #a #I #T1 #T2 #HT12 elim HT12 -T2 - [ /3 width=5/ - | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) - ] -| #V #V2 #_ #HV12 #IHV #a #I #T1 #T2 #HT12 - lapply (tpss_lsubr_trans … HT12 (L. ⓑ{I} V) ?) -HT12 /2 width=1/ #HT12 - lapply (IHV a … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) -] -qed. - -lemma tpss_flat: ∀L,I,V1,V2,T1,T2,d,e. - L ⊢ V1 ▶* [d, e] V2 → L ⊢ T1 ▶* [d, e] T2 → - L ⊢ ⓕ{I} V1. T1 ▶* [d, e] ⓕ{I} V2. T2. -#L #I #V1 #V2 #T1 #T2 #d #e #HV12 elim HV12 -V2 -[ #V2 #HV12 #HT12 elim HT12 -T2 - [ /3 width=1/ - | #T #T2 #_ #HT2 #IHT @step /2 width=5/ (**) (* /3 width=5/ is too slow *) - ] -| #V #V2 #_ #HV12 #IHV #HT12 - lapply (IHV … HT12) -IHV -HT12 #HT12 @step /2 width=5/ (**) (* /3 width=5/ is too slow *) -] -qed. - -lemma tpss_weak: ∀L,T1,T2,d1,e1. L ⊢ T1 ▶* [d1, e1] T2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → - L ⊢ T1 ▶* [d2, e2] T2. -#L #T1 #T2 #d1 #e1 #H #d1 #d2 #Hd21 #Hde12 @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT12 #IHT - lapply (tps_weak … HT12 … Hd21 Hde12) -HT12 -Hd21 -Hde12 /2 width=3/ -] -qed. - -lemma tpss_weak_top: ∀L,T1,T2,d,e. - L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ▶* [d, |L| - d] T2. -#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT12 #IHT - lapply (tps_weak_top … HT12) -HT12 /2 width=3/ -] -qed. - -lemma tpss_weak_full: ∀L,T1,T2,d,e. - L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ▶* [0, |L|] T2. -#L #T1 #T2 #d #e #HT12 -lapply (tpss_weak … HT12 0 (d + e) ? ?) -HT12 // #HT12 -lapply (tpss_weak_top … HT12) // -qed. - -lemma tpss_append: ∀K,T1,T2,d,e. K ⊢ T1 ▶* [d, e] T2 → - ∀L. L @@ K ⊢ T1 ▶* [d, e] T2. -#K #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // /3 width=3/ -qed. - -(* Basic inversion lemmas ***************************************************) - -(* Note: this can be derived from tpss_inv_atom1 *) -lemma tpss_inv_sort1: ∀L,T2,k,d,e. L ⊢ ⋆k ▶* [d, e] T2 → T2 = ⋆k. -#L #T2 #k #d #e #H @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT2 #IHT destruct - >(tps_inv_sort1 … HT2) -HT2 // -] -qed-. - -(* Note: this can be derived from tpss_inv_atom1 *) -lemma tpss_inv_gref1: ∀L,T2,p,d,e. L ⊢ §p ▶* [d, e] T2 → T2 = §p. -#L #T2 #p #d #e #H @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT2 #IHT destruct - >(tps_inv_gref1 … HT2) -HT2 // -] -qed-. - -lemma tpss_inv_bind1: ∀d,e,L,a,I,V1,T1,U2. L ⊢ ⓑ{a,I} V1. T1 ▶* [d, e] U2 → - ∃∃V2,T2. L ⊢ V1 ▶* [d, e] V2 & - L. ⓑ{I} V2 ⊢ T1 ▶* [d + 1, e] T2 & - U2 = ⓑ{a,I} V2. T2. -#d #e #L #a #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 -[ /2 width=5/ -| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct - elim (tps_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H - lapply (tpss_lsubr_trans … HT1 (L. ⓑ{I} V2) ?) -HT1 /2 width=1/ /3 width=5/ -] -qed-. - -lemma tpss_inv_flat1: ∀d,e,L,I,V1,T1,U2. L ⊢ ⓕ{I} V1. T1 ▶* [d, e] U2 → - ∃∃V2,T2. L ⊢ V1 ▶* [d, e] V2 & L ⊢ T1 ▶* [d, e] T2 & - U2 = ⓕ{I} V2. T2. -#d #e #L #I #V1 #T1 #U2 #H @(tpss_ind … H) -U2 -[ /2 width=5/ -| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct - elim (tps_inv_flat1 … HU2) -HU2 /3 width=5/ -] -qed-. - -lemma tpss_inv_refl_O2: ∀L,T1,T2,d. L ⊢ T1 ▶* [d, 0] T2 → T1 = T2. -#L #T1 #T2 #d #H @(tpss_ind … H) -T2 -[ // -| #T #T2 #_ #HT2 #IHT <(tps_inv_refl_O2 … HT2) -HT2 // -] -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma tpss_fwd_tw: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → ♯{T1} ≤ ♯{T2}. -#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT1 -lapply (tps_fwd_tw … HT2) -HT2 #HT2 -@(transitive_le … IHT1) // -qed-. - -lemma tpss_fwd_shift1: ∀L,L1,T1,T,d,e. L ⊢ L1 @@ T1 ▶*[d, e] T → - ∃∃L2,T2. |L1| = |L2| & T = L2 @@ T2. -#L #L1 #T1 #T #d #e #H @(tpss_ind … H) -T -[ /2 width=4/ -| #T #X #_ #H0 * #L0 #T0 #HL10 #H destruct - elim (tps_fwd_shift1 … H0) -H0 #L2 #T2 #HL02 #H destruct /2 width=4/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_alt.ma deleted file mode 100644 index f7670eaa8..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_alt.ma +++ /dev/null @@ -1,105 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -notation "hvbox( L ⊢ break term 46 T1 break ▶ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" - non associative with precedence 45 - for @{ 'PSubstStarAlt $L $T1 $d $e $T2 }. - -include "basic_2/unfold/tpss_lift.ma". - -(* PARALLEL UNFOLD ON TERMS *************************************************) - -(* alternative definition of tpss *) -inductive tpssa: nat → nat → lenv → relation term ≝ -| tpssa_atom : ∀L,I,d,e. tpssa d e L (⓪{I}) (⓪{I}) -| tpssa_subst: ∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e → - ⇩[0, i] L ≡ K. ⓓV1 → tpssa 0 (d + e - i - 1) K V1 V2 → - ⇧[0, i + 1] V2 ≡ W2 → tpssa d e L (#i) W2 -| tpssa_bind : ∀L,a,I,V1,V2,T1,T2,d,e. - tpssa d e L V1 V2 → tpssa (d + 1) e (L. ⓑ{I} V2) T1 T2 → - tpssa d e L (ⓑ{a,I} V1. T1) (ⓑ{a,I} V2. T2) -| tpssa_flat : ∀L,I,V1,V2,T1,T2,d,e. - tpssa d e L V1 V2 → tpssa d e L T1 T2 → - tpssa d e L (ⓕ{I} V1. T1) (ⓕ{I} V2. T2) -. - -interpretation "parallel unfold (term) alternative" - 'PSubstStarAlt L T1 d e T2 = (tpssa d e L T1 T2). - -(* Basic properties *********************************************************) - -lemma tpssa_lsubr_trans: ∀L1,T1,T2,d,e. L1 ⊢ T1 ▶▶* [d, e] T2 → - ∀L2. L2 ⊑ [d, e] L1 → L2 ⊢ T1 ▶▶* [d, e] T2. -#L1 #T1 #T2 #d #e #H elim H -L1 -T1 -T2 -d -e -[ // -| #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 - elim (ldrop_lsubr_ldrop2_abbr … HL12 … HLK1 ? ?) -HL12 -HLK1 // /3 width=6/ -| /4 width=1/ -| /3 width=1/ -] -qed. - -lemma tpssa_refl: ∀T,L,d,e. L ⊢ T ▶▶* [d, e] T. -#T elim T -T // -#I elim I -I /2 width=1/ -qed. - -lemma tpssa_tps_trans: ∀L,T1,T,d,e. L ⊢ T1 ▶▶* [d, e] T → - ∀T2. L ⊢ T ▶ [d, e] T2 → L ⊢ T1 ▶▶* [d, e] T2. -#L #T1 #T #d #e #H elim H -L -T1 -T -d -e -[ #L #I #d #e #X #H - elim (tps_inv_atom1 … H) -H // * /2 width=6/ -| #L #K #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK #_ #HVW2 #IHV12 #T2 #H - lapply (ldrop_fwd_ldrop2 … HLK) #H0LK - lapply (tps_weak … H 0 (d+e) ? ?) -H // #H - elim (tps_inv_lift1_be … H … H0LK … HVW2 ? ?) -H -H0LK -HVW2 // /3 width=6/ -| #L #a #I #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H - elim (tps_inv_bind1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct - lapply (tps_lsubr_trans … HT2 (L.ⓑ{I}V) ?) -HT2 /2 width=1/ #HT2 - lapply (IHV1 … HV2) -IHV1 -HV2 #HV12 - lapply (IHT1 … HT2) -IHT1 -HT2 #HT12 - lapply (tpssa_lsubr_trans … HT12 (L.ⓑ{I}V2) ?) -HT12 /2 width=1/ -| #L #I #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H - elim (tps_inv_flat1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct /3 width=1/ -] -qed. - -lemma tpss_tpssa: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ▶▶* [d, e] T2. -#L #T1 #T2 #d #e #H @(tpss_ind … H) -T2 // /2 width=3/ -qed. - -(* Basic inversion lemmas ***************************************************) - -lemma tpssa_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶▶* [d, e] T2 → L ⊢ T1 ▶* [d, e] T2. -#L #T1 #T2 #d #e #H elim H -L -T1 -T2 -d -e // /2 width=6/ -qed-. - -lemma tpss_ind_alt: ∀R:nat→nat→lenv→relation term. - (∀L,I,d,e. R d e L (⓪{I}) (⓪{I})) → - (∀L,K,V1,V2,W2,i,d,e. d ≤ i → i < d + e → - ⇩[O, i] L ≡ K.ⓓV1 → K ⊢ V1 ▶* [O, d + e - i - 1] V2 → - ⇧[O, i + 1] V2 ≡ W2 → R O (d+e-i-1) K V1 V2 → R d e L (#i) W2 - ) → - (∀L,a,I,V1,V2,T1,T2,d,e. L ⊢ V1 ▶* [d, e] V2 → - L.ⓑ{I}V2 ⊢ T1 ▶* [d + 1, e] T2 → R d e L V1 V2 → - R (d+1) e (L.ⓑ{I}V2) T1 T2 → R d e L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) - ) → - (∀L,I,V1,V2,T1,T2,d,e. L ⊢ V1 ▶* [d, e] V2 → - L ⊢ T1 ▶* [d, e] T2 → R d e L V1 V2 → - R d e L T1 T2 → R d e L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) - ) → - ∀d,e,L,T1,T2. L ⊢ T1 ▶* [d, e] T2 → R d e L T1 T2. -#R #H1 #H2 #H3 #H4 #d #e #L #T1 #T2 #H elim (tpss_tpssa … H) -L -T1 -T2 -d -e -// /3 width=1 by tpssa_tpss/ /3 width=7 by tpssa_tpss/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_lift.ma deleted file mode 100644 index a68f86e32..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_lift.ma +++ /dev/null @@ -1,196 +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/tps_lift.ma". -include "basic_2/unfold/tpss.ma". - -(* PARTIAL UNFOLD ON TERMS **************************************************) - -(* Advanced properties ******************************************************) - -lemma tpss_subst: ∀L,K,V,U1,i,d,e. - d ≤ i → i < d + e → - ⇩[0, i] L ≡ K. ⓓV → K ⊢ V ▶* [0, d + e - i - 1] U1 → - ∀U2. ⇧[0, i + 1] U1 ≡ U2 → L ⊢ #i ▶* [d, e] U2. -#L #K #V #U1 #i #d #e #Hdi #Hide #HLK #H @(tpss_ind … H) -U1 -[ /3 width=4/ -| #U #U1 #_ #HU1 #IHU #U2 #HU12 - elim (lift_total U 0 (i+1)) #U0 #HU0 - lapply (IHU … HU0) -IHU #H - lapply (ldrop_fwd_ldrop2 … HLK) -HLK #HLK - lapply (tps_lift_ge … HU1 … HLK HU0 HU12 ?) -HU1 -HLK -HU0 -HU12 // normalize #HU02 - lapply (tps_weak … HU02 d e ? ?) -HU02 [ >minus_plus >commutative_plus /2 width=1/ | /2 width=1/ | /2 width=3/ ] -] -qed. - -(* Advanced inverion lemmas *************************************************) - -lemma tpss_inv_atom1: ∀L,T2,I,d,e. L ⊢ ⓪{I} ▶* [d, e] T2 → - T2 = ⓪{I} ∨ - ∃∃K,V1,V2,i. d ≤ i & i < d + e & - ⇩[O, i] L ≡ K. ⓓV1 & - K ⊢ V1 ▶* [0, d + e - i - 1] V2 & - ⇧[O, i + 1] V2 ≡ T2 & - I = LRef i. -#L #T2 #I #d #e #H @(tpss_ind … H) -T2 -[ /2 width=1/ -| #T #T2 #_ #HT2 * - [ #H destruct - elim (tps_inv_atom1 … HT2) -HT2 [ /2 width=1/ | * /3 width=10/ ] - | * #K #V1 #V #i #Hdi #Hide #HLK #HV1 #HVT #HI - lapply (ldrop_fwd_ldrop2 … HLK) #H - elim (tps_inv_lift1_ge_up … HT2 … H … HVT ? ? ?) normalize -HT2 -H -HVT [2,3,4: /2 width=1/ ] #V2 (lift_mono … HTU1 … H) -H // -| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 - elim (lift_total T d e) #U #HTU - lapply (IHT … HTU) -IHT #HU1 - lapply (tps_lift_le … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ -] -qed. - -lemma tpss_lift_be: ∀K,T1,T2,dt,et. K ⊢ T1 ▶* [dt, et] T2 → - ∀L,U1,d,e. dt ≤ d → d ≤ dt + et → - ⇩[d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → - ∀U2. ⇧[d, e] T2 ≡ U2 → L ⊢ U1 ▶* [dt, et + e] U2. -#K #T1 #T2 #dt #et #H #L #U1 #d #e #Hdtd #Hddet #HLK #HTU1 @(tpss_ind … H) -T2 -[ #U2 #H >(lift_mono … HTU1 … H) -H // -| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 - elim (lift_total T d e) #U #HTU - lapply (IHT … HTU) -IHT #HU1 - lapply (tps_lift_be … HT2 … HLK HTU HTU2 ? ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ -] -qed. - -lemma tpss_lift_ge: ∀K,T1,T2,dt,et. K ⊢ T1 ▶* [dt, et] T2 → - ∀L,U1,d,e. d ≤ dt → ⇩[d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → - L ⊢ U1 ▶* [dt + e, et] U2. -#K #T1 #T2 #dt #et #H #L #U1 #d #e #Hddt #HLK #HTU1 @(tpss_ind … H) -T2 -[ #U2 #H >(lift_mono … HTU1 … H) -H // -| -HTU1 #T #T2 #_ #HT2 #IHT #U2 #HTU2 - elim (lift_total T d e) #U #HTU - lapply (IHT … HTU) -IHT #HU1 - lapply (tps_lift_ge … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 // /2 width=3/ -] -qed. - -lemma tpss_inv_lift1_le: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt + et ≤ d → - ∃∃T2. K ⊢ T1 ▶* [dt, et] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdetd @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_le … HU2 … HLK … HTU ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. - -lemma tpss_inv_lift1_be: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 ▶* [dt, et - e] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hdedet @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_be … HU2 … HLK … HTU ? ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. - -lemma tpss_inv_lift1_ge: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d + e ≤ dt → - ∃∃T2. K ⊢ T1 ▶* [dt - e, et] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdedt @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_ge … HU2 … HLK … HTU ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. - -lemma tpss_inv_lift1_eq: ∀L,U1,U2,d,e. - L ⊢ U1 ▶* [d, e] U2 → ∀T1. ⇧[d, e] T1 ≡ U1 → U1 = U2. -#L #U1 #U2 #d #e #H #T1 #HTU1 @(tpss_ind … H) -U2 // -#U #U2 #_ #HU2 #IHU destruct -<(tps_inv_lift1_eq … HU2 … HTU1) -HU2 -HTU1 // -qed. - -lemma tpss_inv_lift1_ge_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 ▶* [d, dt + et - (d + e)] T2 & - ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_ge_up … HU2 … HLK … HTU ? ? ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. - -lemma tpss_inv_lift1_be_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → dt + et ≤ d + e → - ∃∃T2. K ⊢ T1 ▶* [dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_be_up … HU2 … HLK … HTU ? ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. - -lemma tpss_inv_lift1_le_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → d ≤ dt + et → dt + et ≤ d + e → - ∃∃T2. K ⊢ T1 ▶* [dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #H #K #d #e #HLK #T1 #HTU1 #Hdtd #Hddet #Hdetde @(tpss_ind … H) -U2 -[ /2 width=3/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (tps_inv_lift1_le_up … HU2 … HLK … HTU ? ? ?) -HU2 -HLK -HTU // /3 width=3/ -] -qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_tpss.ma b/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_tpss.ma deleted file mode 100644 index d124bb32c..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/unfold/tpss_tpss.ma +++ /dev/null @@ -1,96 +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/tps_tps.ma". -include "basic_2/unfold/tpss_lift.ma". - -(* PARTIAL UNFOLD ON TERMS **************************************************) - -(* Advanced inversion lemmas ************************************************) - -lemma tpss_inv_SO2: ∀L,T1,T2,d. L ⊢ T1 ▶* [d, 1] T2 → L ⊢ T1 ▶ [d, 1] T2. -#L #T1 #T2 #d #H @(tpss_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT1 -lapply (tps_trans_ge … IHT1 … HT2 ?) // -qed-. - -(* Advanced properties ******************************************************) - -lemma tpss_strip_eq: ∀L,T0,T1,d1,e1. L ⊢ T0 ▶* [d1, e1] T1 → - ∀T2,d2,e2. L ⊢ T0 ▶ [d2, e2] T2 → - ∃∃T. L ⊢ T1 ▶ [d2, e2] T & L ⊢ T2 ▶* [d1, e1] T. -/3 width=3/ qed. - -lemma tpss_strip_neq: ∀L1,T0,T1,d1,e1. L1 ⊢ T0 ▶* [d1, e1] T1 → - ∀L2,T2,d2,e2. L2 ⊢ T0 ▶ [d2, e2] T2 → - (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → - ∃∃T. L2 ⊢ T1 ▶ [d2, e2] T & L1 ⊢ T2 ▶* [d1, e1] T. -/3 width=3/ qed. - -lemma tpss_strap1_down: ∀L,T1,T0,d1,e1. L ⊢ T1 ▶* [d1, e1] T0 → - ∀T2,d2,e2. L ⊢ T0 ▶ [d2, e2] T2 → d2 + e2 ≤ d1 → - ∃∃T. L ⊢ T1 ▶ [d2, e2] T & L ⊢ T ▶* [d1, e1] T2. -/3 width=3/ qed. - -lemma tpss_strap2_down: ∀L,T1,T0,d1,e1. L ⊢ T1 ▶ [d1, e1] T0 → - ∀T2,d2,e2. L ⊢ T0 ▶* [d2, e2] T2 → d2 + e2 ≤ d1 → - ∃∃T. L ⊢ T1 ▶* [d2, e2] T & L ⊢ T ▶ [d1, e1] T2. -/3 width=3/ qed. - -lemma tpss_split_up: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → - ∀i. d ≤ i → i ≤ d + e → - ∃∃T. L ⊢ T1 ▶* [d, i - d] T & L ⊢ T ▶* [i, d + e - i] T2. -#L #T1 #T2 #d #e #H #i #Hdi #Hide @(tpss_ind … H) -T2 -[ /2 width=3/ -| #T #T2 #_ #HT12 * #T3 #HT13 #HT3 - elim (tps_split_up … HT12 … Hdi Hide) -HT12 -Hide #T0 #HT0 #HT02 - elim (tpss_strap1_down … HT3 … HT0 ?) -T [2: >commutative_plus /2 width=1/ ] - /3 width=7 by ex2_intro, step/ (**) (* just /3 width=7/ is too slow *) -] -qed. - -lemma tpss_inv_lift1_up: ∀L,U1,U2,dt,et. L ⊢ U1 ▶* [dt, et] U2 → - ∀K,d,e. ⇩[d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ d + e → d + e ≤ dt + et → - ∃∃T2. K ⊢ T1 ▶* [d, dt + et - (d + e)] T2 & - ⇧[d, e] T2 ≡ U2. -#L #U1 #U2 #dt #et #HU12 #K #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet -elim (tpss_split_up … HU12 (d + e) ? ?) -HU12 // -Hdedet #U #HU1 #HU2 -lapply (tpss_weak … HU1 d e ? ?) -HU1 // [ >commutative_plus /2 width=1/ ] -Hddt -Hdtde #HU1 -lapply (tpss_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -elim (tpss_inv_lift1_ge … HU2 … HLK … HTU1 ?) -HU2 -HLK -HTU1 // + + Reaxiomatized substitution and reduction + commute with respect to subclosure + (anniversary milestone). + Mutual recursive preservation of stratified native validity for hyper computation on closures. diff --git a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl index 239a44fe7..8ed310e1e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl +++ b/matita/matita/contribs/lambdadelta/basic_2/web/basic_2_src.tbl @@ -133,6 +133,11 @@ table { [ "cnf ( ? ⊢ 𝐍⦃?⦄ )" "cnf_lift" + "cnf_cif" * ] } ] + [ { "revised context-sensitive reduction" * } { + [ "lpr ( ? ⊢ ➡ ? )" "lpr_ldrop" + "lpr_cpr" + "lpr_lpr" * ] + [ "cpr ( ? ⊢ ? ➡ ? )" "cpr_lift" * ] + } + ] [ { "context-sensitive reduction" * } { [ "cpr ( ? ⊢ ? ➡ ? )" "cpr_lift" + "cpr_tpss" + "cpr_ltpss_dx" + "cpr_ltpss_sn" + "cpr_delift" + "cpr_aaa" + "cpr_ltpr" + "cpr_cpr" * ] } @@ -158,13 +163,13 @@ table { [ "lpqs ( ? ⊢ ➤* ? )" "lpqs_ldrop" + "lpqs_cpqs" + "lpqs_lpqs" * ] [ "cpqs ( ? ⊢ ? ➤* ? )" "cpqs_lift" * ] } - ] + ] } ] class "green" [ { "unwind" * } { [ { "iterated stratified static type assignment" * } { - [ "sstas ( ⦃?,?⦄ ⊢ ? •*[?] ? )" "sstas_lift" + "sstas_ltpss_dx" + "sstas_ltpss_sn" + "sstas_aaa" + "sstas_sstas" * ] + [ "sstas ( ⦃?,?⦄ ⊢ ? •*[?] ? )" "sstas_lift" + "sstas_lpss" + "sstas_aaa" + "sstas_sstas" * ] } ] } @@ -172,7 +177,7 @@ table { class "grass" [ { "static typing" * } { [ { "stratified static type assignment" * } { - [ "ssta ( ⦃?,?⦄ ⊢ ? •[?,?] ? )" "ssta_lift" + "ssta_ltpss_dx" + "ssta_ltpss_sn" + "ssta_aaa" + "ssta_ssta" * ] + [ "ssta ( ⦃?,?⦄ ⊢ ? •[?,?] ? )" "ssta_lift" + "ssta_lpss" + "ssta_aaa" + "ssta_ssta" * ] } ] [ { "local env. ref. for atomic arity assignment" * } { @@ -180,7 +185,7 @@ table { } ] [ { "atomic arity assignment" * } { - [ "aaa ( ? ⊢ ? ⁝ ? )" "aaa_lift" + "aaa_lifts" + "aaa_ltpss_dx" + "aaa_ltpss_sn" + "aaa_aaa" * ] + [ "aaa ( ? ⊢ ? ⁝ ? )" "aaa_lift" + "aaa_lifts" + "aaa_lpss" + "aaa_aaa" * ] } ] [ { "parameters" * } { @@ -191,7 +196,8 @@ table { ] class "yellow" [ { "unfold" * } { - [ { "basic local env. thinning" * } { +(* + [ { "basic local env. thinning" * } { [ "thin ( ? ▼*[?,?] ≡ ? )" "thin_ldrop" + "thin_delift" * ] } ] @@ -199,18 +205,14 @@ table { [ "delift ( ? ⊢ ? ▼*[?,?] ≡ ? )" "delift_alt ( ? ⊢ ? ▼▼*[?,?] ≡ ? )" "delift_lift" + "delift_tpss" + "delift_ltpss" + "delift_delift" * ] } ] - [ { "revised parallel substitution" * } { +*) + [ { "parallel substitution" * } { [ "lpss ( ? ⊢ ▶* ? )" "lpss_ldrop" + "lpss_cpss" + "lpss_lpss" * ] [ "cpss ( ? ⊢ ? ▶* ? )" "cpss_lift" * ] } ] - [ { "partial unfold" * } { - [ "ltpss_sn ( ? ⊢ ▶*[?,?] ? )" "ltpss_sn_alt ( ? ⊢ ▶▶*[?,?] ? )" "ltpss_sn_ldrop" + "ltpss_sn_tps" + "ltpss_sn_tpss" + "ltpss_sn_ltpss_sn" * ] - [ "ltpss_dx ( ? ▶*[?,?] ? )" "ltpss_dx_ldrop" + "ltpss_dx_tps" + "ltpss_dx_tpss" + "ltpss_dx_ltpss_dx" * ] - [ "tpss ( ? ⊢ ? ▶*[?,?] ? )" "tpss_alt ( ? ⊢ ? ▶▶*[?,?] ? )" "tpss_lift" "tpss_tpss" * ] - } - ] [ { "iterated structural successor for closures" * } { + [ "fsups ( ⦃?,?⦄ ⊃* ⦃?,?⦄ )" "fsups_fsups" * ] [ "fsupp ( ⦃?,?⦄ ⊃+ ⦃?,?⦄ )" "fsupp_fsupp" * ] } ] @@ -231,10 +233,6 @@ table { ] class "orange" [ { "substitution" * } { - [ { "parallel substitution" * } { - [ "tps ( ? ⊢ ? ▶[?,?] ? )" "tps_lift" + "tps_tps" * ] - } - ] [ { "structural successor for closures" * } { [ "fsup ( ⦃?,?⦄ ⊃ ⦃?,?⦄ )" * ] } diff --git a/matita/matita/predefined_virtuals.ml b/matita/matita/predefined_virtuals.ml index 225fe64bf..d8da251bb 100644 --- a/matita/matita/predefined_virtuals.ml +++ b/matita/matita/predefined_virtuals.ml @@ -1509,7 +1509,7 @@ let predefined_classes = [ ["#"; "♯"; "⌘"; ]; ["-"; "÷"; "⊢"; "⊩"; ]; ["="; "≃"; "≈"; "≝"; "≡"; "≅"; "≐"; "≑"; ]; - ["→"; "↦"; "⇝"; "⇾"; "⤍"; "⤏"; "⤳"; ] ; + ["→"; "↦"; "⇝"; "⤞"; "⇾"; "⤍"; "⤏"; "⤳"; ] ; ["⇒"; "⤇"; "➾"; "⇨"; "➡"; "➤"; "➸"; "⇉"; "⥰"; ] ; ["^"; "↑"; ] ; ["⇑"; "⇧"; "⬆"; ] ; @@ -1527,7 +1527,7 @@ let predefined_classes = [ ["□"; "◽"; "▪"; "◾"; ]; ["◊"; "♢"; "⧫"; "♦"; "⟐"; "⟠"; ] ; [">"; "⭃"; "⧁"; "〉"; "»"; "❭"; "❯"; "❱"; "▸"; "►"; "▶"; ] ; - ["≥"; "≽"; "⥸"; ]; + ["≥"; "⪀"; "≽"; "⥸"; ]; ["a"; "α"; "𝕒"; "𝐚"; "𝛂"; "ⓐ"; ] ; ["A"; "ℵ"; "𝔸"; "𝐀"; "Ⓐ"; ] ; ["b"; "β"; "ß"; "𝕓"; "𝐛"; "𝛃"; "ⓑ"; ] ;