From 0f34e7f3ada66fdc9044b5ed5c5f59ea36d3c6a2 Mon Sep 17 00:00:00 2001 From: Ferruccio Guidi Date: Tue, 8 Apr 2014 20:56:32 +0000 Subject: [PATCH] some advances on pointwise union for local environments ... --- .../lambdadelta/basic_2/relocation/llor.ma | 47 ++++++++++++++++++- .../basic_2/relocation/lpx_sn_alt.ma | 2 +- .../lambdadelta/basic_2/web/basic_2_src.tbl | 13 ++--- 3 files changed, 54 insertions(+), 8 deletions(-) diff --git a/matita/matita/contribs/lambdadelta/basic_2/relocation/llor.ma b/matita/matita/contribs/lambdadelta/basic_2/relocation/llor.ma index 8c76ecc02..7dec554c7 100644 --- a/matita/matita/contribs/lambdadelta/basic_2/relocation/llor.ma +++ b/matita/matita/contribs/lambdadelta/basic_2/relocation/llor.ma @@ -15,7 +15,7 @@ include "basic_2/notation/relations/lazyor_4.ma". include "basic_2/relocation/lpx_sn_alt.ma". -(* LAZY UNION FOR LOCAL ENVIRONMENTS ****************************************) +(* POINTWISE UNION FOR LOCAL ENVIRONMENTS ***********************************) inductive clor (T) (L2) (K1) (V1): predicate term ≝ | clor_sn: ∀U. |K1| < |L2| → ⇧[|L2|-|K1|-1, 1] U ≡ T → clor T L2 K1 V1 V1 @@ -53,3 +53,48 @@ lemma llor_total: ∀T,L2,L1. |L1| ≤ |L2| → ∃L. L1 ⩖[T] L2 ≡ L. /5 width=4 by llor_pair_dx, monotonic_lt_minus_l, ex_intro/ ] qed-. + +(* Alternative definition ***************************************************) + +lemma llor_intro_alt_eq: ∀T,L2,L1,L. |L1| = |L2| → |L1| = |L| → + (∀I1,I,K1,K,V1,V,U,i. ⇧[i, 1] U ≡ T → + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L ≡ K.ⓑ{I}V → + ∧∧ I1 = I & V1 = V & K1 ⩖[T] L2 ≡ K + ) → + (∀I1,I2,I,K1,K2,K,V1,V2,V,i. (∀U. ⇧[i, 1] U ≡ T → ⊥) → + ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → + ⇩[i] L ≡ K.ⓑ{I}V → ∧∧ I1 = I & V2 = V & K1 ⩖[T] L2 ≡ K + ) → L1 ⩖[T] L2 ≡ L. +#T #L2 #L1 #L #HL12 #HL1 #IH1 #IH2 @lpx_sn_intro_alt // -HL1 +#I1 #I #K1 #K #V1 #V #i #HLK1 #HLK +lapply (ldrop_fwd_length_minus4 … HLK1) +lapply (ldrop_fwd_length_le4 … HLK1) +normalize >HL12 HL12