]> matita.cs.unibo.it Git - helm.git/commitdiff
bug fix in lazy union for local environments
authorFerruccio Guidi <ferruccio.guidi@unibo.it>
Sun, 8 Jun 2014 18:01:59 +0000 (18:01 +0000)
committerFerruccio Guidi <ferruccio.guidi@unibo.it>
Sun, 8 Jun 2014 18:01:59 +0000 (18:01 +0000)
matita/matita/contribs/lambdadelta/basic_2/multiple/llor.ma
matita/matita/contribs/lambdadelta/basic_2/multiple/llor_ldrop.ma

index 4f3d948decd40b463a7cfdeb2ce5fcd95055ab10..fa3e511af2b0b0f7e5ce87cd43a5220a0af597c6 100644 (file)
@@ -23,7 +23,7 @@ definition llor: ynat → relation4 term lenv lenv lenv ≝ λd,T,L2,L1,L.
                        ⇩[i] L1 ≡ K1.ⓑ{I1}V1 → ⇩[i] L2 ≡ K2.ⓑ{I2}V2 → ⇩[i] L ≡ K.ⓑ{I}V → ∨∨
                        (∧∧ yinj i < d & I1 = I & V1 = V) |
                        (∧∧ (L1 ⊢ i ϵ 𝐅*[d]⦃T⦄ → ⊥) & I1 = I & V1 = V) |
-                       (∧∧ d ≤ yinj i & L1 ⊢ i ϵ 𝐅*[d]⦃T⦄ & I1 = I & V2 = V)
+                       (∧∧ d ≤ yinj i & L1 ⊢ i ϵ 𝐅*[d]⦃T⦄ & I2 = I & V2 = V)
                     ).
 
 interpretation
@@ -32,6 +32,7 @@ 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 //
 #I1 #I2 #I #K1 #K2 #K #V1 #V2 #V #i #HLK1
index f247442917f0a76e15a187fd94cbd78d7881926e..f8458d3c63a4a7e39e3361e766032e25e21a6db3 100644 (file)
@@ -21,7 +21,7 @@ include "basic_2/multiple/llor.ma".
 
 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 -I2 -L2 -K2
+#I1 #I2 #I #K1 #K2 #K #W1 #W2 #W #i #HLK1 #_ #HLK -L2 -K2
 lapply (ldrop_mono … HLK … HLK1) -HLK #H destruct
 lapply (ldrop_fwd_length_lt2 … HLK1) -K1
 /5 width=3 by ylt_yle_trans, ylt_inj, or3_intro0, and3_intro/