1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "Basic_2/grammar/cl_shift.ma".
16 include "Basic_2/unfold/tpss.ma".
17 include "Basic_2/reducibility/tpr.ma".
19 (* CONTEXT-SENSITIVE PARALLEL REDUCTION ON TERMS ****************************)
21 (* Basic_1: includes: pr2_delta1 *)
22 definition cpr: lenv → relation term ≝
23 λL,T1,T2. ∃∃T. T1 ➡ T & L ⊢ T [0, |L|] ▶* T2.
26 "context-sensitive parallel reduction (term)"
27 'PRed L T1 T2 = (cpr L T1 T2).
29 (* Basic properties *********************************************************)
31 lemma cpr_intro: ∀L,T1,T,T2,d,e. T1 ➡ T → L ⊢ T [d, e] ▶* T2 → L ⊢ T1 ➡ T2.
34 (* Basic_1: was by definition: pr2_free *)
35 lemma cpr_pr: ∀T1,T2. T1 ➡ T2 → ∀L. L ⊢ T1 ➡ T2.
38 lemma cpr_tpss: ∀L,T1,T2,d,e. L ⊢ T1 [d, e] ▶* T2 → L ⊢ T1 ➡ T2.
41 lemma cpr_refl: ∀L,T. L ⊢ T ➡ T.
44 (* Note: new property *)
45 (* Basic_1: was only: pr2_thin_dx *)
46 lemma cpr_flat: ∀I,L,V1,V2,T1,T2.
47 L ⊢ V1 ➡ V2 → L ⊢ T1 ➡ T2 → L ⊢ ⓕ{I} V1. T1 ➡ ⓕ{I} V2. T2.
48 #I #L #V1 #V2 #T1 #T2 * #V #HV1 #HV2 * /3 width=5/
51 lemma cpr_cast: ∀L,V,T1,T2.
52 L ⊢ T1 ➡ T2 → L ⊢ ⓣV. T1 ➡ T2.
53 #L #V #T1 #T2 * /3 width=3/
56 (* Note: it does not hold replacing |L1| with |L2| *)
57 (* Basic_1: was only: pr2_change *)
58 lemma cpr_lsubs_conf: ∀L1,T1,T2. L1 ⊢ T1 ➡ T2 →
59 ∀L2. L1 [0, |L1|] ≼ L2 → L2 ⊢ T1 ➡ T2.
60 #L1 #T1 #T2 * #T #HT1 #HT2 #L2 #HL12
61 lapply (tpss_lsubs_conf … HT2 … HL12) -HT2 -HL12 /3 width=4/
64 (* Basic inversion lemmas ***************************************************)
66 (* Basic_1: was: pr2_gen_csort *)
67 lemma cpr_inv_atom: ∀T1,T2. ⋆ ⊢ T1 ➡ T2 → T1 ➡ T2.
68 #T1 #T2 * #T #HT normalize #HT2
69 <(tpss_inv_refl_O2 … HT2) -HT2 //
72 (* Basic_1: was: pr2_gen_sort *)
73 lemma cpr_inv_sort1: ∀L,T2,k. L ⊢ ⋆k ➡ T2 → T2 = ⋆k.
75 >(tpr_inv_atom1 … H) -H #H
76 >(tpss_inv_sort1 … H) -H //
79 (* Basic_1: was pr2_gen_abbr *)
80 lemma cpr_inv_abbr1: ∀L,V1,T1,U2. L ⊢ ⓓV1. T1 ➡ U2 →
81 (∃∃V,V2,T2. V1 ➡ V & L ⊢ V [O, |L|] ▶* V2 &
85 ∃∃T. ⇧[0,1] T ≡ T1 & L ⊢ T ➡ U2.
86 #L #V1 #T1 #Y * #X #H1 #H2
87 elim (tpr_inv_abbr1 … H1) -H1 *
88 [ #V #T0 #T #HV1 #HT10 #HT0 #H destruct
89 elim (tpss_inv_bind1 … H2) -H2 #V2 #T2 #HV2 #HT2 #H destruct
90 lapply (tps_lsubs_conf … HT0 (L. ⓓV) ?) -HT0 /2 width=1/ #HT0
91 lapply (tps_weak_all … HT0) -HT0 #HT0
92 lapply (tpss_lsubs_conf … HT2 (L. ⓓV) ?) -HT2 /2 width=1/ #HT2
93 lapply (tpss_weak_all … HT2) -HT2 #HT2
94 lapply (tpss_strap … HT0 HT2) -T /4 width=7/
99 (* Basic_1: was: pr2_gen_cast *)
100 lemma cpr_inv_cast1: ∀L,V1,T1,U2. L ⊢ ⓣV1. T1 ➡ U2 → (
101 ∃∃V2,T2. L ⊢ V1 ➡ V2 & L ⊢ T1 ➡ T2 &
104 #L #V1 #T1 #U2 * #X #H #HU2
105 elim (tpr_inv_cast1 … H) -H /3 width=3/
106 * #V #T #HV1 #HT1 #H destruct
107 elim (tpss_inv_flat1 … HU2) -HU2 #V2 #T2 #HV2 #HT2 #H destruct /4 width=5/
110 (* Basic_1: removed theorems 5:
111 pr2_head_1 pr2_head_2 pr2_cflat pr2_gen_cflat clear_pr2_trans
112 Basic_1: removed local theorems 3:
113 pr2_free_free pr2_free_delta pr2_delta_delta