]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2/rt_computation/fpbs_fpb.ma
update in basic_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / rt_computation / fpbs_fpb.ma
index b9c657735fb44ab5e58a20637fb189eb12354ad4..632fe7dc6d656672c7727201a3ecd5f5cec83325 100644 (file)
 (*                                                                        *)
 (**************************************************************************)
 
-include "basic_2/reduction/fpbq_alt.ma".
-include "basic_2/computation/fpbs_alt.ma".
+include "basic_2/rt_transition/fpbq_fpb.ma".
+include "basic_2/rt_computation/fpbs.ma".
 
-(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
+(* PARALLEL RST-COMPUTATION FOR CLOSURES ************************************)
 
-(* Properties on extended context-sensitive parallel computation for terms **)
-
-lemma fpbs_cpx_trans_neq: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄ →
-                          ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, o] U2 → (T2 = U2 → ⊥) →
-                          ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, o] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ≥[h, o] ⦃G2, L2, U2⦄.
-#h #o #G1 #G2 #L1 #L2 #T1 #T2 #H #U2 #HTU2 #HnTU2 elim (fpbs_inv_alt … H) -H
-#L00 #L0 #T0 #HT10 #H10 #HL00 #HL02 lapply (lleq_cpx_trans … HTU2 … HL02) -HTU2
-#HTU2 lapply (cpx_lleq_conf_sn … HTU2 … HL02) -HL02
-#HL02 lapply (lpxs_cpx_trans … HTU2 … HL00) -HTU2
-#HTU2 elim (fqus_cpxs_trans_neq … H10 … HTU2 HnTU2) -H10 -HTU2 -HnTU2
-#U0 #HTU0 #HnTU0 #HU02 elim (eq_term_dec T1 T0) #HnT10 destruct
-[ -HT10 elim (cpxs_neq_inv_step_sn … HTU0 HnTU0) -HTU0 -HnTU0
-| -HnTU0 elim (cpxs_neq_inv_step_sn … HT10 HnT10) -HT10 -HnT10
-]
-/4 width=10 by fpbs_intro_alt, cpxs_trans, ex3_intro/
-qed-.
-
-(* Properties on "rst" proper parallel reduction on closures ****************)
+(* Properties with proper parallel rst-reduction on closures ****************)
 
 lemma fpb_fpbs: ∀h,o,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, o] ⦃G2, L2, T2⦄ →
                 ⦃G1, L1, T1⦄ ≥[h, o] ⦃G2, L2, T2⦄.