]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda-delta/Basic-2/reduction/cpr.ma
- the substitution lemma is proved!
[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/reduction/tpr.ma".
17
18 (* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************)
19
20 definition cpr: lenv → term → term → Prop ≝
21    λL,T1,T2. ∃∃T. T1 ⇒ T & L ⊢ T [0, |L|] ≫ T2.
22
23 interpretation
24    "context-sensitive parallel reduction (term)"
25    'PRed L T1 T2 = (cpr L T1 T2).
26
27 (* Basic properties *********************************************************)
28
29 lemma cpr_pr: ∀T1,T2. T1 ⇒ T2 → ∀L. L ⊢ T1 ⇒ T2.
30 /2/ qed.
31
32 lemma cpr_tps: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 ⇒ T2.
33 /3 width=5/ qed.
34
35 lemma cpr_refl: ∀L,T. L ⊢ T ⇒ T.
36 /2/ qed.
37
38 lemma cpr_bind_sn: ∀I,L,V1,V2,T1,T2. L ⊢ V1 ⇒ V2 → T1 ⇒ T2 →
39                    L ⊢ 𝕓{I} V1. T1 ⇒ 𝕓{I} V2. T2.
40 #I #L #V1 #V2 #T1 #T2 * /3 width=5/
41 qed.
42
43 lemma cpr_bind_dx: ∀I,L,V1,V2,T1,T2. V1 ⇒ V2 → L. 𝕓{I} V2 ⊢ T1 ⇒ T2 →
44                    L ⊢ 𝕓{I} V1. T1 ⇒ 𝕓{I} V2. T2.
45 #I #L #V1 #V2 #T1 #T2 #HV12 * #T #HT1 normalize #HT2
46 elim (tps_split_up … HT2 1 ? ?) -HT2 // #T0 <minus_n_O #HT0 normalize <minus_plus_m_m #HT02
47 lapply (tps_leq_repl_dx … HT0 (⋆. 𝕓{I} V2) ?) -HT0 /2/ #HT0
48 /3 width=5/
49 qed.
50
51 (* NOTE: new property *)
52 lemma cpr_flat: ∀I,L,V1,V2,T1,T2.
53                 L ⊢ V1 ⇒ V2 → L ⊢ T1 ⇒ T2 → L ⊢ 𝕗{I} V1. T1 ⇒ 𝕗{I} V2. T2.
54 #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
55 qed.
56
57 lemma cpr_delta: ∀L,K,V,W,i.
58                  ↓[0, i] L ≡ K. 𝕓{Abbr} V → ↑[0, i + 1] V ≡ W → L ⊢ #i ⇒ W.
59 /3/
60 qed.
61
62 lemma cpr_cast: ∀L,V,T1,T2.
63                 L ⊢ T1 ⇒ T2 → L ⊢ 𝕔{Cast} V. T1 ⇒ T2.
64 #L #V #T1 #T2 * /3/
65 qed.
66
67 (* Basic inversion lemmas ***************************************************)
68
69 lemma cpr_inv_lsort: ∀T1,T2. ⋆ ⊢ T1 ⇒ T2 → T1 ⇒ T2.
70 #T1 #T2 * #T #HT normalize #HT2
71 <(tps_inv_refl_O2 … HT2) -HT2 //
72 qed.
73
74 (* Basic forward lemmas *****************************************************)
75
76 lemma cpr_shift_fwd: ∀L,T1,T2. L ⊢ T1 ⇒ T2 → L @ T1 ⇒ L @ T2.
77 #L elim L -L
78 [ /2/
79 | normalize /3/
80 ].
81 qed.