include "RTLabs/syntax.ma". include "RTL/RTL_paolo.ma". include "common/AssocList.ma". include "common/FrontEndOps.ma". include "common/Graphs.ma". include "joint/TranslateUtils_paolo.ma". include "utilities/bindLists.ma". include alias "ASM/BitVector.ma". include alias "arithmetics/nat.ma". definition size_of_sig_type ≝ λsig. match sig with [ ASTint isize sign ⇒ let isize' ≝ match isize with [ I8 ⇒ 8 | I16 ⇒ 16 | I32 ⇒ 32 ] in isize' ÷ (nat_of_bitvector ? int_size) | ASTfloat _ ⇒ ? (* dpm: not implemented *) | ASTptr 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 ≝ λr,lenv,prf. map … Reg (find_local_env r lenv prf). definition m_fresh : ∀tag.state_monad (universe tag) (identifier tag) ≝ λtag,univ.let pr ≝ fresh … univ in 〈\snd pr, \fst pr〉. let rec register_freshes (n: nat) on n : state_monad (universe RegisterTag) (Σl.|l| = n) ≝ match n return λx.state_monad ? (Σl.|l| = x) with [ O ⇒ return «[ ],?» | S n' ⇒ ! r ← m_fresh ?; ! res ← register_freshes n'; return «r::res,?» ].[2: cases res -res #res #EQ 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 rtl_argument. ∀srcrs2 : list rtl_argument. |dests| = |srcrs1| → |srcrs1| = |srcrs2| → list (joint_seq rtl_params 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_params globals op) destrs srcrs1 srcrs2 prf1 prf2. let rec nat_to_args (size : nat) (k : nat) on size : Σl : list rtl_argument.|l| = size ≝ match size with [ O ⇒ [ ] | S size' ⇒ match nat_to_args size' (k ÷ 8) with [ mk_Sig res prf ⇒ imm_nat k :: res ] ]. normalize // 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 r id offset ⇒ 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_params 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 : rtl_argument)) destrs (split_into_bytes size const) ? | Ofloatconst _ _ ⇒ ? | Oaddrsymbol r id offset ⇒ λcst_prf,prf. let 〈r1, r2〉 ≝ make_addr … destrs ? in [ADDRESS rtl_params globals id ? r1 r2] | Oaddrstack offset ⇒ λcst_prf,prf. let 〈r1, r2〉 ≝ make_addr … destrs ? in [(rtl_stack_address r1 r2 : joint_seq rtl_params 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 rtl_argument. |destrs| = |srcrs| → list (joint_seq rtl_params globals) ≝ λglobals,destrs,srcrs,length_eq. map2 … (λdst,src.dst ← src) destrs srcrs length_eq. lemma make_list_length: ∀A: Type[0]. ∀elt: A. ∀n: nat. n = length ? (make_list A elt n). #A #ELT #N elim N normalize // qed. definition sign_mask : ∀globals.register → register → list (joint_seq rtl_params 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 IH // qed. (* using size of lists as size of ints *) definition translate_cast : ∀globals: list ident. signedness → list register → list register → bind_new (localsT rtl_params) (list (joint_seq rtl_params 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_params globals) ≝ λglobals, destrs, srcrs, prf. map2 ??? (OP1 rtl_params globals Cmpl) destrs srcrs prf. definition translate_negint : ∀globals.? → ? → ? → bind_new (localsT rtl_params) (list (joint_seq rtl_params 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 ?? ]. // qed. definition translate_notbool: ∀globals : list ident. list register → list register → bind_new (localsT rtl_params) (list (joint_seq rtl_params globals)) ≝ λglobals,destrs,srcrs. match destrs with [ nil ⇒ [ ] | cons destr destrs' ⇒ translate_fill_with_zero ? destrs' @@ match srcrs return λ_.bind_new ?? with [ nil ⇒ [destr ← 0] | cons srcr srcrs' ⇒ (destr ← srcr) ::: map register (joint_seq rtl_params 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 ← 0 ; 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_params) (list (joint_seq rtl_params 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_params) (list (joint_seq rtl_params 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 r ⇒ λeq1,eq2. translate_cast globals Unsigned destrs srcrs | Ointofptr sz sg r ⇒ λ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: cases not_implemented |*: destruct >prf1 >prf2 // ] 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 plus_minus_commutative [2: @le_plus_to_minus_r EQ //] >plus_minus_commutative [2: @(transitive_le … k_prf) //] @sym_eq @plus_to_minus length_map >tmp_destrs_prf // | >length_append p1 normalize #EQ >EQ % ] qed. definition translate_lt_signed : ∀globals. ∀destrs: list register. ∀srcrs1: list rtl_argument. ∀srcrs2: list rtl_argument. |srcrs1| = |srcrs2| → bind_new (localsT rtl_params) (list (joint_seq rtl_params 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_params) (list (joint_seq rtl_params 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 rtl_argument. |destrs| = size_of_sig_type ty3 → |srcrs1| = size_of_sig_type ty1 → |srcrs2| = size_of_sig_type ty2 → bind_new (localsT rtl_params) (list (joint_seq rtl_params 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 r ⇒ λprf1,prf2,prf3. let is_Oaddp ≝ 0 in translate_op globals Add destrs srcrs1 srcrs2 ?? | Osub sz sg ⇒ λprf1,prf2,prf2. translate_op globals Sub destrs srcrs1 srcrs2 ?? | Osubpi sz r ⇒ λprf1,prf2,prf3. let is_Osubpi ≝ 0 in translate_op globals Sub destrs srcrs1 srcrs2 ?? | Osubpp sz r1 r2 ⇒ λprf1,prf2,prf3. let is_Osubpp ≝ 0 in translate_op 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 r 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 (* pointer arithmetics is broken at the moment *) cases daemon qed. definition translate_cond: ∀globals: list ident. list register → label → bind_seq_block rtl_params globals (joint_step rtl_params globals) ≝ λglobals: list ident. λsrcrs: list register. λlbl_true: label. match srcrs return λ_.bind_seq_block rtl_params ? (joint_step ??) with [ nil ⇒ bret … 〈[ ], NOOP ??〉 | cons srcr srcrs' ⇒ ν tmpr in let f : register → joint_seq rtl_params globals ≝ λr. tmpr ← tmpr .Or. r in bret … 〈MOVE rtl_params 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_params) ? ≝ λglobals: list ident. λaddr : list rtl_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_params globals). LOAD rtl_params globals destr (Reg tmp_addr_l) (Reg tmp_addr_h) :: translate_op … Add [tmp_addr_l ; tmp_addr_h] [Reg tmp_addr_l ; Reg tmp_addr_h] [imm_nat 0 ; Imm int_size] ? ? @ acc in translate_move … [tmp_addr_l ; tmp_addr_h] addr ? @@ foldr … f [ ] destrs. [1: 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 (register_freshes ??) #U'' #rs [ >mem_set_add @orb_Prop_r assumption | @mem_set_add_id ] qed. definition proj1_rewrite : ∀A,B,a,b.∀pr : A×B.〈a,b〉 = pr → a = \fst pr. // qed. definition proj2_rewrite : ∀A,B,a,b.∀pr : A×B.〈a,b〉 = pr → b = \snd pr. // 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 〈runiverse',lenv〉 as pr_eq ≝ initialize_local_env (match f_result def with [ None ⇒ [ ] | Some r_sig ⇒ [r_sig]] @ f_params def @ f_locals def) runiverse in let params' ≝ map_list_local_env lenv (f_params def) ? in let locals' ≝ map_list_local_env lenv (f_locals def) ? in let result' ≝ match (f_result def) return λx.f_result def = x → ? with [ None ⇒ λ_.[ ] | Some r_typ ⇒ λEQ. find_local_env (\fst r_typ) lenv ? ] (refl …) 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 graph' ≝ empty_map LabelTag (joint_statement rtl_params globals) in let graph' ≝ add LabelTag ? graph' entry' (GOTO … exit') in let graph' ≝ add LabelTag ? graph' exit' (RETURN …) in let f_trans ≝ λlbl,stmt,def. let pr ≝ translate_statement … lenv stmt in b_adds_graph … (rtl_ertl_fresh_reg …) (dpi1 … pr) lbl (dpi2 … pr) def in let res ≝ mk_joint_internal_function globals rtl_params luniverse' runiverse' result' params' locals' stack_size' graph' (pi1 … entry') (pi1 … exit') in foldi … f_trans (f_graph def) res. [1,2: >graph_code_has_point @map_mem_prop [@graph_add_lookup] @graph_add | >(proj2_rewrite ????? pr_eq) @initialize_local_env_in >EQ normalize nodelta %1 % |*: >(proj2_rewrite ????? pr_eq) @(All_mp ??? (λpr.initialize_local_env_in ?? (\fst pr))) [ @(All_mp … (λpr.Exists ? (λx.\fst x = \fst pr) (f_locals def))) [ #a #H @Exists_append_r @Exists_append_r @H | generalize in match (f_locals def); ] | @(All_mp … (λpr.Exists ? (λx.\fst x = \fst pr) (f_params def))) [ #a #H @Exists_append_r @Exists_append_l @H | generalize in match (f_params def); ] ] #l elim l [1,3: %] #hd #tl #IH % [1,3: %1 %] @(All_mp … IH) #x #G %2{G} ] 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)).