]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/multiple/llor_alt.ma
- fpbg can be reflexive (example given)
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / multiple / llor_alt.ma
index 430df9936fb6d58c8ceb66ab40a2706ac381fd66..721d18c94e3ed1a2842e14495ded5e09d3880363 100644 (file)
@@ -19,56 +19,56 @@ include "basic_2/multiple/llor.ma".
 
 (* Alternative definition ***************************************************)
 
-lemma llor_tail_frees: â\88\80L1,L2,L,U,d. L1 â©\96[U, d] L2 ≡ L → d ≤ yinj (|L1|) →
+lemma llor_tail_frees: â\88\80L1,L2,L,U,d. L1 â\8b\93[U, d] L2 ≡ L → d ≤ yinj (|L1|) →
                        ∀I1,W1. ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[d]⦃U⦄ →
-                       â\88\80I2,W2. â\93\91{I1}W1.L1 â©\96[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I2}W2.L.
+                       â\88\80I2,W2. â\93\91{I1}W1.L1 â\8b\93[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
-lapply (ldrop_fwd_length_lt2 … HLK1) >ltail_length #H
+lapply (drop_fwd_length_lt2 … HLK1) >ltail_length #H
 lapply (lt_plus_SO_to_le … H) -H #H
 elim (le_to_or_lt_eq … H) -H #H
-[ elim (ldrop_O1_lt (Ⓕ) … H) #Z1 #Y1 #X1 #HLY1
-  elim (ldrop_O1_lt (Ⓕ) L2 i) // #Z2 #Y2 #X2 #HLY2
-  elim (ldrop_O1_lt (Ⓕ) L i) // #Z #Y #X #HLY
-  lapply (ldrop_O1_inv_append1_le … HLK1 … HLY1) /2 width=1 by lt_to_le/ -HLK1 normalize #H destruct
-  lapply (ldrop_O1_inv_append1_le … HLK2 … HLY2) /2 width=1 by lt_to_le/ -HLK2 normalize #H destruct
-  lapply (ldrop_O1_inv_append1_le … HLK … HLY) /2 width=1 by lt_to_le/ -HLK normalize #H destruct
+[ elim (drop_O1_lt (Ⓕ) … H) #Z1 #Y1 #X1 #HLY1
+  elim (drop_O1_lt (Ⓕ) L2 i) // #Z2 #Y2 #X2 #HLY2
+  elim (drop_O1_lt (Ⓕ) L i) // #Z #Y #X #HLY
+  lapply (drop_O1_inv_append1_le … HLK1 … HLY1) /2 width=1 by lt_to_le/ -HLK1 normalize #H destruct
+  lapply (drop_O1_inv_append1_le … HLK2 … HLY2) /2 width=1 by lt_to_le/ -HLK2 normalize #H destruct
+  lapply (drop_O1_inv_append1_le … HLK … HLY) /2 width=1 by lt_to_le/ -HLK normalize #H destruct
   elim (IH … HLY1 HLY2 HLY) -IH -HLY1 -HLY2 -HLY *
   [ /3 width=1 by and3_intro, or3_intro0/
   | /6 width=2 by frees_inv_append, lt_to_le, or3_intro1, and3_intro/
   | /5 width=1 by frees_append, lt_to_le, or3_intro2, and4_intro/
   ]
 | -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
+  lapply (drop_O1_inv_append1_le … HLK2 … (⋆) ?) // -HLK2 normalize #H destruct
+  lapply (drop_O1_inv_append1_le … HLK … (⋆) ?) // -HLK normalize #H destruct
   /3 width=1 by or3_intro2, and4_intro/
 ]
 qed.
 
-lemma llor_tail_cofrees: â\88\80L1,L2,L,U,d. L1 â©\96[U, d] L2 ≡ L →
+lemma llor_tail_cofrees: â\88\80L1,L2,L,U,d. L1 â\8b\93[U, d] L2 ≡ L →
                          ∀I1,W1. (ⓑ{I1}W1.L1 ⊢ |L1| ϵ 𝐅*[d]⦃U⦄ → ⊥) →
-                         â\88\80I2,W2. â\93\91{I1}W1.L1 â©\96[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I1}W1.L.
+                         â\88\80I2,W2. â\93\91{I1}W1.L1 â\8b\93[U, d] ⓑ{I2}W2.L2 ≡ ⓑ{I1}W1.L.
 #L1 #L2 #L #U #d * #HL12 #HL1 #IH #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
-lapply (ldrop_fwd_length_lt2 … HLK1) >ltail_length #H
+lapply (drop_fwd_length_lt2 … HLK1) >ltail_length #H
 lapply (lt_plus_SO_to_le … H) -H #H
 elim (le_to_or_lt_eq … H) -H #H
-[ elim (ldrop_O1_lt (Ⓕ) … H) #Z1 #Y1 #X1 #HLY1
-  elim (ldrop_O1_lt (Ⓕ) L2 i) // #Z2 #Y2 #X2 #HLY2
-  elim (ldrop_O1_lt (Ⓕ) L i) // #Z #Y #X #HLY
-  lapply (ldrop_O1_inv_append1_le … HLK1 … HLY1) /2 width=1 by lt_to_le/ -HLK1 normalize #H destruct
-  lapply (ldrop_O1_inv_append1_le … HLK2 … HLY2) /2 width=1 by lt_to_le/ -HLK2 normalize #H destruct
-  lapply (ldrop_O1_inv_append1_le … HLK … HLY) /2 width=1 by lt_to_le/ -HLK normalize #H destruct
+[ elim (drop_O1_lt (Ⓕ) … H) #Z1 #Y1 #X1 #HLY1
+  elim (drop_O1_lt (Ⓕ) L2 i) // #Z2 #Y2 #X2 #HLY2
+  elim (drop_O1_lt (Ⓕ) L i) // #Z #Y #X #HLY
+  lapply (drop_O1_inv_append1_le … HLK1 … HLY1) /2 width=1 by lt_to_le/ -HLK1 normalize #H destruct
+  lapply (drop_O1_inv_append1_le … HLK2 … HLY2) /2 width=1 by lt_to_le/ -HLK2 normalize #H destruct
+  lapply (drop_O1_inv_append1_le … HLK … HLY) /2 width=1 by lt_to_le/ -HLK normalize #H destruct
   elim (IH … HLY1 HLY2 HLY) -IH -HLY1 -HLY2 -HLY *
   [ /3 width=1 by and3_intro, or3_intro0/
   | /6 width=2 by frees_inv_append, lt_to_le, or3_intro1, and3_intro/
   | /5 width=1 by frees_append, lt_to_le, or3_intro2, and4_intro/
   ]
 | -IH -HLK2 destruct
-  lapply (ldrop_O1_inv_append1_le … HLK1 … (⋆) ?) // -HLK1 normalize #H destruct
-  lapply (ldrop_O1_inv_append1_le … HLK … (⋆) ?) // -HLK normalize #H destruct
+  lapply (drop_O1_inv_append1_le … HLK1 … (⋆) ?) // -HLK1 normalize #H destruct
+  lapply (drop_O1_inv_append1_le … HLK … (⋆) ?) // -HLK normalize #H destruct
   /4 width=1 by or3_intro1, and3_intro/
 ]
 qed.