1 | include "RTLabs/RTLabs_syntax.ma". |
---|
2 | include "RTL/RTL.ma". |
---|
3 | include "common/AssocList.ma". |
---|
4 | include "common/FrontEndOps.ma". |
---|
5 | include "common/Graphs.ma". |
---|
6 | include "joint/TranslateUtils.ma". |
---|
7 | include alias "ASM/BitVector.ma". |
---|
8 | include alias "arithmetics/nat.ma". |
---|
9 | |
---|
10 | definition size_of_sig_type ≝ |
---|
11 | λsig. |
---|
12 | match sig with |
---|
13 | [ ASTint isize sign ⇒ |
---|
14 | match isize with [ I8 ⇒ 1 | I16 ⇒ 2 | I32 ⇒ 4 ] |
---|
15 | | ASTptr ⇒ 2 (* rgn ⇒ nat_of_bitvector ? ptr_size *) |
---|
16 | ]. |
---|
17 | |
---|
18 | inductive register_type: Type[0] ≝ |
---|
19 | | register_int: register → register_type |
---|
20 | | register_ptr: register → register → register_type. |
---|
21 | |
---|
22 | definition local_env ≝ identifier_map RegisterTag (list register). |
---|
23 | |
---|
24 | definition local_env_typed : |
---|
25 | list (register × typ) → local_env → Prop ≝ |
---|
26 | λl,env.All ? |
---|
27 | (λp.let 〈r, ty〉 ≝ p in ∃regs.lookup … env r = Some ? regs ∧ |
---|
28 | |regs| = size_of_sig_type ty) l. |
---|
29 | |
---|
30 | definition find_local_env ≝ λr.λlenv : local_env. |
---|
31 | λprf : r ∈ lenv.opt_safe … (lookup … lenv r) ?. |
---|
32 | lapply (in_map_domain … lenv r) |
---|
33 | >prf * #x #lookup_eq >lookup_eq % #ABS destruct(ABS) |
---|
34 | qed. |
---|
35 | |
---|
36 | lemma find_local_env_elim : ∀P : list register → Prop.∀r. ∀lenv: local_env.∀prf. |
---|
37 | (∀x.lookup … lenv r = Some ? x → P x) → P (find_local_env r lenv prf). |
---|
38 | #P#r#lenv#prf #H |
---|
39 | change with (P (opt_safe ???)) |
---|
40 | @opt_safe_elim assumption |
---|
41 | qed. |
---|
42 | |
---|
43 | definition find_local_env_arg : register → local_env → ? → list psd_argument ≝ |
---|
44 | λr,lenv,prf. map … (Reg ?) (find_local_env r lenv prf). |
---|
45 | |
---|
46 | (* move *) |
---|
47 | let rec m_iter (M : Monad) X (f : X → M X) (n : ℕ) (m : M X) on n : M X ≝ |
---|
48 | match n with |
---|
49 | [ O ⇒ m |
---|
50 | | S k ⇒ |
---|
51 | ! v ← m ; |
---|
52 | m_iter … f k (f v) |
---|
53 | ]. |
---|
54 | |
---|
55 | definition fresh_registers : ∀p,g.ℕ → state_monad (joint_internal_function p g) (list register) ≝ |
---|
56 | λp,g,n. |
---|
57 | let f ≝ λacc.! m ← fresh_register … ; return (m :: acc) in |
---|
58 | m_iter … f n (return [ ]). |
---|
59 | |
---|
60 | include alias "common/Identifiers.ma". |
---|
61 | let rec map_list_local_env |
---|
62 | lenv (regs : list (register×typ)) on regs : |
---|
63 | All ? (λpr.bool_to_Prop (\fst pr ∈ lenv)) regs → list register ≝ |
---|
64 | match regs return λx.All ?? x → ? with |
---|
65 | [ nil ⇒ λ_.[ ] |
---|
66 | | cons hd tl ⇒ λprf.find_local_env (\fst hd) lenv ? @ map_list_local_env lenv tl ? |
---|
67 | ].cases prf #A #B assumption qed. |
---|
68 | |
---|
69 | definition initialize_local_env : |
---|
70 | ∀globals. |
---|
71 | list (register×typ) → |
---|
72 | state_monad (joint_internal_function RTL globals) local_env ≝ |
---|
73 | λglobals,registers. |
---|
74 | let f ≝ |
---|
75 | λr_sig,lenv. |
---|
76 | let 〈r, sig〉 ≝ r_sig in |
---|
77 | let size ≝ size_of_sig_type sig in |
---|
78 | ! regs ← fresh_registers … size ; |
---|
79 | return add … lenv r regs in |
---|
80 | m_fold … f registers (empty_map …). |
---|
81 | |
---|
82 | lemma initialize_local_env_in : ∀globals,l,def,r. |
---|
83 | Exists ? (λx.\fst x = r) l → r ∈ \snd (initialize_local_env globals l def). |
---|
84 | whd in match initialize_local_env; normalize nodelta #globals |
---|
85 | cut (∀l,init,def,r.(Exists ? (λx.\fst x = r) l ∨ bool_to_Prop (r ∈ init)) → |
---|
86 | r ∈ \snd (m_fold (state_monad ?) ??? l init def)) |
---|
87 | [7: #aux #l #def #r #H @aux %1{H} |*:] |
---|
88 | #l elim l -l |
---|
89 | [ #init #def #r * [*] #H @H ] |
---|
90 | * #hd #sig #tl #IH #init #def #r #H |
---|
91 | whd in ⊢ (?(???(???%)?)); normalize nodelta |
---|
92 | whd in ⊢ (?(???(???(match % with [ _ ⇒ ?]))?)); |
---|
93 | inversion (fresh_registers ????) #def' #regs #EQfresh normalize nodelta |
---|
94 | @IH cases H -H [*] #H |
---|
95 | [ destruct %2 @mem_set_add_id |
---|
96 | | %1{H} |
---|
97 | | %2 >mem_set_add @orb_Prop_r @H |
---|
98 | ] |
---|
99 | qed. |
---|
100 | |
---|
101 | example proj1_rewrite : ∀A,B,a,b.∀pr : A×B.〈a,b〉 = pr → a = \fst pr. |
---|
102 | // qed-. |
---|
103 | example proj2_rewrite : ∀A,B,a,b.∀pr : A×B.〈a,b〉 = pr → b = \snd pr. |
---|
104 | // qed-. |
---|
105 | |
---|
106 | definition initialize_locals_params_ret : |
---|
107 | ∀globals. |
---|
108 | (* locals *) list (register×typ) → |
---|
109 | (* params *) list (register×typ) → |
---|
110 | (* return *) option (register×typ) → |
---|
111 | state_monad (joint_internal_function RTL globals) local_env ≝ |
---|
112 | λglobals,locals,params,ret,def. |
---|
113 | let 〈def',lenv〉 as EQ ≝ |
---|
114 | initialize_local_env globals |
---|
115 | ((match ret with |
---|
116 | [ Some r_sig ⇒ [r_sig] |
---|
117 | | None ⇒ [ ] |
---|
118 | ]) @ locals @ params) def in |
---|
119 | let params' ≝ map_list_local_env lenv params ? in |
---|
120 | let ret' ≝ match ret return λx.ret = x → ? with |
---|
121 | [ Some r_sig ⇒ λprf.find_local_env (\fst r_sig) lenv ? |
---|
122 | | None ⇒ λ_.[ ] |
---|
123 | ] (refl …) in |
---|
124 | let def'' ≝ |
---|
125 | mk_joint_internal_function RTL globals |
---|
126 | (joint_if_luniverse … def') (joint_if_runiverse … def') ret' |
---|
127 | params' (joint_if_stacksize … def') |
---|
128 | (joint_if_code … def') (joint_if_entry … def') in |
---|
129 | 〈def'', lenv〉. @hide_prf |
---|
130 | [ >(proj2_rewrite ????? EQ) |
---|
131 | @initialize_local_env_in >prf %1 % |
---|
132 | |*: >(proj2_rewrite ????? EQ) |
---|
133 | @(All_mp ??? (λpr.initialize_local_env_in ??? (\fst pr))) |
---|
134 | @(All_mp … (λpr.Exists ? (λx.\fst x = \fst pr) params)) |
---|
135 | [ #a #H @Exists_append_r @Exists_append_r @H |
---|
136 | | elim params [%] #hd #tl #IH % [ %1 % ] @(All_mp … IH) #x #G %2{G} |
---|
137 | ] |
---|
138 | ] |
---|
139 | qed. |
---|
140 | |
---|
141 | definition make_addr ≝ |
---|
142 | λA. |
---|
143 | λlst: list A. |
---|
144 | λprf: 2 = length A lst.〈nth_safe … 0 lst ?, nth_safe … 1 lst ?〉. <prf [%2] %1 |
---|
145 | qed. |
---|
146 | |
---|
147 | definition find_and_addr ≝ |
---|
148 | λr,lenv,prf. make_addr ? (find_local_env r lenv prf). |
---|
149 | |
---|
150 | definition find_and_addr_arg ≝ |
---|
151 | λr,lenv,prf. make_addr ? (find_local_env_arg r lenv prf). |
---|
152 | |
---|
153 | (*include alias "common/Identifiers.ma".*) |
---|
154 | let rec rtl_args (args : list register) (env : local_env) on args : |
---|
155 | All ? (λr.bool_to_Prop (r∈env)) args → list psd_argument ≝ |
---|
156 | match args return λx.All ?? x → ? with |
---|
157 | [ nil ⇒ λ_.[ ] |
---|
158 | | cons hd tl ⇒ λprf.find_local_env_arg hd env ? @ rtl_args tl env ? |
---|
159 | ]. |
---|
160 | cases prf #H #G assumption |
---|
161 | qed. |
---|
162 | |
---|
163 | include alias "basics/lists/list.ma". |
---|
164 | let rec vrsplit A (m,n : nat) |
---|
165 | on m : Vector A (m*n) → Σs : list (Vector A n).|s| = m ≝ |
---|
166 | match m return λx.Vector A (x*n) → Sig (list ?) ? with |
---|
167 | [ O ⇒ λv.[ ] |
---|
168 | | S k ⇒ λv.let spl ≝ vsplit ? n … v in \fst spl :: vrsplit ? k n (\snd spl) |
---|
169 | ]. |
---|
170 | [ % |
---|
171 | | cases (vrsplit ????) #lst #EQ normalize >EQ % |
---|
172 | ] qed. |
---|
173 | |
---|
174 | definition split_into_bytes: |
---|
175 | ∀size. ∀int: bvint size. Σbytes: list Byte. |bytes| = size_intsize size ≝ |
---|
176 | λsize.vrsplit ? (size_intsize size) 8. |
---|
177 | |
---|
178 | let rec list_inject_All_aux A P (l : list A) on l : All A P l → list (Σx.P x) ≝ |
---|
179 | match l return λx.All A P x → list (Σx.P x) with |
---|
180 | [ nil ⇒ λ_.[ ] |
---|
181 | | cons hd tl ⇒ λprf.«hd, proj1 … prf» :: list_inject_All_aux A P tl (proj2 … prf) |
---|
182 | ]. |
---|
183 | |
---|
184 | definition translate_op: |
---|
185 | ∀globals. Op2 → |
---|
186 | ∀dests : list register. |
---|
187 | ∀srcrs1 : list psd_argument. |
---|
188 | ∀srcrs2 : list psd_argument. |
---|
189 | |dests| = |srcrs1| → |srcrs1| = |srcrs2| → |
---|
190 | bind_new register (list (joint_seq RTL globals)) |
---|
191 | ≝ |
---|
192 | λglobals: list ident. |
---|
193 | λop. |
---|
194 | λdestrs. |
---|
195 | λsrcrs1. |
---|
196 | λsrcrs2. |
---|
197 | λprf1,prf2. |
---|
198 | (* first, clear carry if op relies on it *) |
---|
199 | match op with |
---|
200 | [ Addc ⇒ [CLEAR_CARRY ??] |
---|
201 | | Sub ⇒ [CLEAR_CARRY ??] |
---|
202 | | _ ⇒ [ ] |
---|
203 | ] @ map3 ???? (OP2 RTL globals op) destrs srcrs1 srcrs2 prf1 prf2. |
---|
204 | |
---|
205 | definition cast_list : ∀A.A → ℕ → list A → list A ≝ |
---|
206 | λA,deflt,new_length,l. |
---|
207 | if leb (|l|) new_length then |
---|
208 | l @ make_list ? deflt (new_length - |l|) |
---|
209 | else |
---|
210 | lhd … l new_length. |
---|
211 | |
---|
212 | lemma length_make_list: |
---|
213 | ∀A: Type[0]. |
---|
214 | ∀elt: A. |
---|
215 | ∀n: nat. |
---|
216 | length ? (make_list A elt n) = n. |
---|
217 | #A #ELT #N |
---|
218 | elim N normalize // qed. |
---|
219 | |
---|
220 | lemma length_lhd : ∀A,l,n.|lhd A l n| = min (|l|) n. |
---|
221 | #A #l elim l -l |
---|
222 | [ * // |
---|
223 | | #hd #tl #IH * normalize [%] |
---|
224 | #n >IH normalize elim (leb ??) % |
---|
225 | ] |
---|
226 | qed. |
---|
227 | |
---|
228 | lemma length_cast_list : ∀A,dflt,n,l.|cast_list A dflt n l| = n. |
---|
229 | #A #dflt #n #l |
---|
230 | normalize @leb_elim #H normalize |
---|
231 | [ >length_append >length_make_list |
---|
232 | @sym_eq @minus_to_plus // |
---|
233 | | >length_lhd normalize @leb_elim |
---|
234 | [ #abs elim (absurd ? abs H) ] |
---|
235 | #_ % |
---|
236 | ] |
---|
237 | qed. |
---|
238 | |
---|
239 | definition translate_op_asym_unsigned : |
---|
240 | ∀globals.Op2 → list register → list psd_argument → list psd_argument → |
---|
241 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
242 | λglobals,op,destrs,srcrs1,srcrs2. |
---|
243 | let l ≝ |destrs| in |
---|
244 | let srcrs1' ≝ cast_list ? (zero_byte : psd_argument) l srcrs1 in |
---|
245 | let srcrs2' ≝ cast_list ? (zero_byte : psd_argument) l srcrs2 in |
---|
246 | translate_op globals op destrs srcrs1' srcrs2' ??. |
---|
247 | @hide_prf |
---|
248 | normalize nodelta |
---|
249 | >length_cast_list [2: >length_cast_list ] % |
---|
250 | qed. |
---|
251 | |
---|
252 | let rec nat_to_args (size : nat) (k : nat) on size : Σl : list psd_argument.|l| = size ≝ |
---|
253 | match size with |
---|
254 | [ O ⇒ [ ] |
---|
255 | | S size' ⇒ |
---|
256 | (byte_of_nat k : psd_argument) :: nat_to_args size' (k ÷ 8) |
---|
257 | ]. @hide_prf [ % | cases (nat_to_args ??) #res #EQ normalize >EQ % ] qed. |
---|
258 | |
---|
259 | definition size_of_cst ≝ λtyp.λcst : constant typ.match cst with |
---|
260 | [ Ointconst size _ _ ⇒ size_intsize size |
---|
261 | | _ ⇒ 2 |
---|
262 | ]. |
---|
263 | |
---|
264 | definition cst_well_defd : ∀ty.list ident → constant ty → Prop ≝ λty,globals,cst. |
---|
265 | match cst with |
---|
266 | [ Oaddrsymbol id _ ⇒ bool_to_Prop (id ∈ globals) |
---|
267 | | _ ⇒ True |
---|
268 | ]. |
---|
269 | |
---|
270 | definition translate_cst : |
---|
271 | ∀ty. |
---|
272 | ∀globals: list ident. |
---|
273 | ∀cst_sig: Σcst : constant ty.cst_well_defd ty globals cst. |
---|
274 | ∀destrs: list register. |
---|
275 | |destrs| = size_of_cst ? cst_sig → |
---|
276 | bind_new register (list (joint_seq RTL globals)) |
---|
277 | ≝ |
---|
278 | λty,globals,cst_sig,destrs. |
---|
279 | match pi1 … cst_sig in constant return λty'.λx : constant ty'. |
---|
280 | cst_well_defd ty' ? x → |destrs| = size_of_cst ty' x → ? |
---|
281 | with |
---|
282 | [ Ointconst size sign const ⇒ λcst_prf,prf. |
---|
283 | map2 … (λr.λb : Byte.r ← b) destrs |
---|
284 | (split_into_bytes size const) ? |
---|
285 | | Oaddrsymbol id offset ⇒ λcst_prf,prf. |
---|
286 | let 〈r1, r2〉 ≝ make_addr … destrs ? in |
---|
287 | [ADDRESS RTL globals id ? r1 r2] |
---|
288 | | Oaddrstack offset ⇒ λcst_prf,prf. |
---|
289 | let 〈r1, r2〉 ≝ make_addr … destrs ? in |
---|
290 | [(rtl_stack_address r1 r2 : joint_seq RTL globals)] |
---|
291 | ] (pi2 … cst_sig). |
---|
292 | @hide_prf |
---|
293 | [ cases (split_into_bytes ??) #lst |
---|
294 | #EQ >EQ >prf whd in ⊢ (??%?); cases size % |
---|
295 | | @cst_prf |
---|
296 | |*: >prf % |
---|
297 | ] |
---|
298 | qed. |
---|
299 | |
---|
300 | definition translate_move : |
---|
301 | ∀globals. |
---|
302 | ∀destrs: list register. |
---|
303 | ∀srcrs: list psd_argument. |
---|
304 | |destrs| = |srcrs| → list (joint_seq RTL globals) ≝ |
---|
305 | λglobals,destrs,srcrs,length_eq. |
---|
306 | map2 … (λdst,src.dst ← src) destrs srcrs length_eq. |
---|
307 | |
---|
308 | definition sign_mask : ∀globals.register → psd_argument → |
---|
309 | list (joint_seq RTL globals) ≝ |
---|
310 | (* this sets destr to 0xFF if s is neg, 0x00 o.w. Done like that: |
---|
311 | byte in destr if srcr is: neg | pos |
---|
312 | destr ← srcr | 127 11...1 | 01...1 |
---|
313 | destr ← destr <rot< 1 1...11 | 1...10 |
---|
314 | destr ← INC destr 0....0 | 1....1 |
---|
315 | destr ← CPL destr 1....1 | 0....0 |
---|
316 | *) |
---|
317 | λglobals,destr,srca. |
---|
318 | match srca with |
---|
319 | [ Reg srcr ⇒ |
---|
320 | let byte_127 : Byte ≝ false ::: maximum ? in |
---|
321 | [destr ← srcr .Or. byte_127 ; |
---|
322 | destr ← .Rl. destr ; |
---|
323 | destr ← .Inc. destr ; |
---|
324 | destr ← .Cmpl. destr ] |
---|
325 | | Imm b ⇒ |
---|
326 | if sign_bit … b then |
---|
327 | [ destr ← (maximum … : Byte) ] |
---|
328 | else |
---|
329 | [ destr ← zero_byte ] |
---|
330 | ]. |
---|
331 | |
---|
332 | definition translate_cast_signed : |
---|
333 | ∀globals : list ident. |
---|
334 | list register → psd_argument → |
---|
335 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
336 | λglobals,destrs,srca. |
---|
337 | ν tmp in |
---|
338 | (sign_mask ? tmp srca @ |
---|
339 | translate_move ? destrs (make_list ? (Reg ? tmp) (|destrs|)) ?). |
---|
340 | >length_make_list % qed. |
---|
341 | |
---|
342 | definition translate_fill_with_zero : |
---|
343 | ∀globals : list ident. |
---|
344 | list register → list (joint_seq RTL globals) ≝ |
---|
345 | λglobals,destrs. |
---|
346 | match nat_to_args (|destrs|) 0 with |
---|
347 | [ mk_Sig res prf ⇒ translate_move ? destrs res ?]. |
---|
348 | // qed. |
---|
349 | |
---|
350 | let rec last A (l : list A) on l : option A ≝ |
---|
351 | match l with |
---|
352 | [ nil ⇒ None ? |
---|
353 | | cons hd tl ⇒ |
---|
354 | match tl with |
---|
355 | [ nil ⇒ Some ? hd |
---|
356 | | _ ⇒ last A tl |
---|
357 | ] |
---|
358 | ]. |
---|
359 | |
---|
360 | lemma last_def : ∀A,hd,tl.last A (hd @ [tl]) = Some ? tl. |
---|
361 | #A #hd elim hd -hd [ #tl % ] #hd * [ #_ #last % ] #hd' #tl #IH #last |
---|
362 | @IH qed. |
---|
363 | |
---|
364 | lemma last_last_ne : ∀A,l.last A (pi1 … l) = Some ? (last_ne … l). |
---|
365 | #A * @list_elim_left [*] #hd #tl #_ #prf |
---|
366 | whd in match last_ne; normalize nodelta |
---|
367 | >split_on_last_ne_def @last_def qed. |
---|
368 | |
---|
369 | lemma last_not_empty : ∀A,l.not_empty A l → |
---|
370 | match last A l with |
---|
371 | [ None ⇒ False |
---|
372 | | _ ⇒ True ]. |
---|
373 | #A #l #prf >(last_last_ne … «l, prf») % qed. |
---|
374 | |
---|
375 | definition translate_op_asym_signed : |
---|
376 | ∀globals.Op2 → list register → list psd_argument → list psd_argument → |
---|
377 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
378 | λglobals,op,destrs,srcrs1,srcrs2. |
---|
379 | νtmp1,tmp2 in |
---|
380 | let l ≝ |destrs| in |
---|
381 | let cast_srcrs ≝ λsrcrs,tmp. |
---|
382 | let srcrs_l ≝ |srcrs| in |
---|
383 | if leb srcrs_l l then |
---|
384 | match last ? srcrs with |
---|
385 | [ Some last ⇒ |
---|
386 | 〈srcrs @ make_list … (Reg ? tmp) (l - srcrs_l), |
---|
387 | sign_mask … tmp last〉 |
---|
388 | | None ⇒ |
---|
389 | 〈make_list … (zero_byte : psd_argument) l, [ ]〉 |
---|
390 | ] |
---|
391 | else |
---|
392 | 〈lhd … srcrs l, [ ]〉 in |
---|
393 | let prf : ∀srcrs,tmp.|destrs| = |\fst (cast_srcrs srcrs tmp)| ≝ ? in |
---|
394 | let srcrs1init ≝ cast_srcrs srcrs1 tmp1 in |
---|
395 | let srcrs2init ≝ cast_srcrs srcrs2 tmp2 in |
---|
396 | \snd srcrs1init @@ \snd srcrs2init @@ |
---|
397 | translate_op globals op destrs (\fst srcrs1init) (\fst srcrs2init) ??. |
---|
398 | @hide_prf |
---|
399 | [ @prf | <prf @prf ] |
---|
400 | #srcrs #tmp normalize nodelta |
---|
401 | @leb_elim #H normalize nodelta |
---|
402 | [ cases (last ??) normalize nodelta |
---|
403 | [ >length_make_list % |
---|
404 | | #_ >length_append >length_make_list |
---|
405 | @minus_to_plus // |
---|
406 | ] |
---|
407 | | >length_lhd normalize >(not_le_to_leb_false … H) % |
---|
408 | ] |
---|
409 | qed. |
---|
410 | |
---|
411 | (* using size of lists as size of ints *) |
---|
412 | definition translate_cast : |
---|
413 | ∀globals: list ident. |
---|
414 | signedness → list register → list register → |
---|
415 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
416 | λglobals,src_sign,destrs,srcrs. |
---|
417 | match reduce_strong ?? destrs srcrs with |
---|
418 | [ mk_Sig t prf ⇒ |
---|
419 | let src_common ≝ \fst (\fst t) in |
---|
420 | let src_rest ≝ \snd (\fst t) in |
---|
421 | let dst_common ≝ \fst (\snd t) in |
---|
422 | let dst_rest ≝ \snd (\snd t) in |
---|
423 | (* first, move the common part *) |
---|
424 | translate_move ? src_common (map … (Reg ?) dst_common) ? @@ |
---|
425 | match src_rest return λ_.bind_new ?? with |
---|
426 | [ nil ⇒ (* upcast *) |
---|
427 | match src_sign return λ_.bind_new ?? with |
---|
428 | [ Unsigned ⇒ translate_fill_with_zero ? dst_rest |
---|
429 | | Signed ⇒ |
---|
430 | match last … srcrs (* = src_common *) with |
---|
431 | [ Some src_last ⇒ translate_cast_signed ? dst_rest src_last |
---|
432 | | None ⇒ (* srcrs is empty *) translate_fill_with_zero ? dst_rest |
---|
433 | ] |
---|
434 | ] |
---|
435 | | _ ⇒ (* downcast, nothing else to do *) [ ] |
---|
436 | ] |
---|
437 | ]. |
---|
438 | >length_map @prf qed. |
---|
439 | |
---|
440 | definition translate_notint : |
---|
441 | ∀globals : list ident. |
---|
442 | ∀destrs : list register. |
---|
443 | ∀srcrs_arg : list register. |
---|
444 | |destrs| = |srcrs_arg| → list (joint_seq RTL globals) ≝ |
---|
445 | λglobals, destrs, srcrs, prf. |
---|
446 | map2 ??? (OP1 RTL globals Cmpl) destrs srcrs prf. |
---|
447 | |
---|
448 | definition translate_negint : ∀globals.? → ? → ? → bind_new register (list (joint_seq RTL globals)) ≝ |
---|
449 | λglobals: list ident. |
---|
450 | λdestrs: list register. |
---|
451 | λsrcrs: list register. |
---|
452 | λprf: |destrs| = |srcrs|. (* assert in caml code *) |
---|
453 | translate_notint … destrs srcrs prf @@ |
---|
454 | translate_op ? Add destrs (map … (Reg ?) destrs) (nat_to_args (|destrs|) 1) ??. |
---|
455 | @hide_prf >length_map [ cases (nat_to_args ??) ] // qed. |
---|
456 | |
---|
457 | definition translate_notbool: |
---|
458 | ∀globals : list ident. |
---|
459 | list register → list register → |
---|
460 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
461 | λglobals,destrs,srcrs. |
---|
462 | match destrs with |
---|
463 | [ nil ⇒ [ ] |
---|
464 | | cons destr destrs' ⇒ |
---|
465 | translate_fill_with_zero ? destrs' @@ |
---|
466 | match srcrs return λ_.bind_new ?? with |
---|
467 | [ nil ⇒ [destr ← zero_byte] |
---|
468 | | cons srcr srcrs' ⇒ |
---|
469 | (destr ← srcr) ::: |
---|
470 | map register (joint_seq RTL globals) (λr. destr ← destr .Or. r) srcrs' @@ |
---|
471 | (* now destr is non-null iff srcrs was non-null *) |
---|
472 | CLEAR_CARRY ?? ::: |
---|
473 | (* many uses of 0, better not use immediates *) |
---|
474 | ν tmp in |
---|
475 | [tmp ← zero_byte ; |
---|
476 | destr ← tmp .Sub. tmp ; |
---|
477 | (* now carry bit is set iff destr was non-null *) |
---|
478 | destr ← tmp .Addc. tmp] |
---|
479 | ] |
---|
480 | ]. |
---|
481 | |
---|
482 | definition translate_op1 : ∀globals.? → ? → ? → ? → ? → ? → ? → |
---|
483 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
484 | λglobals. |
---|
485 | λty, ty'. |
---|
486 | λop1: unary_operation ty ty'. |
---|
487 | λdestrs: list register. |
---|
488 | λsrcrs: list register. |
---|
489 | λprf1: |destrs| = size_of_sig_type ty'. |
---|
490 | λprf2: |srcrs| = size_of_sig_type ty. |
---|
491 | match op1 |
---|
492 | return λty'',ty'''.λx : unary_operation ty'' ty'''.ty'' = ty → ty''' = ty' → |
---|
493 | bind_new register (list (joint_seq RTL globals)) with |
---|
494 | [ Ocastint _ src_sign _ _ ⇒ λeq1,eq2. |
---|
495 | translate_cast globals src_sign destrs srcrs |
---|
496 | | Onegint sz sg ⇒ λeq1,eq2. |
---|
497 | translate_negint globals destrs srcrs ? |
---|
498 | | Onotbool _ _ _ _ ⇒ λeq1,eq2. |
---|
499 | translate_notbool globals destrs srcrs |
---|
500 | | Onotint sz sg ⇒ λeq1,eq2. |
---|
501 | translate_notint globals destrs srcrs ? |
---|
502 | | Optrofint sz sg ⇒ λeq1,eq2. |
---|
503 | translate_cast globals Unsigned destrs srcrs |
---|
504 | | Ointofptr sz sg ⇒ λeq1,eq2. |
---|
505 | translate_cast globals Unsigned destrs srcrs |
---|
506 | | Oid t ⇒ λeq1,eq2. |
---|
507 | translate_move globals destrs (map … (Reg ?) srcrs) ? |
---|
508 | | _ ⇒ λeq1,eq2.? (* float operations implemented in runtime *) |
---|
509 | ] (refl …) (refl …). |
---|
510 | @hide_prf |
---|
511 | destruct >prf1 >prf2 [3: >length_map ] // |
---|
512 | qed. |
---|
513 | |
---|
514 | include alias "arithmetics/nat.ma". |
---|
515 | |
---|
516 | definition translate_mul_i : |
---|
517 | ∀globals. |
---|
518 | register → register → |
---|
519 | (* size of destination and sources *) |
---|
520 | ∀n : ℕ. |
---|
521 | (* the temporary destination, with a dummy register at the end *) |
---|
522 | ∀tmp_destrs_dummy : list register. |
---|
523 | ∀srcrs1,srcrs2 : list psd_argument. |
---|
524 | |tmp_destrs_dummy| = S n → |
---|
525 | n = |srcrs1| → |
---|
526 | |srcrs1| = |srcrs2| → |
---|
527 | (* the position of the least significant byte of the result we compute at |
---|
528 | this stage (goes from 0 to n in the main function) *) |
---|
529 | ∀k : ℕ. |
---|
530 | lt k n → |
---|
531 | (* the position of the byte in the first source we will use in this stage. |
---|
532 | the position in the other source will be k - i *) |
---|
533 | (Σi.i<S k) → |
---|
534 | (* the accumulator *) |
---|
535 | bind_new register (list (joint_seq RTL globals)) → |
---|
536 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
537 | λglobals,a,b,n,tmp_destrs_dummy,srcrs1,srcrs2, |
---|
538 | tmp_destrs_dummy_prf,srcrs1_prf,srcrs2_prf,k,k_prf,i_sig,acc. |
---|
539 | (* the following will expand to |
---|
540 | a, b ← srcrs1[i] * srcrs2[k-i] |
---|
541 | tmp_destrs_dummy[k] ← tmp_destrs_dummy[k] + a |
---|
542 | tmp_destrs_dummy[k+1] ← tmp_destrs_dummy[k+1] + b + C |
---|
543 | tmp_destrs_dummy[k+2] ← tmp_destrs_dummy[k+2] + 0 + C |
---|
544 | ... |
---|
545 | tmp_destrs_dummy[n] ← tmp_destrs_dummy[n] + 0 + C |
---|
546 | ( all calculations on tmp_destrs_dummy[n] will be eliminated with |
---|
547 | liveness analysis) *) |
---|
548 | match i_sig with |
---|
549 | [ mk_Sig i i_prf ⇒ |
---|
550 | (* we pad the result of a byte multiplication with zeros in order |
---|
551 | for the bit to be carried. Redundant calculations will be eliminated |
---|
552 | by constant propagation. *) |
---|
553 | let args : list psd_argument ≝ |
---|
554 | [Reg ? a;Reg ? b] @ make_list ? (zero_byte : psd_argument) (n - 1 - k) in |
---|
555 | let tmp_destrs_view : list register ≝ |
---|
556 | ltl ? tmp_destrs_dummy k in |
---|
557 | [❮a, b❯ ← (nth_safe ? i srcrs1 ?) .Mul. (nth_safe ? (k - i) srcrs2 ?)] @@ |
---|
558 | translate_op … Add tmp_destrs_view (map … (Reg ?) tmp_destrs_view) args ?? @@ |
---|
559 | acc |
---|
560 | ]. |
---|
561 | @hide_prf |
---|
562 | [ <srcrs1_prf |
---|
563 | @(transitive_le … i_prf k_prf) |
---|
564 | | @lt_plus_to_minus [ @le_S_S_to_le assumption | <srcrs2_prf <srcrs1_prf |
---|
565 | whd >(plus_n_O (S k)) @le_plus // ] |
---|
566 | | >length_map % |
---|
567 | | >length_map |
---|
568 | >length_ltl |
---|
569 | >tmp_destrs_dummy_prf >length_append |
---|
570 | >length_make_list |
---|
571 | normalize in ⊢ (???(?%?)); |
---|
572 | cases n in k_prf; [ #ABS cases (absurd ? ABS ?) /2 by le_to_not_lt/ ] |
---|
573 | #n' #k_prf >minus_S_S <minus_n_O |
---|
574 | >plus_minus_commutative [%] @le_S_S_to_le assumption |
---|
575 | ] qed. |
---|
576 | |
---|
577 | definition translate_mul : ∀globals.?→?→?→?→?→bind_new register (list (joint_seq RTL globals)) ≝ |
---|
578 | λglobals : list ident. |
---|
579 | λdestrs : list register. |
---|
580 | λsrcrs1 : list psd_argument. |
---|
581 | λsrcrs2 : list psd_argument. |
---|
582 | λsrcrs1_prf : |destrs| = |srcrs1|. |
---|
583 | λsrcrs2_prf : |srcrs1| = |srcrs2|. |
---|
584 | (* needed fresh registers *) |
---|
585 | νa in |
---|
586 | νb in |
---|
587 | (* temporary registers for the result are created, so to avoid overwriting |
---|
588 | sources *) |
---|
589 | νν |destrs| as tmp_destrs with tmp_destrs_prf in |
---|
590 | νdummy in |
---|
591 | (* the step calculating all products with least significant byte going in the |
---|
592 | k-th position of the result *) |
---|
593 | let translate_mul_k : (Σk.k<|destrs|) → bind_new register (list (joint_seq RTL globals)) → |
---|
594 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
595 | λk_sig,acc.match k_sig with |
---|
596 | [ mk_Sig k k_prf ⇒ |
---|
597 | foldr … (translate_mul_i ? a b (|destrs|) |
---|
598 | (tmp_destrs @ [dummy]) srcrs1 srcrs2 |
---|
599 | ? srcrs1_prf srcrs2_prf k k_prf) acc (range_strong (S k)) |
---|
600 | ] in |
---|
601 | (* initializing tmp_destrs to zero |
---|
602 | dummy is intentionally uninitialized *) |
---|
603 | translate_fill_with_zero … tmp_destrs @@ |
---|
604 | (* the main body, roughly: |
---|
605 | for k in 0 ... n-1 do |
---|
606 | for i in 0 ... k do |
---|
607 | translate_mul_i … k … i *) |
---|
608 | foldr … translate_mul_k [ ] (range_strong (|destrs|)) @@ |
---|
609 | (* epilogue: saving the result *) |
---|
610 | translate_move … destrs (map … (Reg ?) tmp_destrs) ?. |
---|
611 | @hide_prf |
---|
612 | [ >length_map >tmp_destrs_prf // |
---|
613 | | >length_append <plus_n_Sm <plus_n_O // |
---|
614 | ] |
---|
615 | qed. |
---|
616 | |
---|
617 | definition translate_divumodu8 : ∀globals.?→?→?→?→?→?→ |
---|
618 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
619 | λglobals: list ident. |
---|
620 | λdiv_not_mod: bool. |
---|
621 | λdestrs: list register. |
---|
622 | λsrcrs1: list psd_argument. |
---|
623 | λsrcrs2: list psd_argument. |
---|
624 | λsrcrs1_prf : |destrs| = |srcrs1|. |
---|
625 | λsrcrs2_prf : |srcrs1| = |srcrs2|. |
---|
626 | match destrs return λx.x = destrs → bind_new ?? with |
---|
627 | [ nil ⇒ λ_.[ ] |
---|
628 | | cons destr destrs' ⇒ λeq_destrs. |
---|
629 | match destrs' with |
---|
630 | [ nil ⇒ |
---|
631 | match srcrs1 return λx.x = srcrs1 → bind_new ?? with |
---|
632 | [ nil ⇒ λeq_srcrs1.⊥ |
---|
633 | | cons srcr1 srcrs1' ⇒ λeq_srcrs1. |
---|
634 | match srcrs2 return λx.x = srcrs2 → bind_new ?? with |
---|
635 | [ nil ⇒ λeq_srcrs2.⊥ |
---|
636 | | cons srcr2 srcrs2' ⇒ λeq_srcrs2. |
---|
637 | νdummy in |
---|
638 | let 〈destr1, destr2〉 ≝ |
---|
639 | if div_not_mod then 〈destr, dummy〉 else 〈dummy, destr〉 in |
---|
640 | [❮destr1, destr2❯ ← srcr1 .DivuModu. srcr2] |
---|
641 | ] (refl …) |
---|
642 | ] (refl …) |
---|
643 | | _ ⇒ ? (* not implemented *) |
---|
644 | ] |
---|
645 | ] (refl …). |
---|
646 | [3: elim not_implemented] |
---|
647 | @hide_prf |
---|
648 | destruct normalize in srcrs1_prf srcrs2_prf; destruct qed. |
---|
649 | |
---|
650 | (* Paolo: to be moved elsewhere *) |
---|
651 | let rec foldr2 (A : Type[0]) (B : Type[0]) (C : Type[0]) (f : A→B→C→C) (init : C) (l1 : list A) (l2 : list B) |
---|
652 | (prf : |l1| = |l2|) on l1 : C ≝ |
---|
653 | match l1 return λx.x = l1 → C with |
---|
654 | [ nil ⇒ λ_.init |
---|
655 | | cons a l1' ⇒ λeq_l1. |
---|
656 | match l2 return λy.y = l2 → C with |
---|
657 | [ nil ⇒ λeq_l2.⊥ |
---|
658 | | cons b l2' ⇒ λeq_l2. |
---|
659 | f a b (foldr2 A B C f init l1' l2' ?) |
---|
660 | ] (refl …) |
---|
661 | ] (refl …). |
---|
662 | @hide_prf |
---|
663 | destruct normalize in prf; [destruct|//] |
---|
664 | qed. |
---|
665 | |
---|
666 | definition translate_ne: ∀globals: list ident.?→?→?→?→ |
---|
667 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
668 | λglobals: list ident. |
---|
669 | λdestrs: list register. |
---|
670 | λsrcrs1: list psd_argument. |
---|
671 | λsrcrs2: list psd_argument. |
---|
672 | match destrs return λ_.|srcrs1| = |srcrs2| → bind_new ?? with |
---|
673 | [ nil ⇒ λ_.[ ] |
---|
674 | | cons destr destrs' ⇒ λEQ. |
---|
675 | translate_fill_with_zero … destrs' @@ |
---|
676 | match srcrs1 return λx.|x| = |srcrs2| → bind_new ?? with |
---|
677 | [ nil ⇒ λ_.[destr ← zero_byte] |
---|
678 | | cons srcr1 srcrs1' ⇒ |
---|
679 | match srcrs2 return λx.S (|srcrs1'|) = |x| → bind_new ?? with |
---|
680 | [ nil ⇒ λEQ.⊥ |
---|
681 | | cons srcr2 srcrs2' ⇒ λEQ. |
---|
682 | νtmpr in |
---|
683 | let f : psd_argument → psd_argument → list (joint_seq RTL globals) → list (joint_seq RTL globals) ≝ |
---|
684 | λs1,s2,acc. |
---|
685 | tmpr ← s1 .Xor. s2 :: |
---|
686 | destr ← destr .Or. tmpr :: |
---|
687 | acc in |
---|
688 | let epilogue : list (joint_seq RTL globals) ≝ |
---|
689 | [ CLEAR_CARRY ?? ; |
---|
690 | tmpr ← zero_byte .Sub. destr ; |
---|
691 | (* now carry bit is 1 iff destrs != 0 *) |
---|
692 | destr ← zero_byte .Addc. zero_byte ] in |
---|
693 | destr ← srcr1 .Xor. srcr2 :: |
---|
694 | foldr2 ??? f epilogue srcrs1' srcrs2' ? |
---|
695 | ] |
---|
696 | ] EQ |
---|
697 | ]. |
---|
698 | @hide_prf normalize in EQ; destruct(EQ) assumption qed. |
---|
699 | |
---|
700 | (* if destrs is 0 or 1, it inverses it. To be used after operations that |
---|
701 | ensure this. *) |
---|
702 | definition translate_toggle_bool : ∀globals.?→bind_new register (list (joint_seq RTL globals)) ≝ |
---|
703 | λglobals: list ident. |
---|
704 | λdestrs: list register. |
---|
705 | match destrs with |
---|
706 | [ nil ⇒ [ ] |
---|
707 | | cons destr _ ⇒ [destr ← .Cmpl. destr] |
---|
708 | ]. |
---|
709 | |
---|
710 | definition translate_lt_unsigned : |
---|
711 | ∀globals. |
---|
712 | ∀destrs: list register. |
---|
713 | ∀srcrs1: list psd_argument. |
---|
714 | ∀srcrs2: list psd_argument. |
---|
715 | |srcrs1| = |srcrs2| → |
---|
716 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
717 | λglobals,destrs,srcrs1,srcrs2,srcrs_prf. |
---|
718 | match destrs with |
---|
719 | [ nil ⇒ [ ] |
---|
720 | | cons destr destrs' ⇒ |
---|
721 | ν tmpr in |
---|
722 | (translate_fill_with_zero … destrs' @@ |
---|
723 | (* I perform a subtraction, but the only interest is in the carry bit *) |
---|
724 | translate_op ? Sub (make_list … tmpr (|srcrs1|)) srcrs1 srcrs2 ? srcrs_prf @@ |
---|
725 | [ destr ← zero_byte .Addc. zero_byte ]) |
---|
726 | ]. |
---|
727 | @hide_prf |
---|
728 | >length_make_list % qed. |
---|
729 | |
---|
730 | (* shifts signed integers by adding 128 to the most significant byte |
---|
731 | it replaces it with a fresh register which must be provided *) |
---|
732 | let rec shift_signed globals |
---|
733 | (tmp : register) |
---|
734 | (srcrs : list psd_argument) on srcrs : |
---|
735 | Σt : (list psd_argument) × (list (joint_seq RTL globals)).|\fst t| = |srcrs| ≝ |
---|
736 | let byte_128 : Byte ≝ true ::: bv_zero ? in |
---|
737 | match srcrs with |
---|
738 | [ nil ⇒ 〈[ ],[ ]〉 |
---|
739 | | cons srcr srcrs' ⇒ |
---|
740 | match srcrs' with |
---|
741 | [ nil ⇒ 〈[ Reg ? tmp ], [ tmp ← srcr .Add. byte_128 ]〉 |
---|
742 | | _ ⇒ |
---|
743 | let re ≝ shift_signed globals tmp srcrs' in |
---|
744 | 〈srcr :: \fst re, \snd re〉 |
---|
745 | ] |
---|
746 | ]. |
---|
747 | @hide_prf |
---|
748 | [1,2: % |
---|
749 | |*: cases re * #a #b >p1 normalize #EQ >EQ % |
---|
750 | ] qed. |
---|
751 | |
---|
752 | definition translate_lt_signed : |
---|
753 | ∀globals. |
---|
754 | ∀destrs: list register. |
---|
755 | ∀srcrs1: list psd_argument. |
---|
756 | ∀srcrs2: list psd_argument. |
---|
757 | |srcrs1| = |srcrs2| → |
---|
758 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
759 | λglobals,destrs,srcrs1,srcrs2,srcrs_prf. |
---|
760 | νtmp_last_s1 in |
---|
761 | νtmp_last_s2 in |
---|
762 | let p1 ≝ shift_signed globals tmp_last_s1 srcrs1 in |
---|
763 | let new_srcrs1 ≝ \fst p1 in |
---|
764 | let shift_srcrs1 ≝ \snd p1 in |
---|
765 | let p2 ≝ shift_signed globals tmp_last_s2 srcrs2 in |
---|
766 | let new_srcrs2 ≝ \fst p2 in |
---|
767 | let shift_srcrs2 ≝ \snd p2 in |
---|
768 | shift_srcrs1 @@ shift_srcrs2 @@ |
---|
769 | translate_lt_unsigned globals destrs new_srcrs1 new_srcrs2 ?. |
---|
770 | @hide_prf |
---|
771 | whd in match new_srcrs1; whd in match new_srcrs2; |
---|
772 | cases p1 |
---|
773 | cases p2 |
---|
774 | // |
---|
775 | qed. |
---|
776 | |
---|
777 | definition translate_lt : bool→∀globals.?→?→?→?→bind_new register (list (joint_seq RTL globals)) ≝ |
---|
778 | λis_unsigned,globals,destrs,srcrs1,srcrs2,srcrs_prf. |
---|
779 | if is_unsigned then |
---|
780 | translate_lt_unsigned globals destrs srcrs1 srcrs2 srcrs_prf |
---|
781 | else |
---|
782 | translate_lt_signed globals destrs srcrs1 srcrs2 srcrs_prf. |
---|
783 | |
---|
784 | definition translate_cmp ≝ |
---|
785 | λis_unsigned,globals,cmp,destrs,srcrs1,srcrs2,srcrs_prf. |
---|
786 | match cmp with |
---|
787 | [ Ceq ⇒ |
---|
788 | translate_ne globals destrs srcrs1 srcrs2 srcrs_prf @@ |
---|
789 | translate_toggle_bool globals destrs |
---|
790 | | Cne ⇒ |
---|
791 | translate_ne globals destrs srcrs1 srcrs2 srcrs_prf |
---|
792 | | Clt ⇒ |
---|
793 | translate_lt is_unsigned globals destrs srcrs1 srcrs2 srcrs_prf |
---|
794 | | Cgt ⇒ |
---|
795 | translate_lt is_unsigned globals destrs srcrs2 srcrs1 ? |
---|
796 | | Cle ⇒ |
---|
797 | translate_lt is_unsigned globals destrs srcrs2 srcrs1 ? @@ |
---|
798 | translate_toggle_bool globals destrs |
---|
799 | | Cge ⇒ |
---|
800 | translate_lt is_unsigned globals destrs srcrs1 srcrs2 srcrs_prf @@ |
---|
801 | translate_toggle_bool globals destrs |
---|
802 | ]. @sym_eq assumption qed. |
---|
803 | |
---|
804 | definition translate_op2 : |
---|
805 | ∀globals.∀ty1,ty2,ty3.∀op : binary_operation ty1 ty2 ty3. |
---|
806 | ∀destrs : list register. |
---|
807 | ∀srcrs1,srcrs2 : list psd_argument. |
---|
808 | |destrs| = size_of_sig_type ty3 → |
---|
809 | |srcrs1| = size_of_sig_type ty1 → |
---|
810 | |srcrs2| = size_of_sig_type ty2 → |
---|
811 | bind_new register (list (joint_seq RTL globals)) ≝ |
---|
812 | λglobals,ty1,ty2,ty3,op2,destrs,srcrs1,srcrs2. |
---|
813 | match op2 return λty1,ty2,ty3.λx : binary_operation ty1 ty2 ty3. |
---|
814 | ? = size_of_sig_type ty3 → ? = size_of_sig_type ty1 → ? = size_of_sig_type ty2 → |
---|
815 | bind_new ?? with |
---|
816 | [ Oadd sz sg ⇒ λprf1,prf2,prf3. |
---|
817 | translate_op globals Add destrs srcrs1 srcrs2 ?? |
---|
818 | | Oaddpi sz ⇒ λprf1,prf2,prf3. |
---|
819 | translate_op_asym_signed globals Add destrs srcrs1 srcrs2 |
---|
820 | | Oaddip sz ⇒ λprf1,prf2,prf3. |
---|
821 | translate_op_asym_signed globals Add destrs srcrs2 srcrs1 |
---|
822 | | Osub sz sg ⇒ λprf1,prf2,prf2. |
---|
823 | translate_op globals Sub destrs srcrs1 srcrs2 ?? |
---|
824 | | Osubpi sz ⇒ λprf1,prf2,prf3. |
---|
825 | translate_op_asym_signed globals Add destrs srcrs1 srcrs2 |
---|
826 | | Osubpp sz ⇒ λprf1,prf2,prf3. |
---|
827 | translate_op_asym_unsigned globals Sub destrs srcrs1 srcrs2 |
---|
828 | | Omul sz sg ⇒ λprf1,prf2,prf3. |
---|
829 | translate_mul globals destrs srcrs1 srcrs2 ?? |
---|
830 | | Odivu sz ⇒ λprf1,prf2,prf3. |
---|
831 | translate_divumodu8 globals true destrs srcrs1 srcrs2 ?? |
---|
832 | | Omodu sz ⇒ λprf1,prf2,prf3. |
---|
833 | translate_divumodu8 globals false destrs srcrs1 srcrs2 ?? |
---|
834 | | Oand sz sg ⇒ λprf1,prf2,prf3. |
---|
835 | translate_op globals And destrs srcrs1 srcrs2 ?? |
---|
836 | | Oor sz sg ⇒ λprf1,prf2,prf3. |
---|
837 | translate_op globals Or destrs srcrs1 srcrs2 ?? |
---|
838 | | Oxor sz sg ⇒ λprf1,prf2,prf3. |
---|
839 | translate_op globals Xor destrs srcrs1 srcrs2 ?? |
---|
840 | | Ocmp sz sg1 sg2 c ⇒ λprf1,prf2,prf3. |
---|
841 | translate_cmp false globals c destrs srcrs1 srcrs2 ? |
---|
842 | | Ocmpu sz sg c ⇒ λprf1,prf2,prf3. |
---|
843 | translate_cmp true globals c destrs srcrs1 srcrs2 ? |
---|
844 | | Ocmpp sg c ⇒ λprf1,prf2,prf3. |
---|
845 | let is_Ocmpp ≝ 0 in |
---|
846 | translate_cmp true globals c destrs srcrs1 srcrs2 ? |
---|
847 | | _ ⇒ ⊥ (* assert false, implemented in run time or float op *) |
---|
848 | ]. try @not_implemented @hide_prf // |
---|
849 | qed. |
---|
850 | |
---|
851 | definition translate_cond: ∀globals: list ident. list register → label → |
---|
852 | bind_step_block RTL globals ≝ |
---|
853 | λglobals: list ident. |
---|
854 | λsrcrs: list register. |
---|
855 | λlbl_true: label. |
---|
856 | match srcrs return λ_.bind_step_block RTL ? with |
---|
857 | [ nil ⇒ bret … [ ] |
---|
858 | | cons srcr srcrs' ⇒ |
---|
859 | ν tmpr in |
---|
860 | let f : register → joint_seq RTL globals ≝ |
---|
861 | λr. tmpr ← tmpr .Or. r in |
---|
862 | bret … ((MOVE RTL globals 〈tmpr,srcr〉 :: |
---|
863 | map ?? f srcrs' : list (joint_step ??)) @ [COND … tmpr lbl_true]) |
---|
864 | ]. |
---|
865 | |
---|
866 | (* Paolo: to be controlled (original seemed overly complex) *) |
---|
867 | definition translate_load : ∀globals.?→?→?→bind_new register ? ≝ |
---|
868 | λglobals: list ident. |
---|
869 | λaddr : list psd_argument. |
---|
870 | λaddr_prf: 2 = |addr|. |
---|
871 | λdestrs: list register. |
---|
872 | ν tmp_addr_l in |
---|
873 | ν tmp_addr_h in |
---|
874 | (translate_move … [tmp_addr_l ; tmp_addr_h] addr ? @@ |
---|
875 | let f ≝ λdestr : register.λacc : bind_new register (list (joint_seq RTL globals)). |
---|
876 | [LOAD RTL globals destr (Reg ? tmp_addr_l) (Reg ? tmp_addr_h)] @@ |
---|
877 | translate_op ? Add |
---|
878 | [tmp_addr_l ; tmp_addr_h] |
---|
879 | [tmp_addr_l ; tmp_addr_h] |
---|
880 | [zero_byte ; (int_size : Byte)] ? ? @@ acc in |
---|
881 | foldr … f [ ] destrs). |
---|
882 | @hide_prf |
---|
883 | [ <addr_prf ] % qed. |
---|
884 | |
---|
885 | definition translate_store : ∀globals.?→?→?→bind_new register ? ≝ |
---|
886 | λglobals: list ident. |
---|
887 | λaddr : list psd_argument. |
---|
888 | λaddr_prf: 2 = |addr|. |
---|
889 | λsrcrs: list psd_argument. |
---|
890 | ν tmp_addr_l in |
---|
891 | ν tmp_addr_h in |
---|
892 | (translate_move … [tmp_addr_l ; tmp_addr_h] addr ? @@ |
---|
893 | let f ≝ λsrcr : psd_argument.λacc : bind_new register (list (joint_seq RTL globals)). |
---|
894 | [STORE RTL globals (Reg ? tmp_addr_l) (Reg ? tmp_addr_h) srcr] @@ |
---|
895 | translate_op … Add |
---|
896 | [tmp_addr_l ; tmp_addr_h] |
---|
897 | [tmp_addr_l ; tmp_addr_h] |
---|
898 | [zero_byte ; (int_size : Byte)] ? ? @@ acc in |
---|
899 | foldr … f [ ] srcrs). |
---|
900 | @hide_prf [ <addr_prf ] % qed. |
---|
901 | |
---|
902 | lemma lenv_typed_reg_typed_ok1 : |
---|
903 | ∀locals,env,r,ty. |
---|
904 | local_env_typed locals env → |
---|
905 | Exists ? (λx:register×typ.〈r,ty〉=x) locals → |
---|
906 | ∀prf. |
---|
907 | |find_local_env r env prf| = size_of_sig_type ty. |
---|
908 | #locals #env #r #ty #env_typed #r_ok |
---|
909 | elim (Exists_All … r_ok env_typed) |
---|
910 | * #r' #ty' * #EQ destruct(EQ) * #rs * #EQ1 #EQ2 #prf |
---|
911 | whd in match find_local_env; normalize nodelta |
---|
912 | @opt_safe_elim #rs' >EQ1 #EQ' destruct assumption |
---|
913 | qed. |
---|
914 | |
---|
915 | lemma lenv_typed_arg_typed_ok1 : |
---|
916 | ∀locals,env,r,ty. |
---|
917 | local_env_typed locals env → |
---|
918 | Exists ? (λx:register×typ.〈r,ty〉=x) locals → |
---|
919 | ∀prf. |
---|
920 | |find_local_env_arg r env prf| = size_of_sig_type ty. |
---|
921 | #locals #env #r #ty #env_typed #r_ok #prf |
---|
922 | whd in match find_local_env_arg; normalize nodelta |
---|
923 | >length_map @lenv_typed_reg_typed_ok1 assumption |
---|
924 | qed. |
---|
925 | |
---|
926 | lemma lenv_typed_reg_typed_ok2 : |
---|
927 | ∀locals,env,r,ty. |
---|
928 | local_env_typed locals env → |
---|
929 | Exists ? (λx:register×typ.〈r,ty〉=x) locals → |
---|
930 | r ∈ env. |
---|
931 | #locals #env #r #ty #env_typed #r_ok |
---|
932 | elim (Exists_All … r_ok env_typed) |
---|
933 | * #r' #ty' * #EQ destruct(EQ) * #rs * #EQ1 #EQ2 |
---|
934 | whd in ⊢ (?%); |
---|
935 | >EQ1 % |
---|
936 | qed. |
---|
937 | |
---|
938 | lemma cst_size_ok : ∀ty,cst.size_of_sig_type ty=size_of_cst ty cst. |
---|
939 | #ty * -ty [*] // |
---|
940 | qed. |
---|
941 | |
---|
942 | definition ensure_bind_step_block : ∀p : params.∀g. |
---|
943 | bind_new register (list (joint_step p g)) → bind_step_block p g ≝ |
---|
944 | λp,g,b.! l ← b; bret ? (step_block p g) l. |
---|
945 | |
---|
946 | definition ensure_bind_step_block_from_seq : ∀p : params.∀g. |
---|
947 | bind_new register (list (joint_seq p g)) → bind_step_block p g ≝ |
---|
948 | λp,g,b.! l ← b; bret ? (step_block p g) l. |
---|
949 | |
---|
950 | coercion bind_step_block_from_bind_list nocomposites : |
---|
951 | ∀p : params.∀g. |
---|
952 | ∀b : bind_new register (list (joint_step p g)).bind_step_block p g ≝ |
---|
953 | ensure_bind_step_block on _b : bind_new register (list (joint_step ??)) to |
---|
954 | bind_step_block ??. |
---|
955 | |
---|
956 | coercion bind_step_block_from_bind_seq_list nocomposites : |
---|
957 | ∀p : params.∀g. |
---|
958 | ∀b : bind_new register (list (joint_seq p g)).bind_step_block p g ≝ |
---|
959 | ensure_bind_step_block_from_seq on _b : bind_new register (list (joint_seq ??)) to |
---|
960 | bind_step_block ??. |
---|
961 | |
---|
962 | definition translate_statement : ∀globals, locals.∀env. |
---|
963 | local_env_typed locals env → |
---|
964 | ∀stmt : statement.statement_typed locals stmt → |
---|
965 | 𝚺b : |
---|
966 | bind_step_block RTL globals + |
---|
967 | bind_fin_block RTL globals. |
---|
968 | match b with [ inl _ ⇒ label | _ ⇒ unit] ≝ |
---|
969 | λglobals,locals,lenv,lenv_typed,stmt. |
---|
970 | match stmt return λstmt.statement_typed locals stmt → 𝚺b: bind_step_block RTL globals ⊎ bind_fin_block RTL globals.match b with [ inl _ ⇒ label | _ ⇒ unit ] with |
---|
971 | [ St_skip lbl' ⇒ λstmt_typed. |
---|
972 | ❬inl … (bret … [ ]), lbl'❭ |
---|
973 | | St_cost cost_lbl lbl' ⇒ λstmt_typed. |
---|
974 | ❬inl … (bret … [COST_LABEL … cost_lbl]), lbl'❭ |
---|
975 | | St_const ty destr cst lbl' ⇒ λstmt_typed. |
---|
976 | ❬inl … (translate_cst ty globals cst (find_local_env destr lenv ?) ?), lbl'❭ |
---|
977 | | St_op1 ty ty' op1 destr srcr lbl' ⇒ λstmt_typed. |
---|
978 | ❬inl … (translate_op1 globals ty' ty op1 |
---|
979 | (find_local_env destr lenv ?) |
---|
980 | (find_local_env srcr lenv ?) ??), lbl'❭ |
---|
981 | | St_op2 ty1 ty2 ty3 op2 destr srcr1 srcr2 lbl' ⇒ λstmt_typed. |
---|
982 | ❬inl … (translate_op2 globals … op2 |
---|
983 | (find_local_env destr lenv ?) |
---|
984 | (find_local_env_arg srcr1 lenv ?) |
---|
985 | (find_local_env_arg srcr2 lenv ?) ???), lbl'❭ |
---|
986 | (* XXX: should we be ignoring this? *) |
---|
987 | | St_load ignore addr destr lbl' ⇒ λstmt_typed. |
---|
988 | ❬inl … (translate_load globals |
---|
989 | (find_local_env_arg addr lenv ?) ? (find_local_env destr lenv ?)), lbl'❭ |
---|
990 | (* XXX: should we be ignoring this? *) |
---|
991 | | St_store ignore addr srcr lbl' ⇒ λstmt_typed. |
---|
992 | ❬inl … (translate_store globals (find_local_env_arg addr lenv ?) ? |
---|
993 | (find_local_env_arg srcr lenv ?)), lbl'❭ |
---|
994 | | St_call_id f args retr lbl' ⇒ λstmt_typed. |
---|
995 | ❬inl ?? (bret ?? [CALL RTL ? (inl … f) (rtl_args args lenv ?) |
---|
996 | (match retr with |
---|
997 | [ Some retr ⇒ find_local_env retr lenv ? |
---|
998 | | None ⇒ [ ] |
---|
999 | ])]), lbl'❭ |
---|
1000 | | St_call_ptr f args retr lbl' ⇒ λstmt_typed. |
---|
1001 | let fs ≝ find_and_addr_arg f lenv ?? in |
---|
1002 | ❬inl … (bret … [CALL RTL ? (inr ?? fs) (rtl_args args lenv ?) |
---|
1003 | (match retr with |
---|
1004 | [ Some retr ⇒ |
---|
1005 | find_local_env retr lenv ? |
---|
1006 | | None ⇒ [ ] |
---|
1007 | ])]), lbl'❭ |
---|
1008 | | St_cond r lbl_true lbl_false ⇒ λstmt_typed. |
---|
1009 | ❬inl … (translate_cond globals (find_local_env r lenv ?) lbl_true), lbl_false❭ |
---|
1010 | | St_return ⇒ λ_. ❬inr … (bret … 〈[ ], RETURN ?〉),it❭ |
---|
1011 | ]. |
---|
1012 | @hide_prf |
---|
1013 | [ cases daemon (* needs more hypotheses *) |
---|
1014 | | @(lenv_typed_reg_typed_ok2 … lenv_typed stmt_typed) |
---|
1015 | | >(lenv_typed_reg_typed_ok1 … lenv_typed stmt_typed) |
---|
1016 | @cst_size_ok |
---|
1017 | |4,5,6,7: cases stmt_typed #dstr_typed #srcr_typed |
---|
1018 | [3,4: @(lenv_typed_reg_typed_ok1 … lenv_typed) assumption |
---|
1019 | |2: @(lenv_typed_reg_typed_ok2 … lenv_typed srcr_typed) |
---|
1020 | |1: @(lenv_typed_reg_typed_ok2 … lenv_typed dstr_typed) |
---|
1021 | ] |
---|
1022 | |8,9,10,11,12,13: |
---|
1023 | cases stmt_typed * #srcrs1_prf #srcrs2_prf #dstr_prf |
---|
1024 | [5,6: @(lenv_typed_arg_typed_ok1 … lenv_typed) assumption |
---|
1025 | |4: @(lenv_typed_reg_typed_ok1 … lenv_typed) assumption |
---|
1026 | |3: @(lenv_typed_reg_typed_ok2 … lenv_typed srcrs2_prf) |
---|
1027 | |2: @(lenv_typed_reg_typed_ok2 … lenv_typed srcrs1_prf) |
---|
1028 | |1: @(lenv_typed_reg_typed_ok2 … lenv_typed dstr_prf) |
---|
1029 | ] |
---|
1030 | |*: cases daemon (* TODO *) |
---|
1031 | ] |
---|
1032 | qed. |
---|
1033 | |
---|
1034 | definition translate_internal : |
---|
1035 | ∀globals.internal_function → (* insert here more properties *) |
---|
1036 | joint_closed_internal_function RTL globals ≝ |
---|
1037 | λglobals: list ident. |
---|
1038 | λdef. |
---|
1039 | let runiverse' ≝ f_reggen def in |
---|
1040 | let luniverse' ≝ f_labgen def in |
---|
1041 | let stack_size' ≝ f_stacksize def in |
---|
1042 | let entry' ≝ f_entry def in |
---|
1043 | let init ≝ init_graph_if RTL globals luniverse' runiverse' [ ] |
---|
1044 | [ ] stack_size' (pi1 … entry') in |
---|
1045 | let 〈init',lenv〉 as pr_eq ≝ initialize_locals_params_ret globals |
---|
1046 | (f_locals def) (f_params def) (f_result def) init in |
---|
1047 | let vars ≝ (f_locals def) @ (f_params def) @ |
---|
1048 | match f_result def with [ Some x ⇒ [x] | _ ⇒ [ ] ] in |
---|
1049 | let f_trans ≝ λlbl,stmt,def. |
---|
1050 | let pr ≝ translate_statement … vars lenv ? stmt ? in |
---|
1051 | match dpi1 … pr return λx.match x with [ inl _ ⇒ label | _ ⇒ unit ] → ? with |
---|
1052 | [ inl instrs ⇒ λlbl'.b_adds_graph … instrs lbl lbl' def |
---|
1053 | | inr instrs ⇒ λ_.b_fin_adds_graph … instrs lbl def |
---|
1054 | ] (dpi2 … pr) in |
---|
1055 | foldi … f_trans (f_graph def) init'. (* TODO *) cases daemon |
---|
1056 | qed. |
---|
1057 | |
---|
1058 | (*CSC: here we are not doing any alignment on variables, like it is done in OCaml |
---|
1059 | because of CompCert heritage *) |
---|
1060 | definition rtlabs_to_rtl: RTLabs_program → rtl_program ≝ |
---|
1061 | λp. transform_program … p (λvarnames. transf_fundef … (translate_internal varnames)). |
---|