]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/computation/fsb_csx.ma
- lazy extended reduction parked
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / fsb_csx.ma
index 453383dc483bb300eb4414c40cd93742b824f71f..e094ff57fc78a2462734f5e9bcdf1a4d0e7de6ea 100644 (file)
@@ -14,7 +14,7 @@
 
 include "basic_2/computation/fpbs_ext.ma".
 include "basic_2/computation/csx_fpbs.ma".
-include "basic_2/computation/llsx_csx.ma".
+include "basic_2/computation/lsx_csx.ma".
 include "basic_2/computation/fsb_alt.ma".
 
 (* "BIG TREE" STRONGLY NORMALIZING TERMS ************************************)
@@ -27,19 +27,19 @@ lemma csx_fsb_fpbs: ∀h,g,G1,L1,T1. ⦃G1, L1⦄ ⊢ ⬊*[h, g] T1 →
 #T1 #HT1 #IHc #G2 #L2 #T2 @(fqup_wf_ind … G2 L2 T2) -G2 -L2 -T2
 #G0 #L0 #T0 #IHu #H10 lapply (csx_fpbs_conf … H10) // -HT1
 #HT0 generalize in match IHu; -IHu generalize in match H10; -H10
-@(llsx_ind_alt … (csx_llsx … HT0 0)) -L0
+@(lsx_ind_alt … (csx_lsx … HT0 0)) -L0
 #L0 #_ #IHl #H10 #IHu @fsb_intro
 #G2 #L2 #T2 * -G2 -L2 -T2 [ -IHl -IHc | -IHu -IHl |  ]
 [ /3 width=5 by fpbs_fqup_trans/
 | #T2 #HT02 #HnT02 elim (fpbs_cpxs_trans_neq … H10 … HT02 HnT02) -T0
   /3 width=4 by/
 | #L2 #HL02 #HnL02 @(IHl … HL02 HnL02) -IHl -HnL02 [ -IHu -IHc | ]
-  [ /2 width=3 by fpbs_llpxs_trans/
-  | #G3 #L3 #T3 #H03 #_ elim (llpxs_fqup_trans … H03 … HL02) -L2
+  [ /2 width=3 by fpbs_lpxs_trans/
+  | #G3 #L3 #T3 #H03 #_ elim (lpxs_fqup_trans … H03 … HL02) -L2
     #L4 #T4 elim (eq_term_dec T0 T4) [ -IHc | -IHu ]
-    [ #H destruct /4 width=5 by fsb_fpbs_trans, llpxs_fpbs, fpbs_fqup_trans/
+    [ #H destruct /4 width=5 by fsb_fpbs_trans, lpxs_fpbs, fpbs_fqup_trans/
     | #HnT04 #HT04 #H04 elim (fpbs_cpxs_trans_neq … H10 … HT04 HnT04) -T0
-      /4 width=8 by fpbs_fqup_trans, fpbs_llpxs_trans/
+      /4 width=8 by fpbs_fqup_trans, fpbs_lpxs_trans/
     ]
   ]
 ]