]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambdadelta/basic_2A/computation/lprs_cprs.ma
update in lambdadelta
[helm.git] / matita / matita / contribs / lambdadelta / basic_2A / 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_2A/computation/cprs_cprs.ma".
16 include "basic_2A/computation/lprs.ma".
17
18 (* SN PARALLEL COMPUTATION ON LOCAL ENVIRONMENTS ****************************)
19
20 (* Advanced properties ******************************************************)
21
22 lemma lprs_pair: ∀I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 →
23                  ∀V1,V2. ⦃G, L1⦄ ⊢ V1 ➡* V2 → ⦃G, 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,G,K1,L2,V1. ⦃G, K1.ⓑ{I}V1⦄ ⊢ ➡* L2 →
29                       ∃∃K2,V2. ⦃G, K1⦄ ⊢ ➡* K2 & ⦃G, K1⦄ ⊢ V1 ➡* V2 &
30                                L2 = K2.ⓑ{I}V2.
31 /3 width=3 by TC_lpx_sn_inv_pair1, lpr_cprs_trans/ qed-.
32
33 lemma lprs_inv_pair2: ∀I,G,L1,K2,V2. ⦃G, L1⦄ ⊢ ➡* K2.ⓑ{I}V2 →
34                       ∃∃K1,V1. ⦃G, K1⦄ ⊢ ➡* K2 & ⦃G, K1⦄ ⊢ V1 ➡* V2 &
35                                L1 = K1.ⓑ{I}V1.
36 /3 width=3 by TC_lpx_sn_inv_pair2, lpr_cprs_trans/ qed-.
37
38 (* Advanced eliminators *****************************************************)
39
40 lemma lprs_ind_alt: ∀G. ∀R:relation lenv.
41                     R (⋆) (⋆) → (
42                        ∀I,K1,K2,V1,V2.
43                        ⦃G, K1⦄ ⊢ ➡* K2 → ⦃G, K1⦄ ⊢ V1 ➡* V2 →
44                        R K1 K2 → R (K1.ⓑ{I}V1) (K2.ⓑ{I}V2)
45                     ) →
46                     ∀L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 → R L1 L2.
47 /3 width=4 by TC_lpx_sn_ind, lpr_cprs_trans/ qed-.
48
49 (* Properties on context-sensitive parallel computation for terms ***********)
50
51 lemma lprs_cpr_trans: ∀G. s_r_transitive … (cpr G) (λ_. lprs G).
52 /3 width=5 by s_r_trans_LTC2, lpr_cprs_trans/ qed-.
53
54 (* Basic_1: was just: pr3_pr3_pr3_t *)
55 (* Note: alternative proof /3 width=5 by s_r_trans_LTC1, lprs_cpr_trans/ *)
56 lemma lprs_cprs_trans: ∀G. s_rs_transitive … (cpr G) (λ_. lprs G).
57 #G @s_r_to_s_rs_trans @s_r_trans_LTC2
58 @s_rs_trans_TC1 /2 width=3 by lpr_cprs_trans/ (**) (* full auto too slow *)
59 qed-.
60
61 lemma lprs_cprs_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
62                          ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
63                          ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
64 #G #L0 #T0 #T1 #HT01 #L1 #H @(lprs_ind … H) -L1 /2 width=3 by ex2_intro/
65 #L #L1 #_ #HL1 * #T #HT1 #HT0 -L0
66 elim (cprs_lpr_conf_dx … HT1 … HL1) -HT1 #T2 #HT2
67 elim (cprs_lpr_conf_dx … HT0 … HL1) -L #T3 #HT3
68 elim (cprs_conf … HT2 … HT3) -T
69 /3 width=5 by cprs_trans, ex2_intro/
70 qed-.
71
72 lemma lprs_cpr_conf_dx: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
73                         ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
74                         ∃∃T. ⦃G, L1⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
75 /3 width=3 by lprs_cprs_conf_dx, cpr_cprs/ qed-.
76
77 (* Note: this can be proved on its own using lprs_ind_dx *)
78 lemma lprs_cprs_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡* T1 →
79                          ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
80                          ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
81 #G #L0 #T0 #T1 #HT01 #L1 #HL01
82 elim (lprs_cprs_conf_dx … HT01 … HL01) -HT01
83 /3 width=3 by lprs_cprs_trans, ex2_intro/
84 qed-.
85
86 lemma lprs_cpr_conf_sn: ∀G,L0,T0,T1. ⦃G, L0⦄ ⊢ T0 ➡ T1 →
87                         ∀L1. ⦃G, L0⦄ ⊢ ➡* L1 →
88                         ∃∃T. ⦃G, L0⦄ ⊢ T1 ➡* T & ⦃G, L1⦄ ⊢ T0 ➡* T.
89 /3 width=3 by lprs_cprs_conf_sn, cpr_cprs/ qed-.
90
91 lemma cprs_bind2: ∀G,L,V1,V2. ⦃G, L⦄ ⊢ V1 ➡* V2 →
92                   ∀I,T1,T2. ⦃G, L.ⓑ{I}V2⦄ ⊢ T1 ➡* T2 →
93                   ∀a. ⦃G, L⦄ ⊢ ⓑ{a,I}V1.T1 ➡* ⓑ{a,I}V2.T2.
94 /4 width=5 by lprs_cprs_trans, lprs_pair, cprs_bind/ qed.
95
96 (* Inversion lemmas on context-sensitive parallel computation for terms *****)
97
98 (* Basic_1: was: pr3_gen_abst *)
99 lemma cprs_inv_abst1: ∀a,G,L,W1,T1,U2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* U2 →
100                       ∃∃W2,T2. ⦃G, L⦄ ⊢ W1 ➡* W2 & ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2 &
101                                U2 = ⓛ{a}W2.T2.
102 #a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /2 width=5 by ex3_2_intro/
103 #U0 #U2 #_ #HU02 * #V0 #T0 #HV10 #HT10 #H destruct
104 elim (cpr_inv_abst1 … HU02) -HU02 #V2 #T2 #HV02 #HT02 #H destruct
105 lapply (lprs_cpr_trans … HT02 (L.ⓛV1) ?)
106 /3 width=5 by lprs_pair, cprs_trans, cprs_strap1, ex3_2_intro/
107 qed-.
108
109 lemma cprs_inv_abst: ∀a,G,L,W1,W2,T1,T2. ⦃G, L⦄ ⊢ ⓛ{a}W1.T1 ➡* ⓛ{a}W2.T2 →
110                      ⦃G, L⦄ ⊢ W1 ➡* W2 ∧ ⦃G, L.ⓛW1⦄ ⊢ T1 ➡* T2.
111 #a #G #L #W1 #W2 #T1 #T2 #H elim (cprs_inv_abst1 … H) -H
112 #W #T #HW1 #HT1 #H destruct /2 width=1 by conj/
113 qed-.
114
115 (* Basic_1: was pr3_gen_abbr *)
116 lemma cprs_inv_abbr1: ∀a,G,L,V1,T1,U2. ⦃G, L⦄ ⊢ ⓓ{a}V1.T1 ➡* U2 → (
117                       ∃∃V2,T2. ⦃G, L⦄ ⊢ V1 ➡* V2 & ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 &
118                                U2 = ⓓ{a}V2.T2
119                       ) ∨
120                       ∃∃T2. ⦃G, L.ⓓV1⦄ ⊢ T1 ➡* T2 & ⬆[0, 1] U2 ≡ T2 & a = true.
121 #a #G #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5 by ex3_2_intro, or_introl/
122 #U0 #U2 #_ #HU02 * *
123 [ #V0 #T0 #HV10 #HT10 #H destruct
124   elim (cpr_inv_abbr1 … HU02) -HU02 *
125   [ #V2 #T2 #HV02 #HT02 #H destruct
126     lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?)
127     /4 width=5 by lprs_pair, cprs_trans, cprs_strap1, ex3_2_intro, or_introl/
128   | #T2 #HT02 #HUT2
129     lapply (lprs_cpr_trans … HT02 (L.ⓓV1) ?) -HT02
130     /4 width=3 by lprs_pair, cprs_trans, ex3_intro, or_intror/
131   ]
132 | #U1 #HTU1 #HU01 elim (lift_total U2 0 1)
133   #U #HU2 lapply (cpr_lift … HU02 (L.ⓓV1) … HU01 … HU2) -U0
134   /4 width=3 by cprs_strap1, drop_drop, ex3_intro, or_intror/
135 ]
136 qed-.
137
138 (* More advanced properties *************************************************)
139
140 lemma lprs_pair2: ∀I,G,L1,L2. ⦃G, L1⦄ ⊢ ➡* L2 →
141                   ∀V1,V2. ⦃G, L2⦄ ⊢ V1 ➡* V2 → ⦃G, L1.ⓑ{I}V1⦄ ⊢ ➡* L2.ⓑ{I}V2.
142 /3 width=3 by lprs_pair, lprs_cprs_trans/ qed.