aboutsummaryrefslogtreecommitdiffstats
path: root/arm
diff options
context:
space:
mode:
Diffstat (limited to 'arm')
-rw-r--r--arm/Machregs.v34
-rw-r--r--arm/Machregsaux.ml35
-rw-r--r--arm/Machregsaux.mli1
-rw-r--r--arm/Op.v94
-rw-r--r--arm/PrintOp.ml21
-rw-r--r--arm/TargetPrinter.ml8
6 files changed, 115 insertions, 78 deletions
diff --git a/arm/Machregs.v b/arm/Machregs.v
index f373b434..f46f2904 100644
--- a/arm/Machregs.v
+++ b/arm/Machregs.v
@@ -10,6 +10,7 @@
(* *)
(* *********************************************************************)
+Require Import String.
Require Import Coqlib.
Require Import Maps.
Require Import AST.
@@ -75,6 +76,28 @@ End IndexedMreg.
Definition is_stack_reg (r: mreg) : bool := false.
+(** ** Names of registers *)
+
+Local Open Scope string_scope.
+
+Definition register_names :=
+ ("R0", R0) :: ("R1", R1) :: ("R2", R2) :: ("R3", R3) ::
+ ("R4", R4) :: ("R5", R5) :: ("R6", R6) :: ("R7", R7) ::
+ ("R8", R8) :: ("R9", R9) :: ("R10", R10) :: ("R11", R11) ::
+ ("R12", R12) ::
+ ("F0", F0) :: ("F1", F1) :: ("F2", F2) :: ("F3", F3) ::
+ ("F4", F4) :: ("F5", F5) :: ("F6", F6) :: ("F7", F7) ::
+ ("F8", F8) :: ("F9", F9) :: ("F10", F10) :: ("F11", F11) ::
+ ("F12", F12) ::("F13", F13) ::("F14", F14) :: ("F15", F15) :: nil.
+
+Definition register_by_name (s: string) : option mreg :=
+ let fix assoc (l: list (string * mreg)) : option mreg :=
+ match l with
+ | nil => None
+ | (s1, r1) :: l' => if string_dec s s1 then Some r1 else assoc l'
+ end
+ in assoc register_names.
+
(** ** Destroyed registers, preferred registers *)
Definition destroyed_by_op (op: operation): list mreg :=
@@ -95,9 +118,20 @@ Definition destroyed_by_cond (cond: condition): list mreg :=
Definition destroyed_by_jumptable: list mreg :=
nil.
+Fixpoint destroyed_by_clobber (cl: list string): list mreg :=
+ match cl with
+ | nil => nil
+ | c1 :: cl =>
+ match register_by_name c1 with
+ | Some r => r :: destroyed_by_clobber cl
+ | None => destroyed_by_clobber cl
+ end
+ end.
+
Definition destroyed_by_builtin (ef: external_function): list mreg :=
match ef with
| EF_memcpy sz al => if zle sz 32 then F7 :: nil else R2 :: R3 :: R12 :: nil
+ | EF_inline_asm txt sg clob => destroyed_by_clobber clob
| _ => nil
end.
diff --git a/arm/Machregsaux.ml b/arm/Machregsaux.ml
index 44e6b192..d59ec8b8 100644
--- a/arm/Machregsaux.ml
+++ b/arm/Machregsaux.ml
@@ -12,41 +12,24 @@
(** Auxiliary functions on machine registers *)
+open Camlcoq
open Machregs
-let register_names = [
- ("R0", R0); ("R1", R1); ("R2", R2); ("R3", R3);
- ("R4", R4); ("R5", R5); ("R6", R6); ("R7", R7);
- ("R8", R8); ("R9", R9); ("R10", R10); ("R11", R11);
- ("R12", R12);
- ("F0", F0); ("F1", F1); ("F2", F2); ("F3", F3);
- ("F4", F4); ("F5", F5); ("F6", F6); ("F7", F7);
- ("F8", F8); ("F9", F9); ("F10", F10); ("F11", F11);
- ("F12", F12);("F13", F13);("F14", F14); ("F15", F15)
-]
+let register_names : (mreg, string) Hashtbl.t = Hashtbl.create 31
+
+let _ =
+ List.iter
+ (fun (s, r) -> Hashtbl.add register_names r (camlstring_of_coqstring s))
+ Machregs.register_names
let scratch_register_names = [ "R14" ]
let name_of_register r =
- let rec rev_assoc = function
- | [] -> None
- | (a, b) :: rem -> if b = r then Some a else rev_assoc rem
- in rev_assoc register_names
+ try Some (Hashtbl.find register_names r) with Not_found -> None
let register_by_name s =
- try
- Some(List.assoc (String.uppercase s) register_names)
- with Not_found ->
- None
+ Machregs.register_by_name (coqstring_of_camlstring (String.uppercase s))
let can_reserve_register r =
List.mem r Conventions1.int_callee_save_regs
|| List.mem r Conventions1.float_callee_save_regs
-
-let mregs_of_clobber idl =
- List.fold_left
- (fun l c ->
- match register_by_name (Camlcoq.extern_atom c) with
- | Some r -> r :: l
- | None -> l)
- [] idl
diff --git a/arm/Machregsaux.mli b/arm/Machregsaux.mli
index f0feec96..d4877a62 100644
--- a/arm/Machregsaux.mli
+++ b/arm/Machregsaux.mli
@@ -16,4 +16,3 @@ val name_of_register: Machregs.mreg -> string option
val register_by_name: string -> Machregs.mreg option
val scratch_register_names: string list
val can_reserve_register: Machregs.mreg -> bool
-val mregs_of_clobber: Camlcoq.atom list -> Machregs.mreg list
diff --git a/arm/Op.v b/arm/Op.v
index bbdcd123..df39b26a 100644
--- a/arm/Op.v
+++ b/arm/Op.v
@@ -808,40 +808,40 @@ Hypothesis valid_different_pointers_inj:
Ltac InvInject :=
match goal with
- | [ H: val_inject _ (Vint _) _ |- _ ] =>
+ | [ H: Val.inject _ (Vint _) _ |- _ ] =>
inv H; InvInject
- | [ H: val_inject _ (Vfloat _) _ |- _ ] =>
+ | [ H: Val.inject _ (Vfloat _) _ |- _ ] =>
inv H; InvInject
- | [ H: val_inject _ (Vsingle _) _ |- _ ] =>
+ | [ H: Val.inject _ (Vsingle _) _ |- _ ] =>
inv H; InvInject
- | [ H: val_inject _ (Vptr _ _) _ |- _ ] =>
+ | [ H: Val.inject _ (Vptr _ _) _ |- _ ] =>
inv H; InvInject
- | [ H: val_list_inject _ nil _ |- _ ] =>
+ | [ H: Val.inject_list _ nil _ |- _ ] =>
inv H; InvInject
- | [ H: val_list_inject _ (_ :: _) _ |- _ ] =>
+ | [ H: Val.inject_list _ (_ :: _) _ |- _ ] =>
inv H; InvInject
| _ => idtac
end.
Remark eval_shift_inj:
- forall s v v', val_inject f v v' -> val_inject f (eval_shift s v) (eval_shift s v').
+ forall s v v', Val.inject f v v' -> Val.inject f (eval_shift s v) (eval_shift s v').
Proof.
intros. inv H; destruct s; simpl; auto; rewrite s_range; auto.
Qed.
Lemma eval_condition_inj:
forall cond vl1 vl2 b,
- val_list_inject f vl1 vl2 ->
+ Val.inject_list f vl1 vl2 ->
eval_condition cond vl1 m1 = Some b ->
eval_condition cond vl2 m2 = Some b.
Proof.
intros. destruct cond; simpl in H0; FuncInv; InvInject; simpl.
- eauto 4 using val_cmp_bool_inject.
- eauto 4 using val_cmpu_bool_inject, Mem.valid_pointer_implies.
- eauto using val_cmp_bool_inject, eval_shift_inj.
- eauto 4 using val_cmpu_bool_inject, Mem.valid_pointer_implies, eval_shift_inj.
- eauto 4 using val_cmp_bool_inject.
- eauto 4 using val_cmpu_bool_inject, Mem.valid_pointer_implies.
+ eauto 4 using Val.cmp_bool_inject.
+ eauto 4 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
+ eauto using Val.cmp_bool_inject, eval_shift_inj.
+ eauto 4 using Val.cmpu_bool_inject, Mem.valid_pointer_implies, eval_shift_inj.
+ eauto 4 using Val.cmp_bool_inject.
+ eauto 4 using Val.cmpu_bool_inject, Mem.valid_pointer_implies.
inv H3; inv H2; simpl in H0; inv H0; auto.
inv H3; inv H2; simpl in H0; inv H0; auto.
inv H3; simpl in H0; inv H0; auto.
@@ -854,7 +854,7 @@ Qed.
Ltac TrivialExists :=
match goal with
- | [ |- exists v2, Some ?v1 = Some v2 /\ val_inject _ _ v2 ] =>
+ | [ |- exists v2, Some ?v1 = Some v2 /\ Val.inject _ _ v2 ] =>
exists v1; split; auto
| _ => idtac
end.
@@ -863,28 +863,28 @@ Lemma eval_operation_inj:
forall op sp1 vl1 sp2 vl2 v1,
(forall id ofs,
In id (globals_operation op) ->
- val_inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
- val_inject f sp1 sp2 ->
- val_list_inject f vl1 vl2 ->
+ Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
+ Val.inject f sp1 sp2 ->
+ Val.inject_list f vl1 vl2 ->
eval_operation ge1 sp1 op vl1 m1 = Some v1 ->
- exists v2, eval_operation ge2 sp2 op vl2 m2 = Some v2 /\ val_inject f v1 v2.
+ exists v2, eval_operation ge2 sp2 op vl2 m2 = Some v2 /\ Val.inject f v1 v2.
Proof.
intros until v1; intros GL; intros. destruct op; simpl in H1; simpl; FuncInv; InvInject; TrivialExists.
apply GL; simpl; auto.
- apply Values.val_add_inject; auto.
+ apply Values.Val.add_inject; auto.
inv H4; simpl; auto.
inv H4; simpl; auto.
- apply Values.val_add_inject; auto.
- apply Values.val_add_inject; auto. apply eval_shift_inj; auto.
- apply Values.val_add_inject; auto.
+ apply Values.Val.add_inject; auto.
+ apply Values.Val.add_inject; auto. apply eval_shift_inj; auto.
+ apply Values.Val.add_inject; auto.
- apply Values.val_sub_inject; auto.
- apply Values.val_sub_inject; auto. apply eval_shift_inj; auto.
- apply Values.val_sub_inject; auto. apply eval_shift_inj; auto.
- apply (@Values.val_sub_inject f (Vint i) (Vint i) v v'); auto.
+ apply Values.Val.sub_inject; auto.
+ apply Values.Val.sub_inject; auto. apply eval_shift_inj; auto.
+ apply Values.Val.sub_inject; auto. apply eval_shift_inj; auto.
+ apply (@Values.Val.sub_inject f (Vint i) (Vint i) v v'); auto.
inv H4; inv H2; simpl; auto.
- apply Values.val_add_inject; auto. inv H4; inv H2; simpl; auto.
+ apply Values.Val.add_inject; auto. inv H4; inv H2; simpl; auto.
inv H4; inv H2; simpl; auto.
inv H4; inv H2; simpl; auto.
inv H4; inv H3; simpl in H1; inv H1. simpl.
@@ -958,17 +958,17 @@ Lemma eval_addressing_inj:
forall addr sp1 vl1 sp2 vl2 v1,
(forall id ofs,
In id (globals_addressing addr) ->
- val_inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
- val_inject f sp1 sp2 ->
- val_list_inject f vl1 vl2 ->
+ Val.inject f (Genv.symbol_address ge1 id ofs) (Genv.symbol_address ge2 id ofs)) ->
+ Val.inject f sp1 sp2 ->
+ Val.inject_list f vl1 vl2 ->
eval_addressing ge1 sp1 addr vl1 = Some v1 ->
- exists v2, eval_addressing ge2 sp2 addr vl2 = Some v2 /\ val_inject f v1 v2.
+ exists v2, eval_addressing ge2 sp2 addr vl2 = Some v2 /\ Val.inject f v1 v2.
Proof.
intros. destruct addr; simpl in H2; simpl; FuncInv; InvInject; TrivialExists.
- apply Values.val_add_inject; auto.
- apply Values.val_add_inject; auto.
- apply Values.val_add_inject; auto. apply eval_shift_inj; auto.
- apply Values.val_add_inject; auto.
+ apply Values.Val.add_inject; auto.
+ apply Values.Val.add_inject; auto.
+ apply Values.Val.add_inject; auto. apply eval_shift_inj; auto.
+ apply Values.Val.add_inject; auto.
Qed.
End EVAL_COMPAT.
@@ -1034,7 +1034,7 @@ Proof.
apply weak_valid_pointer_extends; auto.
apply weak_valid_pointer_no_overflow_extends.
apply valid_different_pointers_extends; auto.
- rewrite <- val_list_inject_lessdef. eauto. auto.
+ rewrite <- val_inject_list_lessdef. eauto. auto.
Qed.
Lemma eval_operation_lessdef:
@@ -1044,10 +1044,10 @@ Lemma eval_operation_lessdef:
eval_operation genv sp op vl1 m1 = Some v1 ->
exists v2, eval_operation genv sp op vl2 m2 = Some v2 /\ Val.lessdef v1 v2.
Proof.
- intros. rewrite val_list_inject_lessdef in H.
+ intros. rewrite val_inject_list_lessdef in H.
assert (exists v2 : val,
eval_operation genv sp op vl2 m2 = Some v2
- /\ val_inject (fun b => Some(b, 0)) v1 v2).
+ /\ Val.inject (fun b => Some(b, 0)) v1 v2).
eapply eval_operation_inj with (m1 := m1) (sp1 := sp).
apply valid_pointer_extends; auto.
apply weak_valid_pointer_extends; auto.
@@ -1065,10 +1065,10 @@ Lemma eval_addressing_lessdef:
eval_addressing genv sp addr vl1 = Some v1 ->
exists v2, eval_addressing genv sp addr vl2 = Some v2 /\ Val.lessdef v1 v2.
Proof.
- intros. rewrite val_list_inject_lessdef in H.
+ intros. rewrite val_inject_list_lessdef in H.
assert (exists v2 : val,
eval_addressing genv sp addr vl2 = Some v2
- /\ val_inject (fun b => Some(b, 0)) v1 v2).
+ /\ Val.inject (fun b => Some(b, 0)) v1 v2).
eapply eval_addressing_inj with (sp1 := sp).
intros. rewrite <- val_inject_lessdef; auto.
rewrite <- val_inject_lessdef; auto.
@@ -1092,7 +1092,7 @@ Variable delta: Z.
Hypothesis sp_inj: f sp1 = Some(sp2, delta).
Remark symbol_address_inject:
- forall id ofs, val_inject f (Genv.symbol_address genv id ofs) (Genv.symbol_address genv id ofs).
+ forall id ofs, Val.inject f (Genv.symbol_address genv id ofs) (Genv.symbol_address genv id ofs).
Proof.
intros. unfold Genv.symbol_address. destruct (Genv.find_symbol genv id) eqn:?; auto.
exploit (proj1 globals); eauto. intros.
@@ -1101,7 +1101,7 @@ Qed.
Lemma eval_condition_inject:
forall cond vl1 vl2 b m1 m2,
- val_list_inject f vl1 vl2 ->
+ Val.inject_list f vl1 vl2 ->
Mem.inject f m1 m2 ->
eval_condition cond vl1 m1 = Some b ->
eval_condition cond vl2 m2 = Some b.
@@ -1115,11 +1115,11 @@ Qed.
Lemma eval_addressing_inject:
forall addr vl1 vl2 v1,
- val_list_inject f vl1 vl2 ->
+ Val.inject_list f vl1 vl2 ->
eval_addressing genv (Vptr sp1 Int.zero) addr vl1 = Some v1 ->
exists v2,
eval_addressing genv (Vptr sp2 Int.zero) (shift_stack_addressing (Int.repr delta) addr) vl2 = Some v2
- /\ val_inject f v1 v2.
+ /\ Val.inject f v1 v2.
Proof.
intros.
rewrite eval_shift_stack_addressing. simpl.
@@ -1129,12 +1129,12 @@ Qed.
Lemma eval_operation_inject:
forall op vl1 vl2 v1 m1 m2,
- val_list_inject f vl1 vl2 ->
+ Val.inject_list f vl1 vl2 ->
Mem.inject f m1 m2 ->
eval_operation genv (Vptr sp1 Int.zero) op vl1 m1 = Some v1 ->
exists v2,
eval_operation genv (Vptr sp2 Int.zero) (shift_stack_operation (Int.repr delta) op) vl2 m2 = Some v2
- /\ val_inject f v1 v2.
+ /\ Val.inject f v1 v2.
Proof.
intros.
rewrite eval_shift_stack_operation. simpl.
diff --git a/arm/PrintOp.ml b/arm/PrintOp.ml
index 96d13943..886f6de3 100644
--- a/arm/PrintOp.ml
+++ b/arm/PrintOp.ml
@@ -52,6 +52,14 @@ let print_condition reg pp = function
fprintf pp "%a %sf 0.0" reg r1 (comparison_name c)
| (Cnotcompfzero c, [r1]) ->
fprintf pp "%a not(%sf) 0.0" reg r1 (comparison_name c)
+ | (Ccompfs c, [r1;r2]) ->
+ fprintf pp "%a %sfs %a" reg r1 (comparison_name c) reg r2
+ | (Cnotcompfs c, [r1;r2]) ->
+ fprintf pp "%a not(%sfs) %a" reg r1 (comparison_name c) reg r2
+ | (Ccompfszero c, [r1]) ->
+ fprintf pp "%a %sfs 0.0" reg r1 (comparison_name c)
+ | (Cnotcompfszero c, [r1]) ->
+ fprintf pp "%a not(%sfs) 0.0" reg r1 (comparison_name c)
| _ ->
fprintf pp "<bad condition>"
@@ -59,6 +67,7 @@ let print_operation reg pp = function
| Omove, [r1] -> reg pp r1
| Ointconst n, [] -> fprintf pp "%ld" (camlint_of_coqint n)
| Ofloatconst n, [] -> fprintf pp "%F" (camlfloat_of_coqfloat n)
+ | Osingleconst n, [] -> fprintf pp "%Ff" (camlfloat_of_coqfloat32 n)
| Oaddrsymbol(id, ofs), [] ->
fprintf pp "\"%s\" + %ld" (extern_atom id) (camlint_of_coqint ofs)
| Oaddrstack ofs, [] ->
@@ -100,9 +109,21 @@ let print_operation reg pp = function
| Osubf, [r1;r2] -> fprintf pp "%a -f %a" reg r1 reg r2
| Omulf, [r1;r2] -> fprintf pp "%a *f %a" reg r1 reg r2
| Odivf, [r1;r2] -> fprintf pp "%a /f %a" reg r1 reg r2
+ | Onegfs, [r1] -> fprintf pp "negfs(%a)" reg r1
+ | Oabsfs, [r1] -> fprintf pp "absfs(%a)" reg r1
+ | Oaddfs, [r1;r2] -> fprintf pp "%a +fs %a" reg r1 reg r2
+ | Osubfs, [r1;r2] -> fprintf pp "%a -fs %a" reg r1 reg r2
+ | Omulfs, [r1;r2] -> fprintf pp "%a *fs %a" reg r1 reg r2
+ | Odivfs, [r1;r2] -> fprintf pp "%a /fs %a" reg r1 reg r2
| Osingleoffloat, [r1] -> fprintf pp "singleoffloat(%a)" reg r1
| Ointoffloat, [r1] -> fprintf pp "intoffloat(%a)" reg r1
+ | Ointuoffloat, [r1] -> fprintf pp "intuoffloat(%a)" reg r1
| Ofloatofint, [r1] -> fprintf pp "floatofint(%a)" reg r1
+ | Ofloatofintu, [r1] -> fprintf pp "floatofintu(%a)" reg r1
+ | Ointofsingle, [r1] -> fprintf pp "intofsingle(%a)" reg r1
+ | Ointuofsingle, [r1] -> fprintf pp "intuofsingle(%a)" reg r1
+ | Osingleofint, [r1] -> fprintf pp "singleofint(%a)" reg r1
+ | Osingleofintu, [r1] -> fprintf pp "singleofintu(%a)" reg r1
| Omakelong, [r1;r2] -> fprintf pp "makelong(%a,%a)" reg r1 reg r2
| Olowlong, [r1] -> fprintf pp "lowlong(%a)" reg r1
| Ohighlong, [r1] -> fprintf pp "highlong(%a)" reg r1
diff --git a/arm/TargetPrinter.ml b/arm/TargetPrinter.ml
index c77572db..505c3265 100644
--- a/arm/TargetPrinter.ml
+++ b/arm/TargetPrinter.ml
@@ -300,7 +300,7 @@ module Target (Opt: PRINTER_OPTIONS) : TARGET =
(* Emit .file / .loc debugging directives *)
let print_file_line oc file line =
- PrintAnnot.print_file_line oc comment file line
+ print_file_line oc comment file line
let print_location oc loc =
if loc <> Cutil.no_loc then print_file_line oc (fst loc) (snd loc)
@@ -320,12 +320,12 @@ module Target (Opt: PRINTER_OPTIONS) : TARGET =
(int_of_string (Str.matched_group 2 txt))
end else begin
fprintf oc "%s annotation: " comment;
- PrintAnnot.print_annot_stmt preg "sp" oc txt targs args
+ print_annot_stmt preg "sp" oc txt targs args
end
let print_annot_val oc txt args res =
fprintf oc "%s annotation: " comment;
- PrintAnnot.print_annot_val preg oc txt args;
+ print_annot_val preg oc txt args;
match args, res with
| [IR src], [IR dst] ->
if dst = src then 0 else (fprintf oc " mov %a, %a\n" ireg dst ireg src; 1)
@@ -1005,7 +1005,7 @@ module Target (Opt: PRINTER_OPTIONS) : TARGET =
print_annot_val oc (extern_atom txt) args res
| EF_inline_asm(txt, sg, clob) ->
fprintf oc "%s begin inline assembly\n\t" comment;
- PrintAnnot.print_inline_asm preg oc (extern_atom txt) sg args res;
+ print_inline_asm preg oc (extern_atom txt) sg args res;
fprintf oc "%s end inline assembly\n" comment;
5 (* hoping this is an upper bound... *)
| _ ->