]> matita.cs.unibo.it Git - helm.git/blob - matita/matita/lib/lambda-delta/reduction/tpr_defs.ma
- inversion lemmas for tpr completed!
[helm.git] / matita / matita / lib / lambda-delta / reduction / tpr_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/ps_defs.ma".
13
14 (* CONTEXT-FREE PARALLEL REDUCTION ON TERMS *********************************)
15
16 inductive tpr: term → term → Prop ≝
17 | tpr_sort : ∀k. tpr (⋆k) (⋆k)
18 | tpr_lref : ∀i. tpr (#i) (#i)
19 | tpr_bind : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 →
20              tpr (𝕓{I} V1. T1) (𝕓{I} V2. T2)
21 | tpr_flat : ∀I,V1,V2,T1,T2. tpr V1 V2 → tpr T1 T2 →
22              tpr (𝕗{I} V1. T1) (𝕗{I} V2. T2)
23 | tpr_beta : ∀V1,V2,W,T1,T2.
24              tpr V1 V2 → tpr T1 T2 →
25              tpr (𝕚{Appl} V1. 𝕚{Abst} W. T1) (𝕚{Abbr} V2. T2)
26 | tpr_delta: ∀V1,V2,T1,T2,T.
27              tpr V1 V2 → tpr T1 T2 → ⋆.  𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T →
28              tpr (𝕚{Abbr} V1. T1) (𝕚{Abbr} V2. T)
29 | tpr_theta: ∀V,V1,V2,W1,W2,T1,T2.
30              tpr V1 V2 → ↑[0,1] V2 ≡ V → tpr W1 W2 → tpr T1 T2 →
31              tpr (𝕚{Appl} V1. 𝕚{Abbr} W1. T1) (𝕚{Abbr} W2. 𝕚{Appl} V. T2)
32 | tpr_zeta : ∀V,T,T1,T2. ↑[0,1] T1 ≡ T → tpr T1 T2 →
33              tpr (𝕚{Abbr} V. T) T2
34 | tpr_tau  : ∀V,T1,T2. tpr T1 T2 → tpr (𝕚{Cast} V. T1) T2
35 .
36
37 interpretation
38    "context-free parallel reduction (term)"
39    'PRed T1 T2 = (tpr T1 T2).
40
41 (* Basic properties *********************************************************)
42
43 lemma tpr_refl: ∀T. T ⇒ T.
44 #T elim T -T //
45 #I elim I -I /2/
46 qed.
47
48 (* Basic inversion lemmas ***************************************************)
49
50 lemma tpr_inv_sort1_aux: ∀U1,U2. U1 ⇒ U2 → ∀k. U1 = ⋆k → U2 = ⋆k.
51 #U1 #U2 * -U1 U2
52 [ #k0 #k #H destruct -k0 //
53 | #i #k #H destruct
54 | #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct
55 | #I #V1 #V2 #T1 #T2 #_ #_ #k #H destruct
56 | #V1 #V2 #W #T1 #T2 #_ #_ #k #H destruct
57 | #V1 #V2 #T1 #T2 #T #_ #_ #_ #k #H destruct
58 | #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #k #H destruct
59 | #V #T #T1 #T2 #_ #_ #k #H destruct
60 | #V #T1 #T2 #_ #k #H destruct
61 ]
62 qed.
63
64 lemma tpr_inv_sort1: ∀k,U2. ⋆k ⇒ U2 → U2 = ⋆k.
65 /2/ qed.
66
67 lemma tpr_inv_lref1_aux: ∀U1,U2. U1 ⇒ U2 → ∀i. U1 = #i → U2 = #i.
68 #U1 #U2 * -U1 U2
69 [ #k #i #H destruct
70 | #j #i #H destruct -j //
71 | #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct
72 | #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct
73 | #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct
74 | #V1 #V2 #T1 #T2 #T #_ #_ #_ #i #H destruct
75 | #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct
76 | #V #T #T1 #T2 #_ #_ #i #H destruct
77 | #V #T1 #T2 #_ #i #H destruct
78 ]
79 qed.
80
81 lemma tpr_inv_lref1: ∀i,U2. #i ⇒ U2 → U2 = #i.
82 /2/ qed.
83
84 lemma tpr_inv_abbr1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Abbr} V1. T1 →
85                          ∨∨ ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abbr} V2. T2
86                           | ∃∃V2,T2,T. V1 ⇒ V2 & T1 ⇒ T2 &
87                                        ⋆.  𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T &
88                                        U2 = 𝕚{Abbr} V2. T
89                           | ∃∃T. ↑[0,1] T ≡ T1 & T ⇒ U2.
90 #U1 #U2 * -U1 U2
91 [ #k #V #T #H destruct
92 | #i #V #T #H destruct
93 | #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/
94 | #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct
95 | #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct
96 | #V1 #V2 #T1 #T2 #T #HV12 #HT12 #HT2 #V0 #T0 #H destruct -V1 T1 /3 width=7/
97 | #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct
98 | #V #T #T1 #T2 #HT1 #HT12 #V0 #T0 #H destruct -V T /3/
99 | #V #T1 #T2 #_ #V0 #T0 #H destruct
100 ]
101 qed.
102
103 lemma tpr_inv_abbr1: ∀V1,T1,U2. 𝕚{Abbr} V1. T1 ⇒ U2 →
104                      ∨∨ ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abbr} V2. T2
105                       | ∃∃V2,T2,T. V1 ⇒ V2 & T1 ⇒ T2 &
106                                    ⋆.  𝕓{Abbr} V2 ⊢ T2 [0, 1] ≫ T &
107                                    U2 = 𝕚{Abbr} V2. T
108                       | ∃∃T. ↑[0,1] T ≡ T1 & tpr T U2.
109 /2/ qed.
110
111 lemma tpr_inv_abst1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Abst} V1. T1 →
112                          ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abst} V2. T2.
113 #U1 #U2 * -U1 U2
114 [ #k #V #T #H destruct
115 | #i #V #T #H destruct
116 | #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /2 width=5/
117 | #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct
118 | #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct
119 | #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct
120 | #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct
121 | #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct
122 | #V #T1 #T2 #_ #V0 #T0 #H destruct
123 ]
124 qed.
125
126 lemma tpr_inv_abst1: ∀V1,T1,U2. 𝕚{Abst} V1. T1 ⇒ U2 →
127                      ∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Abst} V2. T2.
128 /2/ qed.
129
130 lemma tpr_inv_appl1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,U0. U1 = 𝕚{Appl} V1. U0 →
131                          ∨∨ ∃∃V2,T2.            V1 ⇒ V2 & U0 ⇒ T2 &
132                                                 U2 = 𝕚{Appl} V2. T2
133                           | ∃∃V2,W,T1,T2.       V1 ⇒ V2 & T1 ⇒ T2 &
134                                                 U0 = 𝕚{Abst} W. T1 &
135                                                 U2 = 𝕓{Abbr} V2. T2
136                           | ∃∃V2,V,W1,W2,T1,T2. V1 ⇒ V2 & W1 ⇒ W2 & T1 ⇒ T2 &
137                                                 ↑[0,1] V2 ≡ V &
138                                                 U0 = 𝕚{Abbr} W1. T1 &
139                                                 U2 = 𝕚{Abbr} W2. 𝕚{Appl} V. T2.
140 #U1 #U2 * -U1 U2
141 [ #k #V #T #H destruct
142 | #i #V #T #H destruct
143 | #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct
144 | #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/
145 | #V1 #V2 #W #T1 #T2 #HV12 #HT12 #V #T #H destruct -V1 T /3 width=8/
146 | #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct
147 | #V #V1 #V2 #W1 #W2 #T1 #T2 #HV12 #HV2 #HW12 #HT12 #V0 #T0 #H
148   destruct -V1 T0 /3 width=12/
149 | #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct
150 | #V #T1 #T2 #_ #V0 #T0 #H destruct
151 ]
152 qed.
153
154 lemma tpr_inv_appl1: ∀V1,U0,U2. 𝕚{Appl} V1. U0 ⇒ U2 →
155                      ∨∨ ∃∃V2,T2.            V1 ⇒ V2 & U0 ⇒ T2 &
156                                             U2 = 𝕚{Appl} V2. T2
157                       | ∃∃V2,W,T1,T2.       V1 ⇒ V2 & T1 ⇒ T2 &
158                                             U0 = 𝕚{Abst} W. T1 &
159                                             U2 = 𝕓{Abbr} V2. T2
160                       | ∃∃V2,V,W1,W2,T1,T2. V1 ⇒ V2 & W1 ⇒ W2 & T1 ⇒ T2 &
161                                             ↑[0,1] V2 ≡ V &
162                                             U0 = 𝕚{Abbr} W1. T1 &
163                                             U2 = 𝕚{Abbr} W2. 𝕚{Appl} V. T2.
164 /2/ qed.
165
166 lemma tpr_inv_cast1_aux: ∀U1,U2. U1 ⇒ U2 → ∀V1,T1. U1 = 𝕚{Cast} V1. T1 →
167                            (∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Cast} V2. T2)
168                          ∨ T1 ⇒ U2.
169 #U1 #U2 * -U1 U2
170 [ #k #V #T #H destruct
171 | #i #V #T #H destruct
172 | #I #V1 #V2 #T1 #T2 #_ #_ #V #T #H destruct
173 | #I #V1 #V2 #T1 #T2 #HV12 #HT12 #V #T #H destruct -I V1 T1 /3 width=5/
174 | #V1 #V2 #W #T1 #T2 #_ #_ #V #T #H destruct
175 | #V1 #V2 #T1 #T2 #T #_ #_ #_ #V0 #T0 #H destruct
176 | #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #V0 #T0 #H destruct
177 | #V #T #T1 #T2 #_ #_ #V0 #T0 #H destruct
178 | #V #T1 #T2 #HT12 #V0 #T0 #H destruct -V T1 /2/
179 ]
180 qed.
181
182 lemma tpr_inv_cast1: ∀V1,T1,U2. 𝕚{Cast} V1. T1 ⇒ U2 →
183                        (∃∃V2,T2. V1 ⇒ V2 & T1 ⇒ T2 & U2 = 𝕚{Cast} V2. T2)
184                      ∨ T1 ⇒ U2.
185 /2/ qed.
186
187 lemma tpr_inv_lref2_aux: ∀T1,T2. T1 ⇒ T2 → ∀i. T2 = #i →
188                          ∨∨           T1 = #i
189                           | ∃∃V,T,T0. ↑[O,1] T0 ≡ T & T0 ⇒ #i &
190                                       T1 = 𝕚{Abbr} V. T
191                           | ∃∃V,T.    T ⇒ #i & T1 = 𝕚{Cast} V. T.
192 #T1 #T2 * -T1 T2
193 [ #k #i #H destruct
194 | #j #i /2/
195 | #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct
196 | #I #V1 #V2 #T1 #T2 #_ #_ #i #H destruct
197 | #V1 #V2 #W #T1 #T2 #_ #_ #i #H destruct
198 | #V1 #V2 #T1 #T2 #T #_ #_ #_ #i #H destruct
199 | #V #V1 #V2 #W1 #W2 #T1 #T2 #_ #_ #_ #_ #i #H destruct
200 | #V #T #T1 #T2 #HT1 #HT12 #i #H destruct /3 width=6/
201 | #V #T1 #T2 #HT12 #i #H destruct /3/
202 ]
203 qed.
204
205 lemma tpr_inv_lref2: ∀T1,i. T1 ⇒ #i →
206                      ∨∨           T1 = #i
207                       | ∃∃V,T,T0. ↑[O,1] T0 ≡ T & T0 ⇒ #i &
208                                   T1 = 𝕓{Abbr} V. T
209                       | ∃∃V,T.    T ⇒ #i & T1 = 𝕗{Cast} V. T.
210 /2/ qed.