1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "basic_2A/notation/relations/btpredstaralt_8.ma".
16 include "basic_2A/multiple/lleq_fqus.ma".
17 include "basic_2A/computation/cpxs_lleq.ma".
18 include "basic_2A/computation/lpxs_lleq.ma".
19 include "basic_2A/computation/fpbs.ma".
21 (* "QREST" PARALLEL COMPUTATION FOR CLOSURES ********************************)
23 (* Note: alternative definition of fpbs *)
24 definition fpbsa: ∀h. sd h → tri_relation genv lenv term ≝
25 λh,g,G1,L1,T1,G2,L2,T2.
26 ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T &
27 ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ &
28 ⦃G2, L0⦄ ⊢ ➡*[h, g] L & L ≡[T2, 0] L2.
30 interpretation "'big tree' parallel computation (closure) alternative"
31 'BTPRedStarAlt h g G1 L1 T1 G2 L2 T2 = (fpbsa h g G1 L1 T1 G2 L2 T2).
33 (* Basic properties *********************************************************)
35 lemma fpb_fpbsa_trans: ∀h,g,G1,G,L1,L,T1,T. ⦃G1, L1, T1⦄ ≽[h, g] ⦃G, L, T⦄ →
36 ∀G2,L2,T2. ⦃G, L, T⦄ ≥≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, g] ⦃G2, L2, T2⦄.
37 #h #g #G1 #G #L1 #L #T1 #T * -G -L -T [ #G #L #T #HG1 | #T #HT1 | #L #HL1 | #L #HL1 ]
38 #G2 #L2 #T2 * #L00 #L0 #T0 #HT0 #HG2 #HL00 #HL02
39 [ elim (fquq_cpxs_trans … HT0 … HG1) -T
40 /3 width=7 by fqus_strap2, ex4_3_intro/
41 | /3 width=7 by cpxs_strap2, ex4_3_intro/
42 | lapply (lpx_cpxs_trans … HT0 … HL1) -HT0 #HT10
43 elim (lpx_fqus_trans … HG2 … HL1) -L
44 /3 width=7 by lpxs_strap2, cpxs_trans, ex4_3_intro/
45 | lapply (lleq_cpxs_trans … HT0 … HL1) -HT0 #HT0
46 lapply (cpxs_lleq_conf_sn … HT0 … HL1) -HL1 #HL1
47 elim (lleq_fqus_trans … HG2 … HL1) -L #K00 #HG12 #HKL00
48 elim (lleq_lpxs_trans … HL00 … HKL00) -L00
49 /3 width=9 by lleq_trans, ex4_3_intro/
53 (* Main properties **********************************************************)
55 theorem fpbs_fpbsa: ∀h,g,G1,G2,L1,L2,T1,T2.
56 ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥≥[h, g] ⦃G2, L2, T2⦄.
57 #h #g #G1 #G2 #L1 #L2 #T1 #T2 #H @(fpbs_ind_dx … H) -G1 -L1 -T1
58 /2 width=7 by fpb_fpbsa_trans, ex4_3_intro/
61 (* Main inversion lemmas ****************************************************)
63 theorem fpbsa_inv_fpbs: ∀h,g,G1,G2,L1,L2,T1,T2.
64 ⦃G1, L1, T1⦄ ≥≥[h, g] ⦃G2, L2, T2⦄ → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄.
65 #h #g #G1 #G2 #L1 #L2 #T1 #T2 *
66 /3 width=5 by cpxs_fqus_lpxs_fpbs, fpbs_strap1, fpbq_lleq/
69 (* Advanced properties ******************************************************)
71 lemma fpbs_intro_alt: ∀h,g,G1,G2,L1,L0,L,L2,T1,T,T2.
72 ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T → ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ →
73 ⦃G2, L0⦄ ⊢ ➡*[h, g] L → L ≡[T2, 0] L2 → ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ .
74 /3 width=7 by fpbsa_inv_fpbs, ex4_3_intro/ qed.
76 (* Advanced inversion lemmas *************************************************)
78 lemma fpbs_inv_alt: ∀h,g,G1,G2,L1,L2,T1,T2. ⦃G1, L1, T1⦄ ≥[h, g] ⦃G2, L2, T2⦄ →
79 ∃∃L0,L,T. ⦃G1, L1⦄ ⊢ T1 ➡*[h, g] T &
80 ⦃G1, L1, T⦄ ⊐* ⦃G2, L0, T2⦄ &
81 ⦃G2, L0⦄ ⊢ ➡*[h, g] L & L ≡[T2, 0] L2.
82 /2 width=1 by fpbs_fpbsa/ qed-.