]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/computation/cpds_cpds.ma
- update in basic_2 and ground_2
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / cpds_cpds.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 include "basic_2/unfold/lstas_lstas.ma".
16 include "basic_2/computation/lprs_cprs.ma".
17 include "basic_2/computation/cpxs_cpxs.ma".
18 include "basic_2/computation/cpds.ma".
19
20 (* DECOMPOSED EXTENDED PARALLEL COMPUTATION ON TERMS ************************)
21
22 (* Advanced properties ******************************************************)
23
24 lemma cpds_strap2: ∀h,g,G,L,T1,T,T2,l1,l. ⦃G, L⦄ ⊢ T1 ▪[h, g] l1+1 →
25                    ⦃G, L⦄ ⊢ T1 •[h] T → ⦃G, L⦄ ⊢ T •*➡*[h, g, l] T2 →
26                    ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l+1] T2.
27 #h #g #G #L #T1 #T #T2 #l1 #l #Hl1 #HT1 *
28 #T0 #l0 #Hl0 #HTl0 #HT0 #HT02
29 lapply (da_sta_conf … HT1 … Hl1) <minus_plus_m_m #HTl1
30 lapply (da_mono … HTl0 … HTl1) -HTl0 -HTl1 #H destruct
31 /3 width=6 by lstas_step_sn, le_S_S, ex4_2_intro/
32 qed.
33
34 lemma cpds_cprs_trans: ∀h,g,G,L,T1,T,T2,l.
35                        ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T → ⦃G, L⦄ ⊢ T ➡* T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T2.
36 #h #g #G #L #T1 #T #T2 #l * /3 width=8 by cprs_trans, ex4_2_intro/
37 qed-.
38
39 lemma lstas_cpds_trans: ∀h,g,G,L,T1,T,T2,l1,l2,l.
40                         l2 ≤ l1 → ⦃G, L⦄ ⊢ T1 ▪[h, g] l1 →
41                         ⦃G, L⦄ ⊢ T1 •*[h, l2] T → ⦃G, L⦄ ⊢ T •*➡*[h, g, l] T2 → ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l2+l] T2.
42 #h #g #G #L #T1 #T #T2 #l1 #l2 #l #Hl21 #HTl1 #HT1 * #T0 #l0 #Hl0 #HTl0 #HT0 #HT02
43 lapply (lstas_da_conf … HT1 … HTl1) #HTl12
44 lapply (da_mono … HTl12 … HTl0) -HTl12 -HTl0 #H destruct
45 lapply (le_minus_to_plus_r … Hl21 Hl0) -Hl21 -Hl0
46 /3 width=7 by lstas_trans, ex4_2_intro/
47 qed-.
48
49 (* Advanced inversion lemmas ************************************************)
50
51 lemma cpds_inv_abst1: ∀h,g,a,G,L,V1,T1,U2,l. ⦃G, L⦄ ⊢ ⓛ{a}V1.T1 •*➡*[h, g, l] U2 →
52                       ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓛV1⦄ ⊢ T1 •*➡*[h, g, l] T2 &
53                                U2 = ⓛ{a}V2.T2.
54 #h #g #a #G #L #V1 #T1 #U2 #l2 * #X #l1 #Hl21 #Hl1 #H1 #H2
55 lapply (da_inv_bind … Hl1) -Hl1 #Hl1
56 elim (lstas_inv_bind1 … H1) -H1 #U #HTU1 #H destruct
57 elim (cprs_inv_abst1 … H2) -H2 #V2 #T2 #HV12 #HUT2 #H destruct
58 /3 width=6 by ex4_2_intro, ex3_2_intro/
59 qed-.
60
61 lemma cpds_inv_abbr_abst: ∀h,g,a1,a2,G,L,V1,W2,T1,T2,l. ⦃G, L⦄ ⊢ ⓓ{a1}V1.T1 •*➡*[h, g, l] ⓛ{a2}W2.T2 →
62                           ∃∃T. ⦃G, L.ⓓV1⦄ ⊢ T1 •*➡*[h, g, l] T & ⇧[0, 1] ⓛ{a2}W2.T2 ≡ T & a1 = true.
63 #h #g #a1 #a2 #G #L #V1 #W2 #T1 #T2 #l2 * #X #l1 #Hl21 #Hl1 #H1 #H2
64 lapply (da_inv_bind … Hl1) -Hl1 #Hl1
65 elim (lstas_inv_bind1 … H1) -H1 #U1 #HTU1 #H destruct
66 elim (cprs_inv_abbr1 … H2) -H2 *
67 [ #V2 #U2 #HV12 #HU12 #H destruct
68 | /3 width=6 by ex4_2_intro, ex3_intro/
69 ]
70 qed-.
71
72 lemma cpds_inv_lstas_eq: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T2 →
73                          ∀T. ⦃G, L⦄ ⊢ T1 •*[h, l] T → ⦃G, L⦄ ⊢ T ➡* T2.
74 #h #g #G #L #T1 #T2 #l2 *
75 #T0 #l1 #_ #_ #HT10 #HT02 #T #HT1
76 lapply (lstas_mono … HT10 … HT1) #H destruct //
77 qed-.
78
79 (* Advanced forward lemmas **************************************************)
80
81 lemma cpds_fwd_cpxs: ∀h,g,G,L,T1,T2,l. ⦃G, L⦄ ⊢ T1 •*➡*[h, g, l] T2 → ⦃G, L⦄ ⊢ T1 ➡*[h, g] T2.
82 #h #g #G #L #T1 #T2 #l * /3 width=5 by cpxs_trans, lstas_cpxs, cprs_cpxs/
83 qed-.
84
85 (* Main properties **********************************************************)
86
87 theorem cpds_conf_eq: ∀h,g,G,L,T0,T1,l. ⦃G, L⦄ ⊢ T0 •*➡*[h, g, l] T1 →
88                       ∀T2. ⦃G, L⦄ ⊢ T0 •*➡*[h, g, l] T2 →
89                       ∃∃T. ⦃G, L⦄ ⊢ T1 ➡* T & ⦃G, L⦄ ⊢ T2 ➡* T.
90 #h #g #G #L #T0 #T1 #l0 * #U1 #l1 #_ #_ #H1 #HUT1 #T2 * #U2 #l2 #_ #_ #H2 #HUT2 -l1 -l2
91 lapply (lstas_mono … H1 … H2) #H destruct -h -l0 /2 width=3 by cprs_conf/
92 qed-.