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