]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/computation/cprs_cprs.ma
ca3fe09da3316f6e9c6066f10f4beb06710cc0d9
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / cprs_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/reducibility/cpr_lift.ma".
16 include "basic_2/reducibility/cpr_cpr.ma".
17 include "basic_2/reducibility/lcpr_cpr.ma".
18 include "basic_2/computation/cprs_lcpr.ma".
19
20 (* CONTEXT-SENSITIVE PARALLEL COMPUTATION ON TERMS **************************)
21
22 (* Advanced properties ******************************************************)
23
24 lemma cprs_abbr_dx: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 →
25                     L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
26 #L #V1 #V2 #HV12 #T1 #T2 #HT12 @(cprs_ind_dx … HT12) -T1
27 [ /3 width=5/
28 | #T1 #T #HT1 #_ #IHT1
29   @(cprs_strap2 … IHT1) -IHT1 /2 width=1/
30 ]
31 qed.
32
33 lemma cpr_abbr: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV1 ⊢ T1 ➡ T2 →
34                 L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
35 /3 width=1/ qed.
36
37 lemma cpr_abbr2: ∀L,V1,V2. L ⊢ V1 ➡ V2 → ∀T1,T2. L. ⓓV2 ⊢ T1 ➡ T2 →
38                  L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
39 #L #V1 #V2 #HV12 #T1 #T2 #HT12
40 lapply (lcpr_cpr_trans (L. ⓓV1) … HT12) /2 width=1/
41 qed.
42
43 (* Advanced inversion lemmas ************************************************)
44
45 (* Basic_1: was pr3_gen_appl *)
46 lemma cprs_inv_appl1: ∀L,V1,T1,U2. L ⊢ ⓐV1. T1 ➡* U2 →
47                       ∨∨ ∃∃V2,T2.     L ⊢ V1 ➡* V2 & L ⊢ T1 ➡* T2 &
48                                       U2 = ⓐV2. T2
49                        | ∃∃V2,W,T.    L ⊢ V1 ➡* V2 &
50                                       L ⊢ T1 ➡* ⓛW. T & L ⊢ ⓓV2. T ➡* U2
51                        | ∃∃V0,V2,V,T. L ⊢ V1 ➡* V0 & ⇧[0,1] V0 ≡ V2 &
52                                       L ⊢ T1 ➡* ⓓV. T & L ⊢ ⓓV. ⓐV2. T ➡* U2.
53 #L #V1 #T1 #U2 #H @(cprs_ind … H) -U2 /3 width=5/
54 #U #U2 #_ #HU2 * *
55 [ #V0 #T0 #HV10 #HT10 #H destruct
56   elim (cpr_inv_appl1 … HU2) -HU2 *
57   [ #V2 #T2 #HV02 #HT02 #H destruct /4 width=5/
58   | #V2 #W2 #T #T2 #HV02 #HT2 #H1 #H2 destruct /4 width=7/
59   | #V #V2 #W0 #W2 #T #T2 #HV0 #HW02 #HT2 #HV2 #H1 #H2 destruct
60     @or3_intro2 @(ex4_4_intro … HV2 HT10) /2 width=3/ /3 width=1/ (**) (* explicit constructor. /5 width=8/ is too slow because TC_transitive gets in the way *) 
61   ]
62 | /4 width=8/
63 | /4 width=10/
64 ]
65 qed-.
66
67 (* Main propertis ***********************************************************)
68
69 (* Basic_1: was: pr3_confluence *)
70 theorem cprs_conf: ∀L,T1,T. L ⊢ T ➡* T1 → ∀T2. L ⊢ T ➡* T2 →
71                    ∃∃T0. L ⊢ T1 ➡* T0 & L ⊢ T2 ➡* T0.
72 /3 width=3/ qed.
73
74 (* Basic_1: was: pr3_t *)
75 theorem cprs_trans: ∀L,T1,T. L ⊢ T1 ➡* T → ∀T2. L ⊢ T ➡* T2 → L ⊢ T1 ➡* T2.
76 /2 width=3/ qed.
77
78 (* Basic_1: was: pr3_flat *)
79 lemma cprs_flat: ∀I,L,T1,T2. L ⊢ T1 ➡* T2 → ∀V1,V2. L ⊢ V1 ➡* V2 →
80                  L ⊢ ⓕ{I} V1. T1 ➡* ⓕ{I} V2. T2.
81 #I #L #T1 #T2 #HT12 #V1 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/
82 #V #V2 #_ #HV2 #IHV1
83 @(cprs_trans … IHV1) -IHV1 /2 width=1/ 
84 qed.
85
86 lemma cprs_abbr: ∀L,V1,T1,T2. L. ⓓV1 ⊢ T1 ➡* T2 → ∀V2. L ⊢ V1 ➡* V2 →
87                  L ⊢ ⓓV1. T1 ➡* ⓓV2. T2.
88 #L #V1 #T1 #T2 #HT12 #V2 #HV12 @(cprs_ind … HV12) -V2 /2 width=1/
89 #V #V2 #_ #HV2 #IHV1
90 @(cprs_trans … IHV1) -IHV1 /2 width=1/
91 qed.
92
93 (* Basic_1: was only: pr3_pr2_pr3_t pr3_wcpr0_t *)
94 lemma lcpr_cprs_trans: ∀L1,L2. L1 ⊢ ➡ L2 →
95                        ∀T1,T2. L2 ⊢ T1 ➡* T2 → L1 ⊢ T1 ➡* T2.
96 #L1 #L2 #HL12 #T1 #T2 #H @(cprs_ind … H) -T2 //
97 #T #T2 #_ #HT2 #IHT2
98 @(cprs_trans … IHT2) /2 width=3/
99 qed.