]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/computation/fprs_cprs.ma
4e7a633c32e3bbf4df83e573ae1d4656bbe21ed1
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / computation / fprs_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/fpr_cpr.ma".
16 include "basic_2/computation/cprs.ma".
17 include "basic_2/computation/fprs.ma".
18
19 (* CONTEXT-FREE PARALLEL COMPUTATION ON CLOSURES ****************************)
20
21 (* Properties on context-sensitive parallel computation for terms ***********)
22
23 lemma cprs_fprs: ∀L,T1,T2. L ⊢ T1 ➡* T2 → ⦃L, T1⦄ ➡* ⦃L, T2⦄.
24 #L #T1 #T2 #H @(cprs_ind … H) -T2 // /3 width=4/
25 qed.
26 (*
27 (* Advanced propertis *******************************************************)
28
29 lamma fpr_bind_sn: ∀L1,L2,V1,V2. ⦃L1, V1⦄ ➡ ⦃L2, V2⦄ → ∀T1,T2. T1 ➡ T2 →
30                    ∀a,I. ⦃L1, ⓑ{a,I}V1.T1⦄ ➡ ⦃L2, ⓑ{a,I}V2.T2⦄.
31 #L1 #L2 #V1 #V2 #H #T1 #T2 #HT12 #a #I
32 elim (fpr_inv_all … H) /3 width=4/
33 qed.
34
35 (* Advanced forward lemmas **************************************************)
36
37 lamma fpr_fwd_shift_bind_minus: ∀I1,I2,L1,L2,V1,V2,T1,T2.
38                                 ⦃L1, -ⓑ{I1}V1.T1⦄ ➡ ⦃L2, -ⓑ{I2}V2.T2⦄ →
39                                 ⦃L1, V1⦄ ➡ ⦃L2, V2⦄ ∧ I1 = I2.
40 * #I2 #L1 #L2 #V1 #V2 #T1 #T2 #H
41 elim (fpr_inv_all … H) -H #L #HL1 #H #HL2
42 [ elim (cpr_inv_abbr1 … H) -H *
43   [ #V #V0 #T #HV1 #HV0 #_ #H destruct /4 width=4/
44   | #T #_ #_ #H destruct
45   ]
46 | elim (cpr_inv_abst1 … H Abst V2) -H
47   #V #T #HV1 #_ #H destruct /3 width=4/
48 ]
49 qed-.
50
51 (* Advanced inversion lemmas ************************************************)
52
53 lamma fpr_inv_pair1: ∀I,K1,L2,V1,T1,T2. ⦃K1.ⓑ{I}V1, T1⦄ ➡ ⦃L2, T2⦄ →
54                      ∃∃K2,V2. ⦃K1, V1⦄  ➡ ⦃K2, V2⦄ &
55                               ⦃K1, -ⓑ{I}V1.T1⦄ ➡ ⦃K2, -ⓑ{I}V2.T2⦄  &
56                               L2 = K2.ⓑ{I}V2.
57 #I1 #K1 #X #V1 #T1 #T2 #H
58 elim (fpr_fwd_pair1 … H) -H #I2 #K2 #V2 #HT12 #H destruct
59 elim (fpr_fwd_shift_bind_minus … HT12) #HV12 #H destruct /2 width=5/
60 qed-.
61
62 lamma fpr_inv_pair3: ∀I,L1,K2,V2,T1,T2. ⦃L1, T1⦄ ➡ ⦃K2.ⓑ{I}V2, T2⦄ →
63                      ∃∃K1,V1. ⦃K1, V1⦄  ➡ ⦃K2, V2⦄ &
64                               ⦃K1, -ⓑ{I}V1.T1⦄ ➡ ⦃K2, -ⓑ{I}V2.T2⦄  &
65                               L1 = K1.ⓑ{I}V1.
66 #I2 #X #K2 #V2 #T1 #T2 #H
67 elim (fpr_fwd_pair3 … H) -H #I1 #K1 #V1 #HT12 #H destruct
68 elim (fpr_fwd_shift_bind_minus … HT12) #HV12 #H destruct /2 width=5/
69 qed-.
70 *)