]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/basic_2/reducibility/cpr.ma
- lambda_delta: subject reduction for nativa type assignment begins ...
[helm.git] / matita / matita / contribs / lambda_delta / basic_2 / reducibility / 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/unfold/tpss.ma".
16 include "basic_2/reducibility/tpr.ma".
17
18 (* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************)
19
20 (* Basic_1: includes: pr2_delta1 *)
21 definition cpr: lenv → relation term ≝
22    λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T ▶* [0, |L|] T2.
23
24 interpretation
25    "context-sensitive parallel reduction (term)"
26    'PRed L T1 T2 = (cpr L T1 T2).
27
28 (* Basic properties *********************************************************)
29
30 lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T ▶* [d, e] T2 → L ⊢ T1 ➡ T2.
31 /4 width=3/ qed-.
32
33 (* Basic_1: was by definition: pr2_free *)
34 lemma cpr_tpr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2.
35 /2 width=3/ qed.
36
37 lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 ▶* [d, e] T2 → L ⊢ T1 ➡ T2.
38 /3 width=5/ qed.
39
40 lemma cpr_refl: ∀L,T. L ⊢ T ➡ T.
41 /2 width=1/ qed.
42
43 (* Note: new property *)
44 (* Basic_1: was only: pr2_thin_dx *) 
45 lemma cpr_flat: ∀I,L,V1,V2,T1,T2.
46                 L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2.
47 #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
48 qed.
49
50 lemma cpr_cast: ∀L,V,T1,T2.
51                 L ⊢ T1 ➡ T2 → L ⊢ ⓝV. T1 ➡ T2.
52 #L #V #T1 #T2 * /3 width=3/
53 qed.
54
55 (* Note: it does not hold replacing |L1| with |L2| *)
56 (* Basic_1: was only: pr2_change *)
57 lemma cpr_lsubs_trans: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 →
58                        ∀L2. L2 ≼ [0, |L1|] L1 → L2 ⊢ T1 ➡ T2.
59 #L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 
60 lapply (tpss_lsubs_trans … HT2 … HL12) -HT2 -HL12 /3 width=4/
61 qed.
62
63 (* Basic inversion lemmas ***************************************************)
64
65 (* Basic_1: was: pr2_gen_csort *)
66 lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → T1 ➡ T2.
67 #T1 #T2 * #T #HT normalize #HT2
68 <(tpss_inv_refl_O2 … HT2) -HT2 //
69 qed-.
70
71 (* Basic_1: was: pr2_gen_sort *)
72 lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k.
73 #L #T2 #k * #X #H
74 >(tpr_inv_atom1 … H) -H #H
75 >(tpss_inv_sort1 … H) -H //
76 qed-.
77
78 (* Basic_1: was pr2_gen_abbr *)
79 lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 →
80                      (∃∃V,V2,T2. V1 ➡ V & L ⊢ V ▶* [O, |L|] V2 &
81                                  L. ⓓV ⊢ T1 ➡ T2 &
82                                  U2 = ⓓV2. T2
83                       ) ∨
84                       ∃∃T. ⇧[0,1] T ≡ T1 & L ⊢ T ➡ U2.
85 #L #V1 #T1 #Y * #X #H1 #H2
86 elim (tpr_inv_abbr1 … H1) -H1 *
87 [ #V #T0 #T #HV1 #HT10 #HT0 #H destruct
88   elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct
89   lapply (tps_lsubs_trans … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0
90   lapply (tps_weak_all … HT0) -HT0 #HT0
91   lapply (tpss_lsubs_trans … HT2 (L. ⓓV) ?) -HT2 /2 width=1/ #HT2
92   lapply (tpss_weak_all … HT2) -HT2 #HT2
93   lapply (tpss_strap … HT0 HT2) -T /4 width=7/
94 | /4 width=5/
95 ]
96 qed-.
97
98 (* Basic_1: was: pr2_gen_cast *)
99 lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓝV1. T1 ➡ U2 → (
100                         ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 &
101                                  U2 = ⓝV2. T2
102                      ) ∨ L ⊢ T1 ➡ U2.
103 #L #V1 #T1 #U2 * #X #H #HU2
104 elim (tpr_inv_cast1 … H) -H /3 width=3/
105 * #V #T #HV1 #HT1 #H destruct
106 elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/
107 qed-.
108
109 (* Basic_1: removed theorems 6: 
110             pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans
111             pr2_gen_ctail pr2_ctail
112    Basic_1: removed local theorems 3:
113             pr2_free_free pr2_free_delta pr2_delta_delta
114 *)