]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2/computation/lprs_cprs.ma
partial commit: just the components before "static" ...
[helm.git] / matita / matita / contribs / lambdadelta / basic_2 / computation / lprs_cprs.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/computation/cprs_cprs.ma".
16 include "basic_2/computation/lprs.ma".
17
18 (* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
19
20 (* Advanced properties ******************************************************)
21
22 lemma lprs_pair: ∀I,L1,L2. L1 ⊢ ➡* L2 → ∀V1,V2. L1 ⊢ V1 ➡* V2 →
23                  L1. ⓑ{I} V1 ⊢ ➡* L2.ⓑ{I} V2.
24 /2 width=1 by TC_lpx_sn_pair/ qed.
25
26 (* Advanced inversion lemmas ************************************************)
27
28 lemma lprs_inv_pair1: ∀I,K1,L2,V1. K1. ⓑ{I} V1 ⊢ ➡* L2 →
29                       ∃∃K2,V2. K1 ⊢ ➡* K2 & K1 ⊢ V1 ➡* V2 & L2 = K2. ⓑ{I} V2.
30 /3 width=3 by TC_lpx_sn_inv_pair1, lpr_cprs_trans/ qed-.
31
32 lemma lprs_inv_pair2: ∀I,L1,K2,V2. L1 ⊢ ➡* K2. ⓑ{I} V2 →
33                       ∃∃K1,V1. K1 ⊢ ➡* K2 & K1 ⊢ V1 ➡* V2 & L1 = K1. ⓑ{I} V1.
34 /3 width=3 by TC_lpx_sn_inv_pair2, lpr_cprs_trans/ qed-.
35
36 (* Properties on context-sensitive parallel computation for terms ***********)
37
38 lemma lprs_cpr_trans: s_r_trans … cpr lprs.
39 /3 width=5 by s_r_trans_TC2, lpr_cprs_trans/ qed-.
40
41 (* Basic_1: was just: pr3_pr3_pr3_t *)
42 lemma lprs_cprs_trans: s_rs_trans … cpr lprs.
43 /3 width=5 by s_r_trans_TC1, lprs_cpr_trans/ qed-.
44
45 lemma lprs_cprs_conf_dx: ∀L0,T0,T1. L0 ⊢ T0 ➡* T1 → ∀L1. L0 ⊢ ➡* L1 →
46                          ∃∃T. L1 ⊢ T1 ➡* T & L1 ⊢ T0 ➡* T.
47 #L0 #T0 #T1 #HT01 #L1 #H elim H -L1
48 [ #L1 #HL01
49   elim (cprs_lpr_conf_dx … HT01 … HL01) -L0 /2 width=3/
50 | #L #L1 #_ #HL1 * #T #HT1 #HT0 -L0
51   elim (cprs_lpr_conf_dx … HT1 … HL1) -HT1 #T2 #HT2 #HT12
52   elim (cprs_lpr_conf_dx … HT0 … HL1) -L #T3 #HT3 #HT03
53   elim (cprs_conf … HT2 … HT3) -T #T #HT2 #HT3
54   lapply (cprs_trans … HT03 … HT3) -T3
55   lapply (cprs_trans … HT12 … HT2) -T2 /2 width=3/
56 ]
57 qed-.
58
59 lemma lprs_cpr_conf_dx: ∀L0,T0,T1. L0 ⊢ T0 ➡ T1 → ∀L1. L0 ⊢ ➡* L1 →
60                         ∃∃T. L1 ⊢ T1 ➡* T & L1 ⊢ T0 ➡* T.
61 /3 width=3 by lprs_cprs_conf_dx, cpr_cprs/ qed-.
62
63 lemma lprs_cprs_conf_sn: ∀L0,T0,T1. L0 ⊢ T0 ➡* T1 → ∀L1. L0 ⊢ ➡* L1 →
64                          ∃∃T. L0 ⊢ T1 ➡* T & L1 ⊢ T0 ➡* T.
65 #L0 #T0 #T1 #HT01 #L1 #HL01
66 elim (lprs_cprs_conf_dx … HT01 … HL01) -HT01 #T #HT1
67 lapply (lprs_cprs_trans … HT1 … HL01) -HT1 /2 width=3/
68 qed-.
69
70 lemma lprs_cpr_conf_sn: ∀L0,T0,T1. L0 ⊢ T0 ➡ T1 → ∀L1. L0 ⊢ ➡* L1 →
71                         ∃∃T. L0 ⊢ T1 ➡* T & L1 ⊢ T0 ➡* T.
72 /3 width=3 by lprs_cprs_conf_sn, cpr_cprs/ qed-.
73
74 lemma cprs_bind2: ∀L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 → ∀I,T1,T2. L. ⓑ{I}V2 ⊢ T1 ➡* T2 →
75                   ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1. T1 ➡* ⓑ{a,I}V2. T2.
76 #L #V1 #V2 #HV12 #I #T1 #T2 #HT12
77 lapply (lprs_cprs_trans … HT12 (L.ⓑ{I}V1) ?) /2 width=1/
78 qed.
79
80 (* Inversion lemmas on context-sensitive parallel computation for terms *****)
81
82 (* Basic_1: was: pr3_gen_abst *)
83 lemma cprs_inv_abst1: ∀a,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* U2 →
84                       ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & L.ⓛW1 ⊢ T1 ➡* T2 &
85                                U2 = ⓛ{a}W2.T2.
86 #a #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5/
87 #U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
88 elim (cpr_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
89 lapply (lprs_cpr_trans … HT02 (L.ⓛV1) ?) /2 width=1/ -HT02 #HT02
90 lapply (cprs_strap1 … HV10 … HV02) -V0
91 lapply (cprs_trans … HT10 … HT02) -T0 /2 width=5/
92 qed-.
93
94 lemma cprs_inv_abst: ∀a,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 →
95                      ⦃G, L⦄ ⊢ W1 ➡* W2 ∧ L.ⓛW1 ⊢ T1 ➡* T2.
96 #a #L #W1 #W2 #T1 #T2 #H
97 elim (cprs_inv_abst1 … H) -H #W #T #HW1 #HT1 #H destruct /2 width=1/
98 qed-.
99
100 (* Basic_1: was pr3_gen_abbr *)
101 lemma cprs_inv_abbr1: ∀a,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡* U2 → (
102                       ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & L. ⓓV1 ⊢ T1 ➡* T2 &
103                                U2 = ⓓ{a}V2.T2
104                       ) ∨
105                       ∃∃T2. L. ⓓV1 ⊢ T1 ➡* T2 & ⇧[0, 1] U2 ≡ T2 & a = true.
106 #a #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/
107 #U0 #U2 #_ #HU02 * *
108 [ #V0 #T0 #HV10 #HT10 #H destruct
109   elim (cpr_inv_abbr1 … HU02) -HU02 *
110   [ #V2 #T2 #HV02 #HT02 #H destruct
111     lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) /2 width=1/ -HT02 #HT02
112     lapply (cprs_strap1 … HV10 … HV02) -V0
113     lapply (cprs_trans … HT10 … HT02) -T0 /3 width=5/
114   | #T2 #HT02 #HUT2
115     lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) -HT02 /2 width=1/ -V0 #HT02
116     lapply (cprs_trans … HT10 … HT02) -T0 /3 width=3/
117   ]
118 | #U1 #HTU1 #HU01
119   elim (lift_total U2 0 1) #U #HU2
120   lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0 /2 width=1/ /4 width=3/
121 ]
122 qed-.
123
124 (* More advanced properties *************************************************)
125
126 lemma lprs_pair2: ∀I,L1,L2. L1 ⊢ ➡* L2 → ∀V1,V2. L2 ⊢ V1 ➡* V2 →
127                   L1. ⓑ{I} V1 ⊢ ➡* L2. ⓑ{I} V2.
128 /3 width=3 by lprs_pair, lprs_cprs_trans/ qed.