]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/static_2/static/aaa_drops.ma
partial commit in static_2
[helm.git] / matita / matita / contribs / lambdadelta / static_2 / static / aaa_drops.ma
index 866a0a528bf62fc0260430b4308dec0eab28ed31..b501bd3a6ae785c70b309eac5d4a30068805ed08 100644 (file)
@@ -22,11 +22,11 @@ include "static_2/static/aaa.ma".
 (* Advanced properties ******************************************************)
 
 (* Basic_2A1: was: aaa_lref *)
-lemma aaa_lref_drops: ∀I,G,K,V,B,i,L. ⇩*[i] L ≘ K.ⓑ[I]V → ❪G,K❫ ⊢ V ⁝ B → ❪G,L❫ ⊢ #i ⁝ B.
+lemma aaa_lref_drops: ∀I,G,K,V,B,i,L. ⇩[i] L ≘ K.ⓑ[I]V → ❪G,K❫ ⊢ V ⁝ B → ❪G,L❫ ⊢ #i ⁝ B.
 #I #G #K #V #B #i elim i -i
 [ #L #H lapply (drops_fwd_isid … H ?) -H //
   #H destruct /2 width=1 by aaa_zero/
-| #i #IH #L <uni_succ #H #HB lapply (drops_inv_bind2_isuni_next … H) -H // *
+| #i #IH #L <pr_uni_succ #H #HB lapply (drops_inv_bind2_isuni_next … H) -H // *
   #Z #Y #HY #H destruct /3 width=1 by aaa_lref/
 ]
 qed.
@@ -35,7 +35,7 @@ qed.
 
 (* Basic_2A1: was: aaa_inv_lref *)
 lemma aaa_inv_lref_drops: ∀G,A,i,L. ❪G,L❫ ⊢ #i ⁝ A →
-                          ∃∃I,K,V. ⇩*[i] L ≘ K.ⓑ[I]V & ❪G,K❫ ⊢ V ⁝ A.
+                          ∃∃I,K,V. ⇩[i] L ≘ K.ⓑ[I]V & ❪G,K❫ ⊢ V ⁝ A.
 #G #A #i elim i -i
 [ #L #H elim (aaa_inv_zero … H) -H /3 width=5 by drops_refl, ex2_3_intro/
 | #i #IH #L #H elim (aaa_inv_lref … H) -H
@@ -44,7 +44,7 @@ lemma aaa_inv_lref_drops: ∀G,A,i,L. ❪G,L❫ ⊢ #i ⁝ A →
 qed-.
 
 lemma aaa_pair_inv_lref (G) (L) (i):
-      ∀A. ❪G,L❫ ⊢ #i ⁝ A → ∀I,K,V. ⇩*[i] L ≘ K.ⓑ[I]V → ❪G,K❫ ⊢ V ⁝ A.
+      ∀A. ❪G,L❫ ⊢ #i ⁝ A → ∀I,K,V. ⇩[i] L ≘ K.ⓑ[I]V → ❪G,K❫ ⊢ V ⁝ A.
 #G #L #i #A #H #I #K #V #HLK
 elim (aaa_inv_lref_drops … H) -H #J #Y #X #HLY #HX
 lapply (drops_mono … HLY … HLK) -L -i #H destruct //
@@ -64,7 +64,7 @@ lemma aaa_lifts: ∀G,L1,T1,A. ❪G,L1❫ ⊢ T1 ⁝ A → ∀b,f,L2. ⇩*[b,f]
   elim (aaa_inv_lref_drops … H) -H #J #K1 #V1 #HLK1 #HA
   elim (lifts_inv_lref1 … HX) -HX #i2 #Hf #H destruct
   lapply (drops_trans … HL21 … HLK1 ??) -HL21 [1,2: // ] #H
-  elim (drops_split_trans … H) -H [ |*: /2 width=6 by after_uni_dx/ ] #Y #HLK2 #HY
+  elim (drops_split_trans … H) -H [ |*: /2 width=6 by pr_after_nat_uni/ ] #Y #HLK2 #HY
   lapply (drops_tls_at … Hf … HY) -HY #HY -Hf
   elim (drops_inv_skip2 … HY) -HY #Z #K2 #HK21 #HZ #H destruct
   elim (liftsb_inv_pair_sn … HZ) -HZ #V2 #HV12 #H destruct
@@ -103,7 +103,7 @@ lemma aaa_inv_lifts: ∀G,L2,T2,A. ❪G,L2❫ ⊢ T2 ⁝ A → ∀b,f,L1. ⇩*[b
   elim (aaa_inv_lref_drops … H) -H #J #K2 #V2 #HLK2 #HA
   elim (lifts_inv_lref2 … HX) -HX #i1 #Hf #H destruct
   lapply (drops_split_div … HL21 (𝐔❨i1❩) ???) -HL21 [4: * |*: // ] #Y #HLK1 #HY
-  lapply (drops_conf … HLK2 … HY ??) -HY [1,2: /2 width=6 by after_uni_dx/ ] #HY
+  lapply (drops_conf … HLK2 … HY ??) -HY [1,2: /2 width=6 by pr_after_nat_uni/ ] #HY
   lapply (drops_tls_at … Hf … HY) -HY #HY -Hf
   elim (drops_inv_skip1 … HY) -HY #Z #K1 #HK21 #HZ #H destruct
   elim (liftsb_inv_pair_dx … HZ) -HZ #V1 #HV12 #H destruct