]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda-delta/reduction/pr_defs.ma
- new definition of subst based on drop
[helm.git] / matita / matita / lib / lambda-delta / reduction / pr_defs.ma
1 (*
2     ||M||  This file is part of HELM, an Hypertextual, Electronic
3     ||A||  Library of Mathematics, developed at the Computer Science
4     ||T||  Department of the University of Bologna, Italy.
5     ||I||
6     ||T||
7     ||A||  This file is distributed under the terms of the
8     \   /  GNU General Public License Version 2
9      \ /
10       V_______________________________________________________________ *)
11
12 include "lambda-delta/substitution/drop_defs.ma".
13
14 (* SINGLE STEP PARALLEL REDUCTION ON TERMS **********************************)
15
16 inductive pr: lenv → term → term → Prop ≝
17 | pr_sort : ∀L,k. pr L (⋆k) (⋆k)
18 | pr_lref : ∀L,i. pr L (#i) (#i)
19 | pr_bind : ∀L,I,V1,V2,T1,T2. pr L V1 V2 → pr (L. 𝕓{I} V1) T1 T2 →
20             pr L (𝕓{I} V1. T1) (𝕓{I} V2. T2)
21 | pr_flat : ∀L,I,V1,V2,T1,T2. pr L V1 V2 → pr L T1 T2 →
22             pr L (𝕗{I} V1. T1) (𝕗{I} V2. T2)
23 | pr_beta : ∀L,V1,V2,W,T1,T2.
24             pr L V1 V2 → pr (L. 𝕓{Abst} W) T1 T2 → (*𝕓*)
25             pr L (𝕚{Appl} V1. 𝕚{Abst} W. T1) (𝕚{Abbr} V2. T2)
26 | pr_delta: ∀L,K,V1,V2,V,i.
27             ↑[0,i] K. 𝕓{Abbr} V1 ≡ L → pr K V1 V2 → ↑[0,i+1] V2 ≡ V →
28             pr L (#i) V
29 | pr_theta: ∀L,V,V1,V2,W1,W2,T1,T2.
30             pr L V1 V2 → ↑[0,1] V2 ≡ V → pr L W1 W2 → pr (L. 𝕓{Abbr} W1) T1 T2 → (*𝕓*)
31             pr L (𝕚{Appl} V1. 𝕚{Abbr} W1. T1) (𝕚{Abbr} W2. 𝕚{Appl} V. T2)
32 | pr_zeta : ∀L,V,T,T1,T2. ↑[0,1] T1 ≡ T → pr L T1 T2 →
33             pr L (𝕚{Abbr} V. T) T2
34 | pr_tau  : ∀L,V,T1,T2. pr L T1 T2 → pr L (𝕚{Cast} V. T1) T2
35 .
36
37 interpretation
38    "single step parallel reduction (term)"
39    'PR L T1 T2 = (pr L T1 T2).
40
41 (* The basic properties *****************************************************)
42
43 lemma pr_refl: ∀T,L. L ⊢ T ⇒ T.
44 #T elim T -T //
45 #I elim I -I /2/
46 qed.