]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_transition/cnx_drops.ma
update in ground_2, static_2, basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_transition / cnx_drops.ma
index 76f8f77b093cd6cb65a3b1a85b8a33dd6af249ab..20b92902e1e745ebe4b681489fdb259e3a04c481 100644 (file)
@@ -20,12 +20,12 @@ include "basic_2/rt_transition/cnx.ma".
 
 (* Properties with generic slicing ******************************************)
 
-lemma cnx_lref_atom: ∀h,G,L,i. ⇩*[i] L ≘ ⋆ → ❪G,L❫ ⊢ ⬈[h] 𝐍❪#i❫.
+lemma cnx_lref_atom: ∀h,G,L,i. ⇩[i] L ≘ ⋆ → ❪G,L❫ ⊢ ⬈[h] 𝐍❪#i❫.
 #h #G #L #i #Hi #X #H elim (cpx_inv_lref1_drops … H) -H // *
 #I #K #V1 #V2 #HLK lapply (drops_mono … Hi … HLK) -L #H destruct
 qed.
 
-lemma cnx_lref_unit: ∀h,I,G,L,K,i. ⇩*[i] L ≘ K.ⓤ[I] → ❪G,L❫ ⊢ ⬈[h] 𝐍❪#i❫.
+lemma cnx_lref_unit: ∀h,I,G,L,K,i. ⇩[i] L ≘ K.ⓤ[I] → ❪G,L❫ ⊢ ⬈[h] 𝐍❪#i❫.
 #h #I #G #L #K #i #HLK #X #H elim (cpx_inv_lref1_drops … H) -H // *
 #Z #Y #V1 #V2 #HLY lapply (drops_mono … HLK … HLY) -L #H destruct
 qed.
@@ -40,7 +40,7 @@ qed-.
 (* Inversion lemmas with generic slicing ************************************)
 
 (* Basic_2A1: was: cnx_inv_delta *)
-lemma cnx_inv_lref_pair: ∀h,I,G,L,K,V,i. ⇩*[i] L ≘ K.ⓑ[I]V → ❪G,L❫ ⊢ ⬈[h] 𝐍❪#i❫ → ⊥.
+lemma cnx_inv_lref_pair: ∀h,I,G,L,K,V,i. ⇩[i] L ≘ K.ⓑ[I]V → ❪G,L❫ ⊢ ⬈[h] 𝐍❪#i❫ → ⊥.
 #h #I #G #L #K #V #i #HLK #H
 elim (lifts_total V (𝐔❨↑i❩)) #W #HVW
 lapply (H W ?) -H /2 width=7 by cpx_delta_drops/ -HLK