From 87fbbf33fcc2ed91cc8b8a08e1c378ef49ac723d Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Sat, 24 May 2014 17:17:47 +0000 Subject: [PATCH] non-recursive alternative of llpx_sn completed ! --- .../basic_2/substitution/cofrees.ma | 4 - .../basic_2/substitution/cofrees_alt.ma | 6 +- .../basic_2/substitution/lleq_alt.ma | 33 +-- .../basic_2/substitution/lleq_alt_rec.ma | 54 ++++ .../{llpx_sn_alt2.ma => llpx_sn_alt.ma} | 39 ++- .../basic_2/substitution/llpx_sn_alt1.ma | 250 ------------------ .../basic_2/substitution/llpx_sn_alt_rec.ma | 250 ++++++++++++++++++ .../lambdadelta/basic_2/web/basic_2_src.tbl | 4 +- 8 files changed, 349 insertions(+), 291 deletions(-) create mode 100644 matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt_rec.ma rename matita/matita/contribs/lambdadelta/basic_2/substitution/{llpx_sn_alt2.ma => llpx_sn_alt.ma} (55%) delete mode 100644 matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_alt1.ma create mode 100644 matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_alt_rec.ma diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees.ma index 5679994e2..59b8baab8 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees.ma @@ -30,10 +30,6 @@ interpretation lemma cofrees_fwd_lift: ∀L,U,d,i. L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ∃T. ⇧[i, 1] T ≡ U. /2 width=1 by/ qed-. -lemma cofrees_fwd_nlift: ∀L,U,d,i. (∀T. ⇧[i, 1] T ≡ U → ⊥) → (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥). -#L #U #d #i #HnTU #H elim (cofrees_fwd_lift … H) -H /2 width=2 by/ -qed-. - lemma cofrees_fwd_bind_sn: ∀a,I,L,W,U,i,d. L ⊢ i ~ϵ 𝐅*[d]⦃ⓑ{a,I}W.U⦄ → L ⊢ i ~ϵ 𝐅*[d]⦃W⦄. #a #I #L #W1 #U #i #d #H #W2 #HW12 elim (H (ⓑ{a,I}W2.U)) /2 width=1 by cpys_bind/ -W1 diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees_alt.ma index ebe44cf28..7bb9c9ca7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/cofrees_alt.ma @@ -19,6 +19,10 @@ include "basic_2/substitution/cofrees_lift.ma". (* Alternative definition of frees_ge ***************************************) +lemma nlift_frees: ∀L,U,d,i. (∀T. ⇧[i, 1] T ≡ U → ⊥) → (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥). +#L #U #d #i #HnTU #H elim (cofrees_fwd_lift … H) -H /2 width=2 by/ +qed-. + lemma frees_inv_ge: ∀L,U,d,i. d ≤ yinj i → (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ → ⊥) → (∀T. ⇧[i, 1] T ≡ U → ⊥) ∨ ∃∃I,K,W,j. d ≤ yinj j & j < i & ⇩[j]L ≡ K.ⓑ{I}W & @@ -26,7 +30,7 @@ lemma frees_inv_ge: ∀L,U,d,i. d ≤ yinj i → (L ⊢ i ~ϵ 𝐅*[d]⦃U⦄ #L #U #d #i #Hdi #H @(frees_ind … H) -U /3 width=2 by or_introl/ #U1 #U2 #HU12 #HU2 * [ #HnU2 elim (cpy_fwd_nlift2_ge … HU12 … HnU2) -HU12 -HnU2 /3 width=2 by or_introl/ - * /5 width=9 by cofrees_fwd_nlift, ex5_4_intro, or_intror/ + * /5 width=9 by nlift_frees, ex5_4_intro, or_intror/ | * #I2 #K2 #W2 #j2 #Hdj2 #Hj2i #HLK2 #HnW2 #HnU2 elim (cpy_fwd_nlift2_ge … HU12 … HnU2) -HU12 -HnU2 /4 width=9 by ex5_4_intro, or_intror/ * #I1 #K1 #W1 #j1 #Hdj1 #Hj12 #HLK1 #HnW1 #HnU1 lapply (ldrop_conf_ge … HLK1 … HLK2 ?) -HLK2 /2 width=1 by lt_to_le/ diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt.ma index a209206af..aa9bc1b27 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt.ma @@ -12,43 +12,30 @@ (* *) (**************************************************************************) -include "basic_2/substitution/llpx_sn_alt1.ma". +include "basic_2/substitution/llpx_sn_alt.ma". include "basic_2/substitution/lleq.ma". (* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) -(* Alternative definition ***************************************************) +(* Alternative definition (not recursive) ***********************************) theorem lleq_intro_alt: ∀L1,L2,T,d. |L1| = |L2| → - (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⇧[i, 1] U ≡ T → ⊥) → + (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (L1 ⊢ i ~ϵ 𝐅*[d]⦃T⦄ → ⊥) → ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → - ∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2 + I1 = I2 ∧ V1 = V2 ) → L1 ≡[T, d] L2. -#L1 #L2 #T #d #HL12 #IH @llpx_sn_intro_alt1 // -HL12 +#L1 #L2 #T #d #HL12 #IH @llpx_sn_alt_inv_llpx_sn @conj // -HL12 #I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2 -elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /2 width=1 by and3_intro/ +@(IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 // qed. -theorem lleq_ind_alt: ∀S:relation4 ynat term lenv lenv. - (∀L1,L2,T,d. |L1| = |L2| → ( - ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⇧[i, 1] U ≡ T → ⊥) → - ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → - ∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2 & S 0 V1 K1 K2 - ) → S d T L1 L2) → - ∀L1,L2,T,d. L1 ≡[T, d] L2 → S d T L1 L2. -#S #IH1 #L1 #L2 #T #d #H @(llpx_sn_ind_alt1 … H) -L1 -L2 -T -d -#L1 #L2 #T #d #HL12 #IH2 @IH1 -IH1 // -HL12 -#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2 -elim (IH2 … HnT HLK1 HLK2) -IH2 -HnT -HLK1 -HLK2 /2 width=1 by and4_intro/ -qed-. - theorem lleq_inv_alt: ∀L1,L2,T,d. L1 ≡[T, d] L2 → |L1| = |L2| ∧ - ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⇧[i, 1] U ≡ T → ⊥) → + ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (L1 ⊢ i ~ϵ 𝐅*[d]⦃T⦄ → ⊥) → ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → - ∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2. -#L1 #L2 #T #d #H elim (llpx_sn_inv_alt1 … H) -H + I1 = I2 ∧ V1 = V2. +#L1 #L2 #T #d #H elim (llpx_sn_llpx_sn_alt … H) -H #HL12 #IH @conj // #I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2 -elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /2 width=1 by and3_intro/ +@(IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 // qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt_rec.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt_rec.ma new file mode 100644 index 000000000..258e6e9a9 --- /dev/null +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/lleq_alt_rec.ma @@ -0,0 +1,54 @@ +(**************************************************************************) +(* ___ *) +(* ||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/llpx_sn_alt_rec.ma". +include "basic_2/substitution/lleq.ma". + +(* LAZY EQUIVALENCE FOR LOCAL ENVIRONMENTS **********************************) + +(* Alternative definition (recursive) ***************************************) + +theorem lleq_intro_alt_r: ∀L1,L2,T,d. |L1| = |L2| → + (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⇧[i, 1] U ≡ T → ⊥) → + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → + ∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2 + ) → L1 ≡[T, d] L2. +#L1 #L2 #T #d #HL12 #IH @llpx_sn_intro_alt_r // -HL12 +#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2 +elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /2 width=1 by and3_intro/ +qed. + +theorem lleq_ind_alt_r: ∀S:relation4 ynat term lenv lenv. + (∀L1,L2,T,d. |L1| = |L2| → ( + ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⇧[i, 1] U ≡ T → ⊥) → + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → + ∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2 & S 0 V1 K1 K2 + ) → S d T L1 L2) → + ∀L1,L2,T,d. L1 ≡[T, d] L2 → S d T L1 L2. +#S #IH1 #L1 #L2 #T #d #H @(llpx_sn_ind_alt_r … H) -L1 -L2 -T -d +#L1 #L2 #T #d #HL12 #IH2 @IH1 -IH1 // -HL12 +#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2 +elim (IH2 … HnT HLK1 HLK2) -IH2 -HnT -HLK1 -HLK2 /2 width=1 by and4_intro/ +qed-. + +theorem lleq_inv_alt_r: ∀L1,L2,T,d. L1 ≡[T, d] L2 → + |L1| = |L2| ∧ + ∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (∀U. ⇧[i, 1] U ≡ T → ⊥) → + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → + ∧∧ I1 = I2 & V1 = V2 & K1 ≡[V1, 0] K2. +#L1 #L2 #T #d #H elim (llpx_sn_inv_alt_r … H) -H +#HL12 #IH @conj // +#I1 #I2 #K1 #K2 #V1 #V2 #i #Hid #HnT #HLK1 #HLK2 +elim (IH … HnT HLK1 HLK2) -IH -HnT -HLK1 -HLK2 /2 width=1 by and3_intro/ +qed-. diff --git a/matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_alt2.ma b/matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_alt.ma similarity index 55% rename from matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_alt2.ma rename to matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_alt.ma index 8fbc6f8eb..0916edb67 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_alt2.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/substitution/llpx_sn_alt.ma @@ -13,23 +13,23 @@ (**************************************************************************) include "basic_2/substitution/cofrees_alt.ma". -include "basic_2/substitution/llpx_sn_alt1.ma". +include "basic_2/substitution/llpx_sn_alt_rec.ma". (* LAZY SN POINTWISE EXTENSION OF A CONTEXT-SENSITIVE REALTION FOR TERMS ****) (* alternative definition of llpx_sn (not recursive) *) -definition llpx_sn_alt2: relation4 bind2 lenv term term → relation4 ynat term lenv lenv ≝ - λR,d,T,L1,L2. |L1| = |L2| ∧ - (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (L1 ⊢ i ~ϵ 𝐅*[d]⦃T⦄ → ⊥) → - ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → - I1 = I2 ∧ R I1 K1 V1 V2 - ). +definition llpx_sn_alt: relation4 bind2 lenv term term → relation4 ynat term lenv lenv ≝ + λR,d,T,L1,L2. |L1| = |L2| ∧ + (∀I1,I2,K1,K2,V1,V2,i. d ≤ yinj i → (L1 ⊢ i ~ϵ 𝐅*[d]⦃T⦄ → ⊥) → + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → + I1 = I2 ∧ R I1 K1 V1 V2 + ). (* Main properties **********************************************************) -theorem llpx_sn_llpx_sn_alt2: ∀R,T,L1,L2,d. llpx_sn R d T L1 L2 → llpx_sn_alt2 R d T L1 L2. +theorem llpx_sn_llpx_sn_alt: ∀R,T,L1,L2,d. llpx_sn R d T L1 L2 → llpx_sn_alt R d T L1 L2. #R #U #L1 @(f2_ind … rfw … L1 U) -L1 -U -#n #IHn #L1 #U #Hn #L2 #d #H elim (llpx_sn_inv_alt1 … H) -H +#n #IHn #L1 #U #Hn #L2 #d #H elim (llpx_sn_inv_alt_r … H) -H #HL12 #IHU @conj // #I1 #I2 #K1 #K2 #V1 #V2 #i #Hdi #H #HLK1 #HLK2 elim (frees_inv_ge … H) -H // [ -n #HnU elim (IHU … HnU HLK1 HLK2) -IHU -HnU -HLK1 -HLK2 /2 width=1 by conj/ @@ -39,7 +39,24 @@ theorem llpx_sn_llpx_sn_alt2: ∀R,T,L1,L2,d. llpx_sn R d T L1 L2 → llpx_sn_al elim (ldrop_O1_lt (Ⓕ) L2 j) [2: (minus_plus_m_m j (i+1)) in ⊢ (%→?); >commutative_plus