include "RTLabs/syntax.ma". include "RTL/RTL.ma". include "common/AssocList.ma". include "common/FrontEndOps.ma". include "common/Graphs.ma". include "joint/TranslateUtils.ma". include "utilities/bindLists.ma". include alias "ASM/BitVector.ma". include alias "arithmetics/nat.ma". definition rtl_fresh_reg: ∀globals.freshT RTL globals register ≝ λglobals,def. let 〈r, runiverse〉 ≝ fresh … (joint_if_runiverse … def) in 〈set_locals ?? (set_runiverse ?? def runiverse)(r::joint_if_locals ?? def), r〉. definition rtl_fresh_reg_no_local : ∀globals.freshT RTL globals register ≝ λglobals,def. let 〈r, runiverse〉 ≝ fresh … (joint_if_runiverse … def) in 〈set_runiverse ?? def runiverse, r〉. definition size_of_sig_type ≝ λsig. match sig with [ ASTint isize sign ⇒ match isize with [ I8 ⇒ 1 | I16 ⇒ 2 | I32 ⇒ 4 ] | ASTfloat _ ⇒ ? (* dpm: not implemented *) | ASTptr ⇒ 2 (* rgn ⇒ nat_of_bitvector ? ptr_size *) ]. cases not_implemented qed. inductive register_type: Type[0] ≝ | register_int: register → register_type | register_ptr: register → register → register_type. definition local_env ≝ identifier_map RegisterTag (list register). definition local_env_typed : list (register × typ) → local_env → Prop ≝ λl,env.All ? (λp.let 〈r, ty〉 ≝ p in ∃regs.lookup … env r = Some ? regs ∧ |regs| = size_of_sig_type ty) l. definition find_local_env ≝ λr.λlenv : local_env. λprf : r ∈ lenv.opt_safe … (lookup … lenv r) ?. lapply (in_map_domain … lenv r) >prf * #x #lookup_eq >lookup_eq % #ABS destruct(ABS) qed. lemma find_local_env_elim : ∀P : list register → Prop.∀r. ∀lenv: local_env.∀prf. (∀x.lookup … lenv r = Some ? x → P x) → P (find_local_env r lenv prf). #P#r#lenv#prf #H change with (P (opt_safe ???)) @opt_safe_elim assumption qed. definition find_local_env_arg : register → local_env → ? → list psd_argument ≝ λr,lenv,prf. map … (Reg ?) (find_local_env r lenv prf). definition initialize_local_env_internal : ∀globals. ((joint_internal_function RTL globals) × local_env) → (register×typ) → ((joint_internal_function RTL globals) × local_env) ≝ λglobals,def_env,r_sig. let 〈def,lenv〉 ≝ def_env in let 〈r, sig〉 ≝ r_sig in let size ≝ size_of_sig_type sig in let 〈def,rs〉 ≝ repeat_fresh … (rtl_fresh_reg_no_local globals) size def in 〈def,add … lenv r rs〉. include alias "common/Identifiers.ma". let rec map_list_local_env lenv (regs : list (register×typ)) on regs : All ? (λpr.bool_to_Prop (\fst pr ∈ lenv)) regs → list register ≝ match regs return λx.All ?? x → ? with [ nil ⇒ λ_.[ ] | cons hd tl ⇒ λprf.find_local_env (\fst hd) lenv ? @ map_list_local_env lenv tl ? ].cases prf #A #B assumption qed. definition initialize_local_env : ∀globals. list (register×typ) → freshT RTL globals local_env ≝ λglobals,registers,def. foldl ?? (initialize_local_env_internal globals) 〈def,empty_map …〉 registers. lemma initialize_local_env_in : ∀globals,l,def,r. Exists ? (λx.\fst x = r) l → r ∈ \snd (initialize_local_env globals l def). #globals #l #U #r @(list_elim_left … l) [ * | * #tl #sig #hd #IH #G elim (Exists_append … G) -G whd in match initialize_local_env; normalize nodelta >foldl_step change with (initialize_local_env ???) in match (foldl ?????); [ #H lapply (IH H) | * [2: *] #EQ destruct(EQ) ] cases (initialize_local_env ???) #U' #env' [#G] whd in match initialize_local_env_internal; normalize nodelta elim (repeat_fresh ??????) #U'' #rs [ >mem_set_add @orb_Prop_r assumption | @mem_set_add_id ] qed. example proj1_rewrite : ∀A,B,a,b.∀pr : A×B.〈a,b〉 = pr → a = \fst pr. // qed-. example proj2_rewrite : ∀A,B,a,b.∀pr : A×B.〈a,b〉 = pr → b = \snd pr. // qed-. definition initialize_locals_params_ret : ∀globals. (* locals *) list (register×typ) → (* params *) list (register×typ) → (* return *) option (register×typ) → freshT RTL globals local_env ≝ λglobals,locals,params,ret,def. let 〈def',lenv〉 as EQ ≝ initialize_local_env globals ((match ret with [ Some r_sig ⇒ [r_sig] | None ⇒ [ ] ]) @ locals @ params) def in let locals' ≝ map_list_local_env lenv locals ? in let params' ≝ map_list_local_env lenv params ? in let ret' ≝ match ret return λx.ret = x → ? with [ Some r_sig ⇒ λprf.find_local_env (\fst r_sig) lenv ? | None ⇒ λ_.[ ] ] (refl …) in let def'' ≝ mk_joint_internal_function RTL globals (joint_if_luniverse … def') (joint_if_runiverse … def') ret' params' locals' (joint_if_stacksize … def') (joint_if_code … def') (joint_if_entry … def') (joint_if_exit … def') in 〈def'', lenv〉. @hide_prf [ >(proj2_rewrite ????? EQ) @initialize_local_env_in >prf %1 % |*: >(proj2_rewrite ????? EQ) @(All_mp ??? (λpr.initialize_local_env_in ??? (\fst pr))) [ @(All_mp … (λpr.Exists ? (λx.\fst x = \fst pr) params)) [ #a #H @Exists_append_r @Exists_append_r @H | generalize in match params; ] | @(All_mp … (λpr.Exists ? (λx.\fst x = \fst pr) locals)) [ #a #H @Exists_append_r @Exists_append_l @H | generalize in match locals; ] ] #l elim l [1,3: %] #hd #tl #IH % [1,3: %1 %] @(All_mp … IH) #x #G %2{G} ] qed. definition make_addr ≝ λA. λlst: list A. λprf: 2 = length A lst.〈nth_safe … 0 lst ?, nth_safe … 1 lst ?〉. EQ % ] qed. definition split_into_bytes: ∀size. ∀int: bvint size. Σbytes: list Byte. |bytes| = size_intsize size ≝ λsize.vrsplit ? (size_intsize size) 8. let rec list_inject_All_aux A P (l : list A) on l : All A P l → list (Σx.P x) ≝ match l return λx.All A P x → ? with [ nil ⇒ λ_.[ ] | cons hd tl ⇒ λprf.«hd, ?» :: list_inject_All_aux A P tl ? ]. cases prf #H1 #H2 [@H1 | @H2] qed. include alias "basics/lists/list.ma". definition translate_op: ∀globals. Op2 → ∀dests : list register. ∀srcrs1 : list psd_argument. ∀srcrs2 : list psd_argument. |dests| = |srcrs1| → |srcrs1| = |srcrs2| → list (joint_seq RTL globals) ≝ λglobals: list ident. λop. λdestrs. λsrcrs1. λsrcrs2. λprf1,prf2. (* first, clear carry if op relies on it *) match op with [ Addc ⇒ [CLEAR_CARRY ??] | Sub ⇒ [CLEAR_CARRY ??] | _ ⇒ [ ] ] @ map3 ???? (OP2 RTL globals op) destrs srcrs1 srcrs2 prf1 prf2. definition cast_list : ∀A.A → ℕ → list A → list A ≝ λA,deflt,new_length,l. if leb (|l|) new_length then l @ make_list ? deflt (new_length - |l|) else lhd … l new_length. lemma length_make_list: ∀A: Type[0]. ∀elt: A. ∀n: nat. length ? (make_list A elt n) = n. #A #ELT #N elim N normalize // qed. lemma length_lhd : ∀A,l,n.|lhd A l n| = min (|l|) n. #A #l elim l -l [ * // | #hd #tl #IH * normalize [%] #n >IH normalize elim (leb ??) % ] qed. lemma length_cast_list : ∀A,dflt,n,l.|cast_list A dflt n l| = n. #A #dflt #n #l normalize @leb_elim #H normalize [ >length_append >length_make_list @sym_eq @minus_to_plus // | >length_lhd normalize @leb_elim [ #abs elim (absurd ? abs H) ] #_ % ] qed. definition translate_op_asym_unsigned : ∀globals.Op2 → list register → list psd_argument → list psd_argument → list (joint_seq RTL globals) ≝ λglobals,op,destrs,srcrs1,srcrs2. let l ≝ |destrs| in let srcrs1' ≝ cast_list ? (zero_byte : psd_argument) l srcrs1 in let srcrs2' ≝ cast_list ? (zero_byte : psd_argument) l srcrs2 in translate_op globals op destrs srcrs1' srcrs2' ??. normalize nodelta >length_cast_list [2: >length_cast_list ] % qed. let rec nat_to_args (size : nat) (k : nat) on size : Σl : list psd_argument.|l| = size ≝ match size with [ O ⇒ [ ] | S size' ⇒ (byte_of_nat k : psd_argument) :: nat_to_args size' (k ÷ 8) ]. [ % | cases (nat_to_args ??) #res #EQ normalize >EQ % ] qed. definition size_of_cst ≝ λtyp.λcst : constant typ.match cst with [ Ointconst size _ _ ⇒ size_intsize size | Ofloatconst _ _ ⇒ ? (* not implemented *) | _ ⇒ 2 ]. cases not_implemented qed. definition cst_well_defd : ∀ty.list ident → constant ty → Prop ≝ λty,globals,cst. match cst with [ Oaddrsymbol id _ ⇒ member id (eq_identifier ?) globals | _ ⇒ True ]. definition translate_cst : ∀ty. ∀globals: list ident. ∀cst_sig: Σcst : constant ty.cst_well_defd ty globals cst. ∀destrs: list register. |destrs| = size_of_cst ? cst_sig → list (joint_seq RTL globals) ≝ λty,globals,cst_sig,destrs. match pi1 … cst_sig in constant return λty'.λx : constant ty'. cst_well_defd ty' ? x → |destrs| = size_of_cst ty' x → ? with [ Ointconst size sign const ⇒ λcst_prf,prf. map2 … (λr.λb : Byte.r ← b) destrs (split_into_bytes size const) ? | Ofloatconst _ _ ⇒ ? | Oaddrsymbol id offset ⇒ λcst_prf,prf. let 〈r1, r2〉 ≝ make_addr … destrs ? in [ADDRESS RTL globals id ? r1 r2] | Oaddrstack offset ⇒ λcst_prf,prf. let 〈r1, r2〉 ≝ make_addr … destrs ? in [(rtl_stack_address r1 r2 : joint_seq RTL globals)] ] (pi2 … cst_sig). [2: cases not_implemented |1: cases (split_into_bytes ??) #lst #EQ >EQ >prf whd in ⊢ (??%?); cases size % |3: @cst_prf |*: >prf % ] qed. definition translate_move : ∀globals. ∀destrs: list register. ∀srcrs: list psd_argument. |destrs| = |srcrs| → list (joint_seq RTL globals) ≝ λglobals,destrs,srcrs,length_eq. map2 … (λdst,src.dst ← src) destrs srcrs length_eq. definition sign_mask : ∀globals.register → psd_argument → list (joint_seq RTL globals) ≝ (* this sets destr to 0xFF if s is neg, 0x00 o.w. Done like that: byte in destr if srcr is: neg | pos destr ← srcr | 127 11...1 | 01...1 destr ← destr length_make_list % qed. definition translate_fill_with_zero : ∀globals : list ident. list register → list (joint_seq RTL globals) ≝ λglobals,destrs. match nat_to_args (|destrs|) 0 with [ mk_Sig res prf ⇒ translate_move ? destrs res ?]. // qed. let rec last A (l : list A) on l : option A ≝ match l with [ nil ⇒ None ? | cons hd tl ⇒ match tl with [ nil ⇒ Some ? hd | _ ⇒ last A tl ] ]. lemma last_not_empty : ∀A,l. match l with [ nil ⇒ False | _ ⇒ True ] → match last A l with [ None ⇒ False | _ ⇒ True ]. #A #l elim l [ * ] #hd * [ #_ * % ] #hd' #tl #IH * @(IH I) qed. definition translate_op_asym_signed : ∀globals.Op2 → list register → list psd_argument → list psd_argument → bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λglobals,op,destrs,srcrs1,srcrs2. νtmp1,tmp2 in let l ≝ |destrs| in let f ≝ λsrcrs,tmp. let srcrs_l ≝ |srcrs| in if leb srcrs_l l then match last … srcrs with [ Some last ⇒ 〈srcrs @ make_list … (Reg ? tmp) (l - srcrs_l), sign_mask … tmp last〉 | None ⇒ 〈make_list … (zero_byte : psd_argument) l, [ ]〉 ] else 〈lhd … srcrs l, [ ]〉 in let prf : ∀srcrs,tmp.|destrs| = |\fst (f srcrs tmp)| ≝ ? in let srcrs1init ≝ f srcrs1 tmp1 in let srcrs2init ≝ f srcrs2 tmp2 in \snd srcrs1init @ \snd srcrs2init @ translate_op globals op destrs (\fst srcrs1init) (\fst srcrs2init) ??. [ @prf | length_make_list % | #hd #tl #abs elim(abs I) ] | #last #_ normalize nodelta >length_append >length_make_list @minus_to_plus // ] | >length_lhd normalize @leb_elim [ #G elim (absurd … G H) | #_ % ] ] qed. (* using size of lists as size of ints *) definition translate_cast : ∀globals: list ident. signedness → list register → list register → bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λglobals,src_sign,destrs,srcrs. match reduce_strong ?? destrs srcrs with [ mk_Sig t prf ⇒ let src_common ≝ \fst (\fst t) in let src_rest ≝ \snd (\fst t) in let dst_common ≝ \fst (\snd t) in let dst_rest ≝ \snd (\snd t) in (* first, move the common part *) translate_move ? src_common (map … (Reg ?) dst_common) ? @@ match src_rest return λ_.bind_new ?? with [ nil ⇒ (* upcast *) match src_sign return λ_.bind_new ?? with [ Unsigned ⇒ translate_fill_with_zero ? dst_rest | Signed ⇒ match last … srcrs (* = src_common *) with [ Some src_last ⇒ translate_cast_signed ? dst_rest src_last | None ⇒ (* srcrs is empty *) translate_fill_with_zero ? dst_rest ] ] | _ ⇒ (* downcast, nothing else to do *) [ ] ] ]. >length_map @prf qed. definition translate_notint : ∀globals : list ident. ∀destrs : list register. ∀srcrs_arg : list register. |destrs| = |srcrs_arg| → list (joint_seq RTL globals) ≝ λglobals, destrs, srcrs, prf. map2 ??? (OP1 RTL globals Cmpl) destrs srcrs prf. definition translate_negint : ∀globals.? → ? → ? → bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λglobals: list ident. λdestrs: list register. λsrcrs: list register. λprf: |destrs| = |srcrs|. (* assert in caml code *) translate_notint … destrs srcrs prf @ match nat_to_args (|destrs|) 1 with [ mk_Sig res prf' ⇒ translate_op ? Add destrs (map … (Reg ?) destrs) res ?? ]. >length_map // qed. definition translate_notbool: ∀globals : list ident. list register → list register → bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λglobals,destrs,srcrs. match destrs with [ nil ⇒ [ ] | cons destr destrs' ⇒ translate_fill_with_zero ? destrs' @@ match srcrs return λ_.bind_new ?? with [ nil ⇒ [destr ← zero_byte] | cons srcr srcrs' ⇒ (destr ← srcr) ::: map register (joint_seq RTL globals) (λr. destr ← destr .Or. r) srcrs' @@ (* now destr is non-null iff srcrs was non-null *) CLEAR_CARRY ?? ::: (* many uses of 0, better not use immediates *) ν tmp in [tmp ← zero_byte ; destr ← tmp .Sub. tmp ; (* now carry bit is set iff destr was non-null *) destr ← tmp .Addc. tmp] ] ]. definition translate_op1 : ∀globals.? → ? → ? → ? → ? → ? → ? → bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λglobals. λty, ty'. λop1: unary_operation ty ty'. λdestrs: list register. λsrcrs: list register. λprf1: |destrs| = size_of_sig_type ty'. λprf2: |srcrs| = size_of_sig_type ty. match op1 return λty'',ty'''.λx : unary_operation ty'' ty'''.ty'' = ty → ty''' = ty' → bind_new (localsT RTL) (list (joint_seq RTL globals)) with [ Ocastint _ src_sign _ _ ⇒ λeq1,eq2. translate_cast globals src_sign destrs srcrs | Onegint sz sg ⇒ λeq1,eq2. translate_negint globals destrs srcrs ? | Onotbool _ _ _ _ ⇒ λeq1,eq2. translate_notbool globals destrs srcrs | Onotint sz sg ⇒ λeq1,eq2. translate_notint globals destrs srcrs ? | Optrofint sz sg ⇒ λeq1,eq2. translate_cast globals Unsigned destrs srcrs | Ointofptr sz sg ⇒ λeq1,eq2. translate_cast globals Unsigned destrs srcrs | Oid t ⇒ λeq1,eq2. translate_move globals destrs (map … (Reg ?) srcrs) ? | _ ⇒ λeq1,eq2.? (* float operations implemented in runtime *) ] (refl …) (refl …). [3,4,5,6,7,8,9: (* floats *) cases not_implemented |*: destruct >prf1 >prf2 [3: >length_map ] // ] qed. include alias "arithmetics/nat.ma". let rec range_strong_internal (start : ℕ) (count : ℕ) (* Paolo: no notation to avoid ambiguity *) on count : list (Σn : ℕ.lt n (plus start count)) ≝ match count return λx.count = x → list (Σn : ℕ. n < start + count) with [ O ⇒ λ_.[ ] | S count' ⇒ λEQ. let f : (Σn : ℕ. lt n (S start + count')) → Σn : ℕ. lt n (start + count) ≝ λsig.match sig with [mk_Sig n prf ⇒ n] in start :: map … f (range_strong_internal (S start) count') ] (refl …). destruct(EQ) // qed. definition range_strong : ∀end : ℕ. list (Σn.n(plus_n_O (S k)) @le_plus // ] | length_map // | >length_map >length_ltl >tmp_destrs_dummy_prf >length_append >length_make_list normalize in ⊢ (???(?%?)); >plus_minus_commutative [2: @le_plus_to_minus_r EQ %] >plus_minus_commutative [2: @(transitive_le … k_prf) //] @sym_eq @plus_to_minus % ] qed. definition translate_mul : ∀globals.?→?→?→?→?→bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λglobals : list ident. λdestrs : list register. λsrcrs1 : list psd_argument. λsrcrs2 : list psd_argument. λsrcrs1_prf : |destrs| = |srcrs1|. λsrcrs2_prf : |srcrs1| = |srcrs2|. (* needed fresh registers *) νa in νb in (* temporary registers for the result are created, so to avoid overwriting sources *) νν |destrs| as tmp_destrs with tmp_destrs_prf in νdummy in (* the step calculating all products with least significant byte going in the k-th position of the result *) let translate_mul_k : (Σk.k<|destrs|) → list (joint_seq RTL globals) → list (joint_seq RTL globals) ≝ λk_sig,acc.match k_sig with [ mk_Sig k k_prf ⇒ foldr … (translate_mul_i ? a b (|destrs|) (tmp_destrs @ [dummy]) srcrs1 srcrs2 ? srcrs1_prf srcrs2_prf k k_prf) acc (range_strong (S k)) ] in (* initializing tmp_destrs to zero dummy is intentionally uninitialized *) translate_fill_with_zero … tmp_destrs @ (* the main body, roughly: for k in 0 ... n-1 do for i in 0 ... k do translate_mul_i … k … i *) foldr … translate_mul_k [ ] (range_strong (|destrs|)) @ (* epilogue: saving the result *) translate_move … destrs (map … (Reg ?) tmp_destrs) ?. [ >length_map >tmp_destrs_prf // | >length_append length_make_list % qed. (* shifts signed integers by adding 128 to the most significant byte it replaces it with a fresh register which must be provided *) let rec shift_signed globals (tmp : register) (srcrs : list psd_argument) on srcrs : Σt : (list psd_argument) × (list (joint_seq RTL globals)).|\fst t| = |srcrs| ≝ let byte_128 : Byte ≝ true ::: bv_zero ? in match srcrs with [ nil ⇒ 〈[ ],[ ]〉 | cons srcr srcrs' ⇒ match srcrs' with [ nil ⇒ 〈[ Reg ? tmp ], [ tmp ← srcr .Add. byte_128 ]〉 | _ ⇒ let re ≝ shift_signed globals tmp srcrs' in 〈srcr :: \fst re, \snd re〉 ] ]. [1,2: % |*: cases re * #a #b >p1 normalize #EQ >EQ % ] qed. definition translate_lt_signed : ∀globals. ∀destrs: list register. ∀srcrs1: list psd_argument. ∀srcrs2: list psd_argument. |srcrs1| = |srcrs2| → bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λglobals,destrs,srcrs1,srcrs2,srcrs_prf. νtmp_last_s1 in νtmp_last_s2 in let p1 ≝ shift_signed globals tmp_last_s1 srcrs1 in let new_srcrs1 ≝ \fst p1 in let shift_srcrs1 ≝ \snd p1 in let p2 ≝ shift_signed globals tmp_last_s2 srcrs2 in let new_srcrs2 ≝ \fst p2 in let shift_srcrs2 ≝ \snd p2 in shift_srcrs1 @@ shift_srcrs2 @@ translate_lt_unsigned globals destrs new_srcrs1 new_srcrs2 ?. whd in match new_srcrs1; whd in match new_srcrs2; cases p1 cases p2 // qed. definition translate_lt : bool→∀globals.?→?→?→?→bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λis_unsigned,globals,destrs,srcrs1,srcrs2,srcrs_prf. if is_unsigned then translate_lt_unsigned globals destrs srcrs1 srcrs2 srcrs_prf else translate_lt_signed globals destrs srcrs1 srcrs2 srcrs_prf. definition translate_cmp ≝ λis_unsigned,globals,cmp,destrs,srcrs1,srcrs2,srcrs_prf. match cmp with [ Ceq ⇒ translate_ne globals destrs srcrs1 srcrs2 srcrs_prf @@ translate_toggle_bool globals destrs | Cne ⇒ translate_ne globals destrs srcrs1 srcrs2 srcrs_prf | Clt ⇒ translate_lt is_unsigned globals destrs srcrs1 srcrs2 srcrs_prf | Cgt ⇒ translate_lt is_unsigned globals destrs srcrs2 srcrs1 ? | Cle ⇒ translate_lt is_unsigned globals destrs srcrs2 srcrs1 ? @@ translate_toggle_bool globals destrs | Cge ⇒ translate_lt is_unsigned globals destrs srcrs1 srcrs2 srcrs_prf @@ translate_toggle_bool globals destrs ]. // qed. definition translate_op2 : ∀globals.∀ty1,ty2,ty3.∀op : binary_operation ty1 ty2 ty3. ∀destrs : list register. ∀srcrs1,srcrs2 : list psd_argument. |destrs| = size_of_sig_type ty3 → |srcrs1| = size_of_sig_type ty1 → |srcrs2| = size_of_sig_type ty2 → bind_new (localsT RTL) (list (joint_seq RTL globals)) ≝ λglobals,ty1,ty2,ty3,op2,destrs,srcrs1,srcrs2. match op2 return λty1,ty2,ty3.λx : binary_operation ty1 ty2 ty3. ? = size_of_sig_type ty3 → ? = size_of_sig_type ty1 → ? = size_of_sig_type ty2 → bind_new ?? with [ Oadd sz sg ⇒ λprf1,prf2,prf3. translate_op globals Add destrs srcrs1 srcrs2 ?? | Oaddp sz ⇒ λprf1,prf2,prf3. translate_op_asym_signed globals Add destrs srcrs1 srcrs2 | Osub sz sg ⇒ λprf1,prf2,prf2. translate_op globals Sub destrs srcrs1 srcrs2 ?? | Osubpi sz ⇒ λprf1,prf2,prf3. translate_op_asym_signed globals Add destrs srcrs1 srcrs2 | Osubpp sz ⇒ λprf1,prf2,prf3. translate_op_asym_unsigned globals Sub destrs srcrs1 srcrs2 | Omul sz sg ⇒ λprf1,prf2,prf3. translate_mul globals destrs srcrs1 srcrs2 ?? | Odivu sz ⇒ λprf1,prf2,prf3. translate_divumodu8 globals true destrs srcrs1 srcrs2 ?? | Omodu sz ⇒ λprf1,prf2,prf3. translate_divumodu8 globals false destrs srcrs1 srcrs2 ?? | Oand sz sg ⇒ λprf1,prf2,prf3. translate_op globals And destrs srcrs1 srcrs2 ?? | Oor sz sg ⇒ λprf1,prf2,prf3. translate_op globals Or destrs srcrs1 srcrs2 ?? | Oxor sz sg ⇒ λprf1,prf2,prf3. translate_op globals Xor destrs srcrs1 srcrs2 ?? | Ocmp sz sg1 sg2 c ⇒ λprf1,prf2,prf3. translate_cmp false globals c destrs srcrs1 srcrs2 ? | Ocmpu sz sg c ⇒ λprf1,prf2,prf3. translate_cmp true globals c destrs srcrs1 srcrs2 ? | Ocmpp sg c ⇒ λprf1,prf2,prf3. let is_Ocmpp ≝ 0 in translate_cmp true globals c destrs srcrs1 srcrs2 ? | _ ⇒ ⊥ (* assert false, implemented in run time or float op *) ]. try @not_implemented // qed. definition translate_cond: ∀globals: list ident. list register → label → bind_seq_block RTL globals (joint_step RTL globals) ≝ λglobals: list ident. λsrcrs: list register. λlbl_true: label. match srcrs return λ_.bind_seq_block RTL ? (joint_step ??) with [ nil ⇒ bret … 〈[ ], NOOP ??〉 | cons srcr srcrs' ⇒ ν tmpr in let f : register → joint_seq RTL globals ≝ λr. tmpr ← tmpr .Or. r in bret … 〈MOVE RTL globals 〈tmpr,srcr〉 :: map ?? f srcrs', (COND … tmpr lbl_true : joint_step ??) 〉 ]. (* Paolo: to be controlled (original seemed overly complex) *) definition translate_load : ∀globals.?→?→?→bind_new (localsT RTL) ? ≝ λglobals: list ident. λaddr : list psd_argument. λaddr_prf: 2 = |addr|. λdestrs: list register. ν tmp_addr_l in ν tmp_addr_h in let f ≝ λdestr : register.λacc : list (joint_seq RTL globals). LOAD RTL globals destr (Reg ? tmp_addr_l) (Reg ? tmp_addr_h) :: translate_op ? Add [tmp_addr_l ; tmp_addr_h] [tmp_addr_l ; tmp_addr_h] [zero_byte ; (int_size : Byte)] ? ? @ acc in translate_move … [tmp_addr_l ; tmp_addr_h] addr ? @@ foldr … f [ ] destrs. [1: EQ1 #EQ' destruct assumption qed. lemma lenv_typed_arg_typed_ok1 : ∀locals,env,r,ty. local_env_typed locals env → Exists ? (λx:register×typ.〈r,ty〉=x) locals → ∀prf. |find_local_env_arg r env prf| = size_of_sig_type ty. #locals #env #r #ty #env_typed #r_ok #prf whd in match find_local_env_arg; normalize nodelta >length_map @lenv_typed_reg_typed_ok1 assumption qed. lemma lenv_typed_reg_typed_ok2 : ∀locals,env,r,ty. local_env_typed locals env → Exists ? (λx:register×typ.〈r,ty〉=x) locals → r ∈ env. #locals #env #r #ty #env_typed #r_ok elim (Exists_All … r_ok env_typed) * #r' #ty' * #EQ destruct(EQ) * #rs * #EQ1 #EQ2 whd in ⊢ (?%); >EQ1 % qed. lemma cst_size_ok : ∀ty,cst.size_of_sig_type ty=size_of_cst ty cst. #ty * -ty [*] // qed. definition translate_statement : ∀globals, locals.∀env. local_env_typed locals env → ∀stmt : statement.statement_typed locals stmt → 𝚺b : bind_seq_block RTL globals (joint_step RTL globals) + bind_seq_block RTL globals (joint_fin_step RTL). if is_inl … b then label else unit ≝ λglobals,locals,lenv,lenv_typed,stmt. match stmt return λstmt.statement_typed locals stmt → 𝚺b:bind_seq_block ???+bind_seq_block ???.if is_inl … b then label else unit with [ St_skip lbl' ⇒ λstmt_typed. ❬NOOP ??, lbl'❭ | St_cost cost_lbl lbl' ⇒ λstmt_typed. ❬COST_LABEL … cost_lbl, lbl'❭ | St_const ty destr cst lbl' ⇒ λstmt_typed. ❬translate_cst ty globals cst (find_local_env destr lenv ?) ?, lbl'❭ | St_op1 ty ty' op1 destr srcr lbl' ⇒ λstmt_typed. ❬translate_op1 globals ty' ty op1 (find_local_env destr lenv ?) (find_local_env srcr lenv ?) ??, lbl'❭ | St_op2 ty1 ty2 ty3 op2 destr srcr1 srcr2 lbl' ⇒ λstmt_typed. ❬translate_op2 globals … op2 (find_local_env destr lenv ?) (find_local_env_arg srcr1 lenv ?) (find_local_env_arg srcr2 lenv ?) ???, lbl'❭ (* XXX: should we be ignoring this? *) | St_load ignore addr destr lbl' ⇒ λstmt_typed. ❬translate_load globals (find_local_env_arg addr lenv ?) ? (find_local_env destr lenv ?), lbl'❭ (* XXX: should we be ignoring this? *) | St_store ignore addr srcr lbl' ⇒ λstmt_typed. ❬translate_store globals (find_local_env_arg addr lenv ?) ? (find_local_env_arg srcr lenv ?), lbl'❭ | St_call_id f args retr lbl' ⇒ λstmt_typed. ❬(match retr with [ Some retr ⇒ CALL_ID RTL ? f (rtl_args args lenv ?) (find_local_env retr lenv ?) | None ⇒ CALL_ID RTL ? f (rtl_args args lenv ?) [ ] ] : bind_seq_block ???), lbl'❭ | St_call_ptr f args retr lbl' ⇒ λstmt_typed. let fs ≝ find_and_addr f lenv ?? in ❬(rtl_call_ptr (\fst fs) (\snd fs) (rtl_args args lenv ?) (match retr with [ Some retr ⇒ find_local_env retr lenv ? | None ⇒ [ ] ]) : joint_step ??), lbl'❭ | St_cond r lbl_true lbl_false ⇒ λstmt_typed. ❬translate_cond globals (find_local_env r lenv ?) lbl_true, lbl_false❭ | St_return ⇒ λ_. ❬inr … (RETURN ?),it❭ ]. [ >(lenv_typed_reg_typed_ok1 … lenv_typed stmt_typed) @cst_size_ok | @(lenv_typed_reg_typed_ok2 … lenv_typed stmt_typed) | cases daemon (* needs more hypotheses *) |4,5,6,7: cases stmt_typed #dstr_typed #srcr_typed [1,2: @(lenv_typed_reg_typed_ok1 … lenv_typed) assumption | @(lenv_typed_reg_typed_ok2 … lenv_typed srcr_typed) | @(lenv_typed_reg_typed_ok2 … lenv_typed dstr_typed) ] |8,9,10,11,12,13: cases stmt_typed * #srcrs1_prf #srcrs2_prf #dstr_prf [1,2: @(lenv_typed_arg_typed_ok1 … lenv_typed) assumption | @(lenv_typed_reg_typed_ok1 … lenv_typed) assumption | @(lenv_typed_reg_typed_ok2 … lenv_typed srcrs2_prf) | @(lenv_typed_reg_typed_ok2 … lenv_typed srcrs1_prf) | @(lenv_typed_reg_typed_ok2 … lenv_typed dstr_prf) ] |*: whd in stmt_typed; cases daemon (* TODO need more stringent statement_typed *) ] qed. (* XXX: we use a "hack" here to circumvent the proof obligations required to create res, an empty internal_function record. we insert into our graph the start and end nodes to ensure that they are in, and place dummy skip instructions at these nodes. *) definition translate_internal ≝ λglobals: list ident. λdef. let runiverse' ≝ f_reggen def in let luniverse' ≝ f_labgen def in let stack_size' ≝ f_stacksize def in let entry' ≝ f_entry def in let exit' ≝ f_exit def in let init ≝ init_graph_if RTL globals luniverse' runiverse' [ ] [ ] [ ] stack_size' (pi1 … entry') (pi1 … exit') in let 〈init',lenv〉 as pr_eq ≝ initialize_locals_params_ret globals (f_locals def) (f_params def) (f_result def) init in let vars ≝ (f_locals def) @ (f_params def) @ match f_result def with [ Some x ⇒ [x] | _ ⇒ [ ] ] in let f_trans ≝ λlbl,stmt,def. let pr ≝ translate_statement … vars lenv ? stmt ? in b_adds_graph … (rtl_fresh_reg …) (dpi1 … pr) lbl (dpi2 … pr) def in foldi … f_trans (f_graph def) init'. (* TODO *) cases daemon qed. (*CSC: here we are not doing any alignment on variables, like it is done in OCaml because of CompCert heritage *) definition rtlabs_to_rtl: RTLabs_program → rtl_program ≝ λp. transform_program … p (λvarnames. transf_fundef … (translate_internal varnames)).