]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda_delta/Basic_2/reduction/cpr.ma
e70b9c537a2bf75271e51569e318c05aa0ce572b
[helm.git] / matita / matita / contribs / lambda_delta / Basic_2 / reduction / 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/grammar/cl_shift.ma".
16 include "Basic_2/unfold/tpss.ma".
17 include "Basic_2/reduction/tpr.ma".
18
19 (* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************)
20
21 (* Basic_1: includes: pr2_delta1 *)
22 definition cpr: lenv → relation term ≝
23    λL,T1,T2. ∃∃T. T1 ⇒ T & L ⊢ T [0, |L|] ≫* T2.
24
25 interpretation
26    "context-sensitive parallel reduction (term)"
27    'PRed L T1 T2 = (cpr L T1 T2).
28
29 (* Basic properties *********************************************************)
30
31 (* Basic_1: was by definition: pr2_free *)
32 lemma cpr_pr: ∀T1,T2. T1 ⇒ T2 → ∀L. L ⊢ T1 ⇒ T2.
33 /2/ qed.
34
35 lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫* T2 → L ⊢ T1 ⇒ T2.
36 /3 width=5/ qed.
37
38 lemma cpr_refl: ∀L,T. L ⊢ T ⇒ T.
39 /2/ qed.
40
41 (* Note: new property *)
42 (* Basic_1: was only: pr2_thin_dx *) 
43 lemma cpr_flat: ∀I,L,V1,V2,T1,T2.
44                 L ⊢ V1 ⇒ V2 → L ⊢ T1 ⇒ T2 → L ⊢ 𝕗{I} V1. T1 ⇒ 𝕗{I} V2. T2.
45 #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
46 qed.
47
48 lemma cpr_cast: ∀L,V,T1,T2.
49                 L ⊢ T1 ⇒ T2 → L ⊢ 𝕔{Cast} V. T1 ⇒ T2.
50 #L #V #T1 #T2 * /3/
51 qed.
52
53 (* Note: it does not hold replacing |L1| with |L2| *)
54 (* Basic_1: was only: pr2_change *)
55 lemma cpr_lsubs_conf: ∀L1,T1,T2. L1 ⊢ T1 ⇒ T2 →
56                       ∀L2. L1 [0, |L1|] ≼ L2 → L2 ⊢ T1 ⇒ T2.
57 #L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12 
58 lapply (tpss_lsubs_conf … HT2 … HL12) -HT2 HL12 /3/
59 qed.
60
61 (* Basic inversion lemmas ***************************************************)
62
63 (* Basic_1: was: pr2_gen_csort *)
64 lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ⇒ T2 → T1 ⇒ T2.
65 #T1 #T2 * #T #HT normalize #HT2
66 <(tpss_inv_refl_O2 … HT2) -HT2 //
67 qed.
68
69 (* Basic_1: was: pr2_gen_sort *)
70 lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ⇒ T2 → T2 = ⋆k.
71 #L #T2 #k * #X #H
72 >(tpr_inv_atom1 … H) -H #H
73 >(tpss_inv_sort1 … H) -H //
74 qed.
75
76 (* Basic_1: was: pr2_gen_cast *)
77 lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ 𝕔{Cast} V1. T1 ⇒ U2 → (
78                         ∃∃V2,T2. L ⊢ V1 ⇒ V2 & L ⊢ T1 ⇒ T2 &
79                                  U2 = 𝕔{Cast} V2. T2
80                      ) ∨ L ⊢ T1 ⇒ U2.
81 #L #V1 #T1 #U2 * #X #H #HU2
82 elim (tpr_inv_cast1 … H) -H /3/
83 * #V #T #HV1 #HT1 #H destruct -X;
84 elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct -U2 /4 width=5/
85 qed.
86
87 (* Basic_1: removed theorems 5: 
88             pr2_head_1 pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans
89    Basic_1: removed local theorems 3:
90             pr2_free_free pr2_free_delta pr2_delta_delta
91 *)
92
93 (*
94 pr2/fwd pr2_gen_appl
95 pr2/fwd pr2_gen_abbr
96 *)