]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/csx_lift.ma
notational change of lift, drop, and gget
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / csx_lift.ma
index 2d74a929eb868012ff0927f6b0723e95fc2b8f6e..6b2ddff83adeed4e4df0bb6c126f4763afad3145 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "basic_2/reduction/cnx_lift.ma".
-include "basic_2/computation/acp.ma".
+include "basic_2/computation/gcp.ma".
 include "basic_2/computation/csx.ma".
 
 (* CONTEXT-SENSITIVE EXTENDED STRONGLY NORMALIZING TERMS ********************)
@@ -22,7 +22,7 @@ include "basic_2/computation/csx.ma".
 
 (* Basic_1: was just: sn3_lift *)
 lemma csx_lift: ∀h,g,G,L2,L1,T1,s,d,e. ⦃G, L1⦄ ⊢ ⬊*[h, g] T1 →
-                â\88\80T2. â\87©[s, d, e] L2 â\89¡ L1 â\86\92 â\87§[d, e] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
+                â\88\80T2. â¬\87[s, d, e] L2 â\89¡ L1 â\86\92 â¬\86[d, e] T1 ≡ T2 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
 #h #g #G #L2 #L1 #T1 #s #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL21 #HT12
 @csx_intro #T #HLT2 #HT2
 elim (cpx_inv_lift1 … HLT2 … HL21 … HT12) -HLT2 #T0 #HT0 #HLT10
@@ -32,7 +32,7 @@ qed.
 
 (* Basic_1: was just: sn3_gen_lift *)
 lemma csx_inv_lift: ∀h,g,G,L2,L1,T1,s,d,e. ⦃G, L1⦄ ⊢ ⬊*[h, g] T1 →
-                    â\88\80T2. â\87©[s, d, e] L1 â\89¡ L2 â\86\92 â\87§[d, e] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
+                    â\88\80T2. â¬\87[s, d, e] L1 â\89¡ L2 â\86\92 â¬\86[d, e] T2 ≡ T1 → ⦃G, L2⦄ ⊢ ⬊*[h, g] T2.
 #h #g #G #L2 #L1 #T1 #s #d #e #H elim H -T1 #T1 #_ #IHT1 #T2 #HL12 #HT21
 @csx_intro #T #HLT2 #HT2
 elim (lift_total T d e) #T0 #HT0
@@ -44,24 +44,24 @@ qed.
 (* Advanced inversion lemmas ************************************************)
 
 (* Basic_1: was: sn3_gen_def *)
-lemma csx_inv_lref_bind: â\88\80h,g,I,G,L,K,V,i. â\87©[i] L ≡ K.ⓑ{I}V →
+lemma csx_inv_lref_bind: â\88\80h,g,I,G,L,K,V,i. â¬\87[i] L ≡ K.ⓑ{I}V →
                          ⦃G, L⦄ ⊢ ⬊*[h, g] #i → ⦃G, K⦄ ⊢ ⬊*[h, g] V.
 #h #g #I #G #L #K #V #i #HLK #Hi
 elim (lift_total V 0 (i+1))
-/4 width=9 by csx_inv_lift, csx_cpx_trans, cpx_delta, ldrop_fwd_drop2/
+/4 width=9 by csx_inv_lift, csx_cpx_trans, cpx_delta, drop_fwd_drop2/
 qed-.
 
 (* Advanced properties ******************************************************)
 
 (* Basic_1: was just: sn3_abbr *)
-lemma csx_lref_bind: â\88\80h,g,I,G,L,K,V,i. â\87©[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ ⬊*[h, g] V → ⦃G, L⦄ ⊢ ⬊*[h, g] #i.
+lemma csx_lref_bind: â\88\80h,g,I,G,L,K,V,i. â¬\87[i] L ≡ K.ⓑ{I}V → ⦃G, K⦄ ⊢ ⬊*[h, g] V → ⦃G, L⦄ ⊢ ⬊*[h, g] #i.
 #h #g #I #G #L #K #V #i #HLK #HV
 @csx_intro #X #H #Hi
 elim (cpx_inv_lref1 … H) -H
 [ #H destruct elim Hi //
 | -Hi * #I0 #K0 #V0 #V1 #HLK0 #HV01 #HV1
-  lapply (ldrop_mono … HLK0 … HLK) -HLK #H destruct
-  /3 width=8 by csx_lift, csx_cpx_trans, ldrop_fwd_drop2/
+  lapply (drop_mono … HLK0 … HLK) -HLK #H destruct
+  /3 width=8 by csx_lift, csx_cpx_trans, drop_fwd_drop2/
 ]
 qed.
 
@@ -109,11 +109,11 @@ qed-.
 
 (* Main properties **********************************************************)
 
-theorem csx_acp: ∀h,g. acp (cpx h g) (eq …) (csx h g).
-#h #g @mk_acp
-[ /3 width=13 by cnx_lift/
+theorem csx_gcp: ∀h,g. gcp (cpx h g) (eq …) (csx h g).
+#h #g @mk_gcp
+[ normalize /3 width=13 by cnx_lift/
 | #G #L elim (deg_total h g 0) /3 width=8 by cnx_sort_iter, ex_intro/
+| /2 width=8 by csx_lift/
 | /2 width=3 by csx_fwd_flat_dx/
-| /2 width=1 by csx_cast/
 ]
 qed.