]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/contribs/lambda-delta/Basic-2/reduction/cpr.ma
- confluence of parallel substitution (tps) closed! (a bug in an
[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/reduction/tpr.ma".
16
17 (* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************)
18
19 definition cpr: lenv → term → term → Prop ≝
20    λL,T1,T2. ∃∃T. T1 ⇒ T & L ⊢ T [0, |L|] ≫ T2.
21
22 interpretation
23    "context-sensitive parallel reduction (term)"
24    'PRed L T1 T2 = (cpr L T1 T2).
25
26 (* Basic properties *********************************************************)
27
28 lemma cpr_pr: ∀T1,T2. T1 ⇒ T2 → ∀L. L ⊢ T1 ⇒ T2.
29 /2/ qed.
30
31 lemma cpr_tps: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ≫ T2 → L ⊢ T1 ⇒ T2.
32 /3 width=5/ qed.
33
34 lemma cpr_refl: ∀L,T. L ⊢ T ⇒ T.
35 /2/ qed.
36
37 (* NOTE: new property *)
38 lemma cpr_flat: ∀I,L,V1,V2,T1,T2.
39                 L ⊢ V1 ⇒ V2 → L ⊢ T1 ⇒ T2 → L ⊢ 𝕗{I} V1. T1 ⇒ 𝕗{I} V2. T2.
40 #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
41 qed.
42
43 lemma cpr_delta: ∀L,K,V,W,i.
44                  ↓[0, i] L ≡ K. 𝕓{Abbr} V → ↑[0, i + 1] V ≡ W → L ⊢ #i ⇒ W.
45 /3/
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 (* Basic inversion lemmas ***************************************************)