1 | include "common/AST.ma". |
---|
2 | include "common/Globalenvs.ma". |
---|
3 | |
---|
4 | (* |
---|
5 | axiom find_funct_ptr_symbol_inversion: |
---|
6 | ∀F,V,init. ∀p: program F V. |
---|
7 | ∀id: ident. ∀b: block. ∀f: F ?. |
---|
8 | find_symbol ? (globalenv ?? init p) id = Some ? b → |
---|
9 | find_funct_ptr ? (globalenv ?? init p) b = Some ? f → |
---|
10 | In … (prog_funct ?? p) 〈id, f〉. |
---|
11 | |
---|
12 | axiom find_funct_ptr_exists: |
---|
13 | ∀F,V,init. ∀p: program F V. ∀id: ident. ∀f: F ?. |
---|
14 | In … (prog_funct ?? p) 〈id, f〉 → |
---|
15 | ∃b. find_symbol ? (globalenv ?? init p) id = Some ? b |
---|
16 | ∧ find_funct_ptr ? (globalenv ?? init p) b = Some ? f. |
---|
17 | |
---|
18 | axiom find_symbol_exists: |
---|
19 | ∀F,V,init. ∀p: program F V. |
---|
20 | ∀id: ident. ∀r,v. |
---|
21 | In ? (prog_vars ?? p) 〈〈id, r〉, v〉 → |
---|
22 | ∃b. find_symbol ? (globalenv ?? init p) id = Some ? b. |
---|
23 | |
---|
24 | definition is_function ≝ λF.λge : genv_t F. |
---|
25 | λi : ident.∃fd. |
---|
26 | ! bl ← find_symbol … ge i ; |
---|
27 | find_funct_ptr … ge bl = Some ? fd. |
---|
28 | |
---|
29 | definition is_internal_function ≝ λF.λge : genv_t (fundef F). |
---|
30 | λi : ident.∃fd. |
---|
31 | ! bl ← find_symbol … ge i ; |
---|
32 | find_funct_ptr … ge bl = Some ? (Internal ? fd). |
---|
33 | |
---|
34 | definition is_internal_function_of_program : |
---|
35 | ∀A.program (λvars.fundef (A vars)) ℕ → ident → Prop ≝ |
---|
36 | λA,prog,i.is_internal_function … (globalenv_noinit … prog) i. |
---|
37 | |
---|
38 | definition funct_of_ident : ∀F,ge. |
---|
39 | ident → option (Σi.is_function F ge i) |
---|
40 | ≝ λF,ge,i. |
---|
41 | match ? |
---|
42 | return λx.! bl ← find_symbol … ge i ; |
---|
43 | find_funct_ptr … ge bl = x → ? |
---|
44 | with |
---|
45 | [ Some fd ⇒ λprf.return «i, ex_intro ?? fd prf» |
---|
46 | | None ⇒ λ_.None ? |
---|
47 | ] (refl …). |
---|
48 | *) |
---|
49 | |
---|
50 | lemma symbol_of_function_block_ok : |
---|
51 | ∀F,ge,b,prf.symbol_for_block F ge b = Some ? (symbol_of_function_block F ge b prf). |
---|
52 | #F #ge #b #FFP |
---|
53 | cut (∀A,B : Type[0].∀P : B → Prop.∀x : option A.∀c1,c2. |
---|
54 | (∀v.∀prf : x = Some ? v.P (c1 v prf)) → |
---|
55 | (∀prf:x = None ?.P (c2 prf)) → |
---|
56 | P (match x return λy.x = y → ? with |
---|
57 | [ Some v ⇒ λprf.c1 v prf | None ⇒ λprf.c2 prf] |
---|
58 | (refl …))) [ #A #B #P * // ] #aux |
---|
59 | whd in ⊢ (???(??%)); @aux [//] #H |
---|
60 | generalize in match (? : False); * |
---|
61 | qed. |
---|
62 | |
---|
63 | (* |
---|
64 | definition funct_of_block : ∀F,ge. |
---|
65 | block → option (Σi.is_function F ge i) ≝ |
---|
66 | λF,ge,bl. |
---|
67 | match find_funct_ptr … ge bl |
---|
68 | return λx.find_funct_ptr … ge bl = x → ? with |
---|
69 | [ Some fd ⇒ λprf.return mk_Sig |
---|
70 | ident (λi.is_function F ge i) |
---|
71 | (symbol_of_function_block … ge bl ?) |
---|
72 | (ex_intro … fd ?) |
---|
73 | | None ⇒ λ_.None ? |
---|
74 | ] (refl …). |
---|
75 | [ >(symbol_of_block_rev … (symbol_of_function_block_ok ? ge bl ?)) @prf |
---|
76 | | >prf % #ABS destruct(ABS) |
---|
77 | ] |
---|
78 | qed. |
---|
79 | |
---|
80 | definition block_of_funct ≝ λF,ge. |
---|
81 | λi : Σi.is_function F ge i. |
---|
82 | match find_symbol … ge i return λx.(∃fd.!bl ← x; ? = ?) → ? with |
---|
83 | [ Some bl ⇒ λ_.bl |
---|
84 | | None ⇒ λprf.⊥ |
---|
85 | ] (pi2 … i). |
---|
86 | cases prf #fd normalize #ABS destruct(ABS) |
---|
87 | qed. |
---|
88 | |
---|
89 | definition description_of_function : ∀F,ge.(Σi.is_function F ge i) → F ≝ |
---|
90 | λF,ge,i. |
---|
91 | match ! bl ← find_symbol … ge i ; |
---|
92 | find_funct_ptr … ge bl |
---|
93 | return λx.(∃fd.x = ?) → ? |
---|
94 | with |
---|
95 | [ Some fd ⇒ λ_.fd |
---|
96 | | None ⇒ λprf.⊥ |
---|
97 | ] (pi2 … i). |
---|
98 | cases prf #fd #ABS destruct |
---|
99 | qed. |
---|
100 | |
---|
101 | lemma description_of_internal_function : ∀F,ge,i,fn. |
---|
102 | description_of_function ? ge i = Internal F fn → is_internal_function … ge i. |
---|
103 | #F #ge * #i * #fd #EQ |
---|
104 | #fn whd in ⊢ (??%?→?); >EQ normalize nodelta #EQ' >EQ' in EQ; #EQ |
---|
105 | %{EQ} |
---|
106 | qed. |
---|
107 | |
---|
108 | definition int_funct_of_block : ∀F,ge. |
---|
109 | block → option (Σi.is_internal_function F ge i) ≝ |
---|
110 | λF,ge,bl. |
---|
111 | ! f ← funct_of_block … ge bl ; |
---|
112 | match ? |
---|
113 | return λx.description_of_function … f = x → option (Σi.is_internal_function … ge i) |
---|
114 | with |
---|
115 | [ Internal fn ⇒ λprf.return «pi1 … f, ?» |
---|
116 | | External fn ⇒ λ_.None ? |
---|
117 | ] (refl …). |
---|
118 | @(description_of_internal_function … prf) |
---|
119 | qed. |
---|
120 | |
---|
121 | lemma internal_function_is_function : ∀F,ge,i. |
---|
122 | is_internal_function F ge i → is_function … ge i. |
---|
123 | #F #ge #i * #fn #prf %{prf} qed. |
---|
124 | |
---|
125 | definition if_of_function : ∀F,ge.(Σi.is_internal_function F ge i) → F ≝ |
---|
126 | λF,ge,i. |
---|
127 | match ! bl ← find_symbol … ge i ; |
---|
128 | find_funct_ptr … ge bl |
---|
129 | return λx.(∃fn.x = ?) → ? |
---|
130 | with |
---|
131 | [ Some fd ⇒ |
---|
132 | match fd return λx.(∃fn.Some ? x = ?) → ? with |
---|
133 | [ Internal fn ⇒ λ_.fn |
---|
134 | | External _ ⇒ λprf.⊥ |
---|
135 | ] |
---|
136 | | None ⇒ λprf.⊥ |
---|
137 | ] (pi2 … i). |
---|
138 | cases prf #fn #ABS destruct |
---|
139 | qed. |
---|
140 | |
---|
141 | lemma if_propagate : |
---|
142 | ∀A_in,A_out. |
---|
143 | ∀trans. |
---|
144 | ∀prog_in : program (λvars.fundef (A_in vars)) ℕ. |
---|
145 | let prog_out : program (λvars.fundef (A_out vars)) ℕ ≝ |
---|
146 | transform_program … prog_in (λvars.transf_fundef … (trans vars)) in |
---|
147 | ∀i.is_internal_function_of_program … prog_in i → |
---|
148 | is_internal_function_of_program … prog_out i. |
---|
149 | #A_in #A_out #trans #prog_in #ident |
---|
150 | * #ifn |
---|
151 | inversion (find_symbol ???) [ #_ #ABS whd in ABS : (??%%); destruct(ABS) ] |
---|
152 | #bl #EQfind >m_return_bind #EQfunct_ptr |
---|
153 | %{(trans … ifn)} |
---|
154 | >find_symbol_transf >EQfind >m_return_bind |
---|
155 | >(find_funct_ptr_transf … EQfunct_ptr) % |
---|
156 | qed. |
---|
157 | |
---|
158 | lemma block_of_funct_ident_is_code : ∀F,ge,i. |
---|
159 | block_region (block_of_funct F ge i) = Code. |
---|
160 | #F #ge * #i * #fd |
---|
161 | whd in ⊢ (?→??(?%)?); |
---|
162 | cases (find_symbol ???) |
---|
163 | [ #ABS normalize in ABS; destruct(ABS) ] |
---|
164 | #bl normalize nodelta >m_return_bind |
---|
165 | whd in ⊢ (??%?→?); cases (block_region bl) |
---|
166 | [ #ABS normalize in ABS; destruct(ABS) ] |
---|
167 | #_ % |
---|
168 | qed. |
---|
169 | |
---|
170 | lemma prog_if_of_function_transform : |
---|
171 | ∀A,B,trans.∀prog_in : program (λvars.fundef (A vars)) ℕ. |
---|
172 | let prog_out : program (λvars.fundef (B vars)) ℕ ≝ |
---|
173 | transform_program … prog_in (λvars.transf_fundef … (trans vars)) in |
---|
174 | ∀i_out : Σi.is_internal_function_of_program … prog_out i. |
---|
175 | ∀i_in : Σi.is_internal_function_of_program … prog_in i. |
---|
176 | pi1 … i_out = pi1 … i_in → |
---|
177 | if_of_function … i_out = trans … (if_of_function … i_in). |
---|
178 | #A #B #trans #prog * #i1 #i_prf1 * #i2 #i_prf2 #EQ destruct(EQ) |
---|
179 | cut (∀A,B : Type[0].∀Q : option A → Prop.∀P : B → Prop. |
---|
180 | ∀m.∀prf : Q m.∀f1,f2. |
---|
181 | (∀x,prf.m = Some ? x → P (f1 x prf)) → |
---|
182 | (∀prf.m = None ? → P (f2 prf)) → |
---|
183 | P (match m return λx.Q x → ? with [ Some x ⇒ f1 x | None ⇒ f2 ] prf)) |
---|
184 | [ #A #B #Q #P * /2 by / ] #aux |
---|
185 | whd in ⊢ (??%?); @aux |
---|
186 | [2: #prf #EQ generalize in match (? : False); * ] |
---|
187 | #fd * #ifn #EQ1 #EQ destruct normalize nodelta |
---|
188 | whd in ⊢ (???(??%)); @aux |
---|
189 | [2: #prf #EQ generalize in match (? : False); * ] |
---|
190 | #fd' * #ifn' #EQ1' #EQ' destruct normalize nodelta |
---|
191 | inversion (find_symbol ???) in EQ'; [ #_ #ABS whd in ABS: (??%%); destruct ] |
---|
192 | #bl #EQfind >m_return_bind #EQfunct |
---|
193 | >find_symbol_transf in EQ; >EQfind >m_return_bind |
---|
194 | >(find_funct_ptr_transf … EQfunct) #EQ'' destruct % |
---|
195 | qed. |
---|
196 | *) |
---|
197 | |
---|
198 | include alias "common/PositiveMap.ma". |
---|
199 | |
---|
200 | lemma add_functs_functions_miss : ∀F. ∀ge : genv_t F.∀ l.∀ id. |
---|
201 | id < (nextfunction ? ge) → |
---|
202 | lookup_opt … id (functions ? ge) = None ? → |
---|
203 | lookup_opt … id (functions … (add_functs F ge l)) = None ?. |
---|
204 | #F #ge #l #id whd in match add_functs; normalize nodelta |
---|
205 | elim l -l [ #_ normalize //] |
---|
206 | * #id1 #f1 #tl #IND #H #H1 whd in match (foldr ?????); |
---|
207 | >lookup_opt_insert_miss [ inversion(lookup_opt ? ? ?) [ #_ %] |
---|
208 | #f1 #H3 <(drop_fn_lfn … f1 H3) @(IND H H1) |
---|
209 | | cut(nextfunction F ge ≤ nextfunction F (foldr … (add_funct F) ge tl)) |
---|
210 | [elim tl [normalize //] |
---|
211 | #x #tl2 whd in match (foldr ?????) in ⊢ (? → %); #H %2{H} ] |
---|
212 | #H2 lapply(transitive_le … H H2) @lt_to_not_eq |
---|
213 | qed. |
---|
214 | |
---|
215 | lemma add_globals_functions_miss : ∀F,V,init. ∀gem : genv_t F × mem.∀ id,l. |
---|
216 | lookup_opt … id (functions ? (\fst gem)) = None ? → |
---|
217 | lookup_opt … id (functions … (\fst (add_globals F V init gem l))) = None ?. |
---|
218 | #F #V #init * #ge #m #id #l lapply ge -ge lapply m -m elim l [ #ge #m #H @H] |
---|
219 | ** #x1 #x2 #x3 #tl whd in match add_globals; |
---|
220 | normalize nodelta #IND #m #ge #H whd in match (foldl ?????); normalize nodelta |
---|
221 | cases(alloc m ? ? (*x2*)) #m1 #bl1 normalize nodelta @IND whd in match add_symbol; |
---|
222 | normalize nodelta inversion(lookup_opt ? ? ?) [ #_ %] |
---|
223 | #f1 #H3 <(drop_fn_lfn … f1 H3) assumption |
---|
224 | qed. |
---|
225 | |
---|
226 | |
---|
227 | lemma globalenv_no_minus_one : |
---|
228 | ∀F,V,i,p. |
---|
229 | find_funct_ptr … (globalenv F V i p) (mk_block (*Code*) (-1)) = None ?. |
---|
230 | #F #V #i #p |
---|
231 | whd in match find_funct_ptr; normalize nodelta |
---|
232 | whd in match globalenv; |
---|
233 | whd in match globalenv_allocmem; normalize nodelta |
---|
234 | @add_globals_functions_miss @add_functs_functions_miss normalize // |
---|
235 | qed. |
---|
236 | |
---|
237 | lemma globalenv_no_zero : |
---|
238 | ∀F,V,i,p. |
---|
239 | find_funct_ptr … (globalenv F V i p) (mk_block (*Code*) OZ) = None ?. // |
---|
240 | qed. |
---|
241 | |
---|