]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/fsb_fpbg.ma
λδ-2B is released
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / fsb_fpbg.ma
index 0ebebad3aaef9fddce10d502768f9d37a032215d..1fd7ea36b6f77dc068d3e3ad073880d69522774d 100644 (file)
@@ -13,7 +13,7 @@
 (**************************************************************************)
 
 include "basic_2/rt_computation/fpbg_fpbs.ma".
-include "basic_2/rt_computation/fsb_fdeq.ma".
+include "basic_2/rt_computation/fsb_feqx.ma".
 
 (* STRONGLY NORMALIZING CLOSURES FOR PARALLEL RST-TRANSITION ****************)
 
@@ -24,7 +24,7 @@ lemma fsb_fpbs_trans: ∀h,G1,L1,T1. ≥[h] 𝐒⦃G1,L1,T1⦄ →
 #h #G1 #L1 #T1 #H @(fsb_ind_alt … H) -G1 -L1 -T1
 #G1 #L1 #T1 #H1 #IH #G2 #L2 #T2 #H12
 elim (fpbs_inv_fpbg … H12) -H12
-[ -IH /2 width=5 by fsb_fdeq_trans/
+[ -IH /2 width=5 by fsb_feqx_trans/
 | -H1 * /2 width=5 by/
 ]
 qed-.
@@ -43,7 +43,7 @@ lemma fsb_ind_fpbg_fpbs: ∀h. ∀Q:relation3 genv lenv term.
                                      (∀G2,L2,T2. ⦃G1,L1,T1⦄ >[h] ⦃G2,L2,T2⦄ → Q G2 L2 T2) →
                                      Q G1 L1 T1
                          ) →
-                         ∀G1,L1,T1. ≥[h] 𝐒⦃G1,L1,T1⦄ → 
+                         ∀G1,L1,T1. ≥[h] 𝐒⦃G1,L1,T1⦄ →
                          ∀G2,L2,T2. ⦃G1,L1,T1⦄ ≥[h] ⦃G2,L2,T2⦄ → Q G2 L2 T2.
 #h #Q #IH1 #G1 #L1 #T1 #H @(fsb_ind_alt … H) -G1 -L1 -T1
 #G1 #L1 #T1 #H1 #IH #G2 #L2 #T2 #H12