]> matita.cs.unibo.it Git - helm.git/blobdiff - matita/matita/contribs/lambdadelta/basic_2A/computation/fpbs_fpb.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / computation / fpbs_fpb.ma
diff --git a/matita/matita/contribs/lambdadelta/basic_2A/computation/fpbs_fpb.ma b/matita/matita/contribs/lambdadelta/basic_2A/computation/fpbs_fpb.ma
new file mode 100644 (file)
index 0000000..634c21a
--- /dev/null
@@ -0,0 +1,41 @@
+(**************************************************************************)
+(*       ___                                                              *)
+(*      ||M||                                                             *)
+(*      ||A||       A project by Andrea Asperti                           *)
+(*      ||T||                                                             *)
+(*      ||I||       Developers:                                           *)
+(*      ||T||         The HELM team.                                      *)
+(*      ||A||         http://helm.cs.unibo.it                             *)
+(*      \   /                                                             *)
+(*       \ /        This file is distributed under the terms of the       *)
+(*        v         GNU General Public License Version 2                  *)
+(*                                                                        *)
+(**************************************************************************)
+
+include "basic_2A/reduction/fpbq_alt.ma".
+include "basic_2A/computation/fpbs_alt.ma".
+
+(* "QRST" PARALLEL COMPUTATION FOR CLOSURES *********************************)
+
+(* Properties on extended context-sensitive parallel computation for terms **)
+
+lemma fpbs_cpx_trans_neq: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
+                          ∀U2. ⦃G2, L2⦄ ⊢ T2 ➡[h, g] U2 → (T2 = U2 → ⊥) →
+                          ∃∃U1. ⦃G1, L1⦄ ⊢ T1 ➡[h, g] U1 & T1 = U1 → ⊥ & ⦃G1, L1, U1⦄ ≥[h, g] ⦃G2, L2, U2⦄.
+#h #g #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 ****************)
+
+lemma fpb_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≻[h, g] ⦃G2, L2, T2⦄ →
+                ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
+/3 width=1 by fpbq_fpbs, fpb_fpbq/ qed.