]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/lsx_csx.ma
- commit of the "s_computation" component ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lsx_csx.ma
index a4937895565993723302de6d6ffb95328d416e76..fa00a907d909c416d4b3279b482523173bb2ab41 100644 (file)
@@ -19,10 +19,10 @@ include "basic_2/computation/lcosx_cpx.ma".
 
 (* Advanced properties ******************************************************)
 
-lemma lsx_lref_be_lpxs: ∀h,g,I,G,K1,V,i,d. d ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, g] V →
-                        ∀K2. G ⊢ ⬊*[h, g, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, g] K2 →
-                        ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, g, #i, d] L2.
-#h #g #I #G #K1 #V #i #d #Hdi #H @(csx_ind_alt … H) -V
+lemma lsx_lref_be_lpxs: ∀h,o,I,G,K1,V,i,l. l ≤ yinj i → ⦃G, K1⦄ ⊢ ⬊*[h, o] V →
+                        ∀K2. G ⊢ ⬊*[h, o, V, 0] K2 → ⦃G, K1⦄ ⊢ ➡*[h, o] K2 →
+                        ∀L2. ⬇[i] L2 ≡ K2.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L2.
+#h #o #I #G #K1 #V #i #l #Hli #H @(csx_ind_alt … H) -V
 #V0 #_ #IHV0 #K2 #H @(lsx_ind … H) -K2
 #K0 #HK0 #IHK0 #HK10 #L0 #HLK0 @lsx_intro
 #L2 #HL02 #HnL02 elim (lpx_drop_conf … HLK0 … HL02) -HL02
@@ -35,25 +35,25 @@ elim (eq_term_dec V0 V2) #HnV02 destruct [ -IHV0 -HV02 -HK0 | -IHK0 -HnL02 -HLK0
 ]
 qed.
 
-lemma lsx_lref_be: ∀h,g,I,G,K,V,i,d. d ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, g] V →
-                   G ⊢ ⬊*[h, g, V, 0] K →
-                   ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, g, #i, d] L.
+lemma lsx_lref_be: ∀h,o,I,G,K,V,i,l. l ≤ yinj i → ⦃G, K⦄ ⊢ ⬊*[h, o] V →
+                   G ⊢ ⬊*[h, o, V, 0] K →
+                   ∀L. ⬇[i] L ≡ K.ⓑ{I}V → G ⊢ ⬊*[h, o, #i, l] L.
 /2 width=8 by lsx_lref_be_lpxs/ qed.
 
 (* Main properties **********************************************************)
 
-theorem csx_lsx: ∀h,g,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, g] T → ∀d. G ⊢ ⬊*[h, g, T, d] L.
-#h #g #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
+theorem csx_lsx: ∀h,o,G,L,T. ⦃G, L⦄ ⊢ ⬊*[h, o] T → ∀l. G ⊢ ⬊*[h, o, T, l] L.
+#h #o #G #L #T @(fqup_wf_ind_eq … G L T) -G -L -T
 #Z #Y #X #IH #G #L * * //
-[ #i #HG #HL #HT #H #d destruct
+[ #i #HG #HL #HT #H #l destruct
   elim (lt_or_ge i (|L|)) /2 width=1 by lsx_lref_free/
-  elim (ylt_split i d) /2 width=1 by lsx_lref_skip/
-  #Hdi #Hi elim (drop_O1_lt (Ⓕ) … Hi) -Hi
+  elim (ylt_split i l) /2 width=1 by lsx_lref_skip/
+  #Hli #Hi elim (drop_O1_lt (Ⓕ) … Hi) -Hi
   #I #K #V #HLK lapply (csx_inv_lref_bind … HLK … H) -H
   /4 width=6 by lsx_lref_be, fqup_lref/
-| #a #I #V #T #HG #HL #HT #H #d destruct
+| #a #I #V #T #HG #HL #HT #H #l destruct
   elim (csx_fwd_bind … H) -H /3 width=1 by lsx_bind/
-| #I #V #T #HG #HL #HT #H #d destruct
+| #I #V #T #HG #HL #HT #H #l destruct
   elim (csx_fwd_flat … H) -H /3 width=1 by lsx_flat/
 ]
 qed.