From: Ferruccio Guidi Date: Tue, 11 Mar 2014 17:57:05 +0000 (+0000) Subject: reaxiomatized lleq fixes a bug in it and allows to park substitution in etc X-Git-Tag: make_still_working~949 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=8a5a354c9ac3ef20ca01dbeb61f6b99902f172a7;p=helm.git reaxiomatized lleq fixes a bug in it and allows to park substitution in etc --- diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma deleted file mode 100644 index 881030572..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_cpys.ma +++ /dev/null @@ -1,24 +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/reduction/cpx_cpys.ma". -include "basic_2/computation/cpxs.ma". - -(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************) - -(* Properties on local environment refinement for extended substitution *****) - -lemma lsuby_cpxs_trans: ∀h,g,G. lsub_trans … (cpxs h g G) (lsuby 0 (∞)). -/3 width=5 by lsuby_cpx_trans, LTC_lsub_trans/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_leq.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_leq.ma new file mode 100644 index 000000000..1d35deb53 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/cpxs_leq.ma @@ -0,0 +1,24 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reduction/cpx_leq.ma". +include "basic_2/computation/cpxs.ma". + +(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************) + +(* Properties on equivalence for local environments *************************) + +lemma leq_cpxs_trans: ∀h,g,G. lsub_trans … (cpxs h g G) (leq 0 (∞)). +/3 width=5 by leq_cpx_trans, LTC_lsub_trans/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu.ma index 7165feaee..93e837f9e 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/notation/relations/btpredproper_8.ma". -include "basic_2/substitution/lleq.ma". +include "basic_2/relocation/lleq.ma". include "basic_2/computation/fpbs.ma". (* UNITARY "BIG TREE" PROPER PARALLEL COMPUTATION FOR CLOSURES **************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu_fpns.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu_fpns.ma index 9f5111ad2..863b08196 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu_fpns.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fpbu_fpns.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "basic_2/relocation/lleq_lleq.ma". include "basic_2/computation/cpxs_lleq.ma". include "basic_2/computation/lpxs_lleq.ma". include "basic_2/computation/lpxs_lpxs.ma". diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/fpns.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/fpns.ma index 6e68fef14..dce46b065 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/fpns.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fpns.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/notation/relations/btpredsnstar_8.ma". -include "basic_2/substitution/lleq.ma". +include "basic_2/relocation/lleq.ma". include "basic_2/computation/lpxs.ma". (* PARALLEL COMPUTATION FOR "BIG TREE" NORMAL FORMS *************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/fpns_fpns.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/fpns_fpns.ma index 728f48df9..deaa7a967 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/fpns_fpns.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/fpns_fpns.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/substitution/lleq_lleq.ma". +include "basic_2/relocation/lleq_lleq.ma". include "basic_2/computation/lpxs_lpxs.ma". include "basic_2/computation/fpns.ma". diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lcosx_cpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lcosx_cpxs.ma index 3ed17d258..aef5014e7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lcosx_cpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lcosx_cpxs.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "ground_2/ynat/ynat_max.ma". include "basic_2/computation/lsx_ldrop.ma". include "basic_2/computation/lsx_lpxs.ma". include "basic_2/computation/lcosx.ma". @@ -36,7 +37,7 @@ lemma lsx_cpx_trans_lcosx: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → | #a #I #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #HL #H elim (lsx_inv_bind … H) -H #HV1 #HT1 @lsx_bind /2 width=2 by/ (**) (* explicit constructor *) - @(lsx_leqy_conf … (L.ⓑ{I}V1)) /3 width=1 by lcosx_pair, lsuby_succ/ + @(lsx_leq_conf … (L.ⓑ{I}V1)) /3 width=1 by lcosx_pair, leq_succ/ | #I #G #L #V1 #V2 #T1 #T2 #_ #_ #IHV12 #IHT12 #d #HL #H elim (lsx_inv_flat … H) -H /3 width=1 by lsx_flat/ | #G #L #V #U1 #U2 #T2 #_ #HTU2 #IHU12 #d #HL #H @@ -50,13 +51,13 @@ lemma lsx_cpx_trans_lcosx: ∀h,g,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ➡[h, g] T2 → elim (lsx_inv_flat … H) -H #HV1 #H elim (lsx_inv_bind … H) -H #HW1 #HT1 @lsx_bind /3 width=1 by lsx_flat/ (**) (* explicit constructor *) - @(lsx_leqy_conf … (L.ⓛW1)) /3 width=1 by lcosx_pair, lsuby_succ/ + @(lsx_leq_conf … (L.ⓛW1)) /3 width=1 by lcosx_pair, leq_succ/ | #a #G #L #V1 #V2 #U2 #W1 #W2 #T1 #T2 #_ #HVU2 #_ #_ #IHV12 #IHW12 #IHT12 #d #HL #H elim (lsx_inv_flat … H) -H #HV1 #H elim (lsx_inv_bind … H) -H #HW1 #HT1 @lsx_bind /2 width=1 by/ (**) (* explicit constructor *) @lsx_flat [ /3 width=7 by lsx_lift_ge, ldrop_drop/ ] - @(lsx_leqy_conf … (L.ⓓW1)) /3 width=1 by lcosx_pair, lsuby_succ/ + @(lsx_leq_conf … (L.ⓓW1)) /3 width=1 by lcosx_pair, leq_succ/ ] qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma index 2ca1e193d..e21c79912 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lpxs_lleq.ma @@ -12,8 +12,9 @@ (* *) (**************************************************************************) +include "basic_2/relocation/lleq_leq.ma". include "basic_2/reduction/lpx_lleq.ma". -include "basic_2/computation/cpxs_cpys.ma". +include "basic_2/computation/cpxs_leq.ma". include "basic_2/computation/lpxs_ldrop.ma". include "basic_2/computation/lpxs_cpxs.ma". @@ -37,7 +38,7 @@ lemma lpxs_lleq_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, [ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpxs_inv_pair2 … H1) -H1 #K0 #V0 #H1KL1 #_ #H destruct elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 // - #I1 #K1 #H #H2KL1 lapply (ldrop_inv_O2 … H) -H #H destruct + #K1 #H #H2KL1 lapply (ldrop_inv_O2 … H) -H #H destruct /2 width=4 by fqu_lref_O, ex3_intro/ | * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H [ elim (lleq_inv_bind … H) @@ -93,36 +94,32 @@ elim (fqus_inv_gen … H) -H ] qed-. -fact lsuby_lpxs_trans_lleq_aux: ∀h,g,G,L1,L0,d,e. L1 ⊑×[d, e] L0 → e = ∞ → - ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 → - ∃∃L. L ⊑×[d, e] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L & - (∀T. |L1| = |L0| → |L| = |L2| → L0 ⋕[T, d] L2 ↔ L1 ⋕[T, d] L). +fact leq_lpxs_trans_lleq_aux: ∀h,g,G,L1,L0,d,e. L1 ≃[d, e] L0 → e = ∞ → + ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 → + ∃∃L. L ≃[d, e] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L & + (∀T. L0 ⋕[T, d] L2 ↔ L1 ⋕[T, d] L). #h #g #G #L1 #L0 #d #e #H elim H -L1 -L0 -d -e -[ #L1 #d #e #_ #L2 #H >(lpxs_inv_atom1 … H) -H +[ #d #e #_ #L2 #H >(lpxs_inv_atom1 … H) -H /3 width=5 by ex3_intro, conj/ | #I1 #I0 #L1 #L0 #V1 #V0 #_ #_ #He destruct -| #I1 #I0 #L1 #L0 #V1 #e #HL10 #IHL10 #He #Y #H +| #I #L1 #L0 #V1 #e #HL10 #IHL10 #He #Y #H elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct lapply (ysucc_inv_Y_dx … He) -He #He elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH - @(ex3_intro … (L.ⓑ{I1}V2)) /3 width=3 by lpxs_pair, lsuby_cpxs_trans, lsuby_pair/ - #T #H1 #H2 lapply (injective_plus_l … H1) lapply (injective_plus_l … H2) -H1 -H2 - #H1 #H2 elim (IH T) // #HL0dx #HL0sn - @conj #H @(lleq_lsuby_repl … H) -H normalize - /3 width=1 by lsuby_sym, lsuby_pair_O_Y/ + @(ex3_intro … (L.ⓑ{I}V2)) /3 width=3 by lpxs_pair, leq_cpxs_trans, leq_pair/ + #T elim (IH T) #HL0dx #HL0sn + @conj #H @(lleq_leq_repl … H) -H /3 width=1 by leq_sym, leq_pair_O_Y/ | #I1 #I0 #L1 #L0 #V1 #V0 #d #e #HL10 #IHL10 #He #Y #H elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL02 #HV02 #H destruct elim (IHL10 … HL02) // -IHL10 -HL02 #L #HL2 #HL1 #IH - @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpxs_pair, lsuby_succ/ - #T #H1 #H2 lapply (injective_plus_l … H1) lapply (injective_plus_l … H2) -H1 -H2 - #H1 #H2 elim (IH T) // #HL0dx #HL0sn - @conj #H @(lleq_lsuby_repl … H) -H - /3 width=1 by lsuby_sym, lsuby_succ/ normalize // + @(ex3_intro … (L.ⓑ{I1}V1)) /3 width=1 by lpxs_pair, leq_succ/ + #T elim (IH T) #HL0dx #HL0sn + @conj #H @(lleq_leq_repl … H) -H /3 width=1 by leq_sym, leq_succ/ ] qed-. -lemma lsuby_lpxs_trans_lleq: ∀h,g,G,L1,L0,d. L1 ⊑×[d, ∞] L0 → - ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 → - ∃∃L. L ⊑×[d, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L & - (∀T. |L1| = |L0| → |L| = |L2| → L0 ⋕[T, d] L2 ↔ L1 ⋕[T, d] L). -/2 width=1 by lsuby_lpxs_trans_lleq_aux/ qed-. +lemma leq_lpxs_trans_lleq: ∀h,g,G,L1,L0,d. L1 ≃[d, ∞] L0 → + ∀L2. ⦃G, L0⦄ ⊢ ➡*[h, g] L2 → + ∃∃L. L ≃[d, ∞] L2 & ⦃G, L1⦄ ⊢ ➡*[h, g] L & + (∀T. L0 ⋕[T, d] L2 ↔ L1 ⋕[T, d] L). +/2 width=1 by leq_lpxs_trans_lleq_aux/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lsx.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lsx.ma index 5307af43e..59dda4e82 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lsx.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lsx.ma @@ -13,7 +13,7 @@ (**************************************************************************) include "basic_2/notation/relations/lazysn_6.ma". -include "basic_2/substitution/lleq.ma". +include "basic_2/relocation/lleq.ma". include "basic_2/computation/lpxs.ma". (* SN EXTENDED STRONGLY NORMALIZING LOCAL ENVIRONMENTS **********************) @@ -62,10 +62,10 @@ lemma lsx_gref: ∀h,g,G,L,d,p. G ⊢ ⋕⬊*[h, g, §p, d] L. /3 width=4 by lpxs_fwd_length, lleq_gref/ qed. -lemma lsx_be: ∀h,g,G,L,T,U,dt,d,e. yinj d ≤ dt → dt ≤ d + e → - ⇧[d, e] T ≡ U → G ⊢ ⋕⬊*[h, g, U, dt] L → G ⊢ ⋕⬊*[h, g, U, d] L. -#h #g #G #L #T #U #dt #d #e #Hddt #Hdtde #HTU #H @(lsx_ind … H) -L -/5 width=7 by lsx_intro, lleq_be/ +lemma lsx_ge_up: ∀h,g,G,L,T,U,dt,d,e. dt ≤ yinj d + yinj e → + ⇧[d, e] T ≡ U → G ⊢ ⋕⬊*[h, g, U, dt] L → G ⊢ ⋕⬊*[h, g, U, d] L. +#h #g #G #L #T #U #dt #d #e #Hdtde #HTU #H @(lsx_ind … H) -L +/5 width=7 by lsx_intro, lleq_ge_up/ qed-. (* Basic forward lemmas *****************************************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lsx_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lsx_ldrop.ma index 61466ca35..9bd29a5f1 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lsx_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lsx_ldrop.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/substitution/lleq_ldrop.ma". +include "basic_2/relocation/lleq_ldrop.ma". include "basic_2/computation/lpxs_ldrop.ma". include "basic_2/computation/lsx.ma". diff --git a/matita/matita/contribs/lambdadelta/basic_2/computation/lsx_lpxs.ma b/matita/matita/contribs/lambdadelta/basic_2/computation/lsx_lpxs.ma index b3cb36781..9762de28b 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/computation/lsx_lpxs.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/computation/lsx_lpxs.ma @@ -12,6 +12,7 @@ (* *) (**************************************************************************) +include "basic_2/relocation/lleq_lleq.ma". include "basic_2/computation/lpxs_lleq.ma". include "basic_2/computation/lpxs_lpxs.ma". include "basic_2/computation/lsx.ma". @@ -20,15 +21,12 @@ include "basic_2/computation/lsx.ma". (* Advanced properties ******************************************************) -lemma lsx_leqy_conf: ∀h,g,G,L1,T,d. G ⊢ ⋕⬊*[h, g, T, d] L1 → - ∀L2. L1 ⊑×[d, ∞] L2 → |L1| = |L2| → G ⊢ ⋕⬊*[h, g, T, d] L2. +lemma lsx_leq_conf: ∀h,g,G,L1,T,d. G ⊢ ⋕⬊*[h, g, T, d] L1 → + ∀L2. L1 ≃[d, ∞] L2 → G ⊢ ⋕⬊*[h, g, T, d] L2. #h #g #G #L1 #T #d #H @(lsx_ind … H) -L1 -#L1 #_ #IHL1 #L2 #H1L12 #H2L12 @lsx_intro -#L3 #H1L23 #HnL23 lapply (lpxs_fwd_length … H1L23) -#H2L23 elim (lsuby_lpxs_trans_lleq … H1L12 … H1L23) -H1L12 -H1L23 -#L0 #H1L03 #H1L10 #H lapply (lpxs_fwd_length … H1L10) -#H2L10 elim (H T) -H // -#_ #H @(IHL1 … H1L10) -IHL1 -H1L10 /3 width=1 by/ +#L1 #_ #IHL1 #L2 #HL12 @lsx_intro +#L3 #HL23 #HnL23 elim (leq_lpxs_trans_lleq … HL12 … HL23) -HL12 -HL23 +#L0 #HL03 #HL10 #H elim (H T) -H /4 width=4 by/ qed-. lemma lsx_ge: ∀h,g,G,L,T,d1,d2. d1 ≤ d2 → @@ -113,10 +111,10 @@ lemma lsx_fwd_bind_dx: ∀h,g,a,I,G,L,V,T,d. G ⊢ ⋕⬊*[h, g, ⓑ{a,I}V.T, d] #L1 #_ #IHL1 @lsx_intro #Y #H #HT elim (lpxs_inv_pair1 … H) -H #L2 #V2 #HL12 #_ #H destruct -@(lsx_leqy_conf … (L2.ⓑ{I}V1)) /2 width=1 by lsuby_succ/ +@(lsx_leq_conf … (L2.ⓑ{I}V1)) /2 width=1 by leq_succ/ @IHL1 // #H @HT -IHL1 -HL12 -HT -@(lleq_lsuby_trans … (L2.ⓑ{I}V1)) -/2 width=2 by lleq_fwd_bind_dx, lsuby_succ/ +@(lleq_leq_trans … (L2.ⓑ{I}V1)) +/2 width=2 by lleq_fwd_bind_dx, leq_succ/ qed-. (* Advanced inversion lemmas ************************************************) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpx_cpys.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpx_cpys.etc new file mode 100644 index 000000000..3b38ee6ab --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpx_cpys.etc @@ -0,0 +1,42 @@ +(**************************************************************************) +(* ___ *) +(* ||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/cpys_alt.ma". +include "basic_2/reduction/cpx.ma". + +(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************) + +(* Properties on local environment refinement for extended substitution *****) + +lemma lsuby_cpx_trans: ∀h,g,G. lsub_trans … (cpx h g G) (lsuby 0 (∞)). +#h #g #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 +[ // +| /2 width=2 by cpx_sort/ +| #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 + elim (lsuby_ldrop_trans_be … HL12 … HLK1) // -HL12 -HLK1 /3 width=7 by cpx_delta/ +|4,9: /4 width=1 by cpx_bind, cpx_beta, lsuby_pair_O_Y/ +|5,7,8: /3 width=1 by cpx_flat, cpx_tau, cpx_ti/ +|6,10: /4 width=3 by cpx_zeta, cpx_theta, lsuby_pair_O_Y/ +] +qed-. + +(* Properties on context-sensitive extended multiple substitution for terms *) + +lemma cpys_cpx: ∀h,g,G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2. +#h #g #G #L #T1 #T2 #d #e #H @(cpys_ind_alt … H) -G -L -T1 -T2 -d -e +/2 width=7 by cpx_delta, cpx_bind, cpx_flat/ +qed. + +lemma cpy_cpx: ∀h,g,G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2. +/3 width=3 by cpy_cpys, cpys_cpx/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpxs_cpys.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpxs_cpys.etc new file mode 100644 index 000000000..881030572 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpxs_cpys.etc @@ -0,0 +1,24 @@ +(**************************************************************************) +(* ___ *) +(* ||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/reduction/cpx_cpys.ma". +include "basic_2/computation/cpxs.ma". + +(* CONTEXT-SENSITIVE EXTENDED PARALLEL COMPUTATION ON TERMS *****************) + +(* Properties on local environment refinement for extended substitution *****) + +lemma lsuby_cpxs_trans: ∀h,g,G. lsub_trans … (cpxs h g G) (lsuby 0 (∞)). +/3 width=5 by lsuby_cpx_trans, LTC_lsub_trans/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy.etc new file mode 100644 index 000000000..a1e538ae0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy.etc @@ -0,0 +1,296 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/ynat/ynat_max.ma". +include "basic_2/notation/relations/psubst_6.ma". +include "basic_2/grammar/genv.ma". +include "basic_2/relocation/lsuby.ma". + +(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) + +(* activate genv *) +inductive cpy: ynat → ynat → relation4 genv lenv term term ≝ +| cpy_atom : ∀I,G,L,d,e. cpy d e G L (⓪{I}) (⓪{I}) +| cpy_subst: ∀I,G,L,K,V,W,i,d,e. d ≤ yinj i → i < d+e → + ⇩[i] L ≡ K.ⓑ{I}V → ⇧[0, i+1] V ≡ W → cpy d e G L (#i) W +| cpy_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e. + cpy d e G L V1 V2 → cpy (⫯d) e G (L.ⓑ{I}V1) T1 T2 → + cpy d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) +| cpy_flat : ∀I,G,L,V1,V2,T1,T2,d,e. + cpy d e G L V1 V2 → cpy d e G L T1 T2 → + cpy d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) +. + +interpretation "context-sensitive extended ordinary substritution (term)" + 'PSubst G L T1 d e T2 = (cpy d e G L T1 T2). + +(* Basic properties *********************************************************) + +lemma lsuby_cpy_trans: ∀G,d,e. lsub_trans … (cpy d e G) (lsuby d e). +#G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e +[ // +| #I #G #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12 + elim (lsuby_ldrop_trans_be … HL12 … HLK1) -HL12 -HLK1 /2 width=5 by cpy_subst/ +| /4 width=1 by lsuby_succ, cpy_bind/ +| /3 width=1 by cpy_flat/ +] +qed-. + +lemma cpy_refl: ∀G,T,L,d,e. ⦃G, L⦄ ⊢ T ▶[d, e] T. +#G #T elim T -T // * /2 width=1 by cpy_bind, cpy_flat/ +qed. + +(* Basic_1: was: subst1_ex *) +lemma cpy_full: ∀I,G,K,V,T1,L,d. ⇩[d] L ≡ K.ⓑ{I}V → + ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ▶[d, 1] T2 & ⇧[d, 1] T ≡ T2. +#I #G #K #V #T1 elim T1 -T1 +[ * #i #L #d #HLK + /2 width=4 by lift_sort, lift_gref, ex2_2_intro/ + elim (lt_or_eq_or_gt i d) #Hid + /3 width=4 by lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ + destruct + elim (lift_total V 0 (i+1)) #W #HVW + elim (lift_split … HVW i i) + /4 width=5 by cpy_subst, ylt_inj, ex2_2_intro/ +| * [ #a ] #J #W1 #U1 #IHW1 #IHU1 #L #d #HLK + elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 + [ elim (IHU1 (L.ⓑ{J}W1) (d+1)) -IHU1 + /3 width=9 by cpy_bind, ldrop_drop, lift_bind, ex2_2_intro/ + | elim (IHU1 … HLK) -IHU1 -HLK + /3 width=8 by cpy_flat, lift_flat, ex2_2_intro/ + ] +] +qed-. + +lemma cpy_weak: ∀G,L,T1,T2,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → + ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T2. +#G #L #T1 #T2 #d1 #e1 #H elim H -G -L -T1 -T2 -d1 -e1 // +[ /3 width=5 by cpy_subst, ylt_yle_trans, yle_trans/ +| /4 width=3 by cpy_bind, ylt_yle_trans, yle_succ/ +| /3 width=1 by cpy_flat/ +] +qed-. + +lemma cpy_weak_top: ∀G,L,T1,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶[d, |L| - d] T2. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e // +[ #I #G #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW + lapply (ldrop_fwd_length_lt2 … HLK) + /4 width=5 by cpy_subst, ylt_yle_trans, ylt_inj/ +| #a #I #G #L #V1 #V2 normalize in match (|L.ⓑ{I}V2|); (**) (* |?| does not work *) + /2 width=1 by cpy_bind/ +| /2 width=1 by cpy_flat/ +] +qed-. + +lemma cpy_weak_full: ∀G,L,T1,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶[0, |L|] T2. +#G #L #T1 #T2 #d #e #HT12 +lapply (cpy_weak … HT12 0 (d + e) ? ?) -HT12 +/2 width=2 by cpy_weak_top/ +qed-. + +lemma cpy_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀i. i ≤ d + e → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d, i-d] T & ⦃G, L⦄ ⊢ T ▶[i, d+e-i] T2. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e +[ /2 width=3 by ex2_intro/ +| #I #G #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hjde + elim (ylt_split i j) [ -Hide -Hjde | -Hdi ] + /4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/ +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide + elim (IHV12 i) -IHV12 // #V + elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hide + >yplus_SO2 >yplus_succ1 #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 + /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide + elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hide + /3 width=5 by ex2_intro, cpy_flat/ +] +qed-. + +lemma cpy_split_down: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀i. i ≤ d + e → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[i, d+e-i] T & ⦃G, L⦄ ⊢ T ▶[d, i-d] T2. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e +[ /2 width=3 by ex2_intro/ +| #I #G #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hjde + elim (ylt_split i j) [ -Hide -Hjde | -Hdi ] + /4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/ +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide + elim (IHV12 i) -IHV12 // #V + elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hide + >yplus_SO2 >yplus_succ1 #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 + /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide + elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hide + /3 width=5 by ex2_intro, cpy_flat/ +] +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma cpy_fwd_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → + ∀T1,d,e. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → d + e ≤ dt + et → + ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶[d+e, dt+et-(d+e)] U2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et +[ * #i #G #L #dt #et #T1 #d #e #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ + ] +| #I #G #L #K #V #W #i #dt #et #Hdti #Hidet #HLK #HVW #T1 #d #e #H #Hddt #Hdedet + elim (lift_inv_lref2 … H) -H * #Hid #H destruct [ -V -Hidet -Hdedet | -Hdti -Hddt ] + [ elim (ylt_yle_false … Hddt) -Hddt /3 width=3 by yle_ylt_trans, ylt_inj/ + | elim (le_inv_plus_l … Hid) #Hdie #Hei + elim (lift_split … HVW d (i-e+1) ? ? ?) [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hdie + #T2 #_ >plus_minus // ymax_pre_sn_comm // (**) (* explicit constructor *) + ] +| #a #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #X #d #e #H #Hddt #Hdedet + elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HVW1) -V1 -IHW12 // + elim (IHU12 … HTU1) -T1 -IHU12 /2 width=1 by yle_succ/ + yplus_SO2 >yplus_succ1 >yplus_succ1 + /3 width=2 by cpy_bind, lift_bind, ex2_intro/ +| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #X #d #e #H #Hddt #Hdedet + elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HVW1) -V1 -IHW12 // elim (IHU12 … HTU1) -T1 -IHU12 + /3 width=2 by cpy_flat, lift_flat, ex2_intro/ +] +qed-. + +lemma cpy_fwd_tw: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ♯{T1} ≤ ♯{T2}. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e normalize +/3 width=1 by monotonic_le_plus_l, le_plus/ +qed-. + +(* Basic inversion lemmas ***************************************************) + +fact cpy_inv_atom1_aux: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀J. T1 = ⓪{J} → + T2 = ⓪{J} ∨ + ∃∃I,K,V,i. d ≤ yinj i & i < d + e & + ⇩[i] L ≡ K.ⓑ{I}V & + ⇧[O, i+1] V ≡ T2 & + J = LRef i. +#G #L #T1 #T2 #d #e * -G -L -T1 -T2 -d -e +[ #I #G #L #d #e #J #H destruct /2 width=1 by or_introl/ +| #I #G #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #J #H destruct /3 width=9 by ex5_4_intro, or_intror/ +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct +] +qed-. + +lemma cpy_inv_atom1: ∀I,G,L,T2,d,e. ⦃G, L⦄ ⊢ ⓪{I} ▶[d, e] T2 → + T2 = ⓪{I} ∨ + ∃∃J,K,V,i. d ≤ yinj i & i < d + e & + ⇩[i] L ≡ K.ⓑ{J}V & + ⇧[O, i+1] V ≡ T2 & + I = LRef i. +/2 width=4 by cpy_inv_atom1_aux/ qed-. + +(* Basic_1: was: subst1_gen_sort *) +lemma cpy_inv_sort1: ∀G,L,T2,k,d,e. ⦃G, L⦄ ⊢ ⋆k ▶[d, e] T2 → T2 = ⋆k. +#G #L #T2 #k #d #e #H +elim (cpy_inv_atom1 … H) -H // +* #I #K #V #i #_ #_ #_ #_ #H destruct +qed-. + +(* Basic_1: was: subst1_gen_lref *) +lemma cpy_inv_lref1: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶[d, e] T2 → + T2 = #i ∨ + ∃∃I,K,V. d ≤ i & i < d + e & + ⇩[i] L ≡ K.ⓑ{I}V & + ⇧[O, i+1] V ≡ T2. +#G #L #T2 #i #d #e #H +elim (cpy_inv_atom1 … H) -H /2 width=1 by or_introl/ +* #I #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=5 by ex4_3_intro, or_intror/ +qed-. + +lemma cpy_inv_gref1: ∀G,L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶[d, e] T2 → T2 = §p. +#G #L #T2 #p #d #e #H +elim (cpy_inv_atom1 … H) -H // +* #I #K #V #i #_ #_ #_ #_ #H destruct +qed-. + +fact cpy_inv_bind1_aux: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 → + ∀a,I,V1,T1. U1 = ⓑ{a,I}V1.T1 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 & + ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ▶[⫯d, e] T2 & + U2 = ⓑ{a,I}V2.T2. +#G #L #U1 #U2 #d #e * -G -L -U1 -U2 -d -e +[ #I #G #L #d #e #b #J #W1 #U1 #H destruct +| #I #G #L #K #V #W #i #d #e #_ #_ #_ #_ #b #J #W1 #U1 #H destruct +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #b #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ +| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #b #J #W1 #U1 #H destruct +] +qed-. + +lemma cpy_inv_bind1: ∀a,I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓑ{a,I} V1. T1 ▶[d, e] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 & + ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶[⫯d, e] T2 & + U2 = ⓑ{a,I}V2.T2. +/2 width=3 by cpy_inv_bind1_aux/ qed-. + +fact cpy_inv_flat1_aux: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 → + ∀I,V1,T1. U1 = ⓕ{I}V1.T1 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 & + ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 & + U2 = ⓕ{I}V2.T2. +#G #L #U1 #U2 #d #e * -G -L -U1 -U2 -d -e +[ #I #G #L #d #e #J #W1 #U1 #H destruct +| #I #G #L #K #V #W #i #d #e #_ #_ #_ #_ #J #W1 #U1 #H destruct +| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #W1 #U1 #H destruct +| #I #G #L #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ +] +qed-. + +lemma cpy_inv_flat1: ∀I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ▶[d, e] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 & + ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 & + U2 = ⓕ{I}V2.T2. +/2 width=3 by cpy_inv_flat1_aux/ qed-. + + +fact cpy_inv_refl_O2_aux: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → e = 0 → T1 = T2. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e +[ // +| #I #G #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct + elim (ylt_yle_false … Hdi) -Hdi // +| /3 width=1 by eq_f2/ +| /3 width=1 by eq_f2/ +] +qed-. + +lemma cpy_inv_refl_O2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶[d, 0] T2 → T1 = T2. +/2 width=6 by cpy_inv_refl_O2_aux/ qed-. + +(* Basic_1: was: subst1_gen_lift_eq *) +lemma cpy_inv_lift1_eq: ∀G,T1,U1,d,e. ⇧[d, e] T1 ≡ U1 → + ∀L,U2. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 → U1 = U2. +#G #T1 #U1 #d #e #HTU1 #L #U2 #HU12 elim (cpy_fwd_up … HU12 … HTU1) -HU12 -HTU1 +/2 width=4 by cpy_inv_refl_O2/ +qed-. + +(* Basic_1: removed theorems 25: + subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt + subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans + subst0_lift_lt subst0_lift_ge subst0_lift_ge_S subst0_lift_ge_s + subst0_subst0 subst0_subst0_back subst0_weight_le subst0_weight_lt + subst0_confluence_neq subst0_confluence_eq subst0_tlt_head + subst0_confluence_lift subst0_tlt + subst1_head subst1_gen_head subst1_lift_S subst1_confluence_lift +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy_cpy.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy_cpy.etc new file mode 100644 index 000000000..66b0487db --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy_cpy.etc @@ -0,0 +1,122 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/cpy_lift.ma". + +(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) + +(* Main properties **********************************************************) + +(* Basic_1: was: subst1_confluence_eq *) +theorem cpy_conf_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶[d1, e1] T1 → + ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶[d1, e1] T. +#G #L #T0 #T1 #d1 #e1 #H elim H -G -L -T0 -T1 -d1 -e1 +[ /2 width=3 by ex2_intro/ +| #I1 #G #L #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #T2 #d2 #e2 #H + elim (cpy_inv_lref1 … H) -H + [ #HX destruct /3 width=7 by cpy_subst, ex2_intro/ + | -Hd1 -Hde1 * #I2 #K2 #V2 #_ #_ #HLK2 #HVT2 + lapply (ldrop_mono … HLK1 … HLK2) -HLK1 -HLK2 #H destruct + >(lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3 by ex2_intro/ + ] +| #a #I #G #L #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX + elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02) -IHV01 -HV02 #V #HV1 #HV2 + elim (IHT01 … HT02) -T0 #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT1 (L.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/ + lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V2) ?) -HT2 + /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/ +| #I #G #L #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX + elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02) -V0 + elim (IHT01 … HT02) -T0 /3 width=5 by cpy_flat, ex2_intro/ +] +qed-. + +(* Basic_1: was: subst1_confluence_neq *) +theorem cpy_conf_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶[d1, e1] T1 → + ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶[d2, e2] T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶[d1, e1] T. +#G #L1 #T0 #T1 #d1 #e1 #H elim H -G -L1 -T0 -T1 -d1 -e1 +[ /2 width=3 by ex2_intro/ +| #I1 #G #L1 #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #L2 #T2 #d2 #e2 #H1 #H2 + elim (cpy_inv_lref1 … H1) -H1 + [ #H destruct /3 width=7 by cpy_subst, ex2_intro/ + | -HLK1 -HVT1 * #I2 #K2 #V2 #Hd2 #Hde2 #_ #_ elim H2 -H2 #Hded [ -Hd1 -Hde2 | -Hd2 -Hde1 ] + [ elim (ylt_yle_false … Hde1) -Hde1 /2 width=3 by yle_trans/ + | elim (ylt_yle_false … Hde2) -Hde2 /2 width=3 by yle_trans/ + ] + ] +| #a #I #G #L1 #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H + elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -IHV01 -HV02 #V #HV1 #HV2 + elim (IHT01 … HT02) -T0 + [ -H #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT1 (L2.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/ + lapply (lsuby_cpy_trans … HT2 (L1.ⓑ{I}V2) ?) -HT2 /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/ + | -HV1 -HV2 elim H -H /3 width=1 by yle_succ, or_introl, or_intror/ + ] +| #I #G #L1 #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H + elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV01 … HV02 H) -V0 + elim (IHT01 … HT02 H) -T0 -H /3 width=5 by cpy_flat, ex2_intro/ +] +qed-. + +(* Note: the constant 1 comes from cpy_subst *) +(* Basic_1: was: subst1_trans *) +theorem cpy_trans_ge: ∀G,L,T1,T0,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T0 → + ∀T2. ⦃G, L⦄ ⊢ T0 ▶[d, 1] T2 → 1 ≤ e → ⦃G, L⦄ ⊢ T1 ▶[d, e] T2. +#G #L #T1 #T0 #d #e #H elim H -G -L -T1 -T0 -d -e +[ #I #G #L #d #e #T2 #H #He + elim (cpy_inv_atom1 … H) -H + [ #H destruct // + | * #J #K #V #i #Hd2i #Hide2 #HLK #HVT2 #H destruct + lapply (ylt_yle_trans … (d+e) … Hide2) /2 width=5 by cpy_subst, monotonic_yle_plus_dx/ + ] +| #I #G #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He + lapply (cpy_weak … HVT2 0 (i+1) ? ?) -HVT2 /3 width=1 by yle_plus_dx2_trans, yle_succ/ + >yplus_inj #HVT2 <(cpy_inv_lift1_eq … HVW … HVT2) -HVT2 /2 width=5 by cpy_subst/ +| #a #I #G #L #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He + elim (cpy_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct + lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02 + lapply (IHT10 … HT02 He) -T0 /3 width=1 by cpy_bind/ +| #I #G #L #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He + elim (cpy_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1 by cpy_flat/ +] +qed-. + +theorem cpy_trans_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T0 → + ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → d2 + e2 ≤ d1 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T ▶[d1, e1] T2. +#G #L #T1 #T0 #d1 #e1 #H elim H -G -L -T1 -T0 -d1 -e1 +[ /2 width=3 by ex2_intro/ +| #I #G #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1 + lapply (yle_trans … Hde2d1 … Hdi1) -Hde2d1 #Hde2i1 + lapply (cpy_weak … HWT2 0 (i1+1) ? ?) -HWT2 /3 width=1 by yle_succ, yle_pred_sn/ -Hde2i1 + >yplus_inj #HWT2 <(cpy_inv_lift1_eq … HVW … HWT2) -HWT2 /3 width=9 by cpy_subst, ex2_intro/ +| #a #I #G #L #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02 + elim (IHV10 … HV02) -IHV10 -HV02 // #V + elim (IHT10 … HT02) -T0 /2 width=1 by yle_succ/ #T #HT1 #HT2 + lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 /3 width=6 by cpy_bind, lsuby_succ, ex2_intro/ +| #I #G #L #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 + elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct + elim (IHV10 … HV02) -V0 // + elim (IHT10 … HT02) -T0 /3 width=6 by cpy_flat, ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy_lift.etc new file mode 100644 index 000000000..aa0b2416a --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpy_lift.etc @@ -0,0 +1,249 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/ldrop_ldrop.ma". +include "basic_2/relocation/cpy.ma". + +(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) + +(* Properties on relocation *************************************************) + +(* Basic_1: was: subst1_lift_lt *) +lemma cpy_lift_le: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 → + ∀L,U1,U2,s,d,e. ⇩[s, d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + dt + et ≤ d → ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2. +#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et +[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hdetd + lapply (ylt_yle_trans … Hdetd … Hidet) -Hdetd #Hid + lapply (ylt_inv_inj … Hid) -Hid #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=5 by cpy_subst/ +| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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 + /4 width=7 by cpy_bind, ldrop_skip, yle_succ/ +| #G #I #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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=7 by cpy_flat/ +] +qed-. + +lemma cpy_lift_be: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 → + ∀L,U1,U2,s,d,e. ⇩[s, d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + dt ≤ d → d ≤ dt + et → ⦃G, L⦄ ⊢ U1 ▶[dt, et + e] U2. +#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et +[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_ #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hdtd #_ + elim (lift_inv_lref1 … H) -H * #Hid #H destruct + [ -Hdtd + lapply (ylt_yle_trans … (dt+et+e) … Hidet) // -Hidet #Hidete + elim (lift_trans_ge … HVW … HWU2) -W // (lift_mono … HVY … HVW) -V #H destruct /2 width=5 by cpy_subst/ + | -Hdti + elim (yle_inv_inj2 … Hdtd) -Hdtd #dtt #Hdtd #H destruct + lapply (transitive_le … Hdtd Hid) -Hdtd #Hdti + lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid + /4 width=5 by cpy_subst, ldrop_inv_gen, monotonic_ylt_plus_dx, yle_plus_dx1_trans, yle_inj/ + ] +| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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 + /4 width=7 by cpy_bind, ldrop_skip, yle_succ/ +| #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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=7 by cpy_flat/ +] +qed-. + +(* Basic_1: was: subst1_lift_ge *) +lemma cpy_lift_ge: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 → + ∀L,U1,U2,s,d,e. ⇩[s, d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → + d ≤ dt → ⦃G, L⦄ ⊢ U1 ▶[dt+e, et] U2. +#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et +[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_ + >(lift_mono … H1 … H2) -H1 -H2 // +| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hddt + lapply (yle_trans … Hddt … Hdti) -Hddt #Hid + elim (yle_inv_inj2 … Hid) -Hid #dd #Hddi #H0 destruct + lapply (lift_inv_lref1_ge … H … Hddi) -H #H destruct + lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2 + lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hddi + /3 width=5 by cpy_subst, ldrop_inv_gen, monotonic_ylt_plus_dx, monotonic_yle_plus_dx/ +| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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 + /4 width=6 by cpy_bind, ldrop_skip, yle_succ/ +| #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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=6 by cpy_flat/ +] +qed-. + +(* Inversion lemmas on relocation *******************************************) + +(* Basic_1: was: subst1_gen_lift_lt *) +lemma cpy_inv_lift1_le: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt + et ≤ d → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et +[ * #i #G #L #dt #et #K #s #d #e #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ + ] +| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #s #d #e #HLK #T1 #H #Hdetd + lapply (ylt_yle_trans … Hdetd … Hidet) -Hdetd #Hid + lapply (ylt_inv_inj … Hid) -Hid #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 yplus_minus_assoc_inj /2 width=1 by yle_plus_to_minus_inj2/ ] -Hdedet #Hidete + elim (ldrop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV + elim (lift_trans_le … HUV … HVW) -V // >minus_plus plus_minus // yplus_minus_assoc_inj /3 width=1 by monotonic_ylt_minus_dx, yle_inj/ + ] +| #a #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdtd #Hdedet + elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2 + elim (IHU12 … HTU1) -U1 + /3 width=6 by cpy_bind, ldrop_skip, lift_bind, yle_succ, ex2_intro/ +| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdtd #Hdedet + elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HLK … HVW1) -W1 // + elim (IHU12 … HLK … HTU1) -U1 -HLK // + /3 width=5 by cpy_flat, lift_flat, ex2_intro/ +] +qed-. + +(* Basic_1: was: subst1_gen_lift_ge *) +lemma cpy_inv_lift1_ge: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + yinj d + e ≤ dt → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt-e, et] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et +[ * #i #G #L #dt #et #K #s #d #e #_ #T1 #H #_ + [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ + | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ + | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ + ] +| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #s #d #e #HLK #T1 #H #Hdedt + lapply (yle_trans … Hdedt … Hdti) #Hdei + elim (yle_inv_plus_inj2 … Hdedt) -Hdedt #_ #Hedt + elim (yle_inv_plus_inj2 … Hdei) #Hdie #Hei + lapply (lift_inv_lref2_ge … H ?) -H /2 width=1 by yle_inv_inj/ #H destruct + lapply (ldrop_conf_ge … HLK … HLKV ?) -L /2 width=1 by yle_inv_inj/ #HKV + elim (lift_split … HVW d (i-e+1)) -HVW [2,3,4: /3 width=1 by yle_inv_inj, le_S_S, le_S/ ] -Hdei -Hdie + #V0 #HV10 >plus_minus /2 width=1 by yle_inv_inj/ yminus_succ1_inj /3 width=5 by cpy_bind, lift_bind, ex2_intro/ +| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdetd + elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct + elim (IHW12 … HLK … HVW1) -W1 // + elim (IHU12 … HLK … HTU1) -U1 -HLK /3 width=5 by cpy_flat, lift_flat, ex2_intro/ +] +qed-. + +(* Advancd inversion lemmas on relocation ***********************************) + +lemma cpy_inv_lift1_ge_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[d, dt + et - (yinj d + e)] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet +elim (cpy_split_up … HU12 (d + e)) -HU12 // -Hdedet #U #HU1 #HU2 +lapply (cpy_weak … HU1 d e ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hddt -Hdtde #HU1 +lapply (cpy_inv_lift1_eq … HTU1 … HU1) -HU1 #HU1 destruct +elim (cpy_inv_lift1_ge … HU2 … HLK … HTU1) -U -L /2 width=3 by ex2_intro/ +qed-. + +lemma cpy_inv_lift1_be_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → dt + et ≤ yinj d + e → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, d-dt] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde +lapply (cpy_weak … HU12 dt (d+e-dt) ? ?) -HU12 // +[ >ymax_pre_sn_comm /2 width=1 by yle_plus_dx1_trans/ ] -Hdetde #HU12 +elim (cpy_inv_lift1_be … HU12 … HLK … HTU1) -U1 -L /2 width=3 by ex2_intro/ +qed-. + +lemma cpy_inv_lift1_le_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → d ≤ dt + et → dt + et ≤ yinj d + e → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hddet #Hdetde +elim (cpy_split_up … HU12 d) -HU12 // #U #HU1 #HU2 +elim (cpy_inv_lift1_le … HU1 … HLK … HTU1) -U1 +[2: >ymax_pre_sn_comm // ] -Hdtd #T #HT1 #HTU +lapply (cpy_weak … HU2 d e ? ?) -HU2 // +[ >ymax_pre_sn_comm // ] -Hddet -Hdetde #HU2 +lapply (cpy_inv_lift1_eq … HTU … HU2) -L #H destruct /2 width=3 by ex2_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys.etc new file mode 100644 index 000000000..10c59120d --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys.etc @@ -0,0 +1,166 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/psubststar_6.ma". +include "basic_2/relocation/cpy.ma". + +(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************) + +definition cpys: ynat → ynat → relation4 genv lenv term term ≝ + λd,e,G. LTC … (cpy d e G). + +interpretation "context-sensitive extended multiple substritution (term)" + 'PSubstStar G L T1 d e T2 = (cpys d e G L T1 T2). + +(* Basic eliminators ********************************************************) + +lemma cpys_ind: ∀G,L,T1,d,e. ∀R:predicate term. R T1 → + (∀T,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶[d, e] T2 → R T → R T2) → + ∀T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R T2. +#G #L #T1 #d #e #R #HT1 #IHT1 #T2 #HT12 +@(TC_star_ind … HT1 IHT1 … HT12) // +qed-. + +lemma cpys_ind_dx: ∀G,L,T2,d,e. ∀R:predicate term. R T2 → + (∀T1,T. ⦃G, L⦄ ⊢ T1 ▶[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 → R T → R T1) → + ∀T1. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R T1. +#G #L #T2 #d #e #R #HT2 #IHT2 #T1 #HT12 +@(TC_star_ind_dx … HT2 IHT2 … HT12) // +qed-. + +(* Basic properties *********************************************************) + +lemma cpy_cpys: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. +/2 width=1 by inj/ qed. + +lemma cpys_strap1: ∀G,L,T1,T,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. +normalize /2 width=3 by step/ qed-. + +lemma cpys_strap2: ∀G,L,T1,T,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. +normalize /2 width=3 by TC_strap/ qed-. + +lemma lsuby_cpys_trans: ∀G,d,e. lsub_trans … (cpys d e G) (lsuby d e). +/3 width=5 by lsuby_cpy_trans, LTC_lsub_trans/ +qed-. + +lemma cpys_refl: ∀G,L,d,e. reflexive … (cpys d e G L). +/2 width=1 by cpy_cpys/ qed. + +lemma cpys_bind: ∀G,L,V1,V2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 → + ∀I,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 → + ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[d, e] ⓑ{a,I}V2.T2. +#G #L #V1 #V2 #d #e #HV12 @(cpys_ind … HV12) -V2 +[ #I #T1 #T2 #HT12 @(cpys_ind … HT12) -T2 /3 width=5 by cpys_strap1, cpy_bind/ +| /3 width=5 by cpys_strap1, cpy_bind/ +] +qed. + +lemma cpys_flat: ∀G,L,V1,V2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 → + ∀T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → + ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[d, e] ⓕ{I}V2.T2. +#G #L #V1 #V2 #d #e #HV12 @(cpys_ind … HV12) -V2 +[ #T1 #T2 #HT12 @(cpys_ind … HT12) -T2 /3 width=5 by cpys_strap1, cpy_flat/ +| /3 width=5 by cpys_strap1, cpy_flat/ +qed. + +lemma cpys_weak: ∀G,L,T1,T2,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T2 → + ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → + ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T2. +#G #L #T1 #T2 #d1 #e1 #H #d1 #d2 #Hd21 #Hde12 @(cpys_ind … H) -T2 +/3 width=7 by cpys_strap1, cpy_weak/ +qed-. + +lemma cpys_weak_top: ∀G,L,T1,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, |L| - d] T2. +#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2 +/3 width=4 by cpys_strap1, cpy_weak_top/ +qed-. + +lemma cpys_weak_full: ∀G,L,T1,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[0, |L|] T2. +#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2 +/3 width=5 by cpys_strap1, cpy_weak_full/ +qed-. + +(* Basic forward lemmas *****************************************************) + +lemma cpys_fwd_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → + ∀T1,d,e. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → d + e ≤ dt + et → + ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶*[d+e, dt+et-(d+e)] U2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H #T1 #d #e #HTU1 #Hddt #Hdedet @(cpys_ind … H) -U2 +[ /2 width=3 by ex2_intro/ +| -HTU1 #U #U2 #_ #HU2 * #T #HU1 #HTU + elim (cpy_fwd_up … HU2 … HTU) -HU2 -HTU /3 width=3 by cpys_strap1, ex2_intro/ +] +qed-. + +lemma cpys_fwd_tw: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ♯{T1} ≤ ♯{T2}. +#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT1 lapply (cpy_fwd_tw … HT2) -HT2 +/2 width=3 by transitive_le/ +qed-. + +(* Basic inversion lemmas ***************************************************) + +(* Note: this can be derived from cpys_inv_atom1 *) +lemma cpys_inv_sort1: ∀G,L,T2,k,d,e. ⦃G, L⦄ ⊢ ⋆k ▶*[d, e] T2 → T2 = ⋆k. +#G #L #T2 #k #d #e #H @(cpys_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT1 destruct +>(cpy_inv_sort1 … HT2) -HT2 // +qed-. + +(* Note: this can be derived from cpys_inv_atom1 *) +lemma cpys_inv_gref1: ∀G,L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶*[d, e] T2 → T2 = §p. +#G #L #T2 #p #d #e #H @(cpys_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT1 destruct +>(cpy_inv_gref1 … HT2) -HT2 // +qed-. + +lemma cpys_inv_bind1: ∀a,I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[d, e] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 & + ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 & + U2 = ⓑ{a,I}V2.T2. +#a #I #G #L #V1 #T1 #U2 #d #e #H @(cpys_ind … H) -U2 +[ /2 width=5 by ex3_2_intro/ +| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct + elim (cpy_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H + lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V1) ?) -HT2 + /3 width=5 by cpys_strap1, lsuby_succ, ex3_2_intro/ +] +qed-. + +lemma cpys_inv_flat1: ∀I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[d, e] U2 → + ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 & ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 & + U2 = ⓕ{I}V2.T2. +#I #G #L #V1 #T1 #U2 #d #e #H @(cpys_ind … H) -U2 +[ /2 width=5 by ex3_2_intro/ +| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct + elim (cpy_inv_flat1 … HU2) -HU2 + /3 width=5 by cpys_strap1, ex3_2_intro/ +] +qed-. + +lemma cpys_inv_refl_O2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶*[d, 0] T2 → T1 = T2. +#G #L #T1 #T2 #d #H @(cpys_ind … H) -T2 // +#T #T2 #_ #HT2 #IHT1 <(cpy_inv_refl_O2 … HT2) -HT2 // +qed-. + +lemma cpys_inv_lift1_eq: ∀G,L,U1,U2. ∀d,e:nat. + ⦃G, L⦄ ⊢ U1 ▶*[d, e] U2 → ∀T1. ⇧[d, e] T1 ≡ U1 → U1 = U2. +#G #L #U1 #U2 #d #e #H #T1 #HTU1 @(cpys_ind … H) -U2 +/2 width=7 by cpy_inv_lift1_eq/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_alt.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_alt.etc new file mode 100644 index 000000000..e297be698 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_alt.etc @@ -0,0 +1,102 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/psubststaralt_6.ma". +include "basic_2/substitution/cpys_lift.ma". + +(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************) + +(* alternative definition of cpys *) +inductive cpysa: ynat → ynat → relation4 genv lenv term term ≝ +| cpysa_atom : ∀I,G,L,d,e. cpysa d e G L (⓪{I}) (⓪{I}) +| cpysa_subst: ∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d+e → + ⇩[i] L ≡ K.ⓑ{I}V1 → cpysa 0 (⫰(d+e-i)) G K V1 V2 → + ⇧[0, i+1] V2 ≡ W2 → cpysa d e G L (#i) W2 +| cpysa_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e. + cpysa d e G L V1 V2 → cpysa (⫯d) e G (L.ⓑ{I}V1) T1 T2 → + cpysa d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) +| cpysa_flat : ∀I,G,L,V1,V2,T1,T2,d,e. + cpysa d e G L V1 V2 → cpysa d e G L T1 T2 → + cpysa d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) +. + +interpretation + "context-sensitive extended multiple substritution (term) alternative" + 'PSubstStarAlt G L T1 d e T2 = (cpysa d e G L T1 T2). + +(* Basic properties *********************************************************) + +lemma lsuby_cpysa_trans: ∀G,d,e. lsub_trans … (cpysa d e G) (lsuby d e). +#G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e +[ // +| #I #G #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 + elim (lsuby_ldrop_trans_be … HL12 … HLK1) -HL12 -HLK1 /3 width=7 by cpysa_subst/ +| /4 width=1 by lsuby_succ, cpysa_bind/ +| /3 width=1 by cpysa_flat/ +] +qed-. + +lemma cpysa_refl: ∀G,T,L,d,e. ⦃G, L⦄ ⊢ T ▶▶*[d, e] T. +#G #T elim T -T // +#I elim I -I /2 width=1 by cpysa_bind, cpysa_flat/ +qed. + +lemma cpysa_cpy_trans: ∀G,L,T1,T,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T → + ∀T2. ⦃G, L⦄ ⊢ T ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2. +#G #L #T1 #T #d #e #H elim H -G -L -T1 -T -d -e +[ #I #G #L #d #e #X #H + elim (cpy_inv_atom1 … H) -H // * /2 width=7 by cpysa_subst/ +| #I #G #L #K #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK #_ #HVW2 #IHV12 #T2 #H + lapply (ldrop_fwd_drop2 … HLK) #H0LK + lapply (cpy_weak … H 0 (d+e) ? ?) -H // #H + elim (cpy_inv_lift1_be … H … H0LK … HVW2) -H -H0LK -HVW2 + /3 width=7 by cpysa_subst, ylt_fwd_le_succ/ +| #a #I #G #L #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H + elim (cpy_inv_bind1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct + /5 width=5 by cpysa_bind, lsuby_cpy_trans, lsuby_succ/ +| #I #G #L #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H + elim (cpy_inv_flat1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct /3 width=1 by cpysa_flat/ +] +qed-. + +lemma cpys_cpysa: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2. +/3 width=8 by cpysa_cpy_trans, cpys_ind/ qed. + +(* Basic inversion lemmas ***************************************************) + +lemma cpysa_inv_cpys: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. +#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e +/2 width=7 by cpys_subst, cpys_flat, cpys_bind, cpy_cpys/ +qed-. + +(* Advanced eliminators *****************************************************) + +lemma cpys_ind_alt: ∀R:ynat→ynat→relation4 genv lenv term term. + (∀I,G,L,d,e. R d e G L (⓪{I}) (⓪{I})) → + (∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d + e → + ⇩[i] L ≡ K.ⓑ{I}V1 → ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(d+e-i)] V2 → + ⇧[O, i+1] V2 ≡ W2 → R O (⫰(d+e-i)) G K V1 V2 → R d e G L (#i) W2 + ) → + (∀a,I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 → + ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 → R d e G L V1 V2 → + R (⫯d) e G (L.ⓑ{I}V1) T1 T2 → R d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) + ) → + (∀I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 → + ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R d e G L V1 V2 → + R d e G L T1 T2 → R d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) + ) → + ∀d,e,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R d e G L T1 T2. +#R #H1 #H2 #H3 #H4 #d #e #G #L #T1 #T2 #H elim (cpys_cpysa … H) -G -L -T1 -T2 -d -e +/3 width=8 by cpysa_inv_cpys/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_cpys.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_cpys.etc new file mode 100644 index 000000000..52759ca7f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_cpys.etc @@ -0,0 +1,117 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/cpy_cpy.ma". +include "basic_2/substitution/cpys_alt.ma". + +(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************) + +(* Advanced inversion lemmas ************************************************) + +lemma cpys_inv_SO2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶*[d, 1] T2 → ⦃G, L⦄ ⊢ T1 ▶[d, 1] T2. +#G #L #T1 #T2 #d #H @(cpys_ind … H) -T2 /2 width=3 by cpy_trans_ge/ +qed-. + +(* Advanced properties ******************************************************) + +lemma cpys_strip_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶*[d1, e1] T1 → + ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶*[d1, e1] T. +normalize /3 width=3 by cpy_conf_eq, TC_strip1/ qed-. + +lemma cpys_strip_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶*[d1, e1] T1 → + ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶[d2, e2] T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶*[d1, e1] T. +normalize /3 width=3 by cpy_conf_neq, TC_strip1/ qed-. + +lemma cpys_strap1_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T0 → + ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → d2 + e2 ≤ d1 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T ▶*[d1, e1] T2. +normalize /3 width=3 by cpy_trans_down, TC_strap1/ qed. + +lemma cpys_strap2_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T0 → + ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → d2 + e2 ≤ d1 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T ▶[d1, e1] T2. +normalize /3 width=3 by cpy_trans_down, TC_strap2/ qed-. + +lemma cpys_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → + ∀i. d ≤ i → i ≤ d + e → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d, i - d] T & ⦃G, L⦄ ⊢ T ▶*[i, d + e - i] T2. +#G #L #T1 #T2 #d #e #H #i #Hdi #Hide @(cpys_ind … H) -T2 +[ /2 width=3 by ex2_intro/ +| #T #T2 #_ #HT12 * #T3 #HT13 #HT3 + elim (cpy_split_up … HT12 … Hide) -HT12 -Hide #T0 #HT0 #HT02 + elim (cpys_strap1_down … HT3 … HT0) -T /3 width=5 by cpys_strap1, ex2_intro/ + >ymax_pre_sn_comm // +] +qed-. + +lemma cpys_inv_lift1_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[d, dt + et - (yinj d + e)] T2 & + ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet +elim (cpys_split_up … HU12 (d + e)) -HU12 // -Hdedet #U #HU1 #HU2 +lapply (cpys_weak … HU1 d e ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hddt -Hdtde #HU1 +lapply (cpys_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct +elim (cpys_inv_lift1_ge … HU2 … HLK … HTU1) -HU2 -HLK -HTU1 // +>yplus_minus_inj /2 width=3 by ex2_intro/ +qed-. + +(* Main properties **********************************************************) + +theorem cpys_conf_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶*[d1, e1] T1 → + ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶*[d1, e1] T. +normalize /3 width=3 by cpy_conf_eq, TC_confluent2/ qed-. + +theorem cpys_conf_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶*[d1, e1] T1 → + ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶*[d2, e2] T2 → + (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → + ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶*[d1, e1] T. +normalize /3 width=3 by cpy_conf_neq, TC_confluent2/ qed-. + +theorem cpys_trans_eq: ∀G,L,T1,T,T2,d,e. + ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 → + ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. +normalize /2 width=3 by trans_TC/ qed-. + +theorem cpys_trans_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T0 → + ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → d2 + e2 ≤ d1 → + ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T ▶*[d1, e1] T2. +normalize /3 width=3 by cpy_trans_down, TC_transitive2/ qed-. + +theorem cpys_antisym_eq: ∀G,L1,T1,T2,d,e. ⦃G, L1⦄ ⊢ T1 ▶*[d, e] T2 → + ∀L2. ⦃G, L2⦄ ⊢ T2 ▶*[d, e] T1 → T1 = T2. +#G #L1 #T1 #T2 #d #e #H @(cpys_ind_alt … H) -G -L1 -T1 -T2 // +[ #I1 #G #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #_ #_ #HVW2 #_ #L2 #HW2 + elim (lt_or_ge (|L2|) (i+1)) #Hi [ -Hdi -Hide | ] + [ lapply (cpys_weak_full … HW2) -HW2 #HW2 + lapply (cpys_weak … HW2 0 (i+1) ? ?) -HW2 // + [ >yplus_O1 >yplus_O1 /3 width=1 by ylt_fwd_le, ylt_inj/ ] -Hi + #HW2 >(cpys_inv_lift1_eq … HW2) -HW2 // + | elim (ldrop_O1_le … Hi) -Hi #K2 #HLK2 + elim (cpys_inv_lift1_ge_up … HW2 … HLK2 … HVW2 ? ? ?) -HW2 -HLK2 -HVW2 + /2 width=1 by ylt_fwd_le_succ, yle_succ_dx/ -Hdi -Hide + #X #_ #H elim (lift_inv_lref2_be … H) -H // + ] +| #a #I #G #L1 #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_bind1 … H) -H + #V #T #HV2 #HT2 #H destruct + lapply (IHV12 … HV2) #H destruct -IHV12 -HV2 /3 width=2 by eq_f2/ +| #I #G #L1 #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_flat1 … H) -H + #V #T #HV2 #HT2 #H destruct /3 width=2 by eq_f2/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_lift.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_lift.etc new file mode 100644 index 000000000..b6d9e45a1 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/cpys_lift.etc @@ -0,0 +1,218 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/cpy_lift.ma". +include "basic_2/substitution/cpys.ma". + +(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************) + +(* Advanced properties ******************************************************) + +lemma cpys_subst: ∀I,G,L,K,V,U1,i,d,e. + d ≤ yinj i → i < d + e → + ⇩[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ▶*[0, ⫰(d+e-i)] U1 → + ∀U2. ⇧[0, i+1] U1 ≡ U2 → ⦃G, L⦄ ⊢ #i ▶*[d, e] U2. +#I #G #L #K #V #U1 #i #d #e #Hdi #Hide #HLK #H @(cpys_ind … H) -U1 +[ /3 width=5 by cpy_cpys, cpy_subst/ +| #U #U1 #_ #HU1 #IHU #U2 #HU12 + elim (lift_total U 0 (i+1)) #U0 #HU0 + lapply (IHU … HU0) -IHU #H + lapply (ldrop_fwd_drop2 … HLK) -HLK #HLK + lapply (cpy_lift_ge … HU1 … HLK HU0 HU12 ?) -HU1 -HLK -HU0 -HU12 // #HU02 + lapply (cpy_weak … HU02 d e ? ?) -HU02 + [2,3: /2 width=3 by cpys_strap1, yle_succ_dx/ ] + >yplus_O1 ymax_pre_sn_comm /2 width=1 by ylt_fwd_le_succ/ +] +qed. + +lemma cpys_subst_Y2: ∀I,G,L,K,V,U1,i,d. + d ≤ yinj i → + ⇩[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ▶*[0, ∞] U1 → + ∀U2. ⇧[0, i+1] U1 ≡ U2 → ⦃G, L⦄ ⊢ #i ▶*[d, ∞] U2. +#I #G #L #K #V #U1 #i #d #Hdi #HLK #HVU1 #U2 #HU12 +@(cpys_subst … HLK … HU12) >yminus_Y_inj // +qed. + +(* Advanced inverion lemmas *************************************************) + +lemma cpys_inv_atom1: ∀I,G,L,T2,d,e. ⦃G, L⦄ ⊢ ⓪{I} ▶*[d, e] T2 → + T2 = ⓪{I} ∨ + ∃∃J,K,V1,V2,i. d ≤ yinj i & i < d + e & + ⇩[i] L ≡ K.ⓑ{J}V1 & + ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2 & + ⇧[O, i+1] V2 ≡ T2 & + I = LRef i. +#I #G #L #T2 #d #e #H @(cpys_ind … H) -T2 +[ /2 width=1 by or_introl/ +| #T #T2 #_ #HT2 * + [ #H destruct + elim (cpy_inv_atom1 … HT2) -HT2 [ /2 width=1 by or_introl/ | * /3 width=11 by ex6_5_intro, or_intror/ ] + | * #J #K #V1 #V #i #Hdi #Hide #HLK #HV1 #HVT #HI + lapply (ldrop_fwd_drop2 … HLK) #H + elim (cpy_inv_lift1_ge_up … HT2 … H … HVT) -HT2 -H -HVT + [2,3,4: /2 width=1 by ylt_fwd_le_succ, yle_succ_dx/ ] + /4 width=11 by cpys_strap1, ex6_5_intro, or_intror/ + ] +] +qed-. + +lemma cpys_inv_lref1: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶*[d, e] T2 → + T2 = #i ∨ + ∃∃I,K,V1,V2. d ≤ i & i < d + e & + ⇩[i] L ≡ K.ⓑ{I}V1 & + ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2 & + ⇧[O, i+1] V2 ≡ T2. +#G #L #T2 #i #d #e #H elim (cpys_inv_atom1 … H) -H /2 width=1 by or_introl/ +* #I #K #V1 #V2 #j #Hdj #Hjde #HLK #HV12 #HVT2 #H destruct /3 width=7 by ex5_4_intro, or_intror/ +qed-. + +lemma cpys_inv_lref1_ldrop: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶*[d, e] T2 → + ∀I,K,V1. ⇩[i] L ≡ K.ⓑ{I}V1 → + ∀V2. ⇧[O, i+1] V2 ≡ T2 → + ∧∧ ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2 + & d ≤ i + & i < d + e. +#G #L #T2 #i #d #e #H #I #K #V1 #HLK #V2 #HVT2 elim (cpys_inv_lref1 … H) -H +[ #H destruct elim (lift_inv_lref2_be … HVT2) -HVT2 -HLK // +| * #Z #Y #X1 #X2 #Hdi #Hide #HLY #HX12 #HXT2 + lapply (lift_inj … HXT2 … HVT2) -T2 #H destruct + lapply (ldrop_mono … HLY … HLK) -L #H destruct + /2 width=1 by and3_intro/ +] +qed-. + +(* Properties on relocation *************************************************) + +lemma cpys_lift_le: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 → + ∀L,U1,s,d,e. dt + et ≤ yinj d → ⇩[s, d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → + ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2. +#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hdetd #HLK #HTU1 @(cpys_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 (cpy_lift_le … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/ +] +qed-. + +lemma cpys_lift_be: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 → + ∀L,U1,s,d,e. dt ≤ yinj d → d ≤ dt + et → + ⇩[s, d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → + ∀U2. ⇧[d, e] T2 ≡ U2 → ⦃G, L⦄ ⊢ U1 ▶*[dt, et + e] U2. +#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hdtd #Hddet #HLK #HTU1 @(cpys_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 (cpy_lift_be … HT2 … HLK HTU HTU2 ? ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/ +] +qed-. + +lemma cpys_lift_ge: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 → + ∀L,U1,s,d,e. yinj d ≤ dt → ⇩[s, d, e] L ≡ K → + ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → + ⦃G, L⦄ ⊢ U1 ▶*[dt+e, et] U2. +#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hddt #HLK #HTU1 @(cpys_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 (cpy_lift_ge … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/ +] +qed-. + +(* Inversion lemmas for relocation ******************************************) + +lemma cpys_inv_lift1_le: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt + et ≤ d → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdetd @(cpys_ind … H) -U2 +[ /2 width=3 by ex2_intro/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (cpy_inv_lift1_le … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ +] +qed-. + +lemma cpys_inv_lift1_be: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → yinj d + e ≤ dt + et → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, et - e] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdedet @(cpys_ind … H) -U2 +[ /2 width=3 by ex2_intro/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (cpy_inv_lift1_be … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ +] +qed-. + +lemma cpys_inv_lift1_ge: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + yinj d + e ≤ dt → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt - e, et] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdedt @(cpys_ind … H) -U2 +[ /2 width=3 by ex2_intro/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (cpy_inv_lift1_ge … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ +] +qed-. + +(* Advanced inversion lemmas on relocation **********************************) + +lemma cpys_inv_lift1_ge_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[d, dt + et - (yinj d + e)] T2 & + ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet @(cpys_ind … H) -U2 +[ /2 width=3 by ex2_intro/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (cpy_inv_lift1_ge_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ +] +qed-. + +lemma cpys_inv_lift1_be_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → dt + et ≤ yinj d + e → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde @(cpys_ind … H) -U2 +[ /2 width=3 by ex2_intro/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (cpy_inv_lift1_be_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ +] +qed-. + +lemma cpys_inv_lift1_le_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → + ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → + dt ≤ d → d ≤ dt + et → dt + et ≤ yinj d + e → + ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. +#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hddet #Hdetde @(cpys_ind … H) -U2 +[ /2 width=3 by ex2_intro/ +| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU + elim (cpy_inv_lift1_le_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ +] +qed-. + +lemma cpys_inv_lift1_subst: ∀G,L,W1,W2,d,e. ⦃G, L⦄ ⊢ W1 ▶*[d, e] W2 → + ∀K,V1,i. ⇩[i+1] L ≡ K → ⇧[O, i+1] V1 ≡ W1 → + d ≤ yinj i → i < d + e → + ∃∃V2. ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(d+e-i)] V2 & ⇧[O, i+1] V2 ≡ W2. +#G #L #W1 #W2 #d #e #HW12 #K #V1 #i #HLK #HVW1 #Hdi #Hide +elim (cpys_inv_lift1_ge_up … HW12 … HLK … HVW1 ? ? ?) // +>yplus_O1 yplus_SO2 +[ >yminus_succ2 /2 width=3 by ex2_intro/ +| /2 width=1 by ylt_fwd_le_succ1/ +| /2 width=3 by yle_trans/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/extlrsubeq_4.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/extlrsubeq_4.etc new file mode 100644 index 000000000..8a9714dd8 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/extlrsubeq_4.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( L1 break ⊑ × [ term 46 d , break term 46 e ] break term 46 L2 )" + non associative with precedence 45 + for @{ 'ExtLRSubEq $L1 $d $e $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lazyeqalt_4.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lazyeqalt_4.etc new file mode 100644 index 000000000..b89d7844c --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lazyeqalt_4.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( L1 ⋕ ⋕ break [ term 46 T , break term 46 d ] break term 46 L2 )" + non associative with precedence 45 + for @{ 'LazyEqAlt $T $d $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq.etc new file mode 100644 index 000000000..72a5346ce --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq.etc @@ -0,0 +1,160 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/lazyeq_4.ma". +include "basic_2/substitution/cpys.ma". + +(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) + +definition lleq: relation4 ynat term lenv lenv ≝ + λd,T,L1,L2. |L1| = |L2| ∧ + (∀U. ⦃⋆, L1⦄ ⊢ T ▶*[d, ∞] U ↔ ⦃⋆, L2⦄ ⊢ T ▶*[d, ∞] U). + +interpretation + "lazy equivalence (local environment)" + 'LazyEq T d L1 L2 = (lleq d T L1 L2). + +(* Basic properties *********************************************************) + +lemma lleq_refl: ∀d,T. reflexive … (lleq d T). +/3 width=1 by conj/ qed. + +lemma lleq_sym: ∀d,T. symmetric … (lleq d T). +#d #T #L1 #L2 * /3 width=1 by iff_sym, conj/ +qed-. + +lemma lleq_sort: ∀L1,L2,d,k. |L1| = |L2| → L1 ⋕[⋆k, d] L2. +#L1 #L2 #d #k #HL12 @conj // -HL12 +#U @conj #H >(cpys_inv_sort1 … H) -H // +qed. + +lemma lleq_gref: ∀L1,L2,d,p. |L1| = |L2| → L1 ⋕[§p, d] L2. +#L1 #L2 #d #k #HL12 @conj // -HL12 +#U @conj #H >(cpys_inv_gref1 … H) -H // +qed. + +lemma lleq_bind: ∀a,I,L1,L2,V,T,d. + L1 ⋕[V, d] L2 → L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V → + L1 ⋕[ⓑ{a,I}V.T, d] L2. +#a #I #L1 #L2 #V #T #d * #HL12 #IHV * #_ #IHT @conj // -HL12 +#X @conj #H elim (cpys_inv_bind1 … H) -H +#W #U #HVW #HTU #H destruct +elim (IHV W) -IHV elim (IHT U) -IHT /3 width=1 by cpys_bind/ +qed. + +lemma lleq_flat: ∀I,L1,L2,V,T,d. + L1 ⋕[V, d] L2 → L1 ⋕[T, d] L2 → L1 ⋕[ⓕ{I}V.T, d] L2. +#I #L1 #L2 #V #T #d * #HL12 #IHV * #_ #IHT @conj // -HL12 +#X @conj #H elim (cpys_inv_flat1 … H) -H +#W #U #HVW #HTU #H destruct +elim (IHV W) -IHV elim (IHT U) -IHT +/3 width=1 by cpys_flat/ +qed. + +lemma lleq_be: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → ∀T,d,e. ⇧[d, e] T ≡ U → + d ≤ dt → dt ≤ d + e → L1 ⋕[U, d] L2. +#L1 #L2 #U #dt * #HL12 #IH #T #d #e #HTU #Hddt #Hdtde @conj // -HL12 +#U0 elim (IH U0) -IH #H12 #H21 @conj +#HU0 elim (cpys_fwd_up … HU0 … HTU) // -HU0 /4 width=5 by cpys_weak/ +qed-. + +lemma lsuby_lleq_trans: ∀L2,L,T,d. L2 ⋕[T, d] L → + ∀L1. L1 ⊑×[d, ∞] L2 → |L1| = |L2| → L1 ⋕[T, d] L. +#L2 #L #T #d * #HL2 #IH #L1 #HL12 #H @conj // -HL2 +#U elim (IH U) -IH #Hdx #Hsn @conj #HTU +[ @Hdx -Hdx -Hsn @(lsuby_cpys_trans … HTU) -HTU + /2 width=1 by lsuby_sym/ (**) (* full auto does not work *) +| -H -Hdx /3 width=3 by lsuby_cpys_trans/ +] +qed-. + +lemma lleq_lsuby_trans: ∀L,L1,T,d. L ⋕[T, d] L1 → + ∀L2. L1 ⊑×[d, ∞] L2 → |L1| = |L2| → L ⋕[T, d] L2. +/5 width=4 by lsuby_lleq_trans, lleq_sym, lsuby_sym/ qed-. + +lemma lleq_lsuby_repl: ∀L1,L2,T,d. L1 ⋕[T, d] L2 → + ∀K1. K1 ⊑×[d, ∞] L1 → |K1| = |L1| → + ∀K2. L2 ⊑×[d, ∞] K2 → |L2| = |K2| → + K1 ⋕[T, d] K2. +/3 width=4 by lleq_lsuby_trans, lsuby_lleq_trans/ qed-. + +(* Basic forward lemmas *****************************************************) + +lemma lleq_fwd_length: ∀L1,L2,T,d. L1 ⋕[T, d] L2 → |L1| = |L2|. +#L1 #L2 #T #d * // +qed-. + +lemma lleq_fwd_ldrop_sn: ∀L1,L2,T,d. L1 ⋕[d, T] L2 → ∀K1,i. ⇩[i] L1 ≡ K1 → + ∃K2. ⇩[i] L2 ≡ K2. +#L1 #L2 #T #d #H #K1 #i #HLK1 lapply (lleq_fwd_length … H) -H +#HL12 lapply (ldrop_fwd_length_le2 … HLK1) -HLK1 /2 width=1 by ldrop_O1_le/ +qed-. + +lemma lleq_fwd_ldrop_dx: ∀L1,L2,T,d. L1 ⋕[d, T] L2 → ∀K2,i. ⇩[i] L2 ≡ K2 → + ∃K1. ⇩[i] L1 ≡ K1. +/3 width=6 by lleq_fwd_ldrop_sn, lleq_sym/ qed-. + +lemma lleq_fwd_bind_sn: ∀a,I,L1,L2,V,T,d. + L1 ⋕[ⓑ{a,I}V.T, d] L2 → L1 ⋕[V, d] L2. +#a #I #L1 #L2 #V #T #d * #HL12 #H @conj // -HL12 +#U elim (H (ⓑ{a,I}U.T)) -H +#H1 #H2 @conj +#H [ lapply (H1 ?) | lapply (H2 ?) ] -H1 -H2 +/2 width=1 by cpys_bind/ -H +#H elim (cpys_inv_bind1 … H) -H +#X #Y #H1 #H2 #H destruct // +qed-. + +lemma lleq_fwd_bind_dx: ∀a,I,L1,L2,V,T,d. + L1 ⋕[ⓑ{a,I}V.T, d] L2 → L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V. +#a #I #L1 #L2 #V #T #d * #HL12 #H @conj [ normalize // ] -HL12 +#U elim (H (ⓑ{a,I}V.U)) -H +#H1 #H2 @conj +#H [ lapply (H1 ?) | lapply (H2 ?) ] -H1 -H2 +/2 width=1 by cpys_bind/ -H +#H elim (cpys_inv_bind1 … H) -H +#X #Y #H1 #H2 #H destruct // +qed-. + +lemma lleq_fwd_flat_sn: ∀I,L1,L2,V,T,d. + L1 ⋕[ⓕ{I}V.T, d] L2 → L1 ⋕[V, d] L2. +#I #L1 #L2 #V #T #d * #HL12 #H @conj // -HL12 +#U elim (H (ⓕ{I}U.T)) -H +#H1 #H2 @conj +#H [ lapply (H1 ?) | lapply (H2 ?) ] -H1 -H2 +/2 width=1 by cpys_flat/ -H +#H elim (cpys_inv_flat1 … H) -H +#X #Y #H1 #H2 #H destruct // +qed-. + +lemma lleq_fwd_flat_dx: ∀I,L1,L2,V,T,d. + L1 ⋕[ⓕ{I}V.T, d] L2 → L1 ⋕[T, d] L2. +#I #L1 #L2 #V #T #d * #HL12 #H @conj // -HL12 +#U elim (H (ⓕ{I}V.U)) -H +#H1 #H2 @conj +#H [ lapply (H1 ?) | lapply (H2 ?) ] -H1 -H2 +/2 width=1 by cpys_flat/ -H +#H elim (cpys_inv_flat1 … H) -H +#X #Y #H1 #H2 #H destruct // +qed-. + +(* Basic inversion lemmas ***************************************************) + +lemma lleq_inv_bind: ∀a,I,L1,L2,V,T,d. L1 ⋕[ⓑ{a,I}V.T, d] L2 → + L1 ⋕[V, d] L2 ∧ L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V. +/3 width=4 by lleq_fwd_bind_sn, lleq_fwd_bind_dx, conj/ qed-. + +lemma lleq_inv_flat: ∀I,L1,L2,V,T,d. L1 ⋕[ⓕ{I}V.T, d] L2 → + L1 ⋕[V, d] L2 ∧ L1 ⋕[T, d] L2. +/3 width=3 by lleq_fwd_flat_sn, lleq_fwd_flat_dx, conj/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_alt.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_alt.etc new file mode 100644 index 000000000..a8fb526e5 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_alt.etc @@ -0,0 +1,85 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/notation/relations/lazyeqalt_4.ma". +include "basic_2/substitution/lleq_lleq.ma". + +(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) + +(* Note: alternative definition of lleq *) +inductive lleqa: relation4 ynat term lenv lenv ≝ +| lleqa_sort: ∀L1,L2,d,k. |L1| = |L2| → lleqa d (⋆k) L1 L2 +| lleqa_skip: ∀L1,L2,d,i. |L1| = |L2| → yinj i < d → lleqa d (#i) L1 L2 +| lleqa_lref: ∀I1,I2,L1,L2,K1,K2,V,d,i. d ≤ yinj i → + ⇩[i] L1 ≡ K1.ⓑ{I1}V → ⇩[i] L2 ≡ K2.ⓑ{I2}V → + lleqa (yinj 0) V K1 K2 → lleqa d (#i) L1 L2 +| lleqa_free: ∀L1,L2,d,i. |L1| = |L2| → |L1| ≤ i → |L2| ≤ i → lleqa d (#i) L1 L2 +| lleqa_gref: ∀L1,L2,d,p. |L1| = |L2| → lleqa d (§p) L1 L2 +| lleqa_bind: ∀a,I,L1,L2,V,T,d. + lleqa d V L1 L2 → lleqa (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → + lleqa d (ⓑ{a,I}V.T) L1 L2 +| lleqa_flat: ∀I,L1,L2,V,T,d. + lleqa d V L1 L2 → lleqa d T L1 L2 → lleqa d (ⓕ{I}V.T) L1 L2 +. + +interpretation + "lazy equivalence (local environment) alternative" + 'LazyEqAlt T d L1 L2 = (lleqa d T L1 L2). + +(* Main inversion lemmas ****************************************************) + +theorem lleqa_inv_lleq: ∀L1,L2,T,d. L1 ⋕⋕[T, d] L2 → L1 ⋕[T, d] L2. +#L1 #L2 #T #d #H elim H -L1 -L2 -T -d +/2 width=8 by lleq_flat, lleq_bind, lleq_gref, lleq_free, lleq_lref, lleq_skip, lleq_sort/ +qed-. + +(* Main properties **********************************************************) + +theorem lleq_lleqa: ∀L1,T,L2,d. L1 ⋕[T, d] L2 → L1 ⋕⋕[T, d] L2. +#L1 #T @(f2_ind … rfw … L1 T) -L1 -T +#n #IH #L1 * * /3 width=3 by lleqa_gref, lleqa_sort, lleq_fwd_length/ +[ #i #Hn #L2 #d #H elim (lleq_fwd_lref … H) [ * || * ] + /4 width=9 by lleqa_free, lleqa_lref, lleqa_skip, lleq_fwd_length, ldrop_fwd_rfw/ +| #a #I #V #T #Hn #L2 #d #H elim (lleq_inv_bind … H) -H /3 width=1 by lleqa_bind/ +| #I #V #T #Hn #L2 #d #H elim (lleq_inv_flat … H) -H /3 width=1 by lleqa_flat/ +] +qed. + +(* Advanced eliminators *****************************************************) + +lemma lleq_ind_alt: ∀R:relation4 ynat term lenv lenv. ( + ∀L1,L2,d,k. |L1| = |L2| → R d (⋆k) L1 L2 + ) → ( + ∀L1,L2,d,i. |L1| = |L2| → yinj i < d → R d (#i) L1 L2 + ) → ( + ∀I1,I2,L1,L2,K1,K2,V,d,i. d ≤ yinj i → + ⇩[i] L1 ≡ K1.ⓑ{I1}V → ⇩[i] L2 ≡ K2.ⓑ{I2}V → + K1 ⋕[V, yinj O] K2 → R (yinj O) V K1 K2 → R d (#i) L1 L2 + ) → ( + ∀L1,L2,d,i. |L1| = |L2| → |L1| ≤ i → |L2| ≤ i → R d (#i) L1 L2 + ) → ( + ∀L1,L2,d,p. |L1| = |L2| → R d (§p) L1 L2 + ) → ( + ∀a,I,L1,L2,V,T,d. + L1 ⋕[V, d]L2 → L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V → + R d V L1 L2 → R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → R d (ⓑ{a,I}V.T) L1 L2 + ) → ( + ∀I,L1,L2,V,T,d. + L1 ⋕[V, d]L2 → L1 ⋕[T, d] L2 → + R d V L1 L2 → R d T L1 L2 → R d (ⓕ{I}V.T) L1 L2 + ) → + ∀d,T,L1,L2. L1 ⋕[T, d] L2 → R d T L1 L2. +#R #H1 #H2 #H3 #H4 #H5 #H6 #H7 #d #T #L1 #L2 #H elim (lleq_lleqa … H) -H +/3 width=9 by lleqa_inv_lleq/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_ext.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_ext.etc new file mode 100644 index 000000000..d6bb03fa7 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_ext.etc @@ -0,0 +1,91 @@ +(**************************************************************************) +(* ___ *) +(* ||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/lleq_alt.ma". + +(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) + +(* Advanced inversion lemmas ************************************************) + +fact lleq_inv_S_aux: ∀L1,L2,T,d0. L1 ⋕[T, d0] L2 → ∀d. d0 = d + 1 → + ∀K1,K2,I,V. ⇩[d] L1 ≡ K1.ⓑ{I}V → ⇩[d] L2 ≡ K2.ⓑ{I}V → + K1 ⋕[V, 0] K2 → L1 ⋕[T, d] L2. +#L1 #L2 #T #d0 #H @(lleq_ind_alt … H) -L1 -L2 -T -d0 +/2 width=1 by lleq_gref, lleq_free, lleq_sort/ +[ #L1 #L2 #d0 #i #HL12 #Hid #d #H #K1 #K2 #I #V #HLK1 #HLK2 #HV destruct + elim (yle_split_eq i d) /2 width=1 by lleq_skip, ylt_fwd_succ2/ -HL12 -Hid + #H destruct /2 width=8 by lleq_lref/ +| #I1 #I2 #L1 #L2 #K11 #K22 #V #d0 #i #Hd0i #HLK11 #HLK22 #HV #_ #d #H #K1 #K2 #J #W #_ #_ #_ destruct + /3 width=8 by lleq_lref, yle_pred_sn/ +| #a #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W #HLK1 #HLK2 destruct + /4 width=7 by lleq_bind, ldrop_drop/ +| #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W #HLK1 #HLK2 destruct + /3 width=7 by lleq_flat/ +] +qed-. + +lemma lleq_inv_S: ∀T,L1,L2,d. L1 ⋕[T, d+1] L2 → + ∀K1,K2,I,V. ⇩[d] L1 ≡ K1.ⓑ{I}V → ⇩[d] L2 ≡ K2.ⓑ{I}V → + K1 ⋕[V, 0] K2 → L1 ⋕[T, d] L2. +/2 width=7 by lleq_inv_S_aux/ qed-. + +lemma lleq_inv_bind_O: ∀a,I,L1,L2,V,T. L1 ⋕[ⓑ{a,I}V.T, 0] L2 → + L1 ⋕[V, 0] L2 ∧ L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V. +#a #I #L1 #L2 #V #T #H elim (lleq_inv_bind … H) -H +/3 width=7 by ldrop_pair, conj, lleq_inv_S/ +qed-. + +(* Advanced forward lemmas **************************************************) + +lemma lleq_fwd_bind_O: ∀a,I,L1,L2,V,T. L1 ⋕[ⓑ{a,I}V.T, 0] L2 → + L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V. +#a #I #L1 #L2 #V #T #H elim (lleq_inv_bind_O … H) -H // +qed-. + +(* Advanced properties ******************************************************) + +lemma lleq_ge: ∀L1,L2,T,d1. L1 ⋕[T, d1] L2 → ∀d2. d1 ≤ d2 → L1 ⋕[T, d2] L2. +#L1 #L2 #T #d1 #H @(lleq_ind_alt … H) -L1 -L2 -T -d1 +/4 width=1 by lleq_sort, lleq_free, lleq_gref, lleq_bind, lleq_flat, yle_succ/ +[ /3 width=3 by lleq_skip, ylt_yle_trans/ +| #I1 #I2 #L1 #L2 #K1 #K2 #V #d1 #i #Hi #HLK1 #HLK2 #HV #IHV #d2 #Hd12 elim (ylt_split i d2) + [ lapply (lleq_fwd_length … HV) #HK12 #Hid2 + lapply (ldrop_fwd_length … HLK1) lapply (ldrop_fwd_length … HLK2) + normalize in ⊢ (%→%→?); -I1 -I2 -V -d1 /2 width=1 by lleq_skip/ + | /3 width=8 by lleq_lref, yle_trans/ + ] +] +qed-. + +lemma lleq_bind_O: ∀a,I,L1,L2,V,T. L1 ⋕[V, 0] L2 → L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V → + L1 ⋕[ⓑ{a,I}V.T, 0] L2. +/3 width=3 by lleq_ge, lleq_bind/ qed. + +lemma lleq_bind_repl_SO: ∀I1,I2,L1,L2,V1,V2,T. L1.ⓑ{I1}V1 ⋕[T, 0] L2.ⓑ{I2}V2 → + ∀J1,J2,W1,W2. L1.ⓑ{J1}W1 ⋕[T, 1] L2.ⓑ{J2}W2. +#I1 #I2 #L1 #L2 #V1 #V2 #T #HT #J1 #J2 #W1 #W2 lapply (lleq_ge … HT 1 ?) // -HT +#HT @(lleq_lsuby_repl … HT) /2 width=1 by lsuby_succ/ (**) (* full auto fails *) +qed-. + +lemma lleq_bind_repl_O: ∀I,L1,L2,V,T. L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V → + ∀J,W. L1 ⋕[W, 0] L2 → L1.ⓑ{J}W ⋕[T, 0] L2.ⓑ{J}W. +/3 width=7 by lleq_bind_repl_SO, lleq_inv_S/ qed-. + +(* Inversion lemmas on negated lazy quivalence for local environments *******) + +lemma nlleq_inv_bind_O: ∀a,I,L1,L2,V,T. (L1 ⋕[ⓑ{a,I}V.T, 0] L2 → ⊥) → + (L1 ⋕[V, 0] L2 → ⊥) ∨ (L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V → ⊥). +#a #I #L1 #L2 #V #T #H elim (lleq_dec V L1 L2 0) +/4 width=1 by lleq_bind_O, or_intror, or_introl/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_fqus.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_fqus.etc new file mode 100644 index 000000000..18744a539 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_fqus.etc @@ -0,0 +1,75 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/fqus_alt.ma". +include "basic_2/substitution/lleq_ext.ma". + +(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) + +(* Properties on supclosure and derivatives *********************************) + +lemma lleq_fqu_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃ ⦃G2, K2, U⦄ → + ∀L1. L1 ⋕[T, 0] L2 → + ∃∃K1. ⦃G1, L1, T⦄ ⊃ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. +#G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U +[ #I #G #L2 #V #L1 #H elim (lleq_inv_lref_ge_dx … H … I L2 V) -H // + #I1 #K1 #H1 #H2 lapply (ldrop_inv_O2 … H1) -H1 + #H destruct /2 width=3 by fqu_lref_O, ex2_intro/ +| * [ #a ] #I #G #L2 #V #T #L1 #H + [ elim (lleq_inv_bind … H) + | elim (lleq_inv_flat … H) + ] -H + /2 width=3 by fqu_pair_sn, ex2_intro/ +| #a #I #G #L2 #V #T #L1 #H elim (lleq_inv_bind_O … H) -H + #H3 #H4 /2 width=3 by fqu_bind_dx, ex2_intro/ +| #I #G #L2 #V #T #L1 #H elim (lleq_inv_flat … H) -H + /2 width=3 by fqu_flat_dx, ex2_intro/ +| #G #L2 #K2 #T #U #e #HLK2 #HTU #L1 #HL12 + elim (ldrop_O1_le (e+1) L1) + [ /3 width=12 by fqu_drop, lleq_inv_lift_le, ex2_intro/ + | lapply (ldrop_fwd_length_le2 … HLK2) -K2 + lapply (lleq_fwd_length … HL12) -T -U // + ] +] +qed-. + +lemma lleq_fquq_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃⸮ ⦃G2, K2, U⦄ → + ∀L1. L1 ⋕[T, 0] L2 → + ∃∃K1. ⦃G1, L1, T⦄ ⊃⸮ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. +#G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fquq_inv_gen … H) -H +[ #H elim (lleq_fqu_trans … H … HL12) -L2 /3 width=3 by fqu_fquq, ex2_intro/ +| * #HG #HL #HT destruct /2 width=3 by ex2_intro/ +] +qed-. + +lemma lleq_fqup_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃+ ⦃G2, K2, U⦄ → + ∀L1. L1 ⋕[T, 0] L2 → + ∃∃K1. ⦃G1, L1, T⦄ ⊃+ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. +#G1 #G2 #L2 #K2 #T #U #H @(fqup_ind … H) -G2 -K2 -U +[ #G2 #K2 #U #HTU #L1 #HL12 elim (lleq_fqu_trans … HTU … HL12) -L2 + /3 width=3 by fqu_fqup, ex2_intro/ +| #G #G2 #K #K2 #U #U2 #_ #HU2 #IHTU #L1 #HL12 elim (IHTU … HL12) -L2 + #K1 #HTU #HK1 elim (lleq_fqu_trans … HU2 … HK1) -K + /3 width=5 by fqup_strap1, ex2_intro/ +] +qed-. + +lemma lleq_fqus_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃* ⦃G2, K2, U⦄ → + ∀L1. L1 ⋕[T, 0] L2 → + ∃∃K1. ⦃G1, L1, T⦄ ⊃* ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. +#G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fqus_inv_gen … H) -H +[ #H elim (lleq_fqup_trans … H … HL12) -L2 /3 width=3 by fqup_fqus, ex2_intro/ +| * #HG #HL #HT destruct /2 width=3 by ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_ldrop.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_ldrop.etc new file mode 100644 index 000000000..9cb597c6f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_ldrop.etc @@ -0,0 +1,123 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/substitution/cpys_lift.ma". +include "basic_2/substitution/lleq.ma". + +(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) + +(* Advanced properties ******************************************************) + +lemma lleq_skip: ∀L1,L2,d,i. yinj i < d → |L1| = |L2| → L1 ⋕[#i, d] L2. +#L1 #L2 #d #i #Hid #HL12 @conj // -HL12 +#U @conj #H elim (cpys_inv_lref1 … H) -H // * +#I #Z #Y #X #H elim (ylt_yle_false … Hid … H) +qed. + +lemma lleq_lref: ∀I1,I2,L1,L2,K1,K2,V,d,i. d ≤ yinj i → + ⇩[i] L1 ≡ K1.ⓑ{I1}V → ⇩[i] L2 ≡ K2.ⓑ{I2}V → + K1 ⋕[V, 0] K2 → L1 ⋕[#i, d] L2. +#I1 #I2 #L1 #L2 #K1 #K2 #V #d #i #Hdi #HLK1 #HLK2 * #HK12 #IH @conj [ -IH | -HK12 ] +[ lapply (ldrop_fwd_length … HLK1) -HLK1 #H1 + lapply (ldrop_fwd_length … HLK2) -HLK2 #H2 + >H1 >H2 -H1 -H2 normalize // +| #U @conj #H elim (cpys_inv_lref1 … H) -H // * + >yminus_Y_inj #I #K #X #W #_ #_ #H #HVW #HWU + [ letin HLK ≝ HLK1 | letin HLK ≝ HLK2 ] + lapply (ldrop_mono … H … HLK) -H #H destruct elim (IH W) + /3 width=7 by cpys_subst_Y2/ +] +qed. + +lemma lleq_free: ∀L1,L2,d,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| → L1 ⋕[#i, d] L2. +#L1 #L2 #d #i #HL1 #HL2 #HL12 @conj // -HL12 +#U @conj #H elim (cpys_inv_lref1 … H) -H // * +#I #Z #Y #X #_ #_ #H lapply (ldrop_fwd_length_lt2 … H) -H +#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/ +qed. + +(* Properties on relocation *************************************************) + +lemma lleq_lift_le: ∀K1,K2,T,dt. K1 ⋕[T, dt] K2 → + ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀U. ⇧[d, e] T ≡ U → dt ≤ d → L1 ⋕[U, dt] L2. +#K1 #K2 #T #dt * #HK12 #IHT #L1 #L2 #d #e #HLK1 #HLK2 #U #HTU #Hdtd +lapply (ldrop_fwd_length … HLK1) lapply (ldrop_fwd_length … HLK2) +#H2 #H1 @conj // -HK12 -H1 -H2 #U0 @conj #HU0 +[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 ] +elim (cpys_inv_lift1_be … HU0 … HLKA … HTU) // -HU0 >yminus_Y_inj #T0 #HT0 #HTU0 +elim (IHT T0) [ #H #_ | #_ #H ] -IHT /3 width=12 by cpys_lift_be/ +qed-. + +lemma lleq_lift_ge: ∀K1,K2,T,dt. K1 ⋕[T, dt] K2 → + ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀U. ⇧[d, e] T ≡ U → d ≤ dt → L1 ⋕[U, dt+e] L2. +#K1 #K2 #T #dt * #HK12 #IHT #L1 #L2 #d #e #HLK1 #HLK2 #U #HTU #Hddt +lapply (ldrop_fwd_length … HLK1) lapply (ldrop_fwd_length … HLK2) +#H2 #H1 @conj // -HK12 -H1 -H2 #U0 @conj #HU0 +[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 ] +elim (cpys_inv_lift1_ge … HU0 … HLKA … HTU) /2 width=1 by monotonic_yle_plus_dx/ -HU0 >yplus_minus_inj #T0 #HT0 #HTU0 +elim (IHT T0) [ #H #_ | #_ #H ] -IHT /3 width=10 by cpys_lift_ge/ +qed-. + +(* Inversion lemmas on relocation *******************************************) + +lemma lleq_inv_lift_le: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → dt ≤ d → K1 ⋕[T, dt] K2. +#L1 #L2 #U #dt * #HL12 #IH #K1 #K2 #d #e #HLK1 #HLK2 #T #HTU #Hdtd +lapply (ldrop_fwd_length_minus2 … HLK1) lapply (ldrop_fwd_length_minus2 … HLK2) +#H2 #H1 @conj // -HL12 -H1 -H2 +#T0 elim (lift_total T0 d e) +#U0 #HTU0 elim (IH U0) -IH +#H12 #H21 @conj #HT0 +[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 letin H0 ≝ H12 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 letin H0 ≝ H21 ] +lapply (cpys_lift_be … HT0 … HLKA … HTU … HTU0) // -HT0 +>yplus_Y1 #HU0 elim (cpys_inv_lift1_be … (H0 HU0) … HLKB … HTU) // -L1 -L2 -U -Hdtd +#X #HT0 #HX lapply (lift_inj … HX … HTU0) -U0 // +qed-. + +lemma lleq_inv_lift_ge: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → yinj d + e ≤ dt → K1 ⋕[T, dt-e] K2. +#L1 #L2 #U #dt * #HL12 #IH #K1 #K2 #d #e #HLK1 #HLK2 #T #HTU #Hdedt +lapply (ldrop_fwd_length_minus2 … HLK1) lapply (ldrop_fwd_length_minus2 … HLK2) +#H2 #H1 @conj // -HL12 -H1 -H2 +elim (yle_inv_plus_inj2 … Hdedt) #Hddt #Hedt +#T0 elim (lift_total T0 d e) +#U0 #HTU0 elim (IH U0) -IH +#H12 #H21 @conj #HT0 +[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 letin H0 ≝ H12 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 letin H0 ≝ H21 ] +lapply (cpys_lift_ge … HT0 … HLKA … HTU … HTU0) // -HT0 -Hddt +>ymax_pre_sn // #HU0 elim (cpys_inv_lift1_ge … (H0 HU0) … HLKB … HTU) // -L1 -L2 -U -Hdedt -Hedt +#X #HT0 #HX lapply (lift_inj … HX … HTU0) -U0 // +qed-. + +lemma lleq_inv_lift_be: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → d ≤ dt → dt ≤ yinj d + e → K1 ⋕[T, d] K2. +#L1 #L2 #U #dt * #HL12 #IH #K1 #K2 #d #e #HLK1 #HLK2 #T #HTU #Hddt #Hdtde +lapply (ldrop_fwd_length_minus2 … HLK1) lapply (ldrop_fwd_length_minus2 … HLK2) +#H2 #H1 @conj // -HL12 -H1 -H2 +#T0 elim (lift_total T0 d e) +#U0 #HTU0 elim (IH U0) -IH +#H12 #H21 @conj #HT0 +[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 letin H0 ≝ H12 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 letin H0 ≝ H21 ] +lapply (cpys_lift_ge … HT0 … HLKA … HTU … HTU0) // -HT0 +#HU0 lapply (cpys_weak … HU0 dt (∞) ? ?) // -HU0 +#HU0 lapply (H0 HU0) +#HU0 lapply (cpys_weak … HU0 d (∞) ? ?) // -HU0 +#HU0 elim (cpys_inv_lift1_ge_up … HU0 … HLKB … HTU) // -L1 -L2 -U -Hddt -Hdtde +#X #HT0 #HX lapply (lift_inj … HX … HTU0) -U0 // +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_lleq.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_lleq.etc new file mode 100644 index 000000000..2a8873d9b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lleq_lleq.etc @@ -0,0 +1,175 @@ +(**************************************************************************) +(* ___ *) +(* ||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/cpys_cpys.ma". +include "basic_2/substitution/lleq_ldrop.ma". + +(* Advanced forward lemmas **************************************************) + +lemma lleq_fwd_lref: ∀L1,L2. ∀d:ynat. ∀i:nat. L1 ⋕[#i, d] L2 → + ∨∨ |L1| ≤ i ∧ |L2| ≤ i + | yinj i < d + | ∃∃I1,I2,K1,K2,V. ⇩[i] L1 ≡ K1.ⓑ{I1}V & + ⇩[i] L2 ≡ K2.ⓑ{I2}V & + K1 ⋕[V, yinj 0] K2 & d ≤ yinj i. +#L1 #L2 #d #i * #HL12 #IH elim (lt_or_ge i (|L1|)) /3 width=3 by or3_intro0, conj/ +elim (ylt_split i d) /2 width=1 by or3_intro1/ #Hdi #Hi +elim (ldrop_O1_lt … Hi) #I1 #K1 #V1 #HLK1 +elim (ldrop_O1_lt L2 i) // -Hi #I2 #K2 #V2 #HLK2 +lapply (ldrop_fwd_length_minus2 … HLK2) #H +lapply (ldrop_fwd_length_minus2 … HLK1) >HL12 yminus_Y_inj ] /3 width=7 by cpys_subst_Y2, yle_inj/ +qed-. + +lemma lleq_fwd_lref_dx: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → + ∀I2,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I2}V → + i < d ∨ + ∃∃I1,K1. ⇩[i] L1 ≡ K1.ⓑ{I1}V & K1 ⋕[V, 0] K2 & d ≤ i. +#L1 #L2 #d #i #H #I2 #K2 #V #HLK2 elim (lleq_fwd_lref … H) -H [ * || * ] +[ #_ #H elim (lt_refl_false i) + lapply (ldrop_fwd_length_lt2 … HLK2) -HLK2 + /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *) +| /2 width=1 by or_introl/ +| #I1 #I2 #K11 #K22 #V0 #HLK11 #HLK22 #HV0 #Hdi lapply (ldrop_mono … HLK22 … HLK2) -L2 + #H destruct /3 width=5 by ex3_2_intro, or_intror/ +] +qed-. + +lemma lleq_fwd_lref_sn: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → + ∀I1,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I1}V → + i < d ∨ + ∃∃I2,K2. ⇩[i] L2 ≡ K2.ⓑ{I2}V & K1 ⋕[V, 0] K2 & d ≤ i. +#L1 #L2 #d #i #HL12 #I1 #K1 #V #HLK1 elim (lleq_fwd_lref_dx L2 … d … HLK1) -HLK1 +[2: * ] /4 width=6 by lleq_sym, ex3_2_intro, or_introl, or_intror/ +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma lleq_inv_lref_ge_dx: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → + ∀I2,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I2}V → + ∃∃I1,K1. ⇩[i] L1 ≡ K1.ⓑ{I1}V & K1 ⋕[V, 0] K2. +#L1 #L2 #d #i #H #Hdi #I2 #K2 #V #HLK2 elim (lleq_fwd_lref_dx … H … HLK2) -L2 +[ #H elim (ylt_yle_false … H Hdi) +| * /2 width=4 by ex2_2_intro/ +] +qed-. + +lemma lleq_inv_lref_ge_sn: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → + ∀I1,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I1}V → + ∃∃I2,K2. ⇩[i] L2 ≡ K2.ⓑ{I2}V & K1 ⋕[V, 0] K2. +#L1 #L2 #d #i #HL12 #Hdi #I1 #K1 #V #HLK1 elim (lleq_inv_lref_ge_dx L2 … Hdi … HLK1) -Hdi -HLK1 +/3 width=4 by lleq_sym, ex2_2_intro/ +qed-. + +lemma lleq_inv_lref_ge_gen: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → + ∀I1,I2,K1,K2,V1,V2. + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → + V1 = V2 ∧ K1 ⋕[V2, 0] K2. +#L1 #L2 #d #i #HL12 #Hdi #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2 +elim (lleq_inv_lref_ge_sn … HL12 … HLK1) // -L1 -d +#J #Y #HY lapply (ldrop_mono … HY … HLK2) -L2 -i #H destruct /2 width=1 by conj/ +qed-. + +lemma lleq_inv_lref_ge: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → + ∀I,K1,K2,V. ⇩[i] L1 ≡ K1.ⓑ{I}V → ⇩[i] L2 ≡ K2.ⓑ{I}V → + K1 ⋕[V, 0] K2. +#L1 #L2 #d #i #HL12 #Hdi #I #K1 #K2 #V #HLK1 #HLK2 +elim (lleq_inv_lref_ge_gen … HL12 … HLK1 HLK2) // +qed-. + +(* Advanced properties ******************************************************) + +lemma lleq_dec: ∀T,L1,L2,d. Decidable (L1 ⋕[T, d] L2). +#T #L1 @(f2_ind … rfw … L1 T) -L1 -T +#n #IH #L1 * * +[ #k #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, lleq_sort/ +| #i #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) + [ #HL12 #d elim (ylt_split i d) /3 width=1 by lleq_skip, or_introl/ + #Hdi elim (lt_or_ge i (|L1|)) #HiL1 + elim (lt_or_ge i (|L2|)) #HiL2 /3 width=1 by or_introl, lleq_free/ + elim (ldrop_O1_lt … HiL2) #I2 #K2 #V2 #HLK2 + elim (ldrop_O1_lt … HiL1) #I1 #K1 #V1 #HLK1 + elim (eq_term_dec V2 V1) + [ #H3 elim (IH K1 V1 … K2 0) destruct + /3 width=8 by lleq_lref, ldrop_fwd_rfw, or_introl/ + ] + -IH #H3 @or_intror + #H elim (lleq_fwd_lref … H) -H [1,3,4,6: * ] + [1,3: /3 width=4 by lt_to_le_to_lt, lt_refl_false/ + |5,6: /2 width=4 by ylt_yle_false/ + ] + #Z1 #Z2 #Y1 #Y2 #X #HLY1 #HLY2 #HX #_ + lapply (ldrop_mono … HLY1 … HLK1) -HLY1 -HLK1 + lapply (ldrop_mono … HLY2 … HLK2) -HLY2 -HLK2 + #H2 #H1 destruct /2 width=1 by/ + ] +| #p #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, lleq_gref/ +| #a #I #V #T #Hn #L2 #d destruct + elim (IH L1 V … L2 d) /2 width=1 by/ + elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V) (d+1)) -IH /3 width=1 by or_introl, lleq_bind/ + #H1 #H2 @or_intror + #H elim (lleq_inv_bind … H) -H /2 width=1 by/ +| #I #V #T #Hn #L2 #d destruct + elim (IH L1 V … L2 d) /2 width=1 by/ + elim (IH L1 T … L2 d) -IH /3 width=1 by or_introl, lleq_flat/ + #H1 #H2 @or_intror + #H elim (lleq_inv_flat … H) -H /2 width=1 by/ +] +-n /4 width=3 by lleq_fwd_length, or_intror/ +qed-. + +(* Main properties **********************************************************) + +theorem lleq_trans: ∀d,T. Transitive … (lleq d T). +#d #T #L1 #L * #HL1 #IH1 #L2 * #HL2 #IH2 /3 width=3 by conj, iff_trans/ +qed-. + +theorem lleq_canc_sn: ∀L,L1,L2,T,d. L ⋕[d, T] L1→ L ⋕[d, T] L2 → L1 ⋕[d, T] L2. +/3 width=3 by lleq_trans, lleq_sym/ qed-. + +theorem lleq_canc_dx: ∀L1,L2,L,T,d. L1 ⋕[d, T] L → L2 ⋕[d, T] L → L1 ⋕[d, T] L2. +/3 width=3 by lleq_trans, lleq_sym/ qed-. + +(* Inversion lemmas on negated lazy quivalence for local environments *******) + +lemma nlleq_inv_bind: ∀a,I,L1,L2,V,T,d. (L1 ⋕[ⓑ{a,I}V.T, d] L2 → ⊥) → + (L1 ⋕[V, d] L2 → ⊥) ∨ (L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V → ⊥). +#a #I #L1 #L2 #V #T #d #H elim (lleq_dec V L1 L2 d) +/4 width=1 by lleq_bind, or_intror, or_introl/ +qed-. + +lemma nlleq_inv_flat: ∀I,L1,L2,V,T,d. (L1 ⋕[ⓕ{I}V.T, d] L2 → ⊥) → + (L1 ⋕[V, d] L2 → ⊥) ∨ (L1 ⋕[T, d] L2 → ⊥). +#I #L1 #L2 #V #T #d #H elim (lleq_dec V L1 L2 d) +/4 width=1 by lleq_flat, or_intror, or_introl/ +qed-. + +(* Note: lleq_nlleq_trans: ∀d,T,L1,L. L1⋕[T, d] L → + ∀L2. (L ⋕[T, d] L2 → ⊥) → (L1 ⋕[T, d] L2 → ⊥). +/3 width=3 by lleq_canc_sn/ qed-. +works with /4 width=8/ so lleq_canc_sn is more convenient +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lpx_cpys.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lpx_cpys.etc new file mode 100644 index 000000000..d3c490a89 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lpx_cpys.etc @@ -0,0 +1,69 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/cpy_lift.ma". +include "basic_2/substitution/cpys.ma". +include "basic_2/reduction/lpx_ldrop.ma". + +(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************) + +(* Properties on context-sensitive extended substitution for terms **********) + +lemma cpx_cpy_trans_lpx: ∀h,g,G,L1,T1,T. ⦃G, L1⦄ ⊢ T1 ➡[h, g] T → + ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → + ∀T2,d,e. ⦃G, L2⦄ ⊢ T ▶[d, e] T2 → ⦃G, L1⦄ ⊢ T1 ➡[h, g] T2. +#h #g #G #L1 #T1 #T #H elim H -G -L1 -T1 -T +[ #J #G #L1 #L2 #HL12 #T2 #d #e #H elim (cpy_inv_atom1 … H) -H // + * #I #K2 #V2 #i #_ #_ #HLK2 #HVT2 #H destruct + elim (lpx_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H + elim (lpx_inv_pair2 … H) -H #K1 #V1 #_ #HV12 #H destruct + /2 width=7 by cpx_delta/ +| #G #L1 #k #l #Hkl #L2 #_ #X #d #e #H >(cpy_inv_sort1 … H) -X /2 width=2 by cpx_sort/ +| #I #G #L1 #K1 #V1 #V #T #i #HLK1 #_ #HVT #IHV1 #L2 #HL12 #T2 #d #e #HT2 + elim (lpx_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2 + elim (lpx_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct + lapply (ldrop_fwd_drop2 … HLK2) -V0 #HLK2 + lapply (cpy_weak … HT2 0 (∞) ? ?) -HT2 // #HT2 + elim (cpy_inv_lift1_be … HT2 … HLK2 … HVT) -HT2 -HLK2 -HVT + /3 width=7 by cpx_delta/ +| #a #I #G #L1 #V1 #V #T1 #T #HV1 #_ #IHV1 #IHT1 #L2 #HL12 #X #d #e #H elim (cpy_inv_bind1 … H) -H + #V2 #T2 #HV2 #HT2 #H destruct /4 width=5 by lpx_pair, cpx_bind/ +| #I #G #L1 #V1 #V #T1 #T #_ #_ #IHV1 #IHT1 #L2 #HL12 #X #d #e #H elim (cpy_inv_flat1 … H) -H + #V2 #T2 #HV2 #HT2 #H destruct /3 width=5 by cpx_flat/ +| #G #L1 #V1 #U1 #U #T #_ #HTU #IHU1 #L2 #HL12 #T2 #d #e #HT2 + lapply (cpy_weak … HT2 0 (∞) ? ?) -HT2 // #HT2 + elim (lift_total T2 0 1) #U2 #HTU2 + lapply (cpy_lift_be … HT2 (L2.ⓓV1) … (Ⓕ) … HTU … HTU2 ? ?) -T + /4 width=5 by cpx_zeta, lpx_pair, ldrop_drop/ +| /3 width=5 by cpx_tau/ +| /3 width=5 by cpx_ti/ +| #a #G #L1 #V1 #V #W1 #W #T1 #T #HV1 #HW1 #_ #IHV1 #IHW1 #IHT1 #L2 #HL12 #X #d #e #HX + elim (cpy_inv_bind1 … HX) -HX #Y #T2 #HY #HT2 #H destruct + elim (cpy_inv_flat1 … HY) -HY #W2 #V2 #HW2 #HV2 #H destruct + /5 width=11 by lpx_pair, cpx_beta, lsuby_cpy_trans, lsuby_succ/ +| #a #G #L1 #V1 #V #U #W1 #W #T1 #T #_ #HVU #HW1 #_ #IHV1 #IHW1 #IHT1 #L2 #HL12 #X #d #e #HX + elim (cpy_inv_bind1 … HX) -HX #W2 #Y #HW2 #HY #H destruct + elim (cpy_inv_flat1 … HY) -HY #U2 #T2 #HU2 #HT2 #H destruct + lapply (cpy_weak … HU2 0 (∞) ? ?) -HU2 // #HU2 + elim (cpy_inv_lift1_be … HU2 L2 … HVU) -U + /4 width=7 by lpx_pair, cpx_theta, ldrop_drop/ +] +qed-. + +lemma cpx_cpys_trans_lpx: ∀h,g,G,L1,T1,T. ⦃G, L1⦄ ⊢ T1 ➡[h, g] T → + ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → + ∀T2,d,e. ⦃G, L2⦄ ⊢ T ▶*[d, e] T2 → ⦃G, L1⦄ ⊢ T1 ➡[h, g] T2. +#h #g #G #L1 #T1 #T #HT1 #L2 #HL12 #T2 #d #e #H @(cpys_ind … H) -T2 +/2 width=7 by cpx_cpy_trans_lpx/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lsuby.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lsuby.etc new file mode 100644 index 000000000..6c90b202b --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lsuby.etc @@ -0,0 +1,237 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "ground_2/ynat/ynat_plus.ma". +include "basic_2/notation/relations/extlrsubeq_4.ma". +include "basic_2/relocation/ldrop.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR EXTENDED SUBSTITUTION *******************) + +inductive lsuby: relation4 ynat ynat lenv lenv ≝ +| lsuby_atom: ∀L,d,e. lsuby d e L (⋆) +| lsuby_zero: ∀I1,I2,L1,L2,V1,V2. + lsuby 0 0 L1 L2 → lsuby 0 0 (L1.ⓑ{I1}V1) (L2.ⓑ{I2}V2) +| lsuby_pair: ∀I1,I2,L1,L2,V,e. lsuby 0 e L1 L2 → + lsuby 0 (⫯e) (L1.ⓑ{I1}V) (L2.ⓑ{I2}V) +| lsuby_succ: ∀I1,I2,L1,L2,V1,V2,d,e. + lsuby d e L1 L2 → lsuby (⫯d) e (L1. ⓑ{I1}V1) (L2. ⓑ{I2} V2) +. + +interpretation + "local environment refinement (extended substitution)" + 'ExtLRSubEq L1 d e L2 = (lsuby d e L1 L2). + +(* Basic properties *********************************************************) + +lemma lsuby_pair_lt: ∀I1,I2,L1,L2,V,e. L1 ⊑×[0, ⫰e] L2 → 0 < e → + L1.ⓑ{I1}V ⊑×[0, e] L2.ⓑ{I2}V. +#I1 #I2 #L1 #L2 #V #e #HL12 #He <(ylt_inv_O1 … He) /2 width=1 by lsuby_pair/ +qed. + +lemma lsuby_succ_lt: ∀I1,I2,L1,L2,V1,V2,d,e. L1 ⊑×[⫰d, e] L2 → 0 < d → + L1.ⓑ{I1}V1 ⊑×[d, e] L2. ⓑ{I2}V2. +#I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #Hd <(ylt_inv_O1 … Hd) /2 width=1 by lsuby_succ/ +qed. + +lemma lsuby_pair_O_Y: ∀L1,L2. L1 ⊑×[0, ∞] L2 → + ∀I1,I2,V. L1.ⓑ{I1}V ⊑×[0,∞] L2.ⓑ{I2}V. +#L1 #L2 #HL12 #I1 #I2 #V lapply (lsuby_pair I1 I2 … V … HL12) -HL12 // +qed. + +lemma lsuby_refl: ∀L,d,e. L ⊑×[d, e] L. +#L elim L -L // +#L #I #V #IHL #d elim (ynat_cases … d) [| * #x ] +#Hd destruct /2 width=1 by lsuby_succ/ +#e elim (ynat_cases … e) [| * #x ] +#He destruct /2 width=1 by lsuby_zero, lsuby_pair/ +qed. + +lemma lsuby_O2: ∀L2,L1,d. |L2| ≤ |L1| → L1 ⊑×[d, yinj 0] L2. +#L2 elim L2 -L2 // #L2 #I2 #V2 #IHL2 * normalize +[ #d #H lapply (le_n_O_to_eq … H) -H (length_inv_zero_dx … H) -L1 // +| /2 width=1 by lsuby_O2/ +| #I1 #I2 #L1 #L2 #V #e #_ #IHL12 #H lapply (injective_plus_l … H) + /3 width=1 by lsuby_pair/ +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #H lapply (injective_plus_l … H) + /3 width=1 by lsuby_succ/ +] +qed-. + +(* Basic inversion lemmas ***************************************************) + +fact lsuby_inv_atom1_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → L1 = ⋆ → L2 = ⋆. +#L1 #L2 #d #e * -L1 -L2 -d -e // +[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #H destruct +| #I1 #I2 #L1 #L2 #V #e #_ #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #H destruct +] +qed-. + +lemma lsuby_inv_atom1: ∀L2,d,e. ⋆ ⊑×[d, e] L2 → L2 = ⋆. +/2 width=5 by lsuby_inv_atom1_aux/ qed-. + +fact lsuby_inv_zero1_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → + ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → d = 0 → e = 0 → + L2 = ⋆ ∨ + ∃∃J2,K2,W2. K1 ⊑×[0, 0] K2 & L2 = K2.ⓑ{J2}W2. +#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/ +[ #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J1 #K1 #W1 #H #_ #_ destruct + /3 width=5 by ex2_3_intro, or_intror/ +| #I1 #I2 #L1 #L2 #V #e #_ #J1 #K1 #W1 #_ #_ #H + elim (ysucc_inv_O_dx … H) +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J1 #K1 #W1 #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma lsuby_inv_zero1: ∀I1,K1,L2,V1. K1.ⓑ{I1}V1 ⊑×[0, 0] L2 → + L2 = ⋆ ∨ + ∃∃I2,K2,V2. K1 ⊑×[0, 0] K2 & L2 = K2.ⓑ{I2}V2. +/2 width=9 by lsuby_inv_zero1_aux/ qed-. + +fact lsuby_inv_pair1_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → + ∀J1,K1,W. L1 = K1.ⓑ{J1}W → d = 0 → 0 < e → + L2 = ⋆ ∨ + ∃∃J2,K2. K1 ⊑×[0, ⫰e] K2 & L2 = K2.ⓑ{J2}W. +#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/ +[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W #_ #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #e #HL12 #J1 #K1 #W #H #_ #_ destruct + /3 width=4 by ex2_2_intro, or_intror/ +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J1 #K1 #W #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma lsuby_inv_pair1: ∀I1,K1,L2,V,e. K1.ⓑ{I1}V ⊑×[0, e] L2 → 0 < e → + L2 = ⋆ ∨ + ∃∃I2,K2. K1 ⊑×[0, ⫰e] K2 & L2 = K2.ⓑ{I2}V. +/2 width=6 by lsuby_inv_pair1_aux/ qed-. + +fact lsuby_inv_succ1_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → + ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < d → + L2 = ⋆ ∨ + ∃∃J2,K2,W2. K1 ⊑×[⫰d, e] K2 & L2 = K2.ⓑ{J2}W2. +#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/ +[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W1 #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #e #_ #J1 #K1 #W1 #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #J1 #K1 #W1 #H #_ destruct + /3 width=5 by ex2_3_intro, or_intror/ +] +qed-. + +lemma lsuby_inv_succ1: ∀I1,K1,L2,V1,d,e. K1.ⓑ{I1}V1 ⊑×[d, e] L2 → 0 < d → + L2 = ⋆ ∨ + ∃∃I2,K2,V2. K1 ⊑×[⫰d, e] K2 & L2 = K2.ⓑ{I2}V2. +/2 width=5 by lsuby_inv_succ1_aux/ qed-. + +fact lsuby_inv_zero2_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → + ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → d = 0 → e = 0 → + ∃∃J1,K1,W1. K1 ⊑×[0, 0] K2 & L1 = K1.ⓑ{J1}W1. +#L1 #L2 #d #e * -L1 -L2 -d -e +[ #L1 #d #e #J2 #K2 #W1 #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J2 #K2 #W2 #H #_ #_ destruct + /2 width=5 by ex2_3_intro/ +| #I1 #I2 #L1 #L2 #V #e #_ #J2 #K2 #W2 #_ #_ #H + elim (ysucc_inv_O_dx … H) +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J2 #K2 #W2 #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma lsuby_inv_zero2: ∀I2,K2,L1,V2. L1 ⊑×[0, 0] K2.ⓑ{I2}V2 → + ∃∃I1,K1,V1. K1 ⊑×[0, 0] K2 & L1 = K1.ⓑ{I1}V1. +/2 width=9 by lsuby_inv_zero2_aux/ qed-. + +fact lsuby_inv_pair2_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → + ∀J2,K2,W. L2 = K2.ⓑ{J2}W → d = 0 → 0 < e → + ∃∃J1,K1. K1 ⊑×[0, ⫰e] K2 & L1 = K1.ⓑ{J1}W. +#L1 #L2 #d #e * -L1 -L2 -d -e +[ #L1 #d #e #J2 #K2 #W #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W #_ #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #e #HL12 #J2 #K2 #W #H #_ #_ destruct + /2 width=4 by ex2_2_intro/ +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J2 #K2 #W #_ #H + elim (ysucc_inv_O_dx … H) +] +qed-. + +lemma lsuby_inv_pair2: ∀I2,K2,L1,V,e. L1 ⊑×[0, e] K2.ⓑ{I2}V → 0 < e → + ∃∃I1,K1. K1 ⊑×[0, ⫰e] K2 & L1 = K1.ⓑ{I1}V. +/2 width=6 by lsuby_inv_pair2_aux/ qed-. + +fact lsuby_inv_succ2_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → + ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → 0 < d → + ∃∃J1,K1,W1. K1 ⊑×[⫰d, e] K2 & L1 = K1.ⓑ{J1}W1. +#L1 #L2 #d #e * -L1 -L2 -d -e +[ #L1 #d #e #J2 #K2 #W2 #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W2 #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #e #_ #J2 #K1 #W2 #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #J2 #K2 #W2 #H #_ destruct + /2 width=5 by ex2_3_intro/ +] +qed-. + +lemma lsuby_inv_succ2: ∀I2,K2,L1,V2,d,e. L1 ⊑×[d, e] K2.ⓑ{I2}V2 → 0 < d → + ∃∃I1,K1,V1. K1 ⊑×[⫰d, e] K2 & L1 = K1.ⓑ{I1}V1. +/2 width=5 by lsuby_inv_succ2_aux/ qed-. + +(* Basic forward lemmas *****************************************************) + +lemma lsuby_fwd_length: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → |L2| ≤ |L1|. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize /2 width=1 by le_S_S/ +qed-. + +(* Properties on basic slicing **********************************************) + +lemma lsuby_ldrop_trans_be: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → + ∀I2,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I2}W → + d ≤ i → i < d + e → + ∃∃I1,K1. K1 ⊑×[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I1}W. +#L1 #L2 #d #e #H elim H -L1 -L2 -d -e +[ #L1 #d #e #J2 #K2 #W #s #i #H + elim (ldrop_inv_atom1 … H) -H #H destruct +| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J2 #K2 #W #s #i #_ #_ #H + elim (ylt_yle_false … H) // +| #I1 #I2 #L1 #L2 #V #e #HL12 #IHL12 #J2 #K2 #W #s #i #H #_ >yplus_O1 + elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ] + [ #_ destruct -I2 >ypred_succ + /2 width=4 by ldrop_pair, ex2_2_intro/ + | lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/ + #H yminus_succ yplus_succ1 #H lapply (ylt_inv_succ … H) -H + #Hide lapply (ldrop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/ + #HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 yminus_SO2 + /4 width=4 by ylt_O, ldrop_drop_lt, ex2_2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lsuby_lsuby.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lsuby_lsuby.etc new file mode 100644 index 000000000..24361d3c0 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/lsuby_lsuby.etc @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/lsuby.ma". + +(* LOCAL ENVIRONMENT REFINEMENT FOR EXTENDED SUBSTITUTION *******************) + +(* Main properties **********************************************************) + +theorem lsuby_trans: ∀d,e. Transitive … (lsuby d e). +#d #e #L1 #L2 #H elim H -L1 -L2 -d -e +[ #L1 #d #e #X #H lapply (lsuby_inv_atom1 … H) -H + #H destruct // +| #I1 #I2 #L1 #L #V1 #V #_ #IHL1 #X #H elim (lsuby_inv_zero1 … H) -H // + * #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lsuby_zero/ +| #I1 #I2 #L1 #L2 #V #e #_ #IHL1 #X #H elim (lsuby_inv_pair1 … H) -H // + * #I2 #L2 #HL2 #H destruct /3 width=1 by lsuby_pair/ +| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL1 #X #H elim (lsuby_inv_succ1 … H) -H // + * #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lsuby_succ/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubst_6.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubst_6.etc new file mode 100644 index 000000000..31fe21410 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubst_6.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'PSubst $G $L $T1 $d $e $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubststar_6.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubststar_6.etc new file mode 100644 index 000000000..f2bb4bd78 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubststar_6.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'PSubstStar $G $L $T1 $d $e $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubststaralt_6.etc b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubststaralt_6.etc new file mode 100644 index 000000000..e727d4945 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/etc/lsuby/psubststaralt_6.etc @@ -0,0 +1,19 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +(* NOTATION FOR THE FORMAL SYSTEM λδ ****************************************) + +notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" + non associative with precedence 45 + for @{ 'PSubstStarAlt $G $L $T1 $d $e $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extlrsubeq_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extlrsubeq_4.ma deleted file mode 100644 index 8a9714dd8..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/extlrsubeq_4.ma +++ /dev/null @@ -1,19 +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 FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( L1 break ⊑ × [ term 46 d , break term 46 e ] break term 46 L2 )" - non associative with precedence 45 - for @{ 'ExtLRSubEq $L1 $d $e $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazyeqalt_4.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazyeqalt_4.ma deleted file mode 100644 index b89d7844c..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/lazyeqalt_4.ma +++ /dev/null @@ -1,19 +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 FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( L1 ⋕ ⋕ break [ term 46 T , break term 46 d ] break term 46 L2 )" - non associative with precedence 45 - for @{ 'LazyEqAlt $T $d $L1 $L2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubst_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubst_6.ma deleted file mode 100644 index 31fe21410..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubst_6.ma +++ /dev/null @@ -1,19 +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 FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ [ term 46 d , break term 46 e ] break term 46 T2 )" - non associative with precedence 45 - for @{ 'PSubst $G $L $T1 $d $e $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubststar_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubststar_6.ma deleted file mode 100644 index f2bb4bd78..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubststar_6.ma +++ /dev/null @@ -1,19 +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 FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" - non associative with precedence 45 - for @{ 'PSubstStar $G $L $T1 $d $e $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubststaralt_6.ma b/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubststaralt_6.ma deleted file mode 100644 index e727d4945..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/notation/relations/psubststaralt_6.ma +++ /dev/null @@ -1,19 +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 FOR THE FORMAL SYSTEM λδ ****************************************) - -notation "hvbox( ⦃ term 46 G , break term 46 L ⦄ ⊢ break term 46 T1 break ▶ ▶ * [ term 46 d , break term 46 e ] break term 46 T2 )" - non associative with precedence 45 - for @{ 'PSubstStarAlt $G $L $T1 $d $e $T2 }. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_cpys.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_cpys.ma deleted file mode 100644 index 3b38ee6ab..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_cpys.ma +++ /dev/null @@ -1,42 +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/cpys_alt.ma". -include "basic_2/reduction/cpx.ma". - -(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************) - -(* Properties on local environment refinement for extended substitution *****) - -lemma lsuby_cpx_trans: ∀h,g,G. lsub_trans … (cpx h g G) (lsuby 0 (∞)). -#h #g #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -[ // -| /2 width=2 by cpx_sort/ -| #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 - elim (lsuby_ldrop_trans_be … HL12 … HLK1) // -HL12 -HLK1 /3 width=7 by cpx_delta/ -|4,9: /4 width=1 by cpx_bind, cpx_beta, lsuby_pair_O_Y/ -|5,7,8: /3 width=1 by cpx_flat, cpx_tau, cpx_ti/ -|6,10: /4 width=3 by cpx_zeta, cpx_theta, lsuby_pair_O_Y/ -] -qed-. - -(* Properties on context-sensitive extended multiple substitution for terms *) - -lemma cpys_cpx: ∀h,g,G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2. -#h #g #G #L #T1 #T2 #d #e #H @(cpys_ind_alt … H) -G -L -T1 -T2 -d -e -/2 width=7 by cpx_delta, cpx_bind, cpx_flat/ -qed. - -lemma cpy_cpx: ∀h,g,G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ➡[h, g] T2. -/3 width=3 by cpy_cpys, cpys_cpx/ qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_leq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_leq.ma new file mode 100644 index 000000000..7d86e027d --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_leq.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/ldrop_leq.ma". +include "basic_2/reduction/cpx.ma". + +(* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************) + +(* Properties on equivalence for local environments *************************) + +lemma leq_cpx_trans: ∀h,g,G. lsub_trans … (cpx h g G) (leq 0 (∞)). +#h #g #G #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 +[ // +| /2 width=2 by cpx_sort/ +| #I #G #L1 #K1 #V1 #V2 #W2 #i #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 + elim (leq_ldrop_trans_be … HL12 … HLK1) // -HL12 -HLK1 /3 width=7 by cpx_delta/ +|4,9: /4 width=1 by cpx_bind, cpx_beta, leq_pair_O_Y/ +|5,7,8: /3 width=1 by cpx_flat, cpx_tau, cpx_ti/ +|6,10: /4 width=3 by cpx_zeta, cpx_theta, leq_pair_O_Y/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_lleq.ma index 6ff6c39d0..ae6ddac78 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_lleq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/cpx_lleq.ma @@ -12,7 +12,8 @@ (* *) (**************************************************************************) -include "basic_2/substitution/lleq_ext.ma". +include "basic_2/relocation/lleq_leq.ma". +include "basic_2/relocation/lleq_ldrop.ma". include "basic_2/reduction/cpx.ma". (* CONTEXT-SENSITIVE EXTENDED PARALLEL REDUCTION FOR TERMS ******************) @@ -49,8 +50,8 @@ lemma lleq_cpx_conf_dx: ∀h,g,G,L2,T1,T2. ⦃G, L2⦄ ⊢ T1 ➡[h, g] T2 → #h #g #G #L2 #T1 #T2 #H elim H -G -L2 -T1 -T2 [ // | /3 width=3 by lleq_fwd_length, lleq_sort/ -| #I2 #G #L2 #K2 #V1 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV12 #L1 #H elim (lleq_inv_lref_ge_dx … H … HLK2) // -H - #I1 #K1 #HLK1 #HV1 +| #I #G #L2 #K2 #V1 #V2 #W2 #i #HLK2 #_ #HVW2 #IHV12 #L1 #H elim (lleq_inv_lref_ge_dx … H … HLK2) // -H + #K1 #HLK1 #HV1 lapply (ldrop_fwd_drop2 … HLK1) -HLK1 #HLK1 lapply (ldrop_fwd_drop2 … HLK2) -HLK2 #HLK2 @(lleq_lift_le … HLK1 HLK2 … HVW2) -HLK1 -HLK2 -HVW2 /2 width=1 by/ (**) (* full auto too slow *) diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_cpys.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_cpys.ma deleted file mode 100644 index d3c490a89..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_cpys.ma +++ /dev/null @@ -1,69 +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/relocation/cpy_lift.ma". -include "basic_2/substitution/cpys.ma". -include "basic_2/reduction/lpx_ldrop.ma". - -(* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************) - -(* Properties on context-sensitive extended substitution for terms **********) - -lemma cpx_cpy_trans_lpx: ∀h,g,G,L1,T1,T. ⦃G, L1⦄ ⊢ T1 ➡[h, g] T → - ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → - ∀T2,d,e. ⦃G, L2⦄ ⊢ T ▶[d, e] T2 → ⦃G, L1⦄ ⊢ T1 ➡[h, g] T2. -#h #g #G #L1 #T1 #T #H elim H -G -L1 -T1 -T -[ #J #G #L1 #L2 #HL12 #T2 #d #e #H elim (cpy_inv_atom1 … H) -H // - * #I #K2 #V2 #i #_ #_ #HLK2 #HVT2 #H destruct - elim (lpx_ldrop_trans_O1 … HL12 … HLK2) -L2 #X #HLK1 #H - elim (lpx_inv_pair2 … H) -H #K1 #V1 #_ #HV12 #H destruct - /2 width=7 by cpx_delta/ -| #G #L1 #k #l #Hkl #L2 #_ #X #d #e #H >(cpy_inv_sort1 … H) -X /2 width=2 by cpx_sort/ -| #I #G #L1 #K1 #V1 #V #T #i #HLK1 #_ #HVT #IHV1 #L2 #HL12 #T2 #d #e #HT2 - elim (lpx_ldrop_conf … HLK1 … HL12) -HL12 #X #H #HLK2 - elim (lpx_inv_pair1 … H) -H #K2 #V0 #HK12 #_ #H destruct - lapply (ldrop_fwd_drop2 … HLK2) -V0 #HLK2 - lapply (cpy_weak … HT2 0 (∞) ? ?) -HT2 // #HT2 - elim (cpy_inv_lift1_be … HT2 … HLK2 … HVT) -HT2 -HLK2 -HVT - /3 width=7 by cpx_delta/ -| #a #I #G #L1 #V1 #V #T1 #T #HV1 #_ #IHV1 #IHT1 #L2 #HL12 #X #d #e #H elim (cpy_inv_bind1 … H) -H - #V2 #T2 #HV2 #HT2 #H destruct /4 width=5 by lpx_pair, cpx_bind/ -| #I #G #L1 #V1 #V #T1 #T #_ #_ #IHV1 #IHT1 #L2 #HL12 #X #d #e #H elim (cpy_inv_flat1 … H) -H - #V2 #T2 #HV2 #HT2 #H destruct /3 width=5 by cpx_flat/ -| #G #L1 #V1 #U1 #U #T #_ #HTU #IHU1 #L2 #HL12 #T2 #d #e #HT2 - lapply (cpy_weak … HT2 0 (∞) ? ?) -HT2 // #HT2 - elim (lift_total T2 0 1) #U2 #HTU2 - lapply (cpy_lift_be … HT2 (L2.ⓓV1) … (Ⓕ) … HTU … HTU2 ? ?) -T - /4 width=5 by cpx_zeta, lpx_pair, ldrop_drop/ -| /3 width=5 by cpx_tau/ -| /3 width=5 by cpx_ti/ -| #a #G #L1 #V1 #V #W1 #W #T1 #T #HV1 #HW1 #_ #IHV1 #IHW1 #IHT1 #L2 #HL12 #X #d #e #HX - elim (cpy_inv_bind1 … HX) -HX #Y #T2 #HY #HT2 #H destruct - elim (cpy_inv_flat1 … HY) -HY #W2 #V2 #HW2 #HV2 #H destruct - /5 width=11 by lpx_pair, cpx_beta, lsuby_cpy_trans, lsuby_succ/ -| #a #G #L1 #V1 #V #U #W1 #W #T1 #T #_ #HVU #HW1 #_ #IHV1 #IHW1 #IHT1 #L2 #HL12 #X #d #e #HX - elim (cpy_inv_bind1 … HX) -HX #W2 #Y #HW2 #HY #H destruct - elim (cpy_inv_flat1 … HY) -HY #U2 #T2 #HU2 #HT2 #H destruct - lapply (cpy_weak … HU2 0 (∞) ? ?) -HU2 // #HU2 - elim (cpy_inv_lift1_be … HU2 L2 … HVU) -U - /4 width=7 by lpx_pair, cpx_theta, ldrop_drop/ -] -qed-. - -lemma cpx_cpys_trans_lpx: ∀h,g,G,L1,T1,T. ⦃G, L1⦄ ⊢ T1 ➡[h, g] T → - ∀L2. ⦃G, L1⦄ ⊢ ➡[h, g] L2 → - ∀T2,d,e. ⦃G, L2⦄ ⊢ T ▶*[d, e] T2 → ⦃G, L1⦄ ⊢ T1 ➡[h, g] T2. -#h #g #G #L1 #T1 #T #HT1 #L2 #HL12 #T2 #d #e #H @(cpys_ind … H) -T2 -/2 width=7 by cpx_cpy_trans_lpx/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_lleq.ma index 9daa44656..353f810c3 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_lleq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/reduction/lpx_lleq.ma @@ -12,7 +12,7 @@ (* *) (**************************************************************************) -include "basic_2/substitution/lleq_ext.ma". +include "basic_2/relocation/lleq_ldrop.ma". include "basic_2/reduction/lpx_ldrop.ma". (* SN EXTENDED PARALLEL REDUCTION FOR LOCAL ENVIRONMENTS ********************) @@ -51,7 +51,7 @@ lemma lpx_lleq_fqu_trans: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ⊃ ⦃G2, [ #I #G1 #L1 #V1 #X #H1 #H2 elim (lpx_inv_pair2 … H1) -H1 #K0 #V0 #H1KL1 #_ #H destruct elim (lleq_inv_lref_ge_dx … H2 ? I L1 V1) -H2 // - #I1 #K1 #H #H2KL1 lapply (ldrop_inv_O2 … H) -H #H destruct + #K1 #H #H2KL1 lapply (ldrop_inv_O2 … H) -H #H destruct /2 width=4 by fqu_lref_O, ex3_intro/ | * [ #a ] #I #G1 #L1 #V1 #T1 #K1 #HLK1 #H [ elim (lleq_inv_bind … H) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma deleted file mode 100644 index a1e538ae0..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy.ma +++ /dev/null @@ -1,296 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "ground_2/ynat/ynat_max.ma". -include "basic_2/notation/relations/psubst_6.ma". -include "basic_2/grammar/genv.ma". -include "basic_2/relocation/lsuby.ma". - -(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) - -(* activate genv *) -inductive cpy: ynat → ynat → relation4 genv lenv term term ≝ -| cpy_atom : ∀I,G,L,d,e. cpy d e G L (⓪{I}) (⓪{I}) -| cpy_subst: ∀I,G,L,K,V,W,i,d,e. d ≤ yinj i → i < d+e → - ⇩[i] L ≡ K.ⓑ{I}V → ⇧[0, i+1] V ≡ W → cpy d e G L (#i) W -| cpy_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e. - cpy d e G L V1 V2 → cpy (⫯d) e G (L.ⓑ{I}V1) T1 T2 → - cpy d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) -| cpy_flat : ∀I,G,L,V1,V2,T1,T2,d,e. - cpy d e G L V1 V2 → cpy d e G L T1 T2 → - cpy d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) -. - -interpretation "context-sensitive extended ordinary substritution (term)" - 'PSubst G L T1 d e T2 = (cpy d e G L T1 T2). - -(* Basic properties *********************************************************) - -lemma lsuby_cpy_trans: ∀G,d,e. lsub_trans … (cpy d e G) (lsuby d e). -#G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e -[ // -| #I #G #L1 #K1 #V #W #i #d #e #Hdi #Hide #HLK1 #HVW #L2 #HL12 - elim (lsuby_ldrop_trans_be … HL12 … HLK1) -HL12 -HLK1 /2 width=5 by cpy_subst/ -| /4 width=1 by lsuby_succ, cpy_bind/ -| /3 width=1 by cpy_flat/ -] -qed-. - -lemma cpy_refl: ∀G,T,L,d,e. ⦃G, L⦄ ⊢ T ▶[d, e] T. -#G #T elim T -T // * /2 width=1 by cpy_bind, cpy_flat/ -qed. - -(* Basic_1: was: subst1_ex *) -lemma cpy_full: ∀I,G,K,V,T1,L,d. ⇩[d] L ≡ K.ⓑ{I}V → - ∃∃T2,T. ⦃G, L⦄ ⊢ T1 ▶[d, 1] T2 & ⇧[d, 1] T ≡ T2. -#I #G #K #V #T1 elim T1 -T1 -[ * #i #L #d #HLK - /2 width=4 by lift_sort, lift_gref, ex2_2_intro/ - elim (lt_or_eq_or_gt i d) #Hid - /3 width=4 by lift_lref_ge_minus, lift_lref_lt, ex2_2_intro/ - destruct - elim (lift_total V 0 (i+1)) #W #HVW - elim (lift_split … HVW i i) - /4 width=5 by cpy_subst, ylt_inj, ex2_2_intro/ -| * [ #a ] #J #W1 #U1 #IHW1 #IHU1 #L #d #HLK - elim (IHW1 … HLK) -IHW1 #W2 #W #HW12 #HW2 - [ elim (IHU1 (L.ⓑ{J}W1) (d+1)) -IHU1 - /3 width=9 by cpy_bind, ldrop_drop, lift_bind, ex2_2_intro/ - | elim (IHU1 … HLK) -IHU1 -HLK - /3 width=8 by cpy_flat, lift_flat, ex2_2_intro/ - ] -] -qed-. - -lemma cpy_weak: ∀G,L,T1,T2,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → - ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T2. -#G #L #T1 #T2 #d1 #e1 #H elim H -G -L -T1 -T2 -d1 -e1 // -[ /3 width=5 by cpy_subst, ylt_yle_trans, yle_trans/ -| /4 width=3 by cpy_bind, ylt_yle_trans, yle_succ/ -| /3 width=1 by cpy_flat/ -] -qed-. - -lemma cpy_weak_top: ∀G,L,T1,T2,d,e. - ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶[d, |L| - d] T2. -#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e // -[ #I #G #L #K #V #W #i #d #e #Hdi #_ #HLK #HVW - lapply (ldrop_fwd_length_lt2 … HLK) - /4 width=5 by cpy_subst, ylt_yle_trans, ylt_inj/ -| #a #I #G #L #V1 #V2 normalize in match (|L.ⓑ{I}V2|); (**) (* |?| does not work *) - /2 width=1 by cpy_bind/ -| /2 width=1 by cpy_flat/ -] -qed-. - -lemma cpy_weak_full: ∀G,L,T1,T2,d,e. - ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶[0, |L|] T2. -#G #L #T1 #T2 #d #e #HT12 -lapply (cpy_weak … HT12 0 (d + e) ? ?) -HT12 -/2 width=2 by cpy_weak_top/ -qed-. - -lemma cpy_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀i. i ≤ d + e → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d, i-d] T & ⦃G, L⦄ ⊢ T ▶[i, d+e-i] T2. -#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e -[ /2 width=3 by ex2_intro/ -| #I #G #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hjde - elim (ylt_split i j) [ -Hide -Hjde | -Hdi ] - /4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/ -| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide - elim (IHV12 i) -IHV12 // #V - elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hide - >yplus_SO2 >yplus_succ1 #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 - /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ -| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide - elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hide - /3 width=5 by ex2_intro, cpy_flat/ -] -qed-. - -lemma cpy_split_down: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀i. i ≤ d + e → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[i, d+e-i] T & ⦃G, L⦄ ⊢ T ▶[d, i-d] T2. -#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e -[ /2 width=3 by ex2_intro/ -| #I #G #L #K #V #W #i #d #e #Hdi #Hide #HLK #HVW #j #Hjde - elim (ylt_split i j) [ -Hide -Hjde | -Hdi ] - /4 width=9 by cpy_subst, ylt_yle_trans, ex2_intro/ -| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide - elim (IHV12 i) -IHV12 // #V - elim (IHT12 (i+1)) -IHT12 /2 width=1 by yle_succ/ -Hide - >yplus_SO2 >yplus_succ1 #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 - /3 width=5 by lsuby_succ, ex2_intro, cpy_bind/ -| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #i #Hide - elim (IHV12 i) -IHV12 // elim (IHT12 i) -IHT12 // -Hide - /3 width=5 by ex2_intro, cpy_flat/ -] -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma cpy_fwd_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → - ∀T1,d,e. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → d + e ≤ dt + et → - ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶[d+e, dt+et-(d+e)] U2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et -[ * #i #G #L #dt #et #T1 #d #e #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ - ] -| #I #G #L #K #V #W #i #dt #et #Hdti #Hidet #HLK #HVW #T1 #d #e #H #Hddt #Hdedet - elim (lift_inv_lref2 … H) -H * #Hid #H destruct [ -V -Hidet -Hdedet | -Hdti -Hddt ] - [ elim (ylt_yle_false … Hddt) -Hddt /3 width=3 by yle_ylt_trans, ylt_inj/ - | elim (le_inv_plus_l … Hid) #Hdie #Hei - elim (lift_split … HVW d (i-e+1) ? ? ?) [2,3,4: /2 width=1 by le_S_S, le_S/ ] -Hdie - #T2 #_ >plus_minus // ymax_pre_sn_comm // (**) (* explicit constructor *) - ] -| #a #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #X #d #e #H #Hddt #Hdedet - elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HVW1) -V1 -IHW12 // - elim (IHU12 … HTU1) -T1 -IHU12 /2 width=1 by yle_succ/ - yplus_SO2 >yplus_succ1 >yplus_succ1 - /3 width=2 by cpy_bind, lift_bind, ex2_intro/ -| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #X #d #e #H #Hddt #Hdedet - elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HVW1) -V1 -IHW12 // elim (IHU12 … HTU1) -T1 -IHU12 - /3 width=2 by cpy_flat, lift_flat, ex2_intro/ -] -qed-. - -lemma cpy_fwd_tw: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ♯{T1} ≤ ♯{T2}. -#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e normalize -/3 width=1 by monotonic_le_plus_l, le_plus/ -qed-. - -(* Basic inversion lemmas ***************************************************) - -fact cpy_inv_atom1_aux: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ∀J. T1 = ⓪{J} → - T2 = ⓪{J} ∨ - ∃∃I,K,V,i. d ≤ yinj i & i < d + e & - ⇩[i] L ≡ K.ⓑ{I}V & - ⇧[O, i+1] V ≡ T2 & - J = LRef i. -#G #L #T1 #T2 #d #e * -G -L -T1 -T2 -d -e -[ #I #G #L #d #e #J #H destruct /2 width=1 by or_introl/ -| #I #G #L #K #V #T2 #i #d #e #Hdi #Hide #HLK #HVT2 #J #H destruct /3 width=9 by ex5_4_intro, or_intror/ -| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct -| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #H destruct -] -qed-. - -lemma cpy_inv_atom1: ∀I,G,L,T2,d,e. ⦃G, L⦄ ⊢ ⓪{I} ▶[d, e] T2 → - T2 = ⓪{I} ∨ - ∃∃J,K,V,i. d ≤ yinj i & i < d + e & - ⇩[i] L ≡ K.ⓑ{J}V & - ⇧[O, i+1] V ≡ T2 & - I = LRef i. -/2 width=4 by cpy_inv_atom1_aux/ qed-. - -(* Basic_1: was: subst1_gen_sort *) -lemma cpy_inv_sort1: ∀G,L,T2,k,d,e. ⦃G, L⦄ ⊢ ⋆k ▶[d, e] T2 → T2 = ⋆k. -#G #L #T2 #k #d #e #H -elim (cpy_inv_atom1 … H) -H // -* #I #K #V #i #_ #_ #_ #_ #H destruct -qed-. - -(* Basic_1: was: subst1_gen_lref *) -lemma cpy_inv_lref1: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶[d, e] T2 → - T2 = #i ∨ - ∃∃I,K,V. d ≤ i & i < d + e & - ⇩[i] L ≡ K.ⓑ{I}V & - ⇧[O, i+1] V ≡ T2. -#G #L #T2 #i #d #e #H -elim (cpy_inv_atom1 … H) -H /2 width=1 by or_introl/ -* #I #K #V #j #Hdj #Hjde #HLK #HVT2 #H destruct /3 width=5 by ex4_3_intro, or_intror/ -qed-. - -lemma cpy_inv_gref1: ∀G,L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶[d, e] T2 → T2 = §p. -#G #L #T2 #p #d #e #H -elim (cpy_inv_atom1 … H) -H // -* #I #K #V #i #_ #_ #_ #_ #H destruct -qed-. - -fact cpy_inv_bind1_aux: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 → - ∀a,I,V1,T1. U1 = ⓑ{a,I}V1.T1 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 & - ⦃G, L. ⓑ{I}V1⦄ ⊢ T1 ▶[⫯d, e] T2 & - U2 = ⓑ{a,I}V2.T2. -#G #L #U1 #U2 #d #e * -G -L -U1 -U2 -d -e -[ #I #G #L #d #e #b #J #W1 #U1 #H destruct -| #I #G #L #K #V #W #i #d #e #_ #_ #_ #_ #b #J #W1 #U1 #H destruct -| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #b #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ -| #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #b #J #W1 #U1 #H destruct -] -qed-. - -lemma cpy_inv_bind1: ∀a,I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓑ{a,I} V1. T1 ▶[d, e] U2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 & - ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶[⫯d, e] T2 & - U2 = ⓑ{a,I}V2.T2. -/2 width=3 by cpy_inv_bind1_aux/ qed-. - -fact cpy_inv_flat1_aux: ∀G,L,U1,U2,d,e. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 → - ∀I,V1,T1. U1 = ⓕ{I}V1.T1 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 & - ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 & - U2 = ⓕ{I}V2.T2. -#G #L #U1 #U2 #d #e * -G -L -U1 -U2 -d -e -[ #I #G #L #d #e #J #W1 #U1 #H destruct -| #I #G #L #K #V #W #i #d #e #_ #_ #_ #_ #J #W1 #U1 #H destruct -| #a #I #G #L #V1 #V2 #T1 #T2 #d #e #_ #_ #J #W1 #U1 #H destruct -| #I #G #L #V1 #V2 #T1 #T2 #d #e #HV12 #HT12 #J #W1 #U1 #H destruct /2 width=5 by ex3_2_intro/ -] -qed-. - -lemma cpy_inv_flat1: ∀I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓕ{I} V1. T1 ▶[d, e] U2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶[d, e] V2 & - ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 & - U2 = ⓕ{I}V2.T2. -/2 width=3 by cpy_inv_flat1_aux/ qed-. - - -fact cpy_inv_refl_O2_aux: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → e = 0 → T1 = T2. -#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e -[ // -| #I #G #L #K #V #W #i #d #e #Hdi #Hide #_ #_ #H destruct - elim (ylt_yle_false … Hdi) -Hdi // -| /3 width=1 by eq_f2/ -| /3 width=1 by eq_f2/ -] -qed-. - -lemma cpy_inv_refl_O2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶[d, 0] T2 → T1 = T2. -/2 width=6 by cpy_inv_refl_O2_aux/ qed-. - -(* Basic_1: was: subst1_gen_lift_eq *) -lemma cpy_inv_lift1_eq: ∀G,T1,U1,d,e. ⇧[d, e] T1 ≡ U1 → - ∀L,U2. ⦃G, L⦄ ⊢ U1 ▶[d, e] U2 → U1 = U2. -#G #T1 #U1 #d #e #HTU1 #L #U2 #HU12 elim (cpy_fwd_up … HU12 … HTU1) -HU12 -HTU1 -/2 width=4 by cpy_inv_refl_O2/ -qed-. - -(* Basic_1: removed theorems 25: - subst0_gen_sort subst0_gen_lref subst0_gen_head subst0_gen_lift_lt - subst0_gen_lift_false subst0_gen_lift_ge subst0_refl subst0_trans - subst0_lift_lt subst0_lift_ge subst0_lift_ge_S subst0_lift_ge_s - subst0_subst0 subst0_subst0_back subst0_weight_le subst0_weight_lt - subst0_confluence_neq subst0_confluence_eq subst0_tlt_head - subst0_confluence_lift subst0_tlt - subst1_head subst1_gen_head subst1_lift_S subst1_confluence_lift -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_cpy.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_cpy.ma deleted file mode 100644 index 66b0487db..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_cpy.ma +++ /dev/null @@ -1,122 +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/relocation/cpy_lift.ma". - -(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) - -(* Main properties **********************************************************) - -(* Basic_1: was: subst1_confluence_eq *) -theorem cpy_conf_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶[d1, e1] T1 → - ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶[d1, e1] T. -#G #L #T0 #T1 #d1 #e1 #H elim H -G -L -T0 -T1 -d1 -e1 -[ /2 width=3 by ex2_intro/ -| #I1 #G #L #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #T2 #d2 #e2 #H - elim (cpy_inv_lref1 … H) -H - [ #HX destruct /3 width=7 by cpy_subst, ex2_intro/ - | -Hd1 -Hde1 * #I2 #K2 #V2 #_ #_ #HLK2 #HVT2 - lapply (ldrop_mono … HLK1 … HLK2) -HLK1 -HLK2 #H destruct - >(lift_mono … HVT1 … HVT2) -HVT1 -HVT2 /2 width=3 by ex2_intro/ - ] -| #a #I #G #L #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX - elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02) -IHV01 -HV02 #V #HV1 #HV2 - elim (IHT01 … HT02) -T0 #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT1 (L.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/ - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V2) ?) -HT2 - /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/ -| #I #G #L #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #X #d2 #e2 #HX - elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02) -V0 - elim (IHT01 … HT02) -T0 /3 width=5 by cpy_flat, ex2_intro/ -] -qed-. - -(* Basic_1: was: subst1_confluence_neq *) -theorem cpy_conf_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶[d1, e1] T1 → - ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶[d2, e2] T2 → - (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → - ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶[d1, e1] T. -#G #L1 #T0 #T1 #d1 #e1 #H elim H -G -L1 -T0 -T1 -d1 -e1 -[ /2 width=3 by ex2_intro/ -| #I1 #G #L1 #K1 #V1 #T1 #i0 #d1 #e1 #Hd1 #Hde1 #HLK1 #HVT1 #L2 #T2 #d2 #e2 #H1 #H2 - elim (cpy_inv_lref1 … H1) -H1 - [ #H destruct /3 width=7 by cpy_subst, ex2_intro/ - | -HLK1 -HVT1 * #I2 #K2 #V2 #Hd2 #Hde2 #_ #_ elim H2 -H2 #Hded [ -Hd1 -Hde2 | -Hd2 -Hde1 ] - [ elim (ylt_yle_false … Hde1) -Hde1 /2 width=3 by yle_trans/ - | elim (ylt_yle_false … Hde2) -Hde2 /2 width=3 by yle_trans/ - ] - ] -| #a #I #G #L1 #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H - elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02 H) -IHV01 -HV02 #V #HV1 #HV2 - elim (IHT01 … HT02) -T0 - [ -H #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT1 (L2.ⓑ{I}V1) ?) -HT1 /2 width=1 by lsuby_succ/ - lapply (lsuby_cpy_trans … HT2 (L1.ⓑ{I}V2) ?) -HT2 /3 width=5 by cpy_bind, lsuby_succ, ex2_intro/ - | -HV1 -HV2 elim H -H /3 width=1 by yle_succ, or_introl, or_intror/ - ] -| #I #G #L1 #V0 #V1 #T0 #T1 #d1 #e1 #_ #_ #IHV01 #IHT01 #L2 #X #d2 #e2 #HX #H - elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV01 … HV02 H) -V0 - elim (IHT01 … HT02 H) -T0 -H /3 width=5 by cpy_flat, ex2_intro/ -] -qed-. - -(* Note: the constant 1 comes from cpy_subst *) -(* Basic_1: was: subst1_trans *) -theorem cpy_trans_ge: ∀G,L,T1,T0,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T0 → - ∀T2. ⦃G, L⦄ ⊢ T0 ▶[d, 1] T2 → 1 ≤ e → ⦃G, L⦄ ⊢ T1 ▶[d, e] T2. -#G #L #T1 #T0 #d #e #H elim H -G -L -T1 -T0 -d -e -[ #I #G #L #d #e #T2 #H #He - elim (cpy_inv_atom1 … H) -H - [ #H destruct // - | * #J #K #V #i #Hd2i #Hide2 #HLK #HVT2 #H destruct - lapply (ylt_yle_trans … (d+e) … Hide2) /2 width=5 by cpy_subst, monotonic_yle_plus_dx/ - ] -| #I #G #L #K #V #V2 #i #d #e #Hdi #Hide #HLK #HVW #T2 #HVT2 #He - lapply (cpy_weak … HVT2 0 (i+1) ? ?) -HVT2 /3 width=1 by yle_plus_dx2_trans, yle_succ/ - >yplus_inj #HVT2 <(cpy_inv_lift1_eq … HVW … HVT2) -HVT2 /2 width=5 by cpy_subst/ -| #a #I #G #L #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He - elim (cpy_inv_bind1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct - lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02 - lapply (IHT10 … HT02 He) -T0 /3 width=1 by cpy_bind/ -| #I #G #L #V1 #V0 #T1 #T0 #d #e #_ #_ #IHV10 #IHT10 #X #H #He - elim (cpy_inv_flat1 … H) -H #V2 #T2 #HV02 #HT02 #H destruct /3 width=1 by cpy_flat/ -] -qed-. - -theorem cpy_trans_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T0 → - ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → d2 + e2 ≤ d1 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T ▶[d1, e1] T2. -#G #L #T1 #T0 #d1 #e1 #H elim H -G -L -T1 -T0 -d1 -e1 -[ /2 width=3 by ex2_intro/ -| #I #G #L #K #V #W #i1 #d1 #e1 #Hdi1 #Hide1 #HLK #HVW #T2 #d2 #e2 #HWT2 #Hde2d1 - lapply (yle_trans … Hde2d1 … Hdi1) -Hde2d1 #Hde2i1 - lapply (cpy_weak … HWT2 0 (i1+1) ? ?) -HWT2 /3 width=1 by yle_succ, yle_pred_sn/ -Hde2i1 - >yplus_inj #HWT2 <(cpy_inv_lift1_eq … HVW … HWT2) -HWT2 /3 width=9 by cpy_subst, ex2_intro/ -| #a #I #G #L #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 - elim (cpy_inv_bind1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - lapply (lsuby_cpy_trans … HT02 (L.ⓑ{I}V1) ?) -HT02 /2 width=1 by lsuby_succ/ #HT02 - elim (IHV10 … HV02) -IHV10 -HV02 // #V - elim (IHT10 … HT02) -T0 /2 width=1 by yle_succ/ #T #HT1 #HT2 - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V) ?) -HT2 /3 width=6 by cpy_bind, lsuby_succ, ex2_intro/ -| #I #G #L #V1 #V0 #T1 #T0 #d1 #e1 #_ #_ #IHV10 #IHT10 #X #d2 #e2 #HX #de2d1 - elim (cpy_inv_flat1 … HX) -HX #V2 #T2 #HV02 #HT02 #HX destruct - elim (IHV10 … HV02) -V0 // - elim (IHT10 … HT02) -T0 /3 width=6 by cpy_flat, ex2_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma deleted file mode 100644 index aa0b2416a..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/cpy_lift.ma +++ /dev/null @@ -1,249 +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/relocation/ldrop_ldrop.ma". -include "basic_2/relocation/cpy.ma". - -(* CONTEXT-SENSITIVE EXTENDED ORDINARY SUBSTITUTION FOR TERMS ***************) - -(* Properties on relocation *************************************************) - -(* Basic_1: was: subst1_lift_lt *) -lemma cpy_lift_le: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 → - ∀L,U1,U2,s,d,e. ⇩[s, d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - dt + et ≤ d → ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2. -#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et -[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hdetd - lapply (ylt_yle_trans … Hdetd … Hidet) -Hdetd #Hid - lapply (ylt_inv_inj … Hid) -Hid #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=5 by cpy_subst/ -| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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 - /4 width=7 by cpy_bind, ldrop_skip, yle_succ/ -| #G #I #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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=7 by cpy_flat/ -] -qed-. - -lemma cpy_lift_be: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 → - ∀L,U1,U2,s,d,e. ⇩[s, d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - dt ≤ d → d ≤ dt + et → ⦃G, L⦄ ⊢ U1 ▶[dt, et + e] U2. -#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et -[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_ #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hdtd #_ - elim (lift_inv_lref1 … H) -H * #Hid #H destruct - [ -Hdtd - lapply (ylt_yle_trans … (dt+et+e) … Hidet) // -Hidet #Hidete - elim (lift_trans_ge … HVW … HWU2) -W // (lift_mono … HVY … HVW) -V #H destruct /2 width=5 by cpy_subst/ - | -Hdti - elim (yle_inv_inj2 … Hdtd) -Hdtd #dtt #Hdtd #H destruct - lapply (transitive_le … Hdtd Hid) -Hdtd #Hdti - lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2 - lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hid - /4 width=5 by cpy_subst, ldrop_inv_gen, monotonic_ylt_plus_dx, yle_plus_dx1_trans, yle_inj/ - ] -| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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 - /4 width=7 by cpy_bind, ldrop_skip, yle_succ/ -| #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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=7 by cpy_flat/ -] -qed-. - -(* Basic_1: was: subst1_lift_ge *) -lemma cpy_lift_ge: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 → - ∀L,U1,U2,s,d,e. ⇩[s, d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ⇧[d, e] T2 ≡ U2 → - d ≤ dt → ⦃G, L⦄ ⊢ U1 ▶[dt+e, et] U2. -#G #K #T1 #T2 #dt #et #H elim H -G -K -T1 -T2 -dt -et -[ #I #G #K #dt #et #L #U1 #U2 #s #d #e #_ #H1 #H2 #_ - >(lift_mono … H1 … H2) -H1 -H2 // -| #I #G #K #KV #V #W #i #dt #et #Hdti #Hidet #HKV #HVW #L #U1 #U2 #s #d #e #HLK #H #HWU2 #Hddt - lapply (yle_trans … Hddt … Hdti) -Hddt #Hid - elim (yle_inv_inj2 … Hid) -Hid #dd #Hddi #H0 destruct - lapply (lift_inv_lref1_ge … H … Hddi) -H #H destruct - lapply (lift_trans_be … HVW … HWU2 ? ?) -W /2 width=1 by le_S/ >plus_plus_comm_23 #HVU2 - lapply (ldrop_trans_ge_comm … HLK … HKV ?) -K // -Hddi - /3 width=5 by cpy_subst, ldrop_inv_gen, monotonic_ylt_plus_dx, monotonic_yle_plus_dx/ -| #a #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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 - /4 width=6 by cpy_bind, ldrop_skip, yle_succ/ -| #I #G #K #V1 #V2 #T1 #T2 #dt #et #_ #_ #IHV12 #IHT12 #L #U1 #U2 #s #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=6 by cpy_flat/ -] -qed-. - -(* Inversion lemmas on relocation *******************************************) - -(* Basic_1: was: subst1_gen_lift_lt *) -lemma cpy_inv_lift1_le: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt + et ≤ d → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, et] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et -[ * #i #G #L #dt #et #K #s #d #e #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ - ] -| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #s #d #e #HLK #T1 #H #Hdetd - lapply (ylt_yle_trans … Hdetd … Hidet) -Hdetd #Hid - lapply (ylt_inv_inj … Hid) -Hid #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 yplus_minus_assoc_inj /2 width=1 by yle_plus_to_minus_inj2/ ] -Hdedet #Hidete - elim (ldrop_conf_lt … HLK … HLKV) -L // #L #U #HKL #_ #HUV - elim (lift_trans_le … HUV … HVW) -V // >minus_plus plus_minus // yplus_minus_assoc_inj /3 width=1 by monotonic_ylt_minus_dx, yle_inj/ - ] -| #a #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdtd #Hdedet - elim (lift_inv_bind2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HLK … HVW1) -IHW12 // #V2 #HV12 #HVW2 - elim (IHU12 … HTU1) -U1 - /3 width=6 by cpy_bind, ldrop_skip, lift_bind, yle_succ, ex2_intro/ -| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdtd #Hdedet - elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HLK … HVW1) -W1 // - elim (IHU12 … HLK … HTU1) -U1 -HLK // - /3 width=5 by cpy_flat, lift_flat, ex2_intro/ -] -qed-. - -(* Basic_1: was: subst1_gen_lift_ge *) -lemma cpy_inv_lift1_ge: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - yinj d + e ≤ dt → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt-e, et] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H elim H -G -L -U1 -U2 -dt -et -[ * #i #G #L #dt #et #K #s #d #e #_ #T1 #H #_ - [ lapply (lift_inv_sort2 … H) -H #H destruct /2 width=3 by ex2_intro/ - | elim (lift_inv_lref2 … H) -H * #Hid #H destruct /3 width=3 by lift_lref_ge_minus, lift_lref_lt, ex2_intro/ - | lapply (lift_inv_gref2 … H) -H #H destruct /2 width=3 by ex2_intro/ - ] -| #I #G #L #KV #V #W #i #dt #et #Hdti #Hidet #HLKV #HVW #K #s #d #e #HLK #T1 #H #Hdedt - lapply (yle_trans … Hdedt … Hdti) #Hdei - elim (yle_inv_plus_inj2 … Hdedt) -Hdedt #_ #Hedt - elim (yle_inv_plus_inj2 … Hdei) #Hdie #Hei - lapply (lift_inv_lref2_ge … H ?) -H /2 width=1 by yle_inv_inj/ #H destruct - lapply (ldrop_conf_ge … HLK … HLKV ?) -L /2 width=1 by yle_inv_inj/ #HKV - elim (lift_split … HVW d (i-e+1)) -HVW [2,3,4: /3 width=1 by yle_inv_inj, le_S_S, le_S/ ] -Hdei -Hdie - #V0 #HV10 >plus_minus /2 width=1 by yle_inv_inj/ yminus_succ1_inj /3 width=5 by cpy_bind, lift_bind, ex2_intro/ -| #I #G #L #W1 #W2 #U1 #U2 #dt #et #_ #_ #IHW12 #IHU12 #K #s #d #e #HLK #X #H #Hdetd - elim (lift_inv_flat2 … H) -H #V1 #T1 #HVW1 #HTU1 #H destruct - elim (IHW12 … HLK … HVW1) -W1 // - elim (IHU12 … HLK … HTU1) -U1 -HLK /3 width=5 by cpy_flat, lift_flat, ex2_intro/ -] -qed-. - -(* Advancd inversion lemmas on relocation ***********************************) - -lemma cpy_inv_lift1_ge_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[d, dt + et - (yinj d + e)] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet -elim (cpy_split_up … HU12 (d + e)) -HU12 // -Hdedet #U #HU1 #HU2 -lapply (cpy_weak … HU1 d e ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hddt -Hdtde #HU1 -lapply (cpy_inv_lift1_eq … HTU1 … HU1) -HU1 #HU1 destruct -elim (cpy_inv_lift1_ge … HU2 … HLK … HTU1) -U -L /2 width=3 by ex2_intro/ -qed-. - -lemma cpy_inv_lift1_be_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → dt + et ≤ yinj d + e → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, d-dt] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde -lapply (cpy_weak … HU12 dt (d+e-dt) ? ?) -HU12 // -[ >ymax_pre_sn_comm /2 width=1 by yle_plus_dx1_trans/ ] -Hdetde #HU12 -elim (cpy_inv_lift1_be … HU12 … HLK … HTU1) -U1 -L /2 width=3 by ex2_intro/ -qed-. - -lemma cpy_inv_lift1_le_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → d ≤ dt + et → dt + et ≤ yinj d + e → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶[dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hddet #Hdetde -elim (cpy_split_up … HU12 d) -HU12 // #U #HU1 #HU2 -elim (cpy_inv_lift1_le … HU1 … HLK … HTU1) -U1 -[2: >ymax_pre_sn_comm // ] -Hdtd #T #HT1 #HTU -lapply (cpy_weak … HU2 d e ? ?) -HU2 // -[ >ymax_pre_sn_comm // ] -Hddet -Hdetde #HU2 -lapply (cpy_inv_lift1_eq … HTU … HU2) -L #H destruct /2 width=3 by ex2_intro/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/fqu_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/fqu_lleq.ma new file mode 100644 index 000000000..b4235ba93 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/fqu_lleq.ma @@ -0,0 +1,45 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/lleq_ldrop.ma". +include "basic_2/relocation/fqu.ma". + +(* SUPCLOSURE ***************************************************************) + +(* Properties on lazy equivalence for local environments ********************) + +lemma lleq_fqu_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃ ⦃G2, K2, U⦄ → + ∀L1. L1 ⋕[T, 0] L2 → + ∃∃K1. ⦃G1, L1, T⦄ ⊃ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. +#G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U +[ #I #G #L2 #V #L1 #H elim (lleq_inv_lref_ge_dx … H … I L2 V) -H // + #K1 #H1 #H2 lapply (ldrop_inv_O2 … H1) -H1 + #H destruct /2 width=3 by fqu_lref_O, ex2_intro/ +| * [ #a ] #I #G #L2 #V #T #L1 #H + [ elim (lleq_inv_bind … H) + | elim (lleq_inv_flat … H) + ] -H + /2 width=3 by fqu_pair_sn, ex2_intro/ +| #a #I #G #L2 #V #T #L1 #H elim (lleq_inv_bind_O … H) -H + #H3 #H4 /2 width=3 by fqu_bind_dx, ex2_intro/ +| #I #G #L2 #V #T #L1 #H elim (lleq_inv_flat … H) -H + /2 width=3 by fqu_flat_dx, ex2_intro/ +| #G #L2 #K2 #T #U #e #HLK2 #HTU #L1 #HL12 + elim (ldrop_O1_le (e+1) L1) + [ /3 width=12 by fqu_drop, lleq_inv_lift_le, ex2_intro/ + | lapply (ldrop_fwd_length_le2 … HLK2) -K2 + lapply (lleq_fwd_length … HL12) -T -U // + ] +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/fquq_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/fquq_lleq.ma new file mode 100644 index 000000000..880c387e4 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/fquq_lleq.ma @@ -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/relocation/fqu_lleq.ma". +include "basic_2/relocation/fquq_alt.ma". + +(* OPTIONAL SUPCLOSURE ******************************************************) + +(* Properties on lazy equivalence for local environments ********************) + +lemma lleq_fquq_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃⸮ ⦃G2, K2, U⦄ → + ∀L1. L1 ⋕[T, 0] L2 → + ∃∃K1. ⦃G1, L1, T⦄ ⊃⸮ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. +#G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fquq_inv_gen … H) -H +[ #H elim (lleq_fqu_trans … H … HL12) -L2 /3 width=3 by fqu_fquq, ex2_intro/ +| * #HG #HL #HT destruct /2 width=3 by ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_leq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_leq.ma index 1355827f7..cdca16bb5 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_leq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/ldrop_leq.ma @@ -12,7 +12,6 @@ (* *) (**************************************************************************) -include "ground_2/ynat/ynat_plus.ma". include "basic_2/grammar/leq_leq.ma". include "basic_2/relocation/ldrop.ma". diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq.ma index 7a9ec530e..b431e0d18 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq.ma @@ -25,6 +25,9 @@ interpretation "lazy equivalence (local environment)" 'LazyEq T d L1 L2 = (lleq d T L1 L2). +definition lleq_transitive: predicate (relation3 lenv term term) ≝ + λR. ∀L2,T1,T2. R L2 T1 T2 → ∀L1. L1 ⋕[T1, 0] L2 → R L1 T1 T2. + (* Basic inversion lemmas ***************************************************) lemma lleq_ind: ∀R:relation4 ynat term lenv lenv. ( diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq_ldrop.ma new file mode 100644 index 000000000..299412217 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq_ldrop.ma @@ -0,0 +1,150 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/llpx_sn_ldrop.ma". +include "basic_2/relocation/lleq.ma". + +(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) + +(* Advanced properties ******************************************************) + +lemma lleq_bind_repl_O: ∀I,L1,L2,V,T. L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V → + ∀J,W. L1 ⋕[W, 0] L2 → L1.ⓑ{J}W ⋕[T, 0] L2.ⓑ{J}W. +/2 width=7 by llpx_sn_bind_repl_O/ qed-. + +lemma lleq_dec: ∀T,L1,L2,d. Decidable (L1 ⋕[T, d] L2). +/3 width=1 by llpx_sn_dec, eq_term_dec/ qed-. + +lemma lleq_llpx_sn_trans: ∀R. lleq_transitive R → + ∀L1,L2,T,d. L1 ⋕[T, d] L2 → + ∀L. llpx_sn R d T L2 L → llpx_sn R d T L1 L. +#R #HR #L1 #L2 #T #d #H @(lleq_ind … H) -L1 -L2 -T -d +[1,2,5: /4 width=6 by llpx_sn_fwd_length, llpx_sn_gref, llpx_sn_skip, llpx_sn_sort, trans_eq/ +|4: /4 width=6 by llpx_sn_fwd_length, llpx_sn_free, le_repl_sn_conf_aux, trans_eq/ +| #I #L1 #L2 #K1 #K2 #V #d #i #Hdi #HLK1 #HLK2 #HK12 #IHK12 #L #H elim (llpx_sn_inv_lref_ge_sn … H … HLK2) -H -HLK2 + /3 width=11 by llpx_sn_lref/ +| #a #I #L1 #L2 #V #T #d #_ #_ #IHV #IHT #L #H elim (llpx_sn_inv_bind … H) -H + /3 width=1 by llpx_sn_bind/ +| #I #L1 #L2 #V #T #d #_ #_ #IHV #IHT #L #H elim (llpx_sn_inv_flat … H) -H + /3 width=1 by llpx_sn_flat/ +] +qed-. + +lemma lleq_llpx_sn_conf: ∀R. lleq_transitive R → + ∀L1,L2,T,d. L1 ⋕[T, d] L2 → + ∀L. llpx_sn R d T L1 L → llpx_sn R d T L2 L. +/3 width=3 by lleq_llpx_sn_trans, lleq_sym/ qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma lleq_inv_lref_ge_dx: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → + ∀I,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I}V → + ∃∃K1. ⇩[i] L1 ≡ K1.ⓑ{I}V & K1 ⋕[V, 0] K2. +#L1 #L2 #d #i #H #Hdi #I #K2 #V #HLK2 elim (llpx_sn_inv_lref_ge_dx … H … HLK2) -L2 +/2 width=3 by ex2_intro/ +qed-. + +lemma lleq_inv_lref_ge_sn: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → + ∀I,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I}V → + ∃∃K2. ⇩[i] L2 ≡ K2.ⓑ{I}V & K1 ⋕[V, 0] K2. +#L1 #L2 #d #i #H #Hdi #I1 #K1 #V #HLK1 elim (llpx_sn_inv_lref_ge_sn … H … HLK1) -L1 +/2 width=3 by ex2_intro/ +qed-. + +lemma lleq_inv_lref_ge_bi: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → + ∀I1,I2,K1,K2,V1,V2. + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → + ∧∧ I1 = I2 & K1 ⋕[V1, 0] K2 & V1 = V2. +/2 width=8 by llpx_sn_inv_lref_ge_bi/ qed-. + +lemma lleq_inv_lref_ge: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → + ∀I,K1,K2,V. ⇩[i] L1 ≡ K1.ⓑ{I}V → ⇩[i] L2 ≡ K2.ⓑ{I}V → + K1 ⋕[V, 0] K2. +#L1 #L2 #d #i #HL12 #Hdi #I #K1 #K2 #V #HLK1 #HLK2 +elim (lleq_inv_lref_ge_bi … HL12 … HLK1 HLK2) // +qed-. + +lemma lleq_inv_S: ∀L1,L2,T,d. L1 ⋕[T, d+1] L2 → + ∀I,K1,K2,V. ⇩[d] L1 ≡ K1.ⓑ{I}V → ⇩[d] L2 ≡ K2.ⓑ{I}V → + K1 ⋕[V, 0] K2 → L1 ⋕[T, d] L2. +/2 width=9 by llpx_sn_inv_S/ qed-. + +lemma lleq_inv_bind_O: ∀a,I,L1,L2,V,T. L1 ⋕[ⓑ{a,I}V.T, 0] L2 → + L1 ⋕[V, 0] L2 ∧ L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V. +/2 width=2 by llpx_sn_inv_bind_O/ qed-. + +(* Advanced forward lemmas **************************************************) + +lemma lleq_fwd_lref_dx: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → + ∀I,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I}V → + i < d ∨ + ∃∃K1. ⇩[i] L1 ≡ K1.ⓑ{I}V & K1 ⋕[V, 0] K2 & d ≤ i. +#L1 #L2 #d #i #H #I #K2 #V #HLK2 elim (llpx_sn_fwd_lref_dx … H … HLK2) -L2 +[ | * ] /3 width=3 by ex3_intro, or_intror, or_introl/ +qed-. + +lemma lleq_fwd_lref_sn: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → + ∀I,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I}V → + i < d ∨ + ∃∃K2. ⇩[i] L2 ≡ K2.ⓑ{I}V & K1 ⋕[V, 0] K2 & d ≤ i. +#L1 #L2 #d #i #H #I #K1 #V #HLK1 elim (llpx_sn_fwd_lref_sn … H … HLK1) -L1 +[ | * ] /3 width=3 by ex3_intro, or_intror, or_introl/ +qed-. + +lemma lleq_fwd_bind_O_dx: ∀a,I,L1,L2,V,T. L1 ⋕[ⓑ{a,I}V.T, 0] L2 → + L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V. +/2 width=2 by llpx_sn_fwd_bind_O_dx/ qed-. + +(* Properties on relocation *************************************************) + +lemma lleq_lift_le: ∀K1,K2,T,dt. K1 ⋕[T, dt] K2 → + ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀U. ⇧[d, e] T ≡ U → dt ≤ d → L1 ⋕[U, dt] L2. +/3 width=10 by llpx_sn_lift_le, lift_mono/ qed-. + +lemma lleq_lift_ge: ∀K1,K2,T,dt. K1 ⋕[T, dt] K2 → + ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀U. ⇧[d, e] T ≡ U → d ≤ dt → L1 ⋕[U, dt+e] L2. +/2 width=9 by llpx_sn_lift_ge/ qed-. + +(* Inversion lemmas on relocation *******************************************) + +lemma lleq_inv_lift_le: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → dt ≤ d → K1 ⋕[T, dt] K2. +/3 width=10 by llpx_sn_inv_lift_le, ex2_intro/ qed-. + +lemma lleq_inv_lift_be: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → d ≤ dt → dt ≤ yinj d + e → K1 ⋕[T, d] K2. +/2 width=11 by llpx_sn_inv_lift_be/ qed-. + +lemma lleq_inv_lift_ge: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → yinj d + e ≤ dt → K1 ⋕[T, dt-e] K2. +/2 width=9 by llpx_sn_inv_lift_ge/ qed-. + +(* Inversion lemmas on negated lazy quivalence for local environments *******) + +lemma nlleq_inv_bind: ∀a,I,L1,L2,V,T,d. (L1 ⋕[ⓑ{a,I}V.T, d] L2 → ⊥) → + (L1 ⋕[V, d] L2 → ⊥) ∨ (L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V → ⊥). +/3 width=2 by nllpx_sn_inv_bind, eq_term_dec/ qed-. + +lemma nlleq_inv_flat: ∀I,L1,L2,V,T,d. (L1 ⋕[ⓕ{I}V.T, d] L2 → ⊥) → + (L1 ⋕[V, d] L2 → ⊥) ∨ (L1 ⋕[T, d] L2 → ⊥). +/3 width=2 by nllpx_sn_inv_flat, eq_term_dec/ qed-. + +lemma nlleq_inv_bind_O: ∀a,I,L1,L2,V,T. (L1 ⋕[ⓑ{a,I}V.T, 0] L2 → ⊥) → + (L1 ⋕[V, 0] L2 → ⊥) ∨ (L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V → ⊥). +/3 width=2 by nllpx_sn_inv_bind_O, eq_term_dec/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq_lleq.ma new file mode 100644 index 000000000..a99b1e2f3 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/lleq_lleq.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/lleq_ldrop.ma". + +(* Main properties **********************************************************) + +theorem lleq_trans: ∀d,T. Transitive … (lleq d T). +/2 width=3 by lleq_llpx_sn_trans/ qed-. + +theorem lleq_canc_sn: ∀L,L1,L2,T,d. L ⋕[d, T] L1→ L ⋕[d, T] L2 → L1 ⋕[d, T] L2. +/3 width=3 by lleq_trans, lleq_sym/ qed-. + +theorem lleq_canc_dx: ∀L1,L2,L,T,d. L1 ⋕[d, T] L → L2 ⋕[d, T] L → L1 ⋕[d, T] L2. +/3 width=3 by lleq_trans, lleq_sym/ qed-. + +(* Note: lleq_nlleq_trans: ∀d,T,L1,L. L1⋕[T, d] L → + ∀L2. (L ⋕[T, d] L2 → ⊥) → (L1 ⋕[T, d] L2 → ⊥). +/3 width=3 by lleq_canc_sn/ qed-. +works with /4 width=8/ so lleq_canc_sn is more convenient +*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/llpx_sn_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/llpx_sn_ldrop.ma new file mode 100644 index 000000000..4d90e2332 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/llpx_sn_ldrop.ma @@ -0,0 +1,410 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/ldrop_ldrop.ma". +include "basic_2/relocation/llpx_sn_leq.ma". + +(* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****) + +(* Advanced forward lemmas **************************************************) + +lemma llpx_sn_fwd_lref_dx: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → + ∀I,K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 → + i < d ∨ + ∃∃K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 & llpx_sn R 0 V1 K1 K2 & + R K1 V1 V2 & d ≤ i. +#R #L1 #L2 #d #i #H #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref … H) -H [ * || * ] +[ #_ #H elim (lt_refl_false i) + lapply (ldrop_fwd_length_lt2 … HLK2) -HLK2 + /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *) +| /2 width=1 by or_introl/ +| #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hdi + lapply (ldrop_mono … HLK22 … HLK2) -L2 #H destruct + /3 width=5 by ex4_2_intro, or_intror/ +] +qed-. + +lemma llpx_sn_fwd_lref_sn: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → + ∀I,K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 → + i < d ∨ + ∃∃K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 & llpx_sn R 0 V1 K1 K2 & + R K1 V1 V2 & d ≤ i. +#R #L1 #L2 #d #i #H #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref … H) -H [ * || * ] +[ #H #_ elim (lt_refl_false i) + lapply (ldrop_fwd_length_lt2 … HLK1) -HLK1 + /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *) +| /2 width=1 by or_introl/ +| #I #K11 #K22 #V11 #V22 #HLK11 #HLK22 #HK12 #HV12 #Hdi + lapply (ldrop_mono … HLK11 … HLK1) -L1 #H destruct + /3 width=5 by ex4_2_intro, or_intror/ +] +qed-. + +(* Advanced inversion lemmas ************************************************) + +lemma llpx_sn_inv_lref_ge_dx: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i → + ∀I,K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 → + ∃∃K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 & + llpx_sn R 0 V1 K1 K2 & R K1 V1 V2. +#R #L1 #L2 #d #i #H #Hdi #I #K2 #V2 #HLK2 elim (llpx_sn_fwd_lref_dx … H … HLK2) -L2 +[ #H elim (ylt_yle_false … H Hdi) +| * /2 width=5 by ex3_2_intro/ +] +qed-. + +lemma llpx_sn_inv_lref_ge_sn: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i → + ∀I,K1,V1. ⇩[i] L1 ≡ K1.ⓑ{I}V1 → + ∃∃K2,V2. ⇩[i] L2 ≡ K2.ⓑ{I}V2 & + llpx_sn R 0 V1 K1 K2 & R K1 V1 V2. +#R #L1 #L2 #d #i #H #Hdi #I #K1 #V1 #HLK1 elim (llpx_sn_fwd_lref_sn … H … HLK1) -L1 +[ #H elim (ylt_yle_false … H Hdi) +| * /2 width=5 by ex3_2_intro/ +] +qed-. + +lemma llpx_sn_inv_lref_ge_bi: ∀R,L1,L2,d,i. llpx_sn R d (#i) L1 L2 → d ≤ i → + ∀I1,I2,K1,K2,V1,V2. + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → + ∧∧ I1 = I2 & llpx_sn R 0 V1 K1 K2 & R K1 V1 V2. +#R #L1 #L2 #d #i #HL12 #Hdi #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2 +elim (llpx_sn_inv_lref_ge_sn … HL12 … HLK1) // -L1 -d +#J #Y #HY lapply (ldrop_mono … HY … HLK2) -L2 -i #H destruct /2 width=1 by and3_intro/ +qed-. + +fact llpx_sn_inv_S_aux: ∀R,L1,L2,T,d0. llpx_sn R d0 T L1 L2 → ∀d. d0 = d + 1 → + ∀K1,K2,I,V1,V2. ⇩[d] L1 ≡ K1.ⓑ{I}V1 → ⇩[d] L2 ≡ K2.ⓑ{I}V2 → + llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R d T L1 L2. +#R #L1 #L2 #T #d0 #H elim H -L1 -L2 -T -d0 +/2 width=1 by llpx_sn_gref, llpx_sn_free, llpx_sn_sort/ +[ #L1 #L2 #d0 #i #HL12 #Hid #d #H #K1 #K2 #I #V1 #V2 #HLK1 #HLK2 #HK12 #HV12 destruct + elim (yle_split_eq i d) /2 width=1 by llpx_sn_skip, ylt_fwd_succ2/ -HL12 -Hid + #H destruct /2 width=9 by llpx_sn_lref/ +| #I #L1 #L2 #K11 #K22 #V1 #V2 #d0 #i #Hd0i #HLK11 #HLK22 #HK12 #HV12 #_ #d #H #K1 #K2 #J #W1 #W2 #_ #_ #_ #_ destruct + /3 width=9 by llpx_sn_lref, yle_pred_sn/ +| #a #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W1 #W2 #HLK1 #HLK2 #HK12 #HW12 destruct + /4 width=9 by llpx_sn_bind, ldrop_drop/ +| #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W1 #W2 #HLK1 #HLK2 #HK12 #HW12 destruct + /3 width=9 by llpx_sn_flat/ +] +qed-. + +lemma llpx_sn_inv_S: ∀R,L1,L2,T,d. llpx_sn R (d + 1) T L1 L2 → + ∀K1,K2,I,V1,V2. ⇩[d] L1 ≡ K1.ⓑ{I}V1 → ⇩[d] L2 ≡ K2.ⓑ{I}V2 → + llpx_sn R 0 V1 K1 K2 → R K1 V1 V2 → llpx_sn R d T L1 L2. +/2 width=9 by llpx_sn_inv_S_aux/ qed-. + +lemma llpx_sn_inv_bind_O: ∀R. (∀L. reflexive … (R L)) → + ∀a,I,L1,L2,V,T. llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 → + llpx_sn R 0 V L1 L2 ∧ llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V). +#R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_inv_bind … H) -H +/3 width=9 by ldrop_pair, conj, llpx_sn_inv_S/ +qed-. + +(* More advanced forward lemmas *********************************************) + +lemma llpx_sn_fwd_bind_O_dx: ∀R. (∀L. reflexive … (R L)) → + ∀a,I,L1,L2,V,T. llpx_sn R 0 (ⓑ{a,I}V.T) L1 L2 → + llpx_sn R 0 T (L1.ⓑ{I}V) (L2.ⓑ{I}V). +#R #HR #a #I #L1 #L2 #V #T #H elim (llpx_sn_inv_bind_O … H) -H // +qed-. + +(* Advanced properties ******************************************************) + +lemma llpx_sn_bind_repl_O: ∀R,I,L1,L2,V1,V2,T. llpx_sn R 0 T (L1.ⓑ{I}V1) (L2.ⓑ{I}V2) → + ∀J,W1,W2. llpx_sn R 0 W1 L1 L2 → R L1 W1 W2 → llpx_sn R 0 T (L1.ⓑ{J}W1) (L2.ⓑ{J}W2). +/3 width=9 by llpx_sn_bind_repl_SO, llpx_sn_inv_S/ qed-. + +lemma llpx_sn_dec: ∀R. (∀L,T1,T2. Decidable (R L T1 T2)) → + ∀T,L1,L2,d. Decidable (llpx_sn R d T L1 L2). +#R #HR #T #L1 @(f2_ind … rfw … L1 T) -L1 -T +#n #IH #L1 * * +[ #k #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, llpx_sn_sort/ +| #i #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) + [ #HL12 #d elim (ylt_split i d) /3 width=1 by llpx_sn_skip, or_introl/ + #Hdi elim (lt_or_ge i (|L1|)) #HiL1 + elim (lt_or_ge i (|L2|)) #HiL2 /3 width=1 by or_introl, llpx_sn_free/ + elim (ldrop_O1_lt … HiL2) #I2 #K2 #V2 #HLK2 + elim (ldrop_O1_lt … HiL1) #I1 #K1 #V1 #HLK1 + elim (eq_bind2_dec I2 I1) + [ #H2 elim (HR K1 V1 V2) -HR + [ #H3 elim (IH K1 V1 … K2 0) destruct + /3 width=9 by llpx_sn_lref, ldrop_fwd_rfw, or_introl/ + ] + ] + -IH #H3 @or_intror + #H elim (llpx_sn_fwd_lref … H) -H [1,3,4,6,7,9: * ] + [1,3,5: /3 width=4 by lt_to_le_to_lt, lt_refl_false/ + |7,8,9: /2 width=4 by ylt_yle_false/ + ] + #Z #Y1 #Y2 #X1 #X2 #HLY1 #HLY2 #HY12 #HX12 + lapply (ldrop_mono … HLY1 … HLK1) -HLY1 -HLK1 + lapply (ldrop_mono … HLY2 … HLK2) -HLY2 -HLK2 + #H #H0 destruct /2 width=1 by/ + ] +| #p #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, llpx_sn_gref/ +| #a #I #V #T #Hn #L2 #d destruct + elim (IH L1 V … L2 d) /2 width=1 by/ + elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V) (⫯d)) -IH /3 width=1 by or_introl, llpx_sn_bind/ + #H1 #H2 @or_intror + #H elim (llpx_sn_inv_bind … H) -H /2 width=1 by/ +| #I #V #T #Hn #L2 #d destruct + elim (IH L1 V … L2 d) /2 width=1 by/ + elim (IH L1 T … L2 d) -IH /3 width=1 by or_introl, llpx_sn_flat/ + #H1 #H2 @or_intror + #H elim (llpx_sn_inv_flat … H) -H /2 width=1 by/ +] +-n /4 width=4 by llpx_sn_fwd_length, or_intror/ +qed-. + +(* Properties on relocation *************************************************) + +lemma llpx_sn_lift_le: ∀R. l_liftable R → + ∀K1,K2,T,d0. llpx_sn R d0 T K1 K2 → + ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀U. ⇧[d, e] T ≡ U → d0 ≤ d → llpx_sn R d0 U L1 L2. +#R #HR #K1 #K2 #T #d0 #H elim H -K1 -K2 -T -d0 +[ #K1 #K2 #d0 #k #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X + lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d + /2 width=1 by llpx_sn_sort/ +| #K1 #K2 #d0 #i #HK12 #Hid0 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H + * #Hdi #H destruct + [ lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d + /2 width=1 by llpx_sn_skip/ + | elim (ylt_yle_false … Hid0) -L1 -L2 -K1 -K2 -e -Hid0 + /3 width=3 by yle_trans, yle_inj/ + ] +| #I #K1 #K2 #K11 #K22 #V1 #V2 #d0 #i #Hid0 #HK11 #HK22 #HK12 #HV12 #IHK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H + * #Hdi #H destruct [ -HK12 | -IHK12 ] + [ elim (ldrop_trans_lt … HLK1 … HK11) // -K1 + elim (ldrop_trans_lt … HLK2 … HK22) // -Hdi -K2 + /3 width=18 by llpx_sn_lref/ + | lapply (ldrop_trans_ge_comm … HLK1 … HK11 ?) // -K1 + lapply (ldrop_trans_ge_comm … HLK2 … HK22 ?) // -Hdi -Hd0 -K2 + /3 width=9 by llpx_sn_lref, yle_plus_dx1_trans/ + ] +| #K1 #K2 #d0 #i #HK1 #HK2 #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H + * #Hid #H destruct + lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12 + [ /3 width=7 by llpx_sn_free, ldrop_fwd_be/ + | lapply (ldrop_fwd_length … HLK1) -HLK1 #HLK1 + lapply (ldrop_fwd_length … HLK2) -HLK2 #HLK2 + @llpx_sn_free [ >HLK1 | >HLK2 ] -Hid -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *) + ] +| #K1 #K2 #d0 #p #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X + lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d -e + /2 width=1 by llpx_sn_gref/ +| #a #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind1 … H) -H + #W #U #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, ldrop_skip, yle_succ/ +| #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat1 … H) -H + #W #U #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/ +] +qed-. + +lemma llpx_sn_lift_ge: ∀R,K1,K2,T,d0. llpx_sn R d0 T K1 K2 → + ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀U. ⇧[d, e] T ≡ U → d ≤ d0 → llpx_sn R (d0+e) U L1 L2. +#R #K1 #K2 #T #d0 #H elim H -K1 -K2 -T -d0 +[ #K1 #K2 #d0 #k #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort1 … H) -X + lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d + /2 width=1 by llpx_sn_sort/ +| #K1 #K2 #d0 #i #HK12 #Hid0 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref1 … H) -H + * #_ #H destruct + lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 + [ /3 width=3 by llpx_sn_skip, ylt_plus_dx2_trans/ + | /3 width=3 by llpx_sn_skip, monotonic_ylt_plus_dx/ + ] +| #I #K1 #K2 #K11 #K22 #V1 #V2 #d0 #i #Hid0 #HK11 #HK22 #HK12 #HV12 #_ #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H + * #Hid #H destruct + [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K1 -K2 -K11 -K22 -V1 -V2 -e -Hid0 + /3 width=3 by ylt_yle_trans, ylt_inj/ + | lapply (ldrop_trans_ge_comm … HLK1 … HK11 ?) // -K1 + lapply (ldrop_trans_ge_comm … HLK2 … HK22 ?) // -Hid -Hd0 -K2 + /3 width=9 by llpx_sn_lref, monotonic_yle_plus_dx/ + ] +| #K1 #K2 #d0 #i #HK1 #HK2 #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref1 … H) -H + * #Hid #H destruct + lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -HK12 + [ /3 width=7 by llpx_sn_free, ldrop_fwd_be/ + | lapply (ldrop_fwd_length … HLK1) -HLK1 #HLK1 + lapply (ldrop_fwd_length … HLK2) -HLK2 #HLK2 + @llpx_sn_free [ >HLK1 | >HLK2 ] -Hid -HLK1 -HLK2 /2 width=1 by monotonic_le_plus_r/ (**) (* explicit constructor *) + ] +| #K1 #K2 #d0 #p #HK12 #L1 #L2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref1 … H) -X + lapply (ldrop_fwd_length_eq2 … HLK1 HLK2 HK12) -K1 -K2 -d + /2 width=1 by llpx_sn_gref/ +| #a #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind1 … H) -H + #W #U #HVW #HTU #H destruct /4 width=5 by llpx_sn_bind, ldrop_skip, yle_succ/ +| #I #K1 #K2 #V #T #d0 #_ #_ #IHV #IHT #L1 #L2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat1 … H) -H + #W #U #HVW #HTU #H destruct /3 width=5 by llpx_sn_flat/ +] +qed-. + +(* Inversion lemmas on relocation *******************************************) + +lemma llpx_sn_inv_lift_le: ∀R. l_deliftable_sn R → + ∀L1,L2,U,d0. llpx_sn R d0 U L1 L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → d0 ≤ d → llpx_sn R d0 T K1 K2. +#R #HR #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0 +[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d -e + /2 width=1 by llpx_sn_sort/ +| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ elim (lift_inv_lref2 … H) -H + * #_ #H destruct + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 + [ /2 width=1 by llpx_sn_skip/ + | /3 width=3 by llpx_sn_skip, yle_ylt_trans/ + ] +| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #IHK12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref2 … H) -H + * #Hid #H destruct [ -HK12 | -IHK12 ] + [ elim (ldrop_conf_lt … HLK1 … HLK11) // -L1 #L1 #V1 #HKL1 #HKL11 #HVW1 + elim (ldrop_conf_lt … HLK2 … HLK22) // -Hid -L2 #L2 #V2 #HKL2 #HKL22 #HVW2 + elim (HR … HW12 … HKL11 … HVW1) -HR #V0 #HV0 #HV12 + lapply (lift_inj … HV0 … HVW2) -HV0 -HVW2 #H destruct + /3 width=10 by llpx_sn_lref/ + | lapply (ldrop_conf_ge … HLK1 … HLK11 ?) // -L1 + lapply (ldrop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hid0 + elim (le_inv_plus_l … Hid) -Hid /4 width=9 by llpx_sn_lref, yle_trans, yle_inj/ (**) (* slow *) + ] +| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_lref2 … H) -H + * #_ #H destruct + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) + [ lapply (ldrop_fwd_length_le4 … HLK1) -HLK1 + lapply (ldrop_fwd_length_le4 … HLK2) -HLK2 + #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *) + | lapply (ldrop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H + lapply (ldrop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H + /3 width=1 by llpx_sn_free, le_plus_to_minus_r/ + ] +| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d -e + /2 width=1 by llpx_sn_gref/ +| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_bind2 … H) -H + #V #T #HVW #HTU #H destruct /4 width=6 by llpx_sn_bind, ldrop_skip, yle_succ/ +| #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 elim (lift_inv_flat2 … H) -H + #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/ +] +qed-. + +lemma llpx_sn_inv_lift_be: ∀R,L1,L2,U,d0. llpx_sn R d0 U L1 L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → d ≤ d0 → d0 ≤ yinj d + e → llpx_sn R d T K1 K2. +#R #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0 +[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_sort2 … H) -X + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d0 -e + /2 width=1 by llpx_sn_sort/ +| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H + * #Hid #H destruct + [ lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 + -Hid0 /3 width=1 by llpx_sn_skip, ylt_inj/ + | elim (ylt_yle_false … Hid0) -L1 -L2 -Hd0 -Hid0 + /3 width=3 by yle_trans, yle_inj/ (**) (* slow *) + ] +| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H + * #Hid #H destruct + [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hd0e -Hid0 + /3 width=3 by ylt_yle_trans, ylt_inj/ + | lapply (ldrop_conf_ge … HLK1 … HLK11 ?) // -L1 + lapply (ldrop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hid0 -Hd0 -Hd0e + elim (le_inv_plus_l … Hid) -Hid /3 width=9 by llpx_sn_lref, yle_inj/ + ] +| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_lref2 … H) -H + * #_ #H destruct + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) + [ lapply (ldrop_fwd_length_le4 … HLK1) -HLK1 + lapply (ldrop_fwd_length_le4 … HLK2) -HLK2 + #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *) + | lapply (ldrop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H + lapply (ldrop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H + /3 width=1 by llpx_sn_free, le_plus_to_minus_r/ + ] +| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ #_ >(lift_inv_gref2 … H) -X + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d0 -e + /2 width=1 by llpx_sn_gref/ +| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_bind2 … H) -H + >commutative_plus #V #T #HVW #HTU #H destruct + @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *) + @(IHU … HTU) -IHU -HTU /2 width=1 by ldrop_skip, yle_succ/ +| #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hd0 #Hd0e elim (lift_inv_flat2 … H) -H + #V #T #HVW #HTU #H destruct /3 width=6 by llpx_sn_flat/ +] +qed-. + +lemma llpx_sn_inv_lift_ge: ∀R,L1,L2,U,d0. llpx_sn R d0 U L1 L2 → + ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → + ∀T. ⇧[d, e] T ≡ U → yinj d + e ≤ d0 → llpx_sn R (d0-e) T K1 K2. +#R #L1 #L2 #U #d0 #H elim H -L1 -L2 -U -d0 +[ #L1 #L2 #d0 #k #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_sort2 … H) -X + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d + /2 width=1 by llpx_sn_sort/ +| #L1 #L2 #d0 #i #HL12 #Hid0 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H + * #Hid #H destruct [ -Hid0 | -Hded0 ] + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 + [ /4 width=3 by llpx_sn_skip, yle_plus_to_minus_inj2, ylt_yle_trans, ylt_inj/ + | elim (le_inv_plus_l … Hid) -Hid #_ + /4 width=1 by llpx_sn_skip, monotonic_ylt_minus_dx, yle_inj/ + ] +| #I #L1 #L2 #K11 #K22 #W1 #W2 #d0 #i #Hid0 #HLK11 #HLK22 #HK12 #HW12 #_ #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H + * #Hid #H destruct + [ elim (ylt_yle_false … Hid0) -I -L1 -L2 -K11 -K22 -W1 -W2 -Hid0 + /3 width=3 by yle_fwd_plus_sn1, ylt_yle_trans, ylt_inj/ + | lapply (ldrop_conf_ge … HLK1 … HLK11 ?) // -L1 + lapply (ldrop_conf_ge … HLK2 … HLK22 ?) // -L2 -Hded0 -Hid + /3 width=9 by llpx_sn_lref, monotonic_yle_minus_dx/ + ] +| #L1 #L2 #d0 #i #HL1 #HL2 #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_lref2 … H) -H + * #_ #H destruct + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) + [ lapply (ldrop_fwd_length_le4 … HLK1) -HLK1 + lapply (ldrop_fwd_length_le4 … HLK2) -HLK2 + #HKL2 #HKL1 #HK12 @llpx_sn_free // /2 width=3 by transitive_le/ (**) (* full auto too slow *) + | lapply (ldrop_fwd_length … HLK1) -HLK1 #H >H in HL1; -H + lapply (ldrop_fwd_length … HLK2) -HLK2 #H >H in HL2; -H + /3 width=1 by llpx_sn_free, le_plus_to_minus_r/ + ] +| #L1 #L2 #d0 #p #HL12 #K1 #K2 #d #e #HLK1 #HLK2 #X #H #_ >(lift_inv_gref2 … H) -X + lapply (ldrop_fwd_length_eq1 … HLK1 HLK2 HL12) -L1 -L2 -d + /2 width=1 by llpx_sn_gref/ +| #a #I #L1 #L2 #W #U #d0 #_ #_ #IHW #IHU #K1 #K2 #d #e #HLK1 #HLK2 #X #H #Hded0 elim (lift_inv_bind2 … H) -H + #V #T #HVW #HTU #H destruct + @llpx_sn_bind [ /2 width=5 by/ ] -IHW (**) (* explicit constructor *) + (length_inv_zero_dx … H) -L1 // -| /2 width=1 by lsuby_O2/ -| #I1 #I2 #L1 #L2 #V #e #_ #IHL12 #H lapply (injective_plus_l … H) - /3 width=1 by lsuby_pair/ -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL12 #H lapply (injective_plus_l … H) - /3 width=1 by lsuby_succ/ -] -qed-. - -(* Basic inversion lemmas ***************************************************) - -fact lsuby_inv_atom1_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → L1 = ⋆ → L2 = ⋆. -#L1 #L2 #d #e * -L1 -L2 -d -e // -[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #H destruct -| #I1 #I2 #L1 #L2 #V #e #_ #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #H destruct -] -qed-. - -lemma lsuby_inv_atom1: ∀L2,d,e. ⋆ ⊑×[d, e] L2 → L2 = ⋆. -/2 width=5 by lsuby_inv_atom1_aux/ qed-. - -fact lsuby_inv_zero1_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → - ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → d = 0 → e = 0 → - L2 = ⋆ ∨ - ∃∃J2,K2,W2. K1 ⊑×[0, 0] K2 & L2 = K2.ⓑ{J2}W2. -#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/ -[ #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J1 #K1 #W1 #H #_ #_ destruct - /3 width=5 by ex2_3_intro, or_intror/ -| #I1 #I2 #L1 #L2 #V #e #_ #J1 #K1 #W1 #_ #_ #H - elim (ysucc_inv_O_dx … H) -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J1 #K1 #W1 #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma lsuby_inv_zero1: ∀I1,K1,L2,V1. K1.ⓑ{I1}V1 ⊑×[0, 0] L2 → - L2 = ⋆ ∨ - ∃∃I2,K2,V2. K1 ⊑×[0, 0] K2 & L2 = K2.ⓑ{I2}V2. -/2 width=9 by lsuby_inv_zero1_aux/ qed-. - -fact lsuby_inv_pair1_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → - ∀J1,K1,W. L1 = K1.ⓑ{J1}W → d = 0 → 0 < e → - L2 = ⋆ ∨ - ∃∃J2,K2. K1 ⊑×[0, ⫰e] K2 & L2 = K2.ⓑ{J2}W. -#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/ -[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W #_ #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #e #HL12 #J1 #K1 #W #H #_ #_ destruct - /3 width=4 by ex2_2_intro, or_intror/ -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J1 #K1 #W #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma lsuby_inv_pair1: ∀I1,K1,L2,V,e. K1.ⓑ{I1}V ⊑×[0, e] L2 → 0 < e → - L2 = ⋆ ∨ - ∃∃I2,K2. K1 ⊑×[0, ⫰e] K2 & L2 = K2.ⓑ{I2}V. -/2 width=6 by lsuby_inv_pair1_aux/ qed-. - -fact lsuby_inv_succ1_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → - ∀J1,K1,W1. L1 = K1.ⓑ{J1}W1 → 0 < d → - L2 = ⋆ ∨ - ∃∃J2,K2,W2. K1 ⊑×[⫰d, e] K2 & L2 = K2.ⓑ{J2}W2. -#L1 #L2 #d #e * -L1 -L2 -d -e /2 width=1 by or_introl/ -[ #I1 #I2 #L1 #L2 #V1 #V2 #_ #J1 #K1 #W1 #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #e #_ #J1 #K1 #W1 #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #J1 #K1 #W1 #H #_ destruct - /3 width=5 by ex2_3_intro, or_intror/ -] -qed-. - -lemma lsuby_inv_succ1: ∀I1,K1,L2,V1,d,e. K1.ⓑ{I1}V1 ⊑×[d, e] L2 → 0 < d → - L2 = ⋆ ∨ - ∃∃I2,K2,V2. K1 ⊑×[⫰d, e] K2 & L2 = K2.ⓑ{I2}V2. -/2 width=5 by lsuby_inv_succ1_aux/ qed-. - -fact lsuby_inv_zero2_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → - ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → d = 0 → e = 0 → - ∃∃J1,K1,W1. K1 ⊑×[0, 0] K2 & L1 = K1.ⓑ{J1}W1. -#L1 #L2 #d #e * -L1 -L2 -d -e -[ #L1 #d #e #J2 #K2 #W1 #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #HL12 #J2 #K2 #W2 #H #_ #_ destruct - /2 width=5 by ex2_3_intro/ -| #I1 #I2 #L1 #L2 #V #e #_ #J2 #K2 #W2 #_ #_ #H - elim (ysucc_inv_O_dx … H) -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J2 #K2 #W2 #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma lsuby_inv_zero2: ∀I2,K2,L1,V2. L1 ⊑×[0, 0] K2.ⓑ{I2}V2 → - ∃∃I1,K1,V1. K1 ⊑×[0, 0] K2 & L1 = K1.ⓑ{I1}V1. -/2 width=9 by lsuby_inv_zero2_aux/ qed-. - -fact lsuby_inv_pair2_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → - ∀J2,K2,W. L2 = K2.ⓑ{J2}W → d = 0 → 0 < e → - ∃∃J1,K1. K1 ⊑×[0, ⫰e] K2 & L1 = K1.ⓑ{J1}W. -#L1 #L2 #d #e * -L1 -L2 -d -e -[ #L1 #d #e #J2 #K2 #W #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W #_ #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #e #HL12 #J2 #K2 #W #H #_ #_ destruct - /2 width=4 by ex2_2_intro/ -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #J2 #K2 #W #_ #H - elim (ysucc_inv_O_dx … H) -] -qed-. - -lemma lsuby_inv_pair2: ∀I2,K2,L1,V,e. L1 ⊑×[0, e] K2.ⓑ{I2}V → 0 < e → - ∃∃I1,K1. K1 ⊑×[0, ⫰e] K2 & L1 = K1.ⓑ{I1}V. -/2 width=6 by lsuby_inv_pair2_aux/ qed-. - -fact lsuby_inv_succ2_aux: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → - ∀J2,K2,W2. L2 = K2.ⓑ{J2}W2 → 0 < d → - ∃∃J1,K1,W1. K1 ⊑×[⫰d, e] K2 & L1 = K1.ⓑ{J1}W1. -#L1 #L2 #d #e * -L1 -L2 -d -e -[ #L1 #d #e #J2 #K2 #W2 #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #_ #J2 #K2 #W2 #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #e #_ #J2 #K1 #W2 #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #HL12 #J2 #K2 #W2 #H #_ destruct - /2 width=5 by ex2_3_intro/ -] -qed-. - -lemma lsuby_inv_succ2: ∀I2,K2,L1,V2,d,e. L1 ⊑×[d, e] K2.ⓑ{I2}V2 → 0 < d → - ∃∃I1,K1,V1. K1 ⊑×[⫰d, e] K2 & L1 = K1.ⓑ{I1}V1. -/2 width=5 by lsuby_inv_succ2_aux/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lsuby_fwd_length: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → |L2| ≤ |L1|. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e normalize /2 width=1 by le_S_S/ -qed-. - -(* Properties on basic slicing **********************************************) - -lemma lsuby_ldrop_trans_be: ∀L1,L2,d,e. L1 ⊑×[d, e] L2 → - ∀I2,K2,W,s,i. ⇩[s, 0, i] L2 ≡ K2.ⓑ{I2}W → - d ≤ i → i < d + e → - ∃∃I1,K1. K1 ⊑×[0, ⫰(d+e-i)] K2 & ⇩[s, 0, i] L1 ≡ K1.ⓑ{I1}W. -#L1 #L2 #d #e #H elim H -L1 -L2 -d -e -[ #L1 #d #e #J2 #K2 #W #s #i #H - elim (ldrop_inv_atom1 … H) -H #H destruct -| #I1 #I2 #L1 #L2 #V1 #V2 #_ #_ #J2 #K2 #W #s #i #_ #_ #H - elim (ylt_yle_false … H) // -| #I1 #I2 #L1 #L2 #V #e #HL12 #IHL12 #J2 #K2 #W #s #i #H #_ >yplus_O1 - elim (ldrop_inv_O1_pair1 … H) -H * #Hi #HLK1 [ -IHL12 | -HL12 ] - [ #_ destruct -I2 >ypred_succ - /2 width=4 by ldrop_pair, ex2_2_intro/ - | lapply (ylt_inv_O1 i ?) /2 width=1 by ylt_inj/ - #H yminus_succ yplus_succ1 #H lapply (ylt_inv_succ … H) -H - #Hide lapply (ldrop_inv_drop1_lt … HLK2 ?) -HLK2 /2 width=1 by ylt_O/ - #HLK1 elim (IHL12 … HLK1) -IHL12 -HLK1 yminus_SO2 - /4 width=4 by ylt_O, ldrop_drop_lt, ex2_2_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/lsuby_lsuby.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/lsuby_lsuby.ma deleted file mode 100644 index 24361d3c0..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/lsuby_lsuby.ma +++ /dev/null @@ -1,32 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/lsuby.ma". - -(* LOCAL ENVIRONMENT REFINEMENT FOR EXTENDED SUBSTITUTION *******************) - -(* Main properties **********************************************************) - -theorem lsuby_trans: ∀d,e. Transitive … (lsuby d e). -#d #e #L1 #L2 #H elim H -L1 -L2 -d -e -[ #L1 #d #e #X #H lapply (lsuby_inv_atom1 … H) -H - #H destruct // -| #I1 #I2 #L1 #L #V1 #V #_ #IHL1 #X #H elim (lsuby_inv_zero1 … H) -H // - * #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lsuby_zero/ -| #I1 #I2 #L1 #L2 #V #e #_ #IHL1 #X #H elim (lsuby_inv_pair1 … H) -H // - * #I2 #L2 #HL2 #H destruct /3 width=1 by lsuby_pair/ -| #I1 #I2 #L1 #L2 #V1 #V2 #d #e #_ #IHL1 #X #H elim (lsuby_inv_succ1 … H) -H // - * #I2 #L2 #V2 #HL2 #H destruct /3 width=1 by lsuby_succ/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys.ma deleted file mode 100644 index 10c59120d..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys.ma +++ /dev/null @@ -1,166 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/psubststar_6.ma". -include "basic_2/relocation/cpy.ma". - -(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************) - -definition cpys: ynat → ynat → relation4 genv lenv term term ≝ - λd,e,G. LTC … (cpy d e G). - -interpretation "context-sensitive extended multiple substritution (term)" - 'PSubstStar G L T1 d e T2 = (cpys d e G L T1 T2). - -(* Basic eliminators ********************************************************) - -lemma cpys_ind: ∀G,L,T1,d,e. ∀R:predicate term. R T1 → - (∀T,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶[d, e] T2 → R T → R T2) → - ∀T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R T2. -#G #L #T1 #d #e #R #HT1 #IHT1 #T2 #HT12 -@(TC_star_ind … HT1 IHT1 … HT12) // -qed-. - -lemma cpys_ind_dx: ∀G,L,T2,d,e. ∀R:predicate term. R T2 → - (∀T1,T. ⦃G, L⦄ ⊢ T1 ▶[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 → R T → R T1) → - ∀T1. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R T1. -#G #L #T2 #d #e #R #HT2 #IHT2 #T1 #HT12 -@(TC_star_ind_dx … HT2 IHT2 … HT12) // -qed-. - -(* Basic properties *********************************************************) - -lemma cpy_cpys: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. -/2 width=1 by inj/ qed. - -lemma cpys_strap1: ∀G,L,T1,T,T2,d,e. - ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. -normalize /2 width=3 by step/ qed-. - -lemma cpys_strap2: ∀G,L,T1,T,T2,d,e. - ⦃G, L⦄ ⊢ T1 ▶[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. -normalize /2 width=3 by TC_strap/ qed-. - -lemma lsuby_cpys_trans: ∀G,d,e. lsub_trans … (cpys d e G) (lsuby d e). -/3 width=5 by lsuby_cpy_trans, LTC_lsub_trans/ -qed-. - -lemma cpys_refl: ∀G,L,d,e. reflexive … (cpys d e G L). -/2 width=1 by cpy_cpys/ qed. - -lemma cpys_bind: ∀G,L,V1,V2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 → - ∀I,T1,T2. ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 → - ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[d, e] ⓑ{a,I}V2.T2. -#G #L #V1 #V2 #d #e #HV12 @(cpys_ind … HV12) -V2 -[ #I #T1 #T2 #HT12 @(cpys_ind … HT12) -T2 /3 width=5 by cpys_strap1, cpy_bind/ -| /3 width=5 by cpys_strap1, cpy_bind/ -] -qed. - -lemma cpys_flat: ∀G,L,V1,V2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 → - ∀T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → - ∀I. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[d, e] ⓕ{I}V2.T2. -#G #L #V1 #V2 #d #e #HV12 @(cpys_ind … HV12) -V2 -[ #T1 #T2 #HT12 @(cpys_ind … HT12) -T2 /3 width=5 by cpys_strap1, cpy_flat/ -| /3 width=5 by cpys_strap1, cpy_flat/ -qed. - -lemma cpys_weak: ∀G,L,T1,T2,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T2 → - ∀d2,e2. d2 ≤ d1 → d1 + e1 ≤ d2 + e2 → - ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T2. -#G #L #T1 #T2 #d1 #e1 #H #d1 #d2 #Hd21 #Hde12 @(cpys_ind … H) -T2 -/3 width=7 by cpys_strap1, cpy_weak/ -qed-. - -lemma cpys_weak_top: ∀G,L,T1,T2,d,e. - ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, |L| - d] T2. -#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2 -/3 width=4 by cpys_strap1, cpy_weak_top/ -qed-. - -lemma cpys_weak_full: ∀G,L,T1,T2,d,e. - ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[0, |L|] T2. -#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2 -/3 width=5 by cpys_strap1, cpy_weak_full/ -qed-. - -(* Basic forward lemmas *****************************************************) - -lemma cpys_fwd_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → - ∀T1,d,e. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → d + e ≤ dt + et → - ∃∃T2. ⦃G, L⦄ ⊢ U1 ▶*[d+e, dt+et-(d+e)] U2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H #T1 #d #e #HTU1 #Hddt #Hdedet @(cpys_ind … H) -U2 -[ /2 width=3 by ex2_intro/ -| -HTU1 #U #U2 #_ #HU2 * #T #HU1 #HTU - elim (cpy_fwd_up … HU2 … HTU) -HU2 -HTU /3 width=3 by cpys_strap1, ex2_intro/ -] -qed-. - -lemma cpys_fwd_tw: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ♯{T1} ≤ ♯{T2}. -#G #L #T1 #T2 #d #e #H @(cpys_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT1 lapply (cpy_fwd_tw … HT2) -HT2 -/2 width=3 by transitive_le/ -qed-. - -(* Basic inversion lemmas ***************************************************) - -(* Note: this can be derived from cpys_inv_atom1 *) -lemma cpys_inv_sort1: ∀G,L,T2,k,d,e. ⦃G, L⦄ ⊢ ⋆k ▶*[d, e] T2 → T2 = ⋆k. -#G #L #T2 #k #d #e #H @(cpys_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT1 destruct ->(cpy_inv_sort1 … HT2) -HT2 // -qed-. - -(* Note: this can be derived from cpys_inv_atom1 *) -lemma cpys_inv_gref1: ∀G,L,T2,p,d,e. ⦃G, L⦄ ⊢ §p ▶*[d, e] T2 → T2 = §p. -#G #L #T2 #p #d #e #H @(cpys_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT1 destruct ->(cpy_inv_gref1 … HT2) -HT2 // -qed-. - -lemma cpys_inv_bind1: ∀a,I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ▶*[d, e] U2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 & - ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 & - U2 = ⓑ{a,I}V2.T2. -#a #I #G #L #V1 #T1 #U2 #d #e #H @(cpys_ind … H) -U2 -[ /2 width=5 by ex3_2_intro/ -| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct - elim (cpy_inv_bind1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H - lapply (lsuby_cpy_trans … HT2 (L.ⓑ{I}V1) ?) -HT2 - /3 width=5 by cpys_strap1, lsuby_succ, ex3_2_intro/ -] -qed-. - -lemma cpys_inv_flat1: ∀I,G,L,V1,T1,U2,d,e. ⦃G, L⦄ ⊢ ⓕ{I}V1.T1 ▶*[d, e] U2 → - ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 & ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 & - U2 = ⓕ{I}V2.T2. -#I #G #L #V1 #T1 #U2 #d #e #H @(cpys_ind … H) -U2 -[ /2 width=5 by ex3_2_intro/ -| #U #U2 #_ #HU2 * #V #T #HV1 #HT1 #H destruct - elim (cpy_inv_flat1 … HU2) -HU2 - /3 width=5 by cpys_strap1, ex3_2_intro/ -] -qed-. - -lemma cpys_inv_refl_O2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶*[d, 0] T2 → T1 = T2. -#G #L #T1 #T2 #d #H @(cpys_ind … H) -T2 // -#T #T2 #_ #HT2 #IHT1 <(cpy_inv_refl_O2 … HT2) -HT2 // -qed-. - -lemma cpys_inv_lift1_eq: ∀G,L,U1,U2. ∀d,e:nat. - ⦃G, L⦄ ⊢ U1 ▶*[d, e] U2 → ∀T1. ⇧[d, e] T1 ≡ U1 → U1 = U2. -#G #L #U1 #U2 #d #e #H #T1 #HTU1 @(cpys_ind … H) -U2 -/2 width=7 by cpy_inv_lift1_eq/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.ma deleted file mode 100644 index e297be698..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_alt.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/notation/relations/psubststaralt_6.ma". -include "basic_2/substitution/cpys_lift.ma". - -(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************) - -(* alternative definition of cpys *) -inductive cpysa: ynat → ynat → relation4 genv lenv term term ≝ -| cpysa_atom : ∀I,G,L,d,e. cpysa d e G L (⓪{I}) (⓪{I}) -| cpysa_subst: ∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d+e → - ⇩[i] L ≡ K.ⓑ{I}V1 → cpysa 0 (⫰(d+e-i)) G K V1 V2 → - ⇧[0, i+1] V2 ≡ W2 → cpysa d e G L (#i) W2 -| cpysa_bind : ∀a,I,G,L,V1,V2,T1,T2,d,e. - cpysa d e G L V1 V2 → cpysa (⫯d) e G (L.ⓑ{I}V1) T1 T2 → - cpysa d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) -| cpysa_flat : ∀I,G,L,V1,V2,T1,T2,d,e. - cpysa d e G L V1 V2 → cpysa d e G L T1 T2 → - cpysa d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) -. - -interpretation - "context-sensitive extended multiple substritution (term) alternative" - 'PSubstStarAlt G L T1 d e T2 = (cpysa d e G L T1 T2). - -(* Basic properties *********************************************************) - -lemma lsuby_cpysa_trans: ∀G,d,e. lsub_trans … (cpysa d e G) (lsuby d e). -#G #d #e #L1 #T1 #T2 #H elim H -G -L1 -T1 -T2 -d -e -[ // -| #I #G #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK1 #_ #HVW2 #IHV12 #L2 #HL12 - elim (lsuby_ldrop_trans_be … HL12 … HLK1) -HL12 -HLK1 /3 width=7 by cpysa_subst/ -| /4 width=1 by lsuby_succ, cpysa_bind/ -| /3 width=1 by cpysa_flat/ -] -qed-. - -lemma cpysa_refl: ∀G,T,L,d,e. ⦃G, L⦄ ⊢ T ▶▶*[d, e] T. -#G #T elim T -T // -#I elim I -I /2 width=1 by cpysa_bind, cpysa_flat/ -qed. - -lemma cpysa_cpy_trans: ∀G,L,T1,T,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T → - ∀T2. ⦃G, L⦄ ⊢ T ▶[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2. -#G #L #T1 #T #d #e #H elim H -G -L -T1 -T -d -e -[ #I #G #L #d #e #X #H - elim (cpy_inv_atom1 … H) -H // * /2 width=7 by cpysa_subst/ -| #I #G #L #K #V1 #V2 #W2 #i #d #e #Hdi #Hide #HLK #_ #HVW2 #IHV12 #T2 #H - lapply (ldrop_fwd_drop2 … HLK) #H0LK - lapply (cpy_weak … H 0 (d+e) ? ?) -H // #H - elim (cpy_inv_lift1_be … H … H0LK … HVW2) -H -H0LK -HVW2 - /3 width=7 by cpysa_subst, ylt_fwd_le_succ/ -| #a #I #G #L #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H - elim (cpy_inv_bind1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct - /5 width=5 by cpysa_bind, lsuby_cpy_trans, lsuby_succ/ -| #I #G #L #V1 #V #T1 #T #d #e #_ #_ #IHV1 #IHT1 #X #H - elim (cpy_inv_flat1 … H) -H #V2 #T2 #HV2 #HT2 #H destruct /3 width=1 by cpysa_flat/ -] -qed-. - -lemma cpys_cpysa: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2. -/3 width=8 by cpysa_cpy_trans, cpys_ind/ qed. - -(* Basic inversion lemmas ***************************************************) - -lemma cpysa_inv_cpys: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶▶*[d, e] T2 → ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. -#G #L #T1 #T2 #d #e #H elim H -G -L -T1 -T2 -d -e -/2 width=7 by cpys_subst, cpys_flat, cpys_bind, cpy_cpys/ -qed-. - -(* Advanced eliminators *****************************************************) - -lemma cpys_ind_alt: ∀R:ynat→ynat→relation4 genv lenv term term. - (∀I,G,L,d,e. R d e G L (⓪{I}) (⓪{I})) → - (∀I,G,L,K,V1,V2,W2,i,d,e. d ≤ yinj i → i < d + e → - ⇩[i] L ≡ K.ⓑ{I}V1 → ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(d+e-i)] V2 → - ⇧[O, i+1] V2 ≡ W2 → R O (⫰(d+e-i)) G K V1 V2 → R d e G L (#i) W2 - ) → - (∀a,I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 → - ⦃G, L.ⓑ{I}V1⦄ ⊢ T1 ▶*[⫯d, e] T2 → R d e G L V1 V2 → - R (⫯d) e G (L.ⓑ{I}V1) T1 T2 → R d e G L (ⓑ{a,I}V1.T1) (ⓑ{a,I}V2.T2) - ) → - (∀I,G,L,V1,V2,T1,T2,d,e. ⦃G, L⦄ ⊢ V1 ▶*[d, e] V2 → - ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R d e G L V1 V2 → - R d e G L T1 T2 → R d e G L (ⓕ{I}V1.T1) (ⓕ{I}V2.T2) - ) → - ∀d,e,G,L,T1,T2. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → R d e G L T1 T2. -#R #H1 #H2 #H3 #H4 #d #e #G #L #T1 #T2 #H elim (cpys_cpysa … H) -G -L -T1 -T2 -d -e -/3 width=8 by cpysa_inv_cpys/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_cpys.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_cpys.ma deleted file mode 100644 index 52759ca7f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_cpys.ma +++ /dev/null @@ -1,117 +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/relocation/cpy_cpy.ma". -include "basic_2/substitution/cpys_alt.ma". - -(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************) - -(* Advanced inversion lemmas ************************************************) - -lemma cpys_inv_SO2: ∀G,L,T1,T2,d. ⦃G, L⦄ ⊢ T1 ▶*[d, 1] T2 → ⦃G, L⦄ ⊢ T1 ▶[d, 1] T2. -#G #L #T1 #T2 #d #H @(cpys_ind … H) -T2 /2 width=3 by cpy_trans_ge/ -qed-. - -(* Advanced properties ******************************************************) - -lemma cpys_strip_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶*[d1, e1] T1 → - ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶*[d1, e1] T. -normalize /3 width=3 by cpy_conf_eq, TC_strip1/ qed-. - -lemma cpys_strip_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶*[d1, e1] T1 → - ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶[d2, e2] T2 → - (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → - ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶*[d1, e1] T. -normalize /3 width=3 by cpy_conf_neq, TC_strip1/ qed-. - -lemma cpys_strap1_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T0 → - ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶[d2, e2] T2 → d2 + e2 ≤ d1 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶[d2, e2] T & ⦃G, L⦄ ⊢ T ▶*[d1, e1] T2. -normalize /3 width=3 by cpy_trans_down, TC_strap1/ qed. - -lemma cpys_strap2_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶[d1, e1] T0 → - ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → d2 + e2 ≤ d1 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T ▶[d1, e1] T2. -normalize /3 width=3 by cpy_trans_down, TC_strap2/ qed-. - -lemma cpys_split_up: ∀G,L,T1,T2,d,e. ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2 → - ∀i. d ≤ i → i ≤ d + e → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d, i - d] T & ⦃G, L⦄ ⊢ T ▶*[i, d + e - i] T2. -#G #L #T1 #T2 #d #e #H #i #Hdi #Hide @(cpys_ind … H) -T2 -[ /2 width=3 by ex2_intro/ -| #T #T2 #_ #HT12 * #T3 #HT13 #HT3 - elim (cpy_split_up … HT12 … Hide) -HT12 -Hide #T0 #HT0 #HT02 - elim (cpys_strap1_down … HT3 … HT0) -T /3 width=5 by cpys_strap1, ex2_intro/ - >ymax_pre_sn_comm // -] -qed-. - -lemma cpys_inv_lift1_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[d, dt + et - (yinj d + e)] T2 & - ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #HU12 #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet -elim (cpys_split_up … HU12 (d + e)) -HU12 // -Hdedet #U #HU1 #HU2 -lapply (cpys_weak … HU1 d e ? ?) -HU1 // [ >ymax_pre_sn_comm // ] -Hddt -Hdtde #HU1 -lapply (cpys_inv_lift1_eq … HU1 … HTU1) -HU1 #HU1 destruct -elim (cpys_inv_lift1_ge … HU2 … HLK … HTU1) -HU2 -HLK -HTU1 // ->yplus_minus_inj /2 width=3 by ex2_intro/ -qed-. - -(* Main properties **********************************************************) - -theorem cpys_conf_eq: ∀G,L,T0,T1,d1,e1. ⦃G, L⦄ ⊢ T0 ▶*[d1, e1] T1 → - ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T2 ▶*[d1, e1] T. -normalize /3 width=3 by cpy_conf_eq, TC_confluent2/ qed-. - -theorem cpys_conf_neq: ∀G,L1,T0,T1,d1,e1. ⦃G, L1⦄ ⊢ T0 ▶*[d1, e1] T1 → - ∀L2,T2,d2,e2. ⦃G, L2⦄ ⊢ T0 ▶*[d2, e2] T2 → - (d1 + e1 ≤ d2 ∨ d2 + e2 ≤ d1) → - ∃∃T. ⦃G, L2⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L1⦄ ⊢ T2 ▶*[d1, e1] T. -normalize /3 width=3 by cpy_conf_neq, TC_confluent2/ qed-. - -theorem cpys_trans_eq: ∀G,L,T1,T,T2,d,e. - ⦃G, L⦄ ⊢ T1 ▶*[d, e] T → ⦃G, L⦄ ⊢ T ▶*[d, e] T2 → - ⦃G, L⦄ ⊢ T1 ▶*[d, e] T2. -normalize /2 width=3 by trans_TC/ qed-. - -theorem cpys_trans_down: ∀G,L,T1,T0,d1,e1. ⦃G, L⦄ ⊢ T1 ▶*[d1, e1] T0 → - ∀T2,d2,e2. ⦃G, L⦄ ⊢ T0 ▶*[d2, e2] T2 → d2 + e2 ≤ d1 → - ∃∃T. ⦃G, L⦄ ⊢ T1 ▶*[d2, e2] T & ⦃G, L⦄ ⊢ T ▶*[d1, e1] T2. -normalize /3 width=3 by cpy_trans_down, TC_transitive2/ qed-. - -theorem cpys_antisym_eq: ∀G,L1,T1,T2,d,e. ⦃G, L1⦄ ⊢ T1 ▶*[d, e] T2 → - ∀L2. ⦃G, L2⦄ ⊢ T2 ▶*[d, e] T1 → T1 = T2. -#G #L1 #T1 #T2 #d #e #H @(cpys_ind_alt … H) -G -L1 -T1 -T2 // -[ #I1 #G #L1 #K1 #V1 #V2 #W2 #i #d #e #Hdi #Hide #_ #_ #HVW2 #_ #L2 #HW2 - elim (lt_or_ge (|L2|) (i+1)) #Hi [ -Hdi -Hide | ] - [ lapply (cpys_weak_full … HW2) -HW2 #HW2 - lapply (cpys_weak … HW2 0 (i+1) ? ?) -HW2 // - [ >yplus_O1 >yplus_O1 /3 width=1 by ylt_fwd_le, ylt_inj/ ] -Hi - #HW2 >(cpys_inv_lift1_eq … HW2) -HW2 // - | elim (ldrop_O1_le … Hi) -Hi #K2 #HLK2 - elim (cpys_inv_lift1_ge_up … HW2 … HLK2 … HVW2 ? ? ?) -HW2 -HLK2 -HVW2 - /2 width=1 by ylt_fwd_le_succ, yle_succ_dx/ -Hdi -Hide - #X #_ #H elim (lift_inv_lref2_be … H) -H // - ] -| #a #I #G #L1 #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_bind1 … H) -H - #V #T #HV2 #HT2 #H destruct - lapply (IHV12 … HV2) #H destruct -IHV12 -HV2 /3 width=2 by eq_f2/ -| #I #G #L1 #V1 #V2 #T1 #T2 #d #e #_ #_ #IHV12 #IHT12 #L2 #H elim (cpys_inv_flat1 … H) -H - #V #T #HV2 #HT2 #H destruct /3 width=2 by eq_f2/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_lift.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_lift.ma deleted file mode 100644 index b6d9e45a1..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cpys_lift.ma +++ /dev/null @@ -1,218 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/relocation/cpy_lift.ma". -include "basic_2/substitution/cpys.ma". - -(* CONTEXT-SENSITIVE EXTENDED MULTIPLE SUBSTITUTION FOR TERMS ***************) - -(* Advanced properties ******************************************************) - -lemma cpys_subst: ∀I,G,L,K,V,U1,i,d,e. - d ≤ yinj i → i < d + e → - ⇩[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ▶*[0, ⫰(d+e-i)] U1 → - ∀U2. ⇧[0, i+1] U1 ≡ U2 → ⦃G, L⦄ ⊢ #i ▶*[d, e] U2. -#I #G #L #K #V #U1 #i #d #e #Hdi #Hide #HLK #H @(cpys_ind … H) -U1 -[ /3 width=5 by cpy_cpys, cpy_subst/ -| #U #U1 #_ #HU1 #IHU #U2 #HU12 - elim (lift_total U 0 (i+1)) #U0 #HU0 - lapply (IHU … HU0) -IHU #H - lapply (ldrop_fwd_drop2 … HLK) -HLK #HLK - lapply (cpy_lift_ge … HU1 … HLK HU0 HU12 ?) -HU1 -HLK -HU0 -HU12 // #HU02 - lapply (cpy_weak … HU02 d e ? ?) -HU02 - [2,3: /2 width=3 by cpys_strap1, yle_succ_dx/ ] - >yplus_O1 ymax_pre_sn_comm /2 width=1 by ylt_fwd_le_succ/ -] -qed. - -lemma cpys_subst_Y2: ∀I,G,L,K,V,U1,i,d. - d ≤ yinj i → - ⇩[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ V ▶*[0, ∞] U1 → - ∀U2. ⇧[0, i+1] U1 ≡ U2 → ⦃G, L⦄ ⊢ #i ▶*[d, ∞] U2. -#I #G #L #K #V #U1 #i #d #Hdi #HLK #HVU1 #U2 #HU12 -@(cpys_subst … HLK … HU12) >yminus_Y_inj // -qed. - -(* Advanced inverion lemmas *************************************************) - -lemma cpys_inv_atom1: ∀I,G,L,T2,d,e. ⦃G, L⦄ ⊢ ⓪{I} ▶*[d, e] T2 → - T2 = ⓪{I} ∨ - ∃∃J,K,V1,V2,i. d ≤ yinj i & i < d + e & - ⇩[i] L ≡ K.ⓑ{J}V1 & - ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2 & - ⇧[O, i+1] V2 ≡ T2 & - I = LRef i. -#I #G #L #T2 #d #e #H @(cpys_ind … H) -T2 -[ /2 width=1 by or_introl/ -| #T #T2 #_ #HT2 * - [ #H destruct - elim (cpy_inv_atom1 … HT2) -HT2 [ /2 width=1 by or_introl/ | * /3 width=11 by ex6_5_intro, or_intror/ ] - | * #J #K #V1 #V #i #Hdi #Hide #HLK #HV1 #HVT #HI - lapply (ldrop_fwd_drop2 … HLK) #H - elim (cpy_inv_lift1_ge_up … HT2 … H … HVT) -HT2 -H -HVT - [2,3,4: /2 width=1 by ylt_fwd_le_succ, yle_succ_dx/ ] - /4 width=11 by cpys_strap1, ex6_5_intro, or_intror/ - ] -] -qed-. - -lemma cpys_inv_lref1: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶*[d, e] T2 → - T2 = #i ∨ - ∃∃I,K,V1,V2. d ≤ i & i < d + e & - ⇩[i] L ≡ K.ⓑ{I}V1 & - ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2 & - ⇧[O, i+1] V2 ≡ T2. -#G #L #T2 #i #d #e #H elim (cpys_inv_atom1 … H) -H /2 width=1 by or_introl/ -* #I #K #V1 #V2 #j #Hdj #Hjde #HLK #HV12 #HVT2 #H destruct /3 width=7 by ex5_4_intro, or_intror/ -qed-. - -lemma cpys_inv_lref1_ldrop: ∀G,L,T2,i,d,e. ⦃G, L⦄ ⊢ #i ▶*[d, e] T2 → - ∀I,K,V1. ⇩[i] L ≡ K.ⓑ{I}V1 → - ∀V2. ⇧[O, i+1] V2 ≡ T2 → - ∧∧ ⦃G, K⦄ ⊢ V1 ▶*[0, ⫰(d+e-i)] V2 - & d ≤ i - & i < d + e. -#G #L #T2 #i #d #e #H #I #K #V1 #HLK #V2 #HVT2 elim (cpys_inv_lref1 … H) -H -[ #H destruct elim (lift_inv_lref2_be … HVT2) -HVT2 -HLK // -| * #Z #Y #X1 #X2 #Hdi #Hide #HLY #HX12 #HXT2 - lapply (lift_inj … HXT2 … HVT2) -T2 #H destruct - lapply (ldrop_mono … HLY … HLK) -L #H destruct - /2 width=1 by and3_intro/ -] -qed-. - -(* Properties on relocation *************************************************) - -lemma cpys_lift_le: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 → - ∀L,U1,s,d,e. dt + et ≤ yinj d → ⇩[s, d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → - ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2. -#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hdetd #HLK #HTU1 @(cpys_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 (cpy_lift_le … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/ -] -qed-. - -lemma cpys_lift_be: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 → - ∀L,U1,s,d,e. dt ≤ yinj d → d ≤ dt + et → - ⇩[s, d, e] L ≡ K → ⇧[d, e] T1 ≡ U1 → - ∀U2. ⇧[d, e] T2 ≡ U2 → ⦃G, L⦄ ⊢ U1 ▶*[dt, et + e] U2. -#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hdtd #Hddet #HLK #HTU1 @(cpys_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 (cpy_lift_be … HT2 … HLK HTU HTU2 ? ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/ -] -qed-. - -lemma cpys_lift_ge: ∀G,K,T1,T2,dt,et. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 → - ∀L,U1,s,d,e. yinj d ≤ dt → ⇩[s, d, e] L ≡ K → - ⇧[d, e] T1 ≡ U1 → ∀U2. ⇧[d, e] T2 ≡ U2 → - ⦃G, L⦄ ⊢ U1 ▶*[dt+e, et] U2. -#G #K #T1 #T2 #dt #et #H #L #U1 #s #d #e #Hddt #HLK #HTU1 @(cpys_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 (cpy_lift_ge … HT2 … HLK HTU HTU2 ?) -HT2 -HLK -HTU -HTU2 /2 width=3 by cpys_strap1/ -] -qed-. - -(* Inversion lemmas for relocation ******************************************) - -lemma cpys_inv_lift1_le: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt + et ≤ d → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, et] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdetd @(cpys_ind … H) -U2 -[ /2 width=3 by ex2_intro/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (cpy_inv_lift1_le … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ -] -qed-. - -lemma cpys_inv_lift1_be: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → yinj d + e ≤ dt + et → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, et - e] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdedet @(cpys_ind … H) -U2 -[ /2 width=3 by ex2_intro/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (cpy_inv_lift1_be … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ -] -qed-. - -lemma cpys_inv_lift1_ge: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - yinj d + e ≤ dt → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt - e, et] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdedt @(cpys_ind … H) -U2 -[ /2 width=3 by ex2_intro/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (cpy_inv_lift1_ge … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ -] -qed-. - -(* Advanced inversion lemmas on relocation **********************************) - -lemma cpys_inv_lift1_ge_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - d ≤ dt → dt ≤ yinj d + e → yinj d + e ≤ dt + et → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[d, dt + et - (yinj d + e)] T2 & - ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hddt #Hdtde #Hdedet @(cpys_ind … H) -U2 -[ /2 width=3 by ex2_intro/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (cpy_inv_lift1_ge_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ -] -qed-. - -lemma cpys_inv_lift1_be_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → dt + et ≤ yinj d + e → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hdetde @(cpys_ind … H) -U2 -[ /2 width=3 by ex2_intro/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (cpy_inv_lift1_be_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ -] -qed-. - -lemma cpys_inv_lift1_le_up: ∀G,L,U1,U2,dt,et. ⦃G, L⦄ ⊢ U1 ▶*[dt, et] U2 → - ∀K,s,d,e. ⇩[s, d, e] L ≡ K → ∀T1. ⇧[d, e] T1 ≡ U1 → - dt ≤ d → d ≤ dt + et → dt + et ≤ yinj d + e → - ∃∃T2. ⦃G, K⦄ ⊢ T1 ▶*[dt, d - dt] T2 & ⇧[d, e] T2 ≡ U2. -#G #L #U1 #U2 #dt #et #H #K #s #d #e #HLK #T1 #HTU1 #Hdtd #Hddet #Hdetde @(cpys_ind … H) -U2 -[ /2 width=3 by ex2_intro/ -| -HTU1 #U #U2 #_ #HU2 * #T #HT1 #HTU - elim (cpy_inv_lift1_le_up … HU2 … HLK … HTU) -HU2 -HLK -HTU /3 width=3 by cpys_strap1, ex2_intro/ -] -qed-. - -lemma cpys_inv_lift1_subst: ∀G,L,W1,W2,d,e. ⦃G, L⦄ ⊢ W1 ▶*[d, e] W2 → - ∀K,V1,i. ⇩[i+1] L ≡ K → ⇧[O, i+1] V1 ≡ W1 → - d ≤ yinj i → i < d + e → - ∃∃V2. ⦃G, K⦄ ⊢ V1 ▶*[O, ⫰(d+e-i)] V2 & ⇧[O, i+1] V2 ≡ W2. -#G #L #W1 #W2 #d #e #HW12 #K #V1 #i #HLK #HVW1 #Hdi #Hide -elim (cpys_inv_lift1_ge_up … HW12 … HLK … HVW1 ? ? ?) // ->yplus_O1 yplus_SO2 -[ >yminus_succ2 /2 width=3 by ex2_intro/ -| /2 width=1 by ylt_fwd_le_succ1/ -| /2 width=3 by yle_trans/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/fqup_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/fqup_lleq.ma new file mode 100644 index 000000000..c0bfa41fe --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/fqup_lleq.ma @@ -0,0 +1,32 @@ +(**************************************************************************) +(* ___ *) +(* ||M|| *) +(* ||A|| A project by Andrea Asperti *) +(* ||T|| *) +(* ||I|| Developers: *) +(* ||T|| The HELM team. *) +(* ||A|| http://helm.cs.unibo.it *) +(* \ / *) +(* \ / This file is distributed under the terms of the *) +(* v GNU General Public License Version 2 *) +(* *) +(**************************************************************************) + +include "basic_2/relocation/fqu_lleq.ma". +include "basic_2/substitution/fqup.ma". + +(* PLUS-ITERATED SUPCLOSURE *************************************************) + +(* Properties on lazy equivalence for local environments ********************) + +lemma lleq_fqup_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃+ ⦃G2, K2, U⦄ → + ∀L1. L1 ⋕[T, 0] L2 → + ∃∃K1. ⦃G1, L1, T⦄ ⊃+ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. +#G1 #G2 #L2 #K2 #T #U #H @(fqup_ind … H) -G2 -K2 -U +[ #G2 #K2 #U #HTU #L1 #HL12 elim (lleq_fqu_trans … HTU … HL12) -L2 + /3 width=3 by fqu_fqup, ex2_intro/ +| #G #G2 #K #K2 #U #U2 #_ #HU2 #IHTU #L1 #HL12 elim (IHTU … HL12) -L2 + #K1 #HTU #HK1 elim (lleq_fqu_trans … HU2 … HK1) -K + /3 width=5 by fqup_strap1, ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/fqus_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/fqus_lleq.ma new file mode 100644 index 000000000..ae209399f --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/fqus_lleq.ma @@ -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/substitution/fqup_lleq.ma". +include "basic_2/substitution/fqus_alt.ma". + +(* STAR-ITERATED SUPCLOSURE *************************************************) + +(* Properties on lazy equivalence for local environments ********************) + +lemma lleq_fqus_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃* ⦃G2, K2, U⦄ → + ∀L1. L1 ⋕[T, 0] L2 → + ∃∃K1. ⦃G1, L1, T⦄ ⊃* ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. +#G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fqus_inv_gen … H) -H +[ #H elim (lleq_fqup_trans … H … HL12) -L2 /3 width=3 by fqup_fqus, ex2_intro/ +| * #HG #HL #HT destruct /2 width=3 by ex2_intro/ +] +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq.ma deleted file mode 100644 index 72a5346ce..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq.ma +++ /dev/null @@ -1,160 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/lazyeq_4.ma". -include "basic_2/substitution/cpys.ma". - -(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) - -definition lleq: relation4 ynat term lenv lenv ≝ - λd,T,L1,L2. |L1| = |L2| ∧ - (∀U. ⦃⋆, L1⦄ ⊢ T ▶*[d, ∞] U ↔ ⦃⋆, L2⦄ ⊢ T ▶*[d, ∞] U). - -interpretation - "lazy equivalence (local environment)" - 'LazyEq T d L1 L2 = (lleq d T L1 L2). - -(* Basic properties *********************************************************) - -lemma lleq_refl: ∀d,T. reflexive … (lleq d T). -/3 width=1 by conj/ qed. - -lemma lleq_sym: ∀d,T. symmetric … (lleq d T). -#d #T #L1 #L2 * /3 width=1 by iff_sym, conj/ -qed-. - -lemma lleq_sort: ∀L1,L2,d,k. |L1| = |L2| → L1 ⋕[⋆k, d] L2. -#L1 #L2 #d #k #HL12 @conj // -HL12 -#U @conj #H >(cpys_inv_sort1 … H) -H // -qed. - -lemma lleq_gref: ∀L1,L2,d,p. |L1| = |L2| → L1 ⋕[§p, d] L2. -#L1 #L2 #d #k #HL12 @conj // -HL12 -#U @conj #H >(cpys_inv_gref1 … H) -H // -qed. - -lemma lleq_bind: ∀a,I,L1,L2,V,T,d. - L1 ⋕[V, d] L2 → L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V → - L1 ⋕[ⓑ{a,I}V.T, d] L2. -#a #I #L1 #L2 #V #T #d * #HL12 #IHV * #_ #IHT @conj // -HL12 -#X @conj #H elim (cpys_inv_bind1 … H) -H -#W #U #HVW #HTU #H destruct -elim (IHV W) -IHV elim (IHT U) -IHT /3 width=1 by cpys_bind/ -qed. - -lemma lleq_flat: ∀I,L1,L2,V,T,d. - L1 ⋕[V, d] L2 → L1 ⋕[T, d] L2 → L1 ⋕[ⓕ{I}V.T, d] L2. -#I #L1 #L2 #V #T #d * #HL12 #IHV * #_ #IHT @conj // -HL12 -#X @conj #H elim (cpys_inv_flat1 … H) -H -#W #U #HVW #HTU #H destruct -elim (IHV W) -IHV elim (IHT U) -IHT -/3 width=1 by cpys_flat/ -qed. - -lemma lleq_be: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → ∀T,d,e. ⇧[d, e] T ≡ U → - d ≤ dt → dt ≤ d + e → L1 ⋕[U, d] L2. -#L1 #L2 #U #dt * #HL12 #IH #T #d #e #HTU #Hddt #Hdtde @conj // -HL12 -#U0 elim (IH U0) -IH #H12 #H21 @conj -#HU0 elim (cpys_fwd_up … HU0 … HTU) // -HU0 /4 width=5 by cpys_weak/ -qed-. - -lemma lsuby_lleq_trans: ∀L2,L,T,d. L2 ⋕[T, d] L → - ∀L1. L1 ⊑×[d, ∞] L2 → |L1| = |L2| → L1 ⋕[T, d] L. -#L2 #L #T #d * #HL2 #IH #L1 #HL12 #H @conj // -HL2 -#U elim (IH U) -IH #Hdx #Hsn @conj #HTU -[ @Hdx -Hdx -Hsn @(lsuby_cpys_trans … HTU) -HTU - /2 width=1 by lsuby_sym/ (**) (* full auto does not work *) -| -H -Hdx /3 width=3 by lsuby_cpys_trans/ -] -qed-. - -lemma lleq_lsuby_trans: ∀L,L1,T,d. L ⋕[T, d] L1 → - ∀L2. L1 ⊑×[d, ∞] L2 → |L1| = |L2| → L ⋕[T, d] L2. -/5 width=4 by lsuby_lleq_trans, lleq_sym, lsuby_sym/ qed-. - -lemma lleq_lsuby_repl: ∀L1,L2,T,d. L1 ⋕[T, d] L2 → - ∀K1. K1 ⊑×[d, ∞] L1 → |K1| = |L1| → - ∀K2. L2 ⊑×[d, ∞] K2 → |L2| = |K2| → - K1 ⋕[T, d] K2. -/3 width=4 by lleq_lsuby_trans, lsuby_lleq_trans/ qed-. - -(* Basic forward lemmas *****************************************************) - -lemma lleq_fwd_length: ∀L1,L2,T,d. L1 ⋕[T, d] L2 → |L1| = |L2|. -#L1 #L2 #T #d * // -qed-. - -lemma lleq_fwd_ldrop_sn: ∀L1,L2,T,d. L1 ⋕[d, T] L2 → ∀K1,i. ⇩[i] L1 ≡ K1 → - ∃K2. ⇩[i] L2 ≡ K2. -#L1 #L2 #T #d #H #K1 #i #HLK1 lapply (lleq_fwd_length … H) -H -#HL12 lapply (ldrop_fwd_length_le2 … HLK1) -HLK1 /2 width=1 by ldrop_O1_le/ -qed-. - -lemma lleq_fwd_ldrop_dx: ∀L1,L2,T,d. L1 ⋕[d, T] L2 → ∀K2,i. ⇩[i] L2 ≡ K2 → - ∃K1. ⇩[i] L1 ≡ K1. -/3 width=6 by lleq_fwd_ldrop_sn, lleq_sym/ qed-. - -lemma lleq_fwd_bind_sn: ∀a,I,L1,L2,V,T,d. - L1 ⋕[ⓑ{a,I}V.T, d] L2 → L1 ⋕[V, d] L2. -#a #I #L1 #L2 #V #T #d * #HL12 #H @conj // -HL12 -#U elim (H (ⓑ{a,I}U.T)) -H -#H1 #H2 @conj -#H [ lapply (H1 ?) | lapply (H2 ?) ] -H1 -H2 -/2 width=1 by cpys_bind/ -H -#H elim (cpys_inv_bind1 … H) -H -#X #Y #H1 #H2 #H destruct // -qed-. - -lemma lleq_fwd_bind_dx: ∀a,I,L1,L2,V,T,d. - L1 ⋕[ⓑ{a,I}V.T, d] L2 → L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V. -#a #I #L1 #L2 #V #T #d * #HL12 #H @conj [ normalize // ] -HL12 -#U elim (H (ⓑ{a,I}V.U)) -H -#H1 #H2 @conj -#H [ lapply (H1 ?) | lapply (H2 ?) ] -H1 -H2 -/2 width=1 by cpys_bind/ -H -#H elim (cpys_inv_bind1 … H) -H -#X #Y #H1 #H2 #H destruct // -qed-. - -lemma lleq_fwd_flat_sn: ∀I,L1,L2,V,T,d. - L1 ⋕[ⓕ{I}V.T, d] L2 → L1 ⋕[V, d] L2. -#I #L1 #L2 #V #T #d * #HL12 #H @conj // -HL12 -#U elim (H (ⓕ{I}U.T)) -H -#H1 #H2 @conj -#H [ lapply (H1 ?) | lapply (H2 ?) ] -H1 -H2 -/2 width=1 by cpys_flat/ -H -#H elim (cpys_inv_flat1 … H) -H -#X #Y #H1 #H2 #H destruct // -qed-. - -lemma lleq_fwd_flat_dx: ∀I,L1,L2,V,T,d. - L1 ⋕[ⓕ{I}V.T, d] L2 → L1 ⋕[T, d] L2. -#I #L1 #L2 #V #T #d * #HL12 #H @conj // -HL12 -#U elim (H (ⓕ{I}V.U)) -H -#H1 #H2 @conj -#H [ lapply (H1 ?) | lapply (H2 ?) ] -H1 -H2 -/2 width=1 by cpys_flat/ -H -#H elim (cpys_inv_flat1 … H) -H -#X #Y #H1 #H2 #H destruct // -qed-. - -(* Basic inversion lemmas ***************************************************) - -lemma lleq_inv_bind: ∀a,I,L1,L2,V,T,d. L1 ⋕[ⓑ{a,I}V.T, d] L2 → - L1 ⋕[V, d] L2 ∧ L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V. -/3 width=4 by lleq_fwd_bind_sn, lleq_fwd_bind_dx, conj/ qed-. - -lemma lleq_inv_flat: ∀I,L1,L2,V,T,d. L1 ⋕[ⓕ{I}V.T, d] L2 → - L1 ⋕[V, d] L2 ∧ L1 ⋕[T, d] L2. -/3 width=3 by lleq_fwd_flat_sn, lleq_fwd_flat_dx, conj/ qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt.ma deleted file mode 100644 index a8fb526e5..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt.ma +++ /dev/null @@ -1,85 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/notation/relations/lazyeqalt_4.ma". -include "basic_2/substitution/lleq_lleq.ma". - -(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) - -(* Note: alternative definition of lleq *) -inductive lleqa: relation4 ynat term lenv lenv ≝ -| lleqa_sort: ∀L1,L2,d,k. |L1| = |L2| → lleqa d (⋆k) L1 L2 -| lleqa_skip: ∀L1,L2,d,i. |L1| = |L2| → yinj i < d → lleqa d (#i) L1 L2 -| lleqa_lref: ∀I1,I2,L1,L2,K1,K2,V,d,i. d ≤ yinj i → - ⇩[i] L1 ≡ K1.ⓑ{I1}V → ⇩[i] L2 ≡ K2.ⓑ{I2}V → - lleqa (yinj 0) V K1 K2 → lleqa d (#i) L1 L2 -| lleqa_free: ∀L1,L2,d,i. |L1| = |L2| → |L1| ≤ i → |L2| ≤ i → lleqa d (#i) L1 L2 -| lleqa_gref: ∀L1,L2,d,p. |L1| = |L2| → lleqa d (§p) L1 L2 -| lleqa_bind: ∀a,I,L1,L2,V,T,d. - lleqa d V L1 L2 → lleqa (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → - lleqa d (ⓑ{a,I}V.T) L1 L2 -| lleqa_flat: ∀I,L1,L2,V,T,d. - lleqa d V L1 L2 → lleqa d T L1 L2 → lleqa d (ⓕ{I}V.T) L1 L2 -. - -interpretation - "lazy equivalence (local environment) alternative" - 'LazyEqAlt T d L1 L2 = (lleqa d T L1 L2). - -(* Main inversion lemmas ****************************************************) - -theorem lleqa_inv_lleq: ∀L1,L2,T,d. L1 ⋕⋕[T, d] L2 → L1 ⋕[T, d] L2. -#L1 #L2 #T #d #H elim H -L1 -L2 -T -d -/2 width=8 by lleq_flat, lleq_bind, lleq_gref, lleq_free, lleq_lref, lleq_skip, lleq_sort/ -qed-. - -(* Main properties **********************************************************) - -theorem lleq_lleqa: ∀L1,T,L2,d. L1 ⋕[T, d] L2 → L1 ⋕⋕[T, d] L2. -#L1 #T @(f2_ind … rfw … L1 T) -L1 -T -#n #IH #L1 * * /3 width=3 by lleqa_gref, lleqa_sort, lleq_fwd_length/ -[ #i #Hn #L2 #d #H elim (lleq_fwd_lref … H) [ * || * ] - /4 width=9 by lleqa_free, lleqa_lref, lleqa_skip, lleq_fwd_length, ldrop_fwd_rfw/ -| #a #I #V #T #Hn #L2 #d #H elim (lleq_inv_bind … H) -H /3 width=1 by lleqa_bind/ -| #I #V #T #Hn #L2 #d #H elim (lleq_inv_flat … H) -H /3 width=1 by lleqa_flat/ -] -qed. - -(* Advanced eliminators *****************************************************) - -lemma lleq_ind_alt: ∀R:relation4 ynat term lenv lenv. ( - ∀L1,L2,d,k. |L1| = |L2| → R d (⋆k) L1 L2 - ) → ( - ∀L1,L2,d,i. |L1| = |L2| → yinj i < d → R d (#i) L1 L2 - ) → ( - ∀I1,I2,L1,L2,K1,K2,V,d,i. d ≤ yinj i → - ⇩[i] L1 ≡ K1.ⓑ{I1}V → ⇩[i] L2 ≡ K2.ⓑ{I2}V → - K1 ⋕[V, yinj O] K2 → R (yinj O) V K1 K2 → R d (#i) L1 L2 - ) → ( - ∀L1,L2,d,i. |L1| = |L2| → |L1| ≤ i → |L2| ≤ i → R d (#i) L1 L2 - ) → ( - ∀L1,L2,d,p. |L1| = |L2| → R d (§p) L1 L2 - ) → ( - ∀a,I,L1,L2,V,T,d. - L1 ⋕[V, d]L2 → L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V → - R d V L1 L2 → R (⫯d) T (L1.ⓑ{I}V) (L2.ⓑ{I}V) → R d (ⓑ{a,I}V.T) L1 L2 - ) → ( - ∀I,L1,L2,V,T,d. - L1 ⋕[V, d]L2 → L1 ⋕[T, d] L2 → - R d V L1 L2 → R d T L1 L2 → R d (ⓕ{I}V.T) L1 L2 - ) → - ∀d,T,L1,L2. L1 ⋕[T, d] L2 → R d T L1 L2. -#R #H1 #H2 #H3 #H4 #H5 #H6 #H7 #d #T #L1 #L2 #H elim (lleq_lleqa … H) -H -/3 width=9 by lleqa_inv_lleq/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_ext.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_ext.ma deleted file mode 100644 index d6bb03fa7..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_ext.ma +++ /dev/null @@ -1,91 +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/lleq_alt.ma". - -(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) - -(* Advanced inversion lemmas ************************************************) - -fact lleq_inv_S_aux: ∀L1,L2,T,d0. L1 ⋕[T, d0] L2 → ∀d. d0 = d + 1 → - ∀K1,K2,I,V. ⇩[d] L1 ≡ K1.ⓑ{I}V → ⇩[d] L2 ≡ K2.ⓑ{I}V → - K1 ⋕[V, 0] K2 → L1 ⋕[T, d] L2. -#L1 #L2 #T #d0 #H @(lleq_ind_alt … H) -L1 -L2 -T -d0 -/2 width=1 by lleq_gref, lleq_free, lleq_sort/ -[ #L1 #L2 #d0 #i #HL12 #Hid #d #H #K1 #K2 #I #V #HLK1 #HLK2 #HV destruct - elim (yle_split_eq i d) /2 width=1 by lleq_skip, ylt_fwd_succ2/ -HL12 -Hid - #H destruct /2 width=8 by lleq_lref/ -| #I1 #I2 #L1 #L2 #K11 #K22 #V #d0 #i #Hd0i #HLK11 #HLK22 #HV #_ #d #H #K1 #K2 #J #W #_ #_ #_ destruct - /3 width=8 by lleq_lref, yle_pred_sn/ -| #a #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W #HLK1 #HLK2 destruct - /4 width=7 by lleq_bind, ldrop_drop/ -| #I #L1 #L2 #V #T #d0 #_ #_ #IHV #IHT #d #H #K1 #K2 #J #W #HLK1 #HLK2 destruct - /3 width=7 by lleq_flat/ -] -qed-. - -lemma lleq_inv_S: ∀T,L1,L2,d. L1 ⋕[T, d+1] L2 → - ∀K1,K2,I,V. ⇩[d] L1 ≡ K1.ⓑ{I}V → ⇩[d] L2 ≡ K2.ⓑ{I}V → - K1 ⋕[V, 0] K2 → L1 ⋕[T, d] L2. -/2 width=7 by lleq_inv_S_aux/ qed-. - -lemma lleq_inv_bind_O: ∀a,I,L1,L2,V,T. L1 ⋕[ⓑ{a,I}V.T, 0] L2 → - L1 ⋕[V, 0] L2 ∧ L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V. -#a #I #L1 #L2 #V #T #H elim (lleq_inv_bind … H) -H -/3 width=7 by ldrop_pair, conj, lleq_inv_S/ -qed-. - -(* Advanced forward lemmas **************************************************) - -lemma lleq_fwd_bind_O: ∀a,I,L1,L2,V,T. L1 ⋕[ⓑ{a,I}V.T, 0] L2 → - L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V. -#a #I #L1 #L2 #V #T #H elim (lleq_inv_bind_O … H) -H // -qed-. - -(* Advanced properties ******************************************************) - -lemma lleq_ge: ∀L1,L2,T,d1. L1 ⋕[T, d1] L2 → ∀d2. d1 ≤ d2 → L1 ⋕[T, d2] L2. -#L1 #L2 #T #d1 #H @(lleq_ind_alt … H) -L1 -L2 -T -d1 -/4 width=1 by lleq_sort, lleq_free, lleq_gref, lleq_bind, lleq_flat, yle_succ/ -[ /3 width=3 by lleq_skip, ylt_yle_trans/ -| #I1 #I2 #L1 #L2 #K1 #K2 #V #d1 #i #Hi #HLK1 #HLK2 #HV #IHV #d2 #Hd12 elim (ylt_split i d2) - [ lapply (lleq_fwd_length … HV) #HK12 #Hid2 - lapply (ldrop_fwd_length … HLK1) lapply (ldrop_fwd_length … HLK2) - normalize in ⊢ (%→%→?); -I1 -I2 -V -d1 /2 width=1 by lleq_skip/ - | /3 width=8 by lleq_lref, yle_trans/ - ] -] -qed-. - -lemma lleq_bind_O: ∀a,I,L1,L2,V,T. L1 ⋕[V, 0] L2 → L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V → - L1 ⋕[ⓑ{a,I}V.T, 0] L2. -/3 width=3 by lleq_ge, lleq_bind/ qed. - -lemma lleq_bind_repl_SO: ∀I1,I2,L1,L2,V1,V2,T. L1.ⓑ{I1}V1 ⋕[T, 0] L2.ⓑ{I2}V2 → - ∀J1,J2,W1,W2. L1.ⓑ{J1}W1 ⋕[T, 1] L2.ⓑ{J2}W2. -#I1 #I2 #L1 #L2 #V1 #V2 #T #HT #J1 #J2 #W1 #W2 lapply (lleq_ge … HT 1 ?) // -HT -#HT @(lleq_lsuby_repl … HT) /2 width=1 by lsuby_succ/ (**) (* full auto fails *) -qed-. - -lemma lleq_bind_repl_O: ∀I,L1,L2,V,T. L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V → - ∀J,W. L1 ⋕[W, 0] L2 → L1.ⓑ{J}W ⋕[T, 0] L2.ⓑ{J}W. -/3 width=7 by lleq_bind_repl_SO, lleq_inv_S/ qed-. - -(* Inversion lemmas on negated lazy quivalence for local environments *******) - -lemma nlleq_inv_bind_O: ∀a,I,L1,L2,V,T. (L1 ⋕[ⓑ{a,I}V.T, 0] L2 → ⊥) → - (L1 ⋕[V, 0] L2 → ⊥) ∨ (L1.ⓑ{I}V ⋕[T, 0] L2.ⓑ{I}V → ⊥). -#a #I #L1 #L2 #V #T #H elim (lleq_dec V L1 L2 0) -/4 width=1 by lleq_bind_O, or_intror, or_introl/ -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_fqus.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_fqus.ma deleted file mode 100644 index 18744a539..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_fqus.ma +++ /dev/null @@ -1,75 +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/fqus_alt.ma". -include "basic_2/substitution/lleq_ext.ma". - -(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) - -(* Properties on supclosure and derivatives *********************************) - -lemma lleq_fqu_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃ ⦃G2, K2, U⦄ → - ∀L1. L1 ⋕[T, 0] L2 → - ∃∃K1. ⦃G1, L1, T⦄ ⊃ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. -#G1 #G2 #L2 #K2 #T #U #H elim H -G1 -G2 -L2 -K2 -T -U -[ #I #G #L2 #V #L1 #H elim (lleq_inv_lref_ge_dx … H … I L2 V) -H // - #I1 #K1 #H1 #H2 lapply (ldrop_inv_O2 … H1) -H1 - #H destruct /2 width=3 by fqu_lref_O, ex2_intro/ -| * [ #a ] #I #G #L2 #V #T #L1 #H - [ elim (lleq_inv_bind … H) - | elim (lleq_inv_flat … H) - ] -H - /2 width=3 by fqu_pair_sn, ex2_intro/ -| #a #I #G #L2 #V #T #L1 #H elim (lleq_inv_bind_O … H) -H - #H3 #H4 /2 width=3 by fqu_bind_dx, ex2_intro/ -| #I #G #L2 #V #T #L1 #H elim (lleq_inv_flat … H) -H - /2 width=3 by fqu_flat_dx, ex2_intro/ -| #G #L2 #K2 #T #U #e #HLK2 #HTU #L1 #HL12 - elim (ldrop_O1_le (e+1) L1) - [ /3 width=12 by fqu_drop, lleq_inv_lift_le, ex2_intro/ - | lapply (ldrop_fwd_length_le2 … HLK2) -K2 - lapply (lleq_fwd_length … HL12) -T -U // - ] -] -qed-. - -lemma lleq_fquq_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃⸮ ⦃G2, K2, U⦄ → - ∀L1. L1 ⋕[T, 0] L2 → - ∃∃K1. ⦃G1, L1, T⦄ ⊃⸮ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. -#G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fquq_inv_gen … H) -H -[ #H elim (lleq_fqu_trans … H … HL12) -L2 /3 width=3 by fqu_fquq, ex2_intro/ -| * #HG #HL #HT destruct /2 width=3 by ex2_intro/ -] -qed-. - -lemma lleq_fqup_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃+ ⦃G2, K2, U⦄ → - ∀L1. L1 ⋕[T, 0] L2 → - ∃∃K1. ⦃G1, L1, T⦄ ⊃+ ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. -#G1 #G2 #L2 #K2 #T #U #H @(fqup_ind … H) -G2 -K2 -U -[ #G2 #K2 #U #HTU #L1 #HL12 elim (lleq_fqu_trans … HTU … HL12) -L2 - /3 width=3 by fqu_fqup, ex2_intro/ -| #G #G2 #K #K2 #U #U2 #_ #HU2 #IHTU #L1 #HL12 elim (IHTU … HL12) -L2 - #K1 #HTU #HK1 elim (lleq_fqu_trans … HU2 … HK1) -K - /3 width=5 by fqup_strap1, ex2_intro/ -] -qed-. - -lemma lleq_fqus_trans: ∀G1,G2,L2,K2,T,U. ⦃G1, L2, T⦄ ⊃* ⦃G2, K2, U⦄ → - ∀L1. L1 ⋕[T, 0] L2 → - ∃∃K1. ⦃G1, L1, T⦄ ⊃* ⦃G2, K1, U⦄ & K1 ⋕[U, 0] K2. -#G1 #G2 #L2 #K2 #T #U #H #L1 #HL12 elim(fqus_inv_gen … H) -H -[ #H elim (lleq_fqup_trans … H … HL12) -L2 /3 width=3 by fqup_fqus, ex2_intro/ -| * #HG #HL #HT destruct /2 width=3 by ex2_intro/ -] -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_ldrop.ma deleted file mode 100644 index 9cb597c6f..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_ldrop.ma +++ /dev/null @@ -1,123 +0,0 @@ -(**************************************************************************) -(* ___ *) -(* ||M|| *) -(* ||A|| A project by Andrea Asperti *) -(* ||T|| *) -(* ||I|| Developers: *) -(* ||T|| The HELM team. *) -(* ||A|| http://helm.cs.unibo.it *) -(* \ / *) -(* \ / This file is distributed under the terms of the *) -(* v GNU General Public License Version 2 *) -(* *) -(**************************************************************************) - -include "basic_2/substitution/cpys_lift.ma". -include "basic_2/substitution/lleq.ma". - -(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) - -(* Advanced properties ******************************************************) - -lemma lleq_skip: ∀L1,L2,d,i. yinj i < d → |L1| = |L2| → L1 ⋕[#i, d] L2. -#L1 #L2 #d #i #Hid #HL12 @conj // -HL12 -#U @conj #H elim (cpys_inv_lref1 … H) -H // * -#I #Z #Y #X #H elim (ylt_yle_false … Hid … H) -qed. - -lemma lleq_lref: ∀I1,I2,L1,L2,K1,K2,V,d,i. d ≤ yinj i → - ⇩[i] L1 ≡ K1.ⓑ{I1}V → ⇩[i] L2 ≡ K2.ⓑ{I2}V → - K1 ⋕[V, 0] K2 → L1 ⋕[#i, d] L2. -#I1 #I2 #L1 #L2 #K1 #K2 #V #d #i #Hdi #HLK1 #HLK2 * #HK12 #IH @conj [ -IH | -HK12 ] -[ lapply (ldrop_fwd_length … HLK1) -HLK1 #H1 - lapply (ldrop_fwd_length … HLK2) -HLK2 #H2 - >H1 >H2 -H1 -H2 normalize // -| #U @conj #H elim (cpys_inv_lref1 … H) -H // * - >yminus_Y_inj #I #K #X #W #_ #_ #H #HVW #HWU - [ letin HLK ≝ HLK1 | letin HLK ≝ HLK2 ] - lapply (ldrop_mono … H … HLK) -H #H destruct elim (IH W) - /3 width=7 by cpys_subst_Y2/ -] -qed. - -lemma lleq_free: ∀L1,L2,d,i. |L1| ≤ i → |L2| ≤ i → |L1| = |L2| → L1 ⋕[#i, d] L2. -#L1 #L2 #d #i #HL1 #HL2 #HL12 @conj // -HL12 -#U @conj #H elim (cpys_inv_lref1 … H) -H // * -#I #Z #Y #X #_ #_ #H lapply (ldrop_fwd_length_lt2 … H) -H -#H elim (lt_refl_false i) /2 width=3 by lt_to_le_to_lt/ -qed. - -(* Properties on relocation *************************************************) - -lemma lleq_lift_le: ∀K1,K2,T,dt. K1 ⋕[T, dt] K2 → - ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → - ∀U. ⇧[d, e] T ≡ U → dt ≤ d → L1 ⋕[U, dt] L2. -#K1 #K2 #T #dt * #HK12 #IHT #L1 #L2 #d #e #HLK1 #HLK2 #U #HTU #Hdtd -lapply (ldrop_fwd_length … HLK1) lapply (ldrop_fwd_length … HLK2) -#H2 #H1 @conj // -HK12 -H1 -H2 #U0 @conj #HU0 -[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 ] -elim (cpys_inv_lift1_be … HU0 … HLKA … HTU) // -HU0 >yminus_Y_inj #T0 #HT0 #HTU0 -elim (IHT T0) [ #H #_ | #_ #H ] -IHT /3 width=12 by cpys_lift_be/ -qed-. - -lemma lleq_lift_ge: ∀K1,K2,T,dt. K1 ⋕[T, dt] K2 → - ∀L1,L2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → - ∀U. ⇧[d, e] T ≡ U → d ≤ dt → L1 ⋕[U, dt+e] L2. -#K1 #K2 #T #dt * #HK12 #IHT #L1 #L2 #d #e #HLK1 #HLK2 #U #HTU #Hddt -lapply (ldrop_fwd_length … HLK1) lapply (ldrop_fwd_length … HLK2) -#H2 #H1 @conj // -HK12 -H1 -H2 #U0 @conj #HU0 -[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 ] -elim (cpys_inv_lift1_ge … HU0 … HLKA … HTU) /2 width=1 by monotonic_yle_plus_dx/ -HU0 >yplus_minus_inj #T0 #HT0 #HTU0 -elim (IHT T0) [ #H #_ | #_ #H ] -IHT /3 width=10 by cpys_lift_ge/ -qed-. - -(* Inversion lemmas on relocation *******************************************) - -lemma lleq_inv_lift_le: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → - ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → - ∀T. ⇧[d, e] T ≡ U → dt ≤ d → K1 ⋕[T, dt] K2. -#L1 #L2 #U #dt * #HL12 #IH #K1 #K2 #d #e #HLK1 #HLK2 #T #HTU #Hdtd -lapply (ldrop_fwd_length_minus2 … HLK1) lapply (ldrop_fwd_length_minus2 … HLK2) -#H2 #H1 @conj // -HL12 -H1 -H2 -#T0 elim (lift_total T0 d e) -#U0 #HTU0 elim (IH U0) -IH -#H12 #H21 @conj #HT0 -[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 letin H0 ≝ H12 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 letin H0 ≝ H21 ] -lapply (cpys_lift_be … HT0 … HLKA … HTU … HTU0) // -HT0 ->yplus_Y1 #HU0 elim (cpys_inv_lift1_be … (H0 HU0) … HLKB … HTU) // -L1 -L2 -U -Hdtd -#X #HT0 #HX lapply (lift_inj … HX … HTU0) -U0 // -qed-. - -lemma lleq_inv_lift_ge: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → - ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → - ∀T. ⇧[d, e] T ≡ U → yinj d + e ≤ dt → K1 ⋕[T, dt-e] K2. -#L1 #L2 #U #dt * #HL12 #IH #K1 #K2 #d #e #HLK1 #HLK2 #T #HTU #Hdedt -lapply (ldrop_fwd_length_minus2 … HLK1) lapply (ldrop_fwd_length_minus2 … HLK2) -#H2 #H1 @conj // -HL12 -H1 -H2 -elim (yle_inv_plus_inj2 … Hdedt) #Hddt #Hedt -#T0 elim (lift_total T0 d e) -#U0 #HTU0 elim (IH U0) -IH -#H12 #H21 @conj #HT0 -[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 letin H0 ≝ H12 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 letin H0 ≝ H21 ] -lapply (cpys_lift_ge … HT0 … HLKA … HTU … HTU0) // -HT0 -Hddt ->ymax_pre_sn // #HU0 elim (cpys_inv_lift1_ge … (H0 HU0) … HLKB … HTU) // -L1 -L2 -U -Hdedt -Hedt -#X #HT0 #HX lapply (lift_inj … HX … HTU0) -U0 // -qed-. - -lemma lleq_inv_lift_be: ∀L1,L2,U,dt. L1 ⋕[U, dt] L2 → - ∀K1,K2,d,e. ⇩[Ⓕ, d, e] L1 ≡ K1 → ⇩[Ⓕ, d, e] L2 ≡ K2 → - ∀T. ⇧[d, e] T ≡ U → d ≤ dt → dt ≤ yinj d + e → K1 ⋕[T, d] K2. -#L1 #L2 #U #dt * #HL12 #IH #K1 #K2 #d #e #HLK1 #HLK2 #T #HTU #Hddt #Hdtde -lapply (ldrop_fwd_length_minus2 … HLK1) lapply (ldrop_fwd_length_minus2 … HLK2) -#H2 #H1 @conj // -HL12 -H1 -H2 -#T0 elim (lift_total T0 d e) -#U0 #HTU0 elim (IH U0) -IH -#H12 #H21 @conj #HT0 -[ letin HLKA ≝ HLK1 letin HLKB ≝ HLK2 letin H0 ≝ H12 | letin HLKA ≝ HLK2 letin HLKB ≝ HLK1 letin H0 ≝ H21 ] -lapply (cpys_lift_ge … HT0 … HLKA … HTU … HTU0) // -HT0 -#HU0 lapply (cpys_weak … HU0 dt (∞) ? ?) // -HU0 -#HU0 lapply (H0 HU0) -#HU0 lapply (cpys_weak … HU0 d (∞) ? ?) // -HU0 -#HU0 elim (cpys_inv_lift1_ge_up … HU0 … HLKB … HTU) // -L1 -L2 -U -Hddt -Hdtde -#X #HT0 #HX lapply (lift_inj … HX … HTU0) -U0 // -qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_lleq.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_lleq.ma deleted file mode 100644 index 2a8873d9b..000000000 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_lleq.ma +++ /dev/null @@ -1,175 +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/cpys_cpys.ma". -include "basic_2/substitution/lleq_ldrop.ma". - -(* Advanced forward lemmas **************************************************) - -lemma lleq_fwd_lref: ∀L1,L2. ∀d:ynat. ∀i:nat. L1 ⋕[#i, d] L2 → - ∨∨ |L1| ≤ i ∧ |L2| ≤ i - | yinj i < d - | ∃∃I1,I2,K1,K2,V. ⇩[i] L1 ≡ K1.ⓑ{I1}V & - ⇩[i] L2 ≡ K2.ⓑ{I2}V & - K1 ⋕[V, yinj 0] K2 & d ≤ yinj i. -#L1 #L2 #d #i * #HL12 #IH elim (lt_or_ge i (|L1|)) /3 width=3 by or3_intro0, conj/ -elim (ylt_split i d) /2 width=1 by or3_intro1/ #Hdi #Hi -elim (ldrop_O1_lt … Hi) #I1 #K1 #V1 #HLK1 -elim (ldrop_O1_lt L2 i) // -Hi #I2 #K2 #V2 #HLK2 -lapply (ldrop_fwd_length_minus2 … HLK2) #H -lapply (ldrop_fwd_length_minus2 … HLK1) >HL12 yminus_Y_inj ] /3 width=7 by cpys_subst_Y2, yle_inj/ -qed-. - -lemma lleq_fwd_lref_dx: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → - ∀I2,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I2}V → - i < d ∨ - ∃∃I1,K1. ⇩[i] L1 ≡ K1.ⓑ{I1}V & K1 ⋕[V, 0] K2 & d ≤ i. -#L1 #L2 #d #i #H #I2 #K2 #V #HLK2 elim (lleq_fwd_lref … H) -H [ * || * ] -[ #_ #H elim (lt_refl_false i) - lapply (ldrop_fwd_length_lt2 … HLK2) -HLK2 - /2 width=3 by lt_to_le_to_lt/ (**) (* full auto too slow *) -| /2 width=1 by or_introl/ -| #I1 #I2 #K11 #K22 #V0 #HLK11 #HLK22 #HV0 #Hdi lapply (ldrop_mono … HLK22 … HLK2) -L2 - #H destruct /3 width=5 by ex3_2_intro, or_intror/ -] -qed-. - -lemma lleq_fwd_lref_sn: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → - ∀I1,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I1}V → - i < d ∨ - ∃∃I2,K2. ⇩[i] L2 ≡ K2.ⓑ{I2}V & K1 ⋕[V, 0] K2 & d ≤ i. -#L1 #L2 #d #i #HL12 #I1 #K1 #V #HLK1 elim (lleq_fwd_lref_dx L2 … d … HLK1) -HLK1 -[2: * ] /4 width=6 by lleq_sym, ex3_2_intro, or_introl, or_intror/ -qed-. - -(* Advanced inversion lemmas ************************************************) - -lemma lleq_inv_lref_ge_dx: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → - ∀I2,K2,V. ⇩[i] L2 ≡ K2.ⓑ{I2}V → - ∃∃I1,K1. ⇩[i] L1 ≡ K1.ⓑ{I1}V & K1 ⋕[V, 0] K2. -#L1 #L2 #d #i #H #Hdi #I2 #K2 #V #HLK2 elim (lleq_fwd_lref_dx … H … HLK2) -L2 -[ #H elim (ylt_yle_false … H Hdi) -| * /2 width=4 by ex2_2_intro/ -] -qed-. - -lemma lleq_inv_lref_ge_sn: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → - ∀I1,K1,V. ⇩[i] L1 ≡ K1.ⓑ{I1}V → - ∃∃I2,K2. ⇩[i] L2 ≡ K2.ⓑ{I2}V & K1 ⋕[V, 0] K2. -#L1 #L2 #d #i #HL12 #Hdi #I1 #K1 #V #HLK1 elim (lleq_inv_lref_ge_dx L2 … Hdi … HLK1) -Hdi -HLK1 -/3 width=4 by lleq_sym, ex2_2_intro/ -qed-. - -lemma lleq_inv_lref_ge_gen: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → - ∀I1,I2,K1,K2,V1,V2. - ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → - V1 = V2 ∧ K1 ⋕[V2, 0] K2. -#L1 #L2 #d #i #HL12 #Hdi #I1 #I2 #K1 #K2 #V1 #V2 #HLK1 #HLK2 -elim (lleq_inv_lref_ge_sn … HL12 … HLK1) // -L1 -d -#J #Y #HY lapply (ldrop_mono … HY … HLK2) -L2 -i #H destruct /2 width=1 by conj/ -qed-. - -lemma lleq_inv_lref_ge: ∀L1,L2,d,i. L1 ⋕[#i, d] L2 → d ≤ i → - ∀I,K1,K2,V. ⇩[i] L1 ≡ K1.ⓑ{I}V → ⇩[i] L2 ≡ K2.ⓑ{I}V → - K1 ⋕[V, 0] K2. -#L1 #L2 #d #i #HL12 #Hdi #I #K1 #K2 #V #HLK1 #HLK2 -elim (lleq_inv_lref_ge_gen … HL12 … HLK1 HLK2) // -qed-. - -(* Advanced properties ******************************************************) - -lemma lleq_dec: ∀T,L1,L2,d. Decidable (L1 ⋕[T, d] L2). -#T #L1 @(f2_ind … rfw … L1 T) -L1 -T -#n #IH #L1 * * -[ #k #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, lleq_sort/ -| #i #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) - [ #HL12 #d elim (ylt_split i d) /3 width=1 by lleq_skip, or_introl/ - #Hdi elim (lt_or_ge i (|L1|)) #HiL1 - elim (lt_or_ge i (|L2|)) #HiL2 /3 width=1 by or_introl, lleq_free/ - elim (ldrop_O1_lt … HiL2) #I2 #K2 #V2 #HLK2 - elim (ldrop_O1_lt … HiL1) #I1 #K1 #V1 #HLK1 - elim (eq_term_dec V2 V1) - [ #H3 elim (IH K1 V1 … K2 0) destruct - /3 width=8 by lleq_lref, ldrop_fwd_rfw, or_introl/ - ] - -IH #H3 @or_intror - #H elim (lleq_fwd_lref … H) -H [1,3,4,6: * ] - [1,3: /3 width=4 by lt_to_le_to_lt, lt_refl_false/ - |5,6: /2 width=4 by ylt_yle_false/ - ] - #Z1 #Z2 #Y1 #Y2 #X #HLY1 #HLY2 #HX #_ - lapply (ldrop_mono … HLY1 … HLK1) -HLY1 -HLK1 - lapply (ldrop_mono … HLY2 … HLK2) -HLY2 -HLK2 - #H2 #H1 destruct /2 width=1 by/ - ] -| #p #Hn #L2 elim (eq_nat_dec (|L1|) (|L2|)) /3 width=1 by or_introl, lleq_gref/ -| #a #I #V #T #Hn #L2 #d destruct - elim (IH L1 V … L2 d) /2 width=1 by/ - elim (IH (L1.ⓑ{I}V) T … (L2.ⓑ{I}V) (d+1)) -IH /3 width=1 by or_introl, lleq_bind/ - #H1 #H2 @or_intror - #H elim (lleq_inv_bind … H) -H /2 width=1 by/ -| #I #V #T #Hn #L2 #d destruct - elim (IH L1 V … L2 d) /2 width=1 by/ - elim (IH L1 T … L2 d) -IH /3 width=1 by or_introl, lleq_flat/ - #H1 #H2 @or_intror - #H elim (lleq_inv_flat … H) -H /2 width=1 by/ -] --n /4 width=3 by lleq_fwd_length, or_intror/ -qed-. - -(* Main properties **********************************************************) - -theorem lleq_trans: ∀d,T. Transitive … (lleq d T). -#d #T #L1 #L * #HL1 #IH1 #L2 * #HL2 #IH2 /3 width=3 by conj, iff_trans/ -qed-. - -theorem lleq_canc_sn: ∀L,L1,L2,T,d. L ⋕[d, T] L1→ L ⋕[d, T] L2 → L1 ⋕[d, T] L2. -/3 width=3 by lleq_trans, lleq_sym/ qed-. - -theorem lleq_canc_dx: ∀L1,L2,L,T,d. L1 ⋕[d, T] L → L2 ⋕[d, T] L → L1 ⋕[d, T] L2. -/3 width=3 by lleq_trans, lleq_sym/ qed-. - -(* Inversion lemmas on negated lazy quivalence for local environments *******) - -lemma nlleq_inv_bind: ∀a,I,L1,L2,V,T,d. (L1 ⋕[ⓑ{a,I}V.T, d] L2 → ⊥) → - (L1 ⋕[V, d] L2 → ⊥) ∨ (L1.ⓑ{I}V ⋕[T, ⫯d] L2.ⓑ{I}V → ⊥). -#a #I #L1 #L2 #V #T #d #H elim (lleq_dec V L1 L2 d) -/4 width=1 by lleq_bind, or_intror, or_introl/ -qed-. - -lemma nlleq_inv_flat: ∀I,L1,L2,V,T,d. (L1 ⋕[ⓕ{I}V.T, d] L2 → ⊥) → - (L1 ⋕[V, d] L2 → ⊥) ∨ (L1 ⋕[T, d] L2 → ⊥). -#I #L1 #L2 #V #T #d #H elim (lleq_dec V L1 L2 d) -/4 width=1 by lleq_flat, or_intror, or_introl/ -qed-. - -(* Note: lleq_nlleq_trans: ∀d,T,L1,L. L1⋕[T, d] L → - ∀L2. (L ⋕[T, d] L2 → ⊥) → (L1 ⋕[T, d] L2 → ⊥). -/3 width=3 by lleq_canc_sn/ qed-. -works with /4 width=8/ so lleq_canc_sn is more convenient -*) 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 4bb4189a0..a80f8d3b5 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 @@ -107,7 +107,7 @@ table { ] [ { "context-sensitive extended computation" * } { [ "lpxs ( ⦃?,?⦄ ⊢ ➡*[?,?] ? )" "lpxs_alt ( ⦃?,?⦄ ⊢ ➡➡*[?,?] ? )" "lpxs_ldrop" + "lpxs_lleq" + "lpxs_aaa" + "lpxs_cpxs" + "lpxs_lpxs" * ] - [ "cpxs ( ⦃?,?⦄ ⊢ ? ➡*[?,?] ? )" "cpxs_tstc" + "cpxs_tstc_vector" + "cpxs_lift" + "cpxs_cpys" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ] + [ "cpxs ( ⦃?,?⦄ ⊢ ? ➡*[?,?] ? )" "cpxs_tstc" + "cpxs_tstc_vector" + "cpxs_leq" + "cpxs_lift" + "cpxs_lleq" + "cpxs_aaa" + "cpxs_cpxs" * ] } ] [ { "context-sensitive computation" * } { @@ -136,8 +136,8 @@ table { } ] [ { "context-sensitive extended reduction" * } { - [ "lpx ( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lpx_ldrop" + "lpx_cpys" + "lpx_lleq" + "lpx_aaa" * ] - [ "cpx ( ⦃?,?⦄ ⊢ ? ➡[?,?] ? )" "cpx_lift" + "cpx_cpys" + "cpx_lleq" + "cpx_cix" * ] + [ "lpx ( ⦃?,?⦄ ⊢ ➡[?,?] ? )" "lpx_ldrop" + "lpx_lleq" + "lpx_aaa" * ] + [ "cpx ( ⦃?,?⦄ ⊢ ? ➡[?,?] ? )" "cpx_leq" + "cpx_lift" + "cpx_lleq" + "cpx_cix" * ] } ] [ { "irreducible forms for context-sensitive extended reduction" * } { @@ -213,17 +213,9 @@ table { ] class "yellow" [ { "substitution" * } { - [ { "lazy equivalence for local environments" * } { - [ "lleq ( ? ⋕[?,?] ? )" "lleq_alt ( ? ⋕⋕[?,?] ? )" "lleq_ldrop" + "lleq_fqus" + "lleq_lleq" + "lleq_ext" * ] - } - ] - [ { "contxt-sensitive extended multiple substitution" * } { - [ "cpys ( ⦃?,?⦄ ⊢ ? ▶*[?,?] ? )" "cpys_alt ( ⦃?,?⦄ ⊢ ? ▶▶*[?,?] ? )" "cpys_lift" + "cpys_cpys" * ] - } - ] [ { "iterated structural successor for closures" * } { - [ "fqus ( ⦃?,?,?⦄ ⊃* ⦃?,?,?⦄ )" "fqus_alt" + "fqus_fqus" * ] - [ "fqup ( ⦃?,?,?⦄ ⊃+ ⦃?,?,?⦄ )" "fqup_fqup" * ] + [ "fqus ( ⦃?,?,?⦄ ⊃* ⦃?,?,?⦄ )" "fqus_alt" + "fqus_lleq" + "fqus_fqus" * ] + [ "fqup ( ⦃?,?,?⦄ ⊃+ ⦃?,?,?⦄ )" "fqup_lleq" + "fqup_fqup" * ] } ] [ { "iterated local env. slicing" * } { @@ -243,17 +235,9 @@ table { ] class "orange" [ { "relocation" * } { - [ { "contxt-sensitive extended ordinary substitution" * } { - [ "cpy ( ⦃?,?⦄ ⊢ ? ▶[?,?] ? )" "cpy_lift" + "cpy_cpy" * ] - } - ] - [ { "local env. ref. for extended substitution" * } { - [ "lsuby ( ? ⊑×[?,?] ? )" "lsuby_lsuby" * ] - } - ] [ { "structural successor for closures" * } { - [ "fquq ( ⦃?,?,?⦄ ⊃⸮ ⦃?,?,?⦄ )" "fquq_alt ( ⦃?,?,?⦄ ⊃⊃⸮ ⦃?,?,?⦄ )" * ] - [ "fqu ( ⦃?,?,?⦄ ⊃ ⦃?,?,?⦄ )" * ] + [ "fquq ( ⦃?,?,?⦄ ⊃⸮ ⦃?,?,?⦄ )" "fquq_alt ( ⦃?,?,?⦄ ⊃⊃⸮ ⦃?,?,?⦄ )" "fquq_lleq" * ] + [ "fqu ( ⦃?,?,?⦄ ⊃ ⦃?,?,?⦄ )" "fqu_lleq" * ] } ] [ { "global env. slicing" * } { @@ -261,11 +245,11 @@ table { } ] [ { "lazy equivalence for local environments" * } { - [ "lleq ( ? ⋕[?,?] ? )" "lleq_leq" * ] + [ "lleq ( ? ⋕[?,?] ? )" "lleq_leq" + "lleq_ldrop" + "lleq_lleq" * ] } ] [ { "lazy pointwise extension of a relation" * } { - [ "llpx_sn" "llpx_sn_leq" * ] + [ "llpx_sn" "llpx_sn_leq" + "llpx_sn_ldrop" * ] } ] [ { "basic local env. slicing" * } { diff --git a/matita/matita/predefined_virtuals.ml b/matita/matita/predefined_virtuals.ml index 14fff0a3c..6311d5933 100644 --- a/matita/matita/predefined_virtuals.ml +++ b/matita/matita/predefined_virtuals.ml @@ -1531,6 +1531,7 @@ let predefined_classes = [ ["◊"; "♢"; "⧫"; "♦"; "⟐"; "⟠"; ] ; [">"; "⭃"; "⧁"; "〉"; "»"; "❭"; "❯"; "❱"; "▸"; "►"; "▶"; ] ; ["≥"; "⪀"; "≽"; "⪴"; "⥸"; "⊒"; ]; + ["∨"; "⩖"; "⋓"; ] ; ["a"; "α"; "𝕒"; "𝐚"; "𝛂"; "ⓐ"; ] ; ["A"; "ℵ"; "𝔸"; "𝐀"; "Ⓐ"; ] ; ["b"; "β"; "ß"; "𝕓"; "𝐛"; "𝛃"; "ⓑ"; ] ;