From c88fe30a6900da730a684d7df6522407a388cd67 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Sun, 8 Jun 2014 18:34:23 +0000 Subject: [PATCH] now we use the version of lazy union for local environments in which the arguments have the same length. it should have better invariants --- .../contribs/lambdadelta/basic_2/multiple/llor.ma | 6 +++--- .../lambdadelta/basic_2/multiple/llor_etc.ma | 12 +++++------- .../lambdadelta/basic_2/multiple/llor_ldrop.ma | 4 ++-- 3 files changed, 10 insertions(+), 12 deletions(-) diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma index fa3e511af..7832b4898 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma @@ -18,7 +18,7 @@ include "basic_2/multiple/frees.ma". (* POINTWISE UNION FOR LOCAL ENVIRONMENTS ***********************************) definition llor: ynat → relation4 term lenv lenv lenv ≝ λd,T,L2,L1,L. - ∧∧ |L1| ≤ |L2| & |L1| = |L| + ∧∧ |L1| = |L2| & |L1| = |L| & (∀I1,I2,I,K1,K2,K,V1,V2,V,i. ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → ⇩[i] L ≡ K.ⓑ{I}V → ∨∨ (∧∧ yinj i < d & I1 = I & V1 = V) | @@ -33,8 +33,8 @@ interpretation (* Basic properties *********************************************************) (* Note: this can be proved by llor_skip *) -lemma llor_atom: ∀L2,T,d. ⋆ ⩖[T, d] L2 ≡ ⋆. -#L2 #T #d @and3_intro // +lemma llor_atom: ∀T,d. ⋆ ⩖[T, d] ⋆ ≡ ⋆. +#T #d @and3_intro // #I1 #I2 #I #K1 #K2 #K #V1 #V2 #V #i #HLK1 elim (ldrop_inv_atom1 … HLK1) -HLK1 #H destruct qed. diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_etc.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_etc.ma index 4ed8f333d..a1a80baa3 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_etc.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_etc.ma @@ -20,10 +20,10 @@ include "basic_2/multiple/llor_ldrop.ma". lemma lt_plus_SO_to_le: ∀x,y. x < y + 1 → x ≤ y. /2 width=1 by monotonic_pred/ qed-. -(* + lemma llor_tail_frees: ∀L1,L2,L,U,d. L1 ⩖[U, d] L2 ≡ L → d < yinj (|L1|) → ∀I1,W1. ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[d]⦃U⦄ → - ∀I2,W2. ⓑ{I1}W1.L1 ⩖[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I1}W2.L. + ∀I2,W2. ⓑ{I1}W1.L1 ⩖[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I2}W2.L. #L1 #L2 #L #U #d * #HL12 #HL1 #IH #Hd #I1 #W1 #HU #I2 #W2 @and3_intro [1,2: >ltail_length /2 width=1 by le_S_S/ ] #J1 #J2 #J #K1 #K2 #K #V1 #V2 #V #i #HLK1 #HLK2 #HLK @@ -41,10 +41,8 @@ elim (le_to_or_lt_eq … H) -H #H | #HnU #HZ #HX | #Hdi #H2U #HZ #HX ] -| -IH destruct - lapply (ldrop_O1_inv_append1_le … HLK1 … (⋆) ?) // -HLK1 normalize #H destruct - lapply (ldrop_O1_inv_append1_le … HLK2 … HL12) +| -IH -HLK1 destruct + lapply (ldrop_O1_inv_append1_le … HLK2 … (⋆) ?) // -HLK2 normalize #H destruct lapply (ldrop_O1_inv_append1_le … HLK … (⋆) ?) // -HLK normalize #H destruct - @or3_intro2 @and4_intro /2 width=1 by ylt_fwd_le/ + /4 width=1 by ylt_fwd_le, or3_intro2, and4_intro/ ] -*) diff --git a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_ldrop.ma b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_ldrop.ma index f8458d3c6..afc0924d9 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_ldrop.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/multiple/llor_ldrop.ma @@ -19,7 +19,7 @@ include "basic_2/multiple/llor.ma". (* Advanced properties ******************************************************) -lemma llor_skip: ∀L1,L2,U,d. |L1| ≤ |L2| → yinj (|L1|) ≤ d → L1 ⩖[U, d] L2 ≡ L1. +lemma llor_skip: ∀L1,L2,U,d. |L1| = |L2| → yinj (|L1|) ≤ d → L1 ⩖[U, d] L2 ≡ L1. #L1 #L2 #U #d #HL12 #Hd @and3_intro // -HL12 #I1 #I2 #I #K1 #K2 #K #W1 #W2 #W #i #HLK1 #_ #HLK -L2 -K2 lapply (ldrop_mono … HLK … HLK1) -HLK #H destruct @@ -27,4 +27,4 @@ lapply (ldrop_fwd_length_lt2 … HLK1) -K1 /5 width=3 by ylt_yle_trans, ylt_inj, or3_intro0, and3_intro/ qed. -axiom llor_total: ∀L1,L2,T,d. |L1| ≤ |L2| → ∃L. L1 ⩖[T, d] L2 ≡ L. +axiom llor_total: ∀L1,L2,T,d. |L1| = |L2| → ∃L. L1 ⩖[T, d] L2 ≡ L. -- 2.39.2