aboutsummaryrefslogtreecommitdiffstats
diff options
context:
space:
mode:
-rw-r--r--backend/Allocation.v4
-rw-r--r--backend/Allocproof.v6
-rw-r--r--backend/CSE.v6
-rw-r--r--backend/CSE2.v6
-rw-r--r--backend/Constprop.v6
-rw-r--r--backend/Constpropproof.v4
-rw-r--r--backend/Deadcode.v4
-rw-r--r--backend/Duplicate.v4
-rw-r--r--backend/Duplicateaux.ml526
-rw-r--r--backend/Duplicateproof.v8
-rw-r--r--backend/ForwardMoves.v6
-rw-r--r--backend/Inlining.v4
-rw-r--r--backend/Inliningspec.v6
-rw-r--r--backend/LTL.v8
-rw-r--r--backend/Linearize.v2
-rw-r--r--backend/Linearizeaux.ml178
-rw-r--r--backend/Liveness.v2
-rw-r--r--backend/PrintLTL.ml5
-rw-r--r--backend/PrintRTL.ml5
-rw-r--r--backend/PrintXTL.ml2
-rw-r--r--backend/RTL.v19
-rw-r--r--backend/RTLgen.v2
-rw-r--r--backend/RTLgenspec.v4
-rw-r--r--backend/RTLtyping.v6
-rw-r--r--backend/Regalloc.ml20
-rw-r--r--backend/Renumber.v2
-rw-r--r--backend/Splitting.ml4
-rw-r--r--backend/Tunneling.v4
-rw-r--r--backend/Unusedglob.v2
-rw-r--r--backend/ValueAnalysis.v2
-rw-r--r--backend/XTL.ml6
-rw-r--r--backend/XTL.mli2
-rwxr-xr-xconfigure2
33 files changed, 533 insertions, 334 deletions
diff --git a/backend/Allocation.v b/backend/Allocation.v
index d18b07a9..2323c050 100644
--- a/backend/Allocation.v
+++ b/backend/Allocation.v
@@ -314,10 +314,10 @@ Definition pair_instr_block
Some(BSbuiltin ef args res mv1 args' res' mv2 s)
| _ => None
end
- | Icond cond args s1 s2 =>
+ | Icond cond args s1 s2 i =>
let (mv1, b1) := extract_moves nil b in
match b1 with
- | Lcond cond' args' s1' s2' :: b2 =>
+ | Lcond cond' args' s1' s2' i' :: b2 =>
assertion (eq_condition cond cond');
assertion (peq s1 s1');
assertion (peq s2 s2');
diff --git a/backend/Allocproof.v b/backend/Allocproof.v
index b6880860..3c7df58a 100644
--- a/backend/Allocproof.v
+++ b/backend/Allocproof.v
@@ -169,11 +169,11 @@ Inductive expand_block_shape: block_shape -> RTL.instruction -> LTL.bblock -> Pr
(Ibuiltin ef args res s)
(expand_moves mv1
(Lbuiltin ef args' res' :: expand_moves mv2 (Lbranch s :: k)))
- | ebs_cond: forall cond args mv args' s1 s2 k,
+ | ebs_cond: forall cond args mv args' s1 s2 k i i',
wf_moves mv ->
expand_block_shape (BScond cond args mv args' s1 s2)
- (Icond cond args s1 s2)
- (expand_moves mv (Lcond cond args' s1 s2 :: k))
+ (Icond cond args s1 s2 i)
+ (expand_moves mv (Lcond cond args' s1 s2 i' :: k))
| ebs_jumptable: forall arg mv arg' tbl k,
wf_moves mv ->
expand_block_shape (BSjumptable arg mv arg' tbl)
diff --git a/backend/CSE.v b/backend/CSE.v
index 2827161d..1936d4e4 100644
--- a/backend/CSE.v
+++ b/backend/CSE.v
@@ -496,7 +496,7 @@ Definition transfer (f: function) (approx: PMap.t VA.t) (pc: node) (before: numb
| EF_vload _ | EF_annot _ _ _ | EF_annot_val _ _ _ | EF_debug _ _ _ =>
set_res_unknown before res
end
- | Icond cond args ifso ifnot =>
+ | Icond cond args ifso ifnot _ =>
before
| Ijumptable arg tbl =>
before
@@ -549,10 +549,10 @@ Definition transf_instr (n: numbering) (instr: instruction) :=
let (n1, vl) := valnum_regs n args in
let (addr', args') := reduce _ combine_addr n1 addr args vl in
Istore chunk addr' args' src s
- | Icond cond args s1 s2 =>
+ | Icond cond args s1 s2 i =>
let (n1, vl) := valnum_regs n args in
let (cond', args') := reduce _ combine_cond n1 cond args vl in
- Icond cond' args' s1 s2
+ Icond cond' args' s1 s2 i
| _ =>
instr
end.
diff --git a/backend/CSE2.v b/backend/CSE2.v
index efa70b40..dc206c75 100644
--- a/backend/CSE2.v
+++ b/backend/CSE2.v
@@ -292,7 +292,7 @@ Qed.
Definition apply_instr instr (rel : RELATION.t) : RB.t :=
match instr with
| Inop _
- | Icond _ _ _ _
+ | Icond _ _ _ _ _
| Ijumptable _ _ => Some rel
| Istore chunk addr args _ _ => Some (kill_store chunk addr args rel)
| Iop op args dst _ => Some (gen_oper op dst args rel)
@@ -372,8 +372,8 @@ Definition transf_instr (fmap : option (PMap.t RB.t))
Icall sig ros (subst_args fmap pc args) dst s
| Itailcall sig ros args =>
Itailcall sig ros (subst_args fmap pc args)
- | Icond cond args s1 s2 =>
- Icond cond (subst_args fmap pc args) s1 s2
+ | Icond cond args s1 s2 i =>
+ Icond cond (subst_args fmap pc args) s1 s2 i
| Ijumptable arg tbl =>
Ijumptable (subst_arg fmap pc arg) tbl
| Ireturn (Some arg) =>
diff --git a/backend/Constprop.v b/backend/Constprop.v
index eda41b39..0be9438c 100644
--- a/backend/Constprop.v
+++ b/backend/Constprop.v
@@ -69,7 +69,7 @@ Fixpoint successor_rec (n: nat) (f: function) (ae: AE.t) (pc: node) : node :=
match f.(fn_code)!pc with
| Some (Inop s) =>
successor_rec n' f ae s
- | Some (Icond cond args s1 s2) =>
+ | Some (Icond cond args s1 s2 _) =>
match resolve_branch (eval_static_condition cond (aregs ae args)) with
| Some b => successor_rec n' f ae (if b then s1 else s2)
| None => pc
@@ -217,14 +217,14 @@ Definition transf_instr (f: function) (an: PMap.t VA.t) (rm: romem)
end
| _, _ => dfl
end
- | Icond cond args s1 s2 =>
+ | Icond cond args s1 s2 i =>
let aargs := aregs ae args in
match resolve_branch (eval_static_condition cond aargs) with
| Some b =>
if b then Inop s1 else Inop s2
| None =>
let (cond', args') := cond_strength_reduction cond args aargs in
- Icond cond' args' s1 s2
+ Icond cond' args' s1 s2 i
end
| Ijumptable arg tbl =>
match areg ae arg with
diff --git a/backend/Constpropproof.v b/backend/Constpropproof.v
index 63cfee24..60663503 100644
--- a/backend/Constpropproof.v
+++ b/backend/Constpropproof.v
@@ -142,8 +142,8 @@ Inductive match_pc (f: function) (rs: regset) (m: mem): nat -> node -> node -> P
f.(fn_code)!pc = Some (Inop s) ->
match_pc f rs m n s pcx ->
match_pc f rs m (S n) pc pcx
- | match_pc_cond: forall n pc cond args s1 s2 pcx,
- f.(fn_code)!pc = Some (Icond cond args s1 s2) ->
+ | match_pc_cond: forall n pc cond args s1 s2 pcx i,
+ f.(fn_code)!pc = Some (Icond cond args s1 s2 i) ->
(forall b,
eval_condition cond rs##args m = Some b ->
match_pc f rs m n (if b then s1 else s2) pcx) ->
diff --git a/backend/Deadcode.v b/backend/Deadcode.v
index 1f208a91..3412a6fa 100644
--- a/backend/Deadcode.v
+++ b/backend/Deadcode.v
@@ -142,7 +142,7 @@ Definition transfer (f: function) (approx: PMap.t VA.t)
nmem_dead_stack f.(fn_stacksize))
| Some(Ibuiltin ef args res s) =>
transfer_builtin approx!!pc ef args res ne nm
- | Some(Icond cond args s1 s2) =>
+ | Some(Icond cond args s1 s2 _) =>
if peq s1 s2 then after else
(add_needs args (needs_of_condition cond) ne, nm)
| Some(Ijumptable arg tbl) =>
@@ -192,7 +192,7 @@ Definition transf_instr (approx: PMap.t VA.t) (an: PMap.t NA.t)
if nmem_contains (snd an!!pc) (aaddr_arg approx!!pc dst) sz
then instr
else Inop s
- | Icond cond args s1 s2 =>
+ | Icond cond args s1 s2 _ =>
if peq s1 s2 then Inop s1 else instr
| _ =>
instr
diff --git a/backend/Duplicate.v b/backend/Duplicate.v
index 82c17367..af85efe4 100644
--- a/backend/Duplicate.v
+++ b/backend/Duplicate.v
@@ -134,8 +134,8 @@ Definition verify_match_inst dupmap inst tinst :=
else Error (msg "Different ef in Ibuiltin")
| _ => Error (msg "verify_match_inst Ibuiltin") end
- | Icond cond lr n1 n2 => match tinst with
- | Icond cond' lr' n1' n2' =>
+ | Icond cond lr n1 n2 i => match tinst with
+ | Icond cond' lr' n1' n2' i' =>
if (list_eq_dec Pos.eq_dec lr lr') then
if (eq_condition cond cond') then
do u1 <- verify_is_copy dupmap n1 n1';
diff --git a/backend/Duplicateaux.ml b/backend/Duplicateaux.ml
index a84f9754..89f187da 100644
--- a/backend/Duplicateaux.ml
+++ b/backend/Duplicateaux.ml
@@ -1,25 +1,32 @@
+(* Oracle for Duplicate pass.
+ * - Add static prediction information to Icond nodes
+ * - Performs tail duplication on interesting traces to form superblocks
+ * - (TODO: perform partial loop unrolling inside innermost loops)
+ *)
+
open RTL
open Maps
open Camlcoq
-(* TTL : IR emphasizing the preferred next node *)
-module TTL = struct
- type instruction =
- | Tleaf of RTL.instruction
- | Tnext of node * RTL.instruction
-
- type code = instruction PTree.t
-end;;
-
-open TTL
+let debug_flag = ref false
-(** RTL to TTL *)
+let debug fmt =
+ if !debug_flag then Printf.eprintf fmt
+ else Printf.ifprintf stderr fmt
let get_some = function
| None -> failwith "Did not get some"
| Some thing -> thing
-let bfs code entrypoint =
+let rtl_successors = function
+| Itailcall _ | Ireturn _ -> []
+| Icall(_,_,_,_,n) | Ibuiltin(_,_,_,n) | Inop n | Iop (_,_,_,n)
+| Iload (_,_,_,_,_,n) | Istore (_,_,_,_,n) -> [n]
+| Icond (_,_,n1,n2,_) -> [n1; n2]
+| Ijumptable (_,ln) -> ln
+
+let bfs code entrypoint = begin
+ debug "bfs\n";
let visited = ref (PTree.map (fun n i -> false) code)
and bfs_list = ref []
and to_visit = Queue.create ()
@@ -33,32 +40,24 @@ let bfs code entrypoint =
match PTree.get !node code with
| None -> failwith "No such node"
| Some i ->
- bfs_list := !bfs_list @ [!node];
- match i with
- | Icall(_, _, _, _, n) -> Queue.add n to_visit
- | Ibuiltin(_, _, _, n) -> Queue.add n to_visit
- | Ijumptable(_, ln) -> List.iter (fun n -> Queue.add n to_visit) ln
- | Itailcall _ | Ireturn _ -> ()
- | Icond (_, _, n1, n2) -> Queue.add n1 to_visit; Queue.add n2 to_visit
- | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) -> Queue.add n to_visit
+ bfs_list := !node :: !bfs_list;
+ let succ = rtl_successors i in
+ List.iter (fun n -> Queue.add n to_visit) succ
end
done;
- !bfs_list
+ List.rev !bfs_list
end
+end
let optbool o = match o with Some _ -> true | None -> false
let ptree_get_some n ptree = get_some @@ PTree.get n ptree
-let get_predecessors_rtl code =
+let get_predecessors_rtl code = begin
+ debug "get_predecessors_rtl\n";
let preds = ref (PTree.map (fun n i -> []) code) in
let process_inst (node, i) =
- let succ = match i with
- | Inop n | Iop (_,_,_,n) | Iload (_, _,_,_,_,n) | Istore (_,_,_,_,n)
- | Icall (_,_,_,_,n) | Ibuiltin (_, _, _, n) -> [n]
- | Icond (_,_,n1,n2) -> [n1;n2]
- | Ijumptable (_,ln) -> ln
- | Itailcall _ | Ireturn _ -> []
+ let succ = rtl_successors i
in List.iter (fun s ->
let previous_preds = ptree_get_some s !preds in
if optbool @@ List.find_opt (fun e -> e == node) previous_preds then ()
@@ -67,6 +66,7 @@ let get_predecessors_rtl code =
List.iter process_inst (PTree.elements code);
!preds
end
+end
module PInt = struct
type t = P.t
@@ -80,19 +80,23 @@ let print_intlist l =
| [] -> ()
| n::ln -> (Printf.printf "%d " (P.to_int n); f ln)
in begin
- Printf.printf "[";
- f l;
- Printf.printf "]"
+ if !debug_flag then begin
+ Printf.printf "[";
+ f l;
+ Printf.printf "]"
+ end
end
let print_intset s =
let seq = PSet.to_seq s
in begin
- Printf.printf "{";
- Seq.iter (fun n ->
- Printf.printf "%d " (P.to_int n)
- ) seq;
- Printf.printf "}"
+ if !debug_flag then begin
+ Printf.printf "{";
+ Seq.iter (fun n ->
+ Printf.printf "%d " (P.to_int n)
+ ) seq;
+ Printf.printf "}"
+ end
end
type vstate = Unvisited | Processed | Visited
@@ -104,7 +108,8 @@ type vstate = Unvisited | Processed | Visited
*
* If we come accross an edge to a Processed node, it's a loop!
*)
-let get_loop_headers code entrypoint =
+let get_loop_headers code entrypoint = begin
+ debug "get_loop_headers\n";
let visited = ref (PTree.map (fun n i -> Unvisited) code)
and is_loop_header = ref (PTree.map (fun n i -> false) code)
in let rec dfs_visit code = function
@@ -113,6 +118,7 @@ let get_loop_headers code entrypoint =
match (get_some @@ PTree.get node !visited) with
| Visited -> ()
| Processed -> begin
+ debug "Node %d is a loop header\n" (P.to_int node);
is_loop_header := PTree.set node true !is_loop_header;
visited := PTree.set node Visited !visited
end
@@ -120,13 +126,7 @@ let get_loop_headers code entrypoint =
visited := PTree.set node Processed !visited;
match PTree.get node code with
| None -> failwith "No such node"
- | Some i -> let next_visits = (match i with
- | Icall (_, _, _, _, n) | Ibuiltin (_, _, _, n) | Inop n | Iop (_, _, _, n)
- | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) -> [n]
- | Icond (_, _, n1, n2) -> [n1; n2]
- | Itailcall _ | Ireturn _ -> []
- | Ijumptable (_, ln) -> ln
- ) in dfs_visit code next_visits;
+ | Some i -> let next_visits = rtl_successors i in dfs_visit code next_visits;
visited := PTree.set node Visited !visited;
dfs_visit code ln
end
@@ -134,124 +134,220 @@ let get_loop_headers code entrypoint =
dfs_visit code [entrypoint];
!is_loop_header
end
+end
let ptree_printbool pt =
let elements = PTree.elements pt
in begin
- Printf.printf "[";
- List.iter (fun (n, b) ->
- if b then Printf.printf "%d, " (P.to_int n) else ()
- ) elements;
- Printf.printf "]"
+ if !debug_flag then begin
+ Printf.printf "[";
+ List.iter (fun (n, b) ->
+ if b then Printf.printf "%d, " (P.to_int n) else ()
+ ) elements;
+ Printf.printf "]"
+ end
end
(* Looks ahead (until a branch) to see if a node further down verifies
* the given predicate *)
let rec look_ahead code node is_loop_header predicate =
if (predicate node) then true
- else match (get_some @@ PTree.get node code) with
- | Ireturn _ | Itailcall _ | Icond _ | Ijumptable _ -> false
- | Inop n | Iop (_, _, _, n) | Iload (_, _, _, _, _, n)
- | Istore (_, _, _, _, n) | Icall (_, _, _, _, n)
- | Ibuiltin (_, _, _, n) ->
- if (predicate n) then true
- else (
- if (get_some @@ PTree.get n is_loop_header) then false
- else look_ahead code n is_loop_header predicate
- )
+ else match (rtl_successors @@ get_some @@ PTree.get node code) with
+ | [n] -> if (predicate n) then true
+ else (
+ if (get_some @@ PTree.get n is_loop_header) then false
+ else look_ahead code n is_loop_header predicate
+ )
+ | _ -> false
let do_call_heuristic code cond ifso ifnot is_loop_header =
- let predicate n = (function
- | Icall _ -> true
- | _ -> false) @@ get_some @@ PTree.get n code
- in if (look_ahead code ifso is_loop_header predicate) then Some false
- else if (look_ahead code ifnot is_loop_header predicate) then Some true
- else None
+ begin
+ debug "\tCall heuristic..\n";
+ let predicate n = (function
+ | Icall _ -> true
+ | _ -> false) @@ get_some @@ PTree.get n code
+ in let ifso_call = look_ahead code ifso is_loop_header predicate
+ in let ifnot_call = look_ahead code ifnot is_loop_header predicate
+ in if ifso_call && ifnot_call then None
+ else if ifso_call then Some false
+ else if ifnot_call then Some true
+ else None
+ end
-let do_opcode_heuristic code cond ifso ifnot is_loop_header = DuplicateOpcodeHeuristic.opcode_heuristic code cond ifso ifnot is_loop_header
+let do_opcode_heuristic code cond ifso ifnot is_loop_header =
+ begin
+ debug "\tOpcode heuristic..\n";
+ DuplicateOpcodeHeuristic.opcode_heuristic code cond ifso ifnot is_loop_header
+ end
let do_return_heuristic code cond ifso ifnot is_loop_header =
- let predicate n = (function
- | Ireturn _ -> true
- | _ -> false) @@ get_some @@ PTree.get n code
- in if (look_ahead code ifso is_loop_header predicate) then Some false
- else if (look_ahead code ifnot is_loop_header predicate) then Some true
- else None
+ begin
+ debug "\tReturn heuristic..\n";
+ let predicate n = (function
+ | Ireturn _ -> true
+ | _ -> false) @@ get_some @@ PTree.get n code
+ in let ifso_return = look_ahead code ifso is_loop_header predicate
+ in let ifnot_return = look_ahead code ifnot is_loop_header predicate
+ in if ifso_return && ifnot_return then None
+ else if ifso_return then Some false
+ else if ifnot_return then Some true
+ else None
+ end
let do_store_heuristic code cond ifso ifnot is_loop_header =
- let predicate n = (function
- | Istore _ -> true
- | _ -> false) @@ get_some @@ PTree.get n code
- in if (look_ahead code ifso is_loop_header predicate) then Some false
- else if (look_ahead code ifnot is_loop_header predicate) then Some true
- else None
+ begin
+ debug "\tStore heuristic..\n";
+ let predicate n = (function
+ | Istore _ -> true
+ | _ -> false) @@ get_some @@ PTree.get n code
+ in let ifso_store = look_ahead code ifso is_loop_header predicate
+ in let ifnot_store = look_ahead code ifnot is_loop_header predicate
+ in if ifso_store && ifnot_store then None
+ else if ifso_store then Some false
+ else if ifnot_store then Some true
+ else None
+ end
let do_loop_heuristic code cond ifso ifnot is_loop_header =
- let predicate n = get_some @@ PTree.get n is_loop_header
- in if (look_ahead code ifso is_loop_header predicate) then Some true
- else if (look_ahead code ifnot is_loop_header predicate) then Some false
- else None
-
-let get_directions code entrypoint =
- let bfs_order = bfs code entrypoint
- and is_loop_header = get_loop_headers code entrypoint
- and directions = ref (PTree.map (fun n i -> false) code) (* false <=> fallthru *)
+ begin
+ debug "\tLoop heuristic..\n";
+ let predicate n = get_some @@ PTree.get n is_loop_header in
+ let ifso_loop = look_ahead code ifso is_loop_header predicate in
+ let ifnot_loop = look_ahead code ifnot is_loop_header predicate in
+ if ifso_loop && ifnot_loop then None (* TODO - take the innermost loop ? *)
+ else if ifso_loop then Some true
+ else if ifnot_loop then Some false
+ else None
+ end
+
+let do_loop2_heuristic loop_info n code cond ifso ifnot is_loop_header =
+ begin
+ debug "\tLoop2 heuristic..\n";
+ match get_some @@ PTree.get n loop_info with
+ | None -> None
+ | Some b -> Some b
+ end
+
+(* Returns a PTree of either None or Some b where b determines the node following the loop, for a cb instruction *)
+(* It uses the fact that loops in CompCert are done by a branch (backedge) instruction followed by a cb *)
+let get_loop_info is_loop_header bfs_order code =
+ let loop_info = ref (PTree.map (fun n i -> None) code) in
+ let mark_path s n =
+ let visited = ref (PTree.map (fun n i -> false) code) in
+ let rec explore src dest =
+ if (get_some @@ PTree.get src !visited) then false
+ else if src == dest then true
+ else begin
+ visited := PTree.set src true !visited;
+ match rtl_successors @@ get_some @@ PTree.get src code with
+ | [] -> false
+ | [s] -> explore s dest
+ | [s1; s2] -> (explore s1 dest) || (explore s2 dest)
+ | _ -> false
+ end
+ in let rec advance_to_cb src =
+ if (get_some @@ PTree.get src !visited) then None
+ else begin
+ visited := PTree.set src true !visited;
+ match get_some @@ PTree.get src code with
+ | Inop s | Iop (_, _, _, s) | Iload (_,_,_,_,_,s) | Istore (_,_,_,_,s) | Icall (_,_,_,_,s)
+ | Ibuiltin (_,_,_,s) -> advance_to_cb s
+ | Icond _ -> Some src
+ | Ijumptable _ | Itailcall _ | Ireturn _ -> None
+ end
+ in begin
+ debug "Marking path from %d to %d\n" (P.to_int n) (P.to_int s);
+ match advance_to_cb s with
+ | None -> (debug "Nothing found\n")
+ | Some s -> ( debug "Advancing to %d\n" (P.to_int s);
+ match get_some @@ PTree.get s !loop_info with
+ | None | Some _ -> begin
+ match get_some @@ PTree.get s code with
+ | Icond (_, _, n1, n2, _) ->
+ let b1 = explore n1 n in
+ let b2 = explore n2 n in
+ if (b1 && b2) then (debug "both true\n")
+ else if b1 then (debug "true privileged\n"; loop_info := PTree.set s (Some true) !loop_info)
+ else if b2 then (debug "false privileged\n"; loop_info := PTree.set s (Some false) !loop_info)
+ else (debug "none true\n")
+ | _ -> ( debug "not an icond\n" )
+ end
+ (* | Some _ -> ( debug "already loop info there\n" ) FIXME - we don't know yet whether a branch to a loop head is a backedge or not *)
+ )
+ end
in begin
- (* Printf.printf "Loop headers: "; *)
+ List.iter (fun n ->
+ match get_some @@ PTree.get n code with
+ | Inop s | Iop (_,_,_,s) | Iload (_,_,_,_,_,s) | Istore (_,_,_,_,s) | Icall (_,_,_,_,s)
+ | Ibuiltin (_, _, _, s) ->
+ if get_some @@ PTree.get s is_loop_header then mark_path s n
+ | Icond _ -> () (* loop backedges are never Icond in CompCert RTL.3 *)
+ | Ijumptable _ -> ()
+ | Itailcall _ | Ireturn _ -> ()
+ ) bfs_order;
+ !loop_info
+ end
+
+(* Remark - compared to the original paper, we don't use the store heuristic *)
+let get_directions code entrypoint = begin
+ debug "get_directions\n";
+ let bfs_order = bfs code entrypoint in
+ let is_loop_header = get_loop_headers code entrypoint in
+ let loop_info = get_loop_info is_loop_header bfs_order code in
+ let directions = ref (PTree.map (fun n i -> None) code) in (* None <=> no predicted direction *)
+ begin
(* ptree_printbool is_loop_header; *)
- (* Printf.printf "\n"; *)
+ (* debug "\n"; *)
List.iter (fun n ->
match (get_some @@ PTree.get n code) with
- | Icond (cond, lr, ifso, ifnot) ->
- (* Printf.printf "Analyzing %d.." (P.to_int n); *)
- let heuristics = [ do_call_heuristic; do_opcode_heuristic;
- do_return_heuristic; do_store_heuristic; do_loop_heuristic ] in
+ | Icond (cond, lr, ifso, ifnot, _) ->
+ (* debug "Analyzing %d.." (P.to_int n); *)
+ let heuristics = [ do_opcode_heuristic;
+ do_return_heuristic; do_loop2_heuristic loop_info n; do_loop_heuristic; do_call_heuristic;
+ (* do_store_heuristic *) ] in
let preferred = ref None in
begin
+ debug "Deciding condition for RTL node %d\n" (P.to_int n);
List.iter (fun do_heur ->
match !preferred with
| None -> preferred := do_heur code cond ifso ifnot is_loop_header
| Some _ -> ()
) heuristics;
- (match !preferred with None -> preferred := Some (Random.bool ()) | Some _ -> ());
- directions := PTree.set n (get_some !preferred) !directions
+ directions := PTree.set n !preferred !directions;
+ (match !preferred with | Some false -> debug "\tFALLTHROUGH\n"
+ | Some true -> debug "\tBRANCH\n"
+ | None -> debug "\tUNSURE\n");
+ debug "---------------------------------------\n"
end
| _ -> ()
) bfs_order;
!directions
end
+end
+
+let update_direction direction = function
+| Icond (cond, lr, n, n', _) -> Icond (cond, lr, n, n', direction)
+| i -> i
-let to_ttl_inst direction = function
-| Ireturn o -> Tleaf (Ireturn o)
-| Inop n -> Tnext (n, Inop n)
-| Iop (op, lr, r, n) -> Tnext (n, Iop(op, lr, r, n))
-| Iload (tm, m, a, lr, r, n) -> Tnext (n, Iload(tm, m, a, lr, r, n))
-| Istore (m, a, lr, r, n) -> Tnext (n, Istore(m, a, lr, r, n))
-| Icall (s, ri, lr, r, n) -> Tleaf (Icall(s, ri, lr, r, n))
-| Itailcall (s, ri, lr) -> Tleaf (Itailcall(s, ri, lr))
-| Ibuiltin (ef, lbr, br, n) -> Tleaf (Ibuiltin(ef, lbr, br, n))
-| Icond (cond, lr, n, n') -> (match direction with
- | false -> Tnext (n', Icond(cond, lr, n, n'))
- | true -> Tnext (n, Icond(cond, lr, n, n')))
-| Ijumptable (r, ln) -> Tleaf (Ijumptable(r, ln))
-
-let rec to_ttl_code_rec directions = function
+let rec update_direction_rec directions = function
| [] -> PTree.empty
| m::lm -> let (n, i) = m
in let direction = get_some @@ PTree.get n directions
- in PTree.set n (to_ttl_inst direction i) (to_ttl_code_rec directions lm)
+ in PTree.set n (update_direction direction i) (update_direction_rec directions lm)
-let to_ttl_code code entrypoint =
+(* Uses branch prediction to write prediction annotations in Icond *)
+let update_directions code entrypoint = begin
+ debug "Update_directions\n";
let directions = get_directions code entrypoint
in begin
- (* Printf.printf "Ifso directions: ";
+ (* debug "Ifso directions: ";
ptree_printbool directions;
- Printf.printf "\n"; *)
- Random.init(0); (* using same seed to make it deterministic *)
- to_ttl_code_rec directions (PTree.elements code)
+ debug "\n"; *)
+ update_direction_rec directions (PTree.elements code)
end
+end
-(** Trace selection on TTL *)
+(** Trace selection *)
let rec exists_false_rec = function
| [] -> false
@@ -259,50 +355,29 @@ let rec exists_false_rec = function
let exists_false boolmap = exists_false_rec (PTree.elements boolmap)
-(* DFS on TTL to guide the exploration *)
-let dfs code entrypoint =
+(* DFS using prediction info to guide the exploration *)
+let dfs code entrypoint = begin
+ debug "dfs\n";
let visited = ref (PTree.map (fun n i -> false) code) in
let rec dfs_list code = function
| [] -> []
| node :: ln ->
- let node_dfs =
- if not (get_some @@ PTree.get node !visited) then begin
- visited := PTree.set node true !visited;
- match PTree.get node code with
- | None -> failwith "No such node"
- | Some ti -> [node] @ match ti with
- | Tleaf i -> (match i with
- | Icall(_, _, _, _, n) -> dfs_list code [n]
- | Ibuiltin(_, _, _, n) -> dfs_list code [n]
- | Ijumptable(_, ln) -> dfs_list code ln
- | Itailcall _ | Ireturn _ -> []
- | _ -> failwith "Tleaf case not handled in dfs" )
- | Tnext (n,i) -> (dfs_list code [n]) @ match i with
- | Icond (_, _, n1, n2) -> dfs_list code [n1; n2]
- | Inop _ | Iop _ | Iload _ | Istore _ -> []
- | _ -> failwith "Tnext case not handled in dfs"
- end
- else []
- in node_dfs @ (dfs_list code ln)
+ if get_some @@ PTree.get node !visited then dfs_list code ln
+ else begin
+ visited := PTree.set node true !visited;
+ let next_nodes = (match get_some @@ PTree.get node code with
+ | Icall(_, _, _, _, n) | Ibuiltin (_, _, _, n) | Iop (_, _, _, n)
+ | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) | Inop n -> [n]
+ | Ijumptable (_, ln) -> ln
+ | Itailcall _ | Ireturn _ -> []
+ | Icond (_, _, n1, n2, info) -> (match info with
+ | Some false -> [n2; n1]
+ | _ -> [n1; n2]
+ )
+ ) in node :: dfs_list code (next_nodes @ ln)
+ end
in dfs_list code [entrypoint]
-
-let get_predecessors_ttl code =
- let preds = ref (PTree.map (fun n i -> []) code) in
- let process_inst (node, ti) = match ti with
- | Tleaf _ -> ()
- | Tnext (_, i) -> let succ = match i with
- | Inop n | Iop (_,_,_,n) | Iload (_, _,_,_,_,n) | Istore (_,_,_,_,n)
- | Icall (_,_,_,_,n) | Ibuiltin (_, _, _, n) -> [n]
- | Icond (_,_,n1,n2) -> [n1;n2]
- | Ijumptable (_,ln) -> ln
- | _ -> []
- in List.iter (fun s -> preds := PTree.set s (node::(get_some @@ PTree.get s !preds)) !preds) succ
- in begin
- List.iter process_inst (PTree.elements code);
- !preds
- end
-
-let rtl_proj code = PTree.map (fun n ti -> match ti with Tleaf i | Tnext(_, i) -> i) code
+end
let rec select_unvisited_node is_visited = function
| [] -> failwith "Empty list"
@@ -311,24 +386,87 @@ let rec select_unvisited_node is_visited = function
let best_successor_of node code is_visited =
match (PTree.get node code) with
| None -> failwith "No such node in the code"
- | Some ti -> match ti with
- | Tleaf _ -> None
- | Tnext (n,_) -> if not (ptree_get_some n is_visited) then Some n
- else None
-
-let best_predecessor_of node predecessors order is_visited =
+ | Some i ->
+ let next_node = match i with
+ | Inop n | Iop (_,_,_,n) | Iload (_,_,_,_,_,n) | Istore(_,_,_,_,n)
+ | Icall (_,_,_,_,n) | Ibuiltin (_,_,_,n) -> Some n
+ | Icond (_, _, n1, n2, ob) -> (match ob with None -> None | Some false -> Some n2 | Some true -> Some n1)
+ | _ -> None
+ in match next_node with
+ | None -> None
+ | Some n -> if not (ptree_get_some n is_visited) then Some n else None
+
+(* FIXME - could be improved by selecting in priority the predicted paths *)
+let best_predecessor_of node predecessors code order is_visited =
match (PTree.get node predecessors) with
| None -> failwith "No predecessor list found"
- | Some lp -> try Some (List.find (fun n -> (List.mem n lp) && (not (ptree_get_some n is_visited))) order)
- with Not_found -> None
+ | Some lp ->
+ try Some (List.find (fun n ->
+ if (List.mem n lp) && (not (ptree_get_some n is_visited)) then
+ match ptree_get_some n code with
+ | Icond (_, _, n1, n2, ob) -> (match ob with
+ | None -> false
+ | Some false -> n == n2
+ | Some true -> n == n1
+ )
+ | _ -> true
+ else false
+ ) order)
+ with Not_found -> None
+
+let print_trace t = print_intlist t
+
+let print_traces traces =
+ let rec f = function
+ | [] -> ()
+ | t::lt -> Printf.printf "\n\t"; print_trace t; Printf.printf ",\n"; f lt
+ in begin
+ if !debug_flag then begin
+ Printf.printf "Traces: {";
+ f traces;
+ Printf.printf "}\n";
+ end
+ end
+
+(* Dumb (but linear) trace selection *)
+let select_traces_linear code entrypoint =
+ let is_visited = ref (PTree.map (fun n i -> false) code) in
+ let bfs_order = bfs code entrypoint in
+ let rec go_through node = begin
+ is_visited := PTree.set node true !is_visited;
+ let next_node = match (get_some @@ PTree.get node code) with
+ | Icall(_, _, _, _, n) | Ibuiltin (_, _, _, n) | Iop (_, _, _, n)
+ | Iload (_, _, _, _, _, n) | Istore (_, _, _, _, n) | Inop n -> Some n
+ | Ijumptable _ | Itailcall _ | Ireturn _ -> None
+ | Icond (_, _, n1, n2, info) -> (match info with
+ | Some false -> Some n2
+ | Some true -> Some n1
+ | None -> None
+ )
+ in match next_node with
+ | None -> [node]
+ | Some n ->
+ if not (get_some @@ PTree.get n !is_visited) then node :: go_through n
+ else [node]
+ end
+ in let traces = ref [] in begin
+ List.iter (fun n ->
+ if not (get_some @@ PTree.get n !is_visited) then
+ traces := (go_through n) :: !traces
+ ) bfs_order;
+ !traces
+ end
+
(* Algorithm mostly inspired from Chang and Hwu 1988
* "Trace Selection for Compiling Large C Application Programs to Microcode" *)
-let select_traces code entrypoint =
+let select_traces_chang code entrypoint = begin
+ debug "select_traces\n";
let order = dfs code entrypoint in
- let predecessors = get_predecessors_ttl code in
+ let predecessors = get_predecessors_rtl code in
let traces = ref [] in
let is_visited = ref (PTree.map (fun n i -> false) code) in begin (* mark all nodes visited *)
+ debug "Length: %d\n" (List.length order);
while exists_false !is_visited do (* while (there are unvisited nodes) *)
let seed = select_unvisited_node !is_visited order in
let trace = ref [seed] in
@@ -348,7 +486,7 @@ let select_traces code entrypoint =
current := seed;
quit_loop := false;
while not !quit_loop do
- let s = best_predecessor_of !current predecessors order !is_visited in
+ let s = best_predecessor_of !current predecessors code order !is_visited in
match s with
| None -> quit_loop := true (* if (s==0) exit loop *)
| Some pred -> begin
@@ -361,21 +499,16 @@ let select_traces code entrypoint =
end
end
done;
- (* Printf.printf "DFS: \t"; print_intlist order; Printf.printf "\n"; *)
+ (* debug "DFS: \t"; print_intlist order; debug "\n"; *)
+ debug "Traces: "; print_traces !traces;
!traces
end
+end
-let print_trace t = print_intlist t
-
-let print_traces traces =
- let rec f = function
- | [] -> ()
- | t::lt -> Printf.printf "\n\t"; print_trace t; Printf.printf ",\n"; f lt
- in begin
- Printf.printf "Traces: {";
- f traces;
- Printf.printf "}\n";
- end
+let select_traces code entrypoint =
+ let length = List.length @@ PTree.elements code in
+ if (length < 5000) then select_traces_chang code entrypoint
+ else select_traces_linear code entrypoint
let rec make_identity_ptree_rec = function
| [] -> PTree.empty
@@ -392,10 +525,10 @@ let rec change_pointers code n n' = function
| Ibuiltin(a, b, c, n0) -> assert (n0 == n); Ibuiltin(a, b, c, n')
| Ijumptable(a, ln) -> assert (optbool @@ List.find_opt (fun e -> e == n) ln);
Ijumptable(a, List.map (fun e -> if (e == n) then n' else e) ln)
- | Icond(a, b, n1, n2) -> assert (n1 == n || n2 == n);
+ | Icond(a, b, n1, n2, i) -> assert (n1 == n || n2 == n);
let n1' = if (n1 == n) then n' else n1
in let n2' = if (n2 == n) then n' else n2
- in Icond(a, b, n1', n2')
+ in Icond(a, b, n1', n2', i)
| Inop n0 -> assert (n0 == n); Inop n'
| Iop (a, b, c, n0) -> assert (n0 == n); Iop (a, b, c, n')
| Iload (a, b, c, d, e, n0) -> assert (n0 == n); Iload (a, b, c, d, e, n')
@@ -409,7 +542,7 @@ let rec change_pointers code n n' = function
* n': the integer which should contain the duplicate of n
* returns: new code, new ptree *)
let duplicate code ptree parent n preds n' =
- (* Printf.printf "Duplicating node %d into %d..\n" (P.to_int n) (P.to_int n'); *)
+ debug "Duplicating node %d into %d..\n" (P.to_int n) (P.to_int n');
match PTree.get n' code with
| Some _ -> failwith "The PTree already has a node n'"
| None ->
@@ -473,24 +606,24 @@ let superblockify_traces code preds traces =
| [] -> (code, ptree, 0)
| trace :: traces ->
let new_code, new_ptree, nb_duplicated = tail_duplicate code preds ptree trace
- in if (nb_duplicated < max_nb_duplicated) then f new_code new_ptree traces
- else (Printf.printf "Too many duplicated nodes, aborting tail duplication\n"; (code, ptree, 0))
+ in if (nb_duplicated < max_nb_duplicated)
+ then (debug "End duplication\n"; f new_code new_ptree traces)
+ else (debug "Too many duplicated nodes, aborting tail duplication\n"; (code, ptree, 0))
in let new_code, new_ptree, _ = f code ptree traces
in (new_code, new_ptree)
let rec invert_iconds_trace code = function
| [] -> code
- | n::[] -> code
- | n :: n' :: t ->
+ | n :: ln ->
let code' = match ptree_get_some n code with
- | Icond (c, lr, ifso, ifnot) ->
- assert (n' == ifso || n' == ifnot);
- if (n' == ifso) then (
- (* Printf.printf "Reversing ifso/ifnot for node %d\n" (P.to_int n); *)
- PTree.set n (Icond (Op.negate_condition c, lr, ifnot, ifso)) code )
- else code
+ | Icond (c, lr, ifso, ifnot, info) -> (match info with
+ | Some true -> begin
+ (* debug "Reversing ifso/ifnot for node %d\n" (P.to_int n); *)
+ PTree.set n (Icond (Op.negate_condition c, lr, ifnot, ifso, Some false)) code
+ end
+ | _ -> code)
| _ -> code
- in invert_iconds_trace code' (n'::t)
+ in invert_iconds_trace code' ln
let rec invert_iconds code = function
| [] -> code
@@ -501,11 +634,11 @@ let rec invert_iconds code = function
let duplicate_aux f =
let entrypoint = f.fn_entrypoint in
- let code = f.fn_code in
if !Clflags.option_fduplicate < 0 then
- ((code, entrypoint), make_identity_ptree code)
+ ((f.fn_code, entrypoint), make_identity_ptree f.fn_code)
else
- let traces = select_traces (to_ttl_code code entrypoint) entrypoint in
+ let code = update_directions (f.fn_code) entrypoint in
+ let traces = select_traces code entrypoint in
let icond_code = invert_iconds code traces in
let preds = get_predecessors_rtl icond_code in
if !Clflags.option_fduplicate >= 1 then
@@ -513,4 +646,3 @@ let duplicate_aux f =
((new_code, f.fn_entrypoint), pTreeId)
else
((icond_code, entrypoint), make_identity_ptree code)
-
diff --git a/backend/Duplicateproof.v b/backend/Duplicateproof.v
index 466b4b75..6b598dc7 100644
--- a/backend/Duplicateproof.v
+++ b/backend/Duplicateproof.v
@@ -23,12 +23,12 @@ Inductive match_inst (dupmap: PTree.t node): instruction -> instruction -> Prop
match_inst dupmap (Itailcall s ri lr) (Itailcall s ri lr)
| match_inst_builtin: forall n n' ef la br,
dupmap!n' = (Some n) -> match_inst dupmap (Ibuiltin ef la br n) (Ibuiltin ef la br n')
- | match_inst_cond: forall ifso ifso' ifnot ifnot' c lr,
+ | match_inst_cond: forall ifso ifso' ifnot ifnot' c lr i i',
dupmap!ifso' = (Some ifso) -> dupmap!ifnot' = (Some ifnot) ->
- match_inst dupmap (Icond c lr ifso ifnot) (Icond c lr ifso' ifnot')
- | match_inst_revcond: forall ifso ifso' ifnot ifnot' c lr,
+ match_inst dupmap (Icond c lr ifso ifnot i) (Icond c lr ifso' ifnot' i')
+ | match_inst_revcond: forall ifso ifso' ifnot ifnot' c lr i i',
dupmap!ifso' = (Some ifso) -> dupmap!ifnot' = (Some ifnot) ->
- match_inst dupmap (Icond c lr ifso ifnot) (Icond (negate_condition c) lr ifnot' ifso')
+ match_inst dupmap (Icond c lr ifso ifnot i) (Icond (negate_condition c) lr ifnot' ifso' i')
| match_inst_jumptable: forall ln ln' r,
list_forall2 (fun n n' => (dupmap!n' = (Some n))) ln ln' ->
match_inst dupmap (Ijumptable r ln) (Ijumptable r ln')
diff --git a/backend/ForwardMoves.v b/backend/ForwardMoves.v
index c73b0213..7cfd411f 100644
--- a/backend/ForwardMoves.v
+++ b/backend/ForwardMoves.v
@@ -250,7 +250,7 @@ Fixpoint kill_builtin_res (res : builtin_res reg) (rel : RELATION.t) :=
Definition apply_instr instr x :=
match instr with
| Inop _
- | Icond _ _ _ _
+ | Icond _ _ _ _ _
| Ijumptable _ _
| Istore _ _ _ _ _ => Some x
| Iop Omove (src :: nil) dst _ => Some (move src dst x)
@@ -309,8 +309,8 @@ Definition transf_instr (fmap : option (PMap.t RB.t))
Icall sig ros (subst_args fmap pc args) dst s
| Itailcall sig ros args =>
Itailcall sig ros (subst_args fmap pc args)
- | Icond cond args s1 s2 =>
- Icond cond (subst_args fmap pc args) s1 s2
+ | Icond cond args s1 s2 i =>
+ Icond cond (subst_args fmap pc args) s1 s2 i
| Ijumptable arg tbl =>
Ijumptable (subst_arg fmap pc arg) tbl
| Ireturn (Some arg) =>
diff --git a/backend/Inlining.v b/backend/Inlining.v
index 9cf535b9..8c7e1898 100644
--- a/backend/Inlining.v
+++ b/backend/Inlining.v
@@ -397,9 +397,9 @@ Definition expand_instr (ctx: context) (pc: node) (i: instruction): mon unit :=
| Ibuiltin ef args res s =>
set_instr (spc ctx pc)
(Ibuiltin ef (map (sbuiltinarg ctx) args) (sbuiltinres ctx res) (spc ctx s))
- | Icond cond args s1 s2 =>
+ | Icond cond args s1 s2 info =>
set_instr (spc ctx pc)
- (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2))
+ (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2) info)
| Ijumptable r tbl =>
set_instr (spc ctx pc)
(Ijumptable (sreg ctx r) (List.map (spc ctx) tbl))
diff --git a/backend/Inliningspec.v b/backend/Inliningspec.v
index e20fb373..eba026ec 100644
--- a/backend/Inliningspec.v
+++ b/backend/Inliningspec.v
@@ -312,9 +312,9 @@ Inductive tr_instr: context -> node -> instruction -> code -> Prop :=
match res with BR r => Ple r ctx.(mreg) | _ => True end ->
c!(spc ctx pc) = Some (Ibuiltin ef (map (sbuiltinarg ctx) args) (sbuiltinres ctx res) (spc ctx s)) ->
tr_instr ctx pc (Ibuiltin ef args res s) c
- | tr_cond: forall ctx pc cond args s1 s2 c,
- c!(spc ctx pc) = Some (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2)) ->
- tr_instr ctx pc (Icond cond args s1 s2) c
+ | tr_cond: forall ctx pc cond args s1 s2 c i,
+ c!(spc ctx pc) = Some (Icond cond (sregs ctx args) (spc ctx s1) (spc ctx s2) i) ->
+ tr_instr ctx pc (Icond cond args s1 s2 i) c
| tr_jumptable: forall ctx pc r tbl c,
c!(spc ctx pc) = Some (Ijumptable (sreg ctx r) (List.map (spc ctx) tbl)) ->
tr_instr ctx pc (Ijumptable r tbl) c
diff --git a/backend/LTL.v b/backend/LTL.v
index ee8b4826..3edd60a2 100644
--- a/backend/LTL.v
+++ b/backend/LTL.v
@@ -37,7 +37,7 @@ Inductive instruction: Type :=
| Ltailcall (sg: signature) (ros: mreg + ident)
| Lbuiltin (ef: external_function) (args: list (builtin_arg loc)) (res: builtin_res mreg)
| Lbranch (s: node)
- | Lcond (cond: condition) (args: list mreg) (s1 s2: node)
+ | Lcond (cond: condition) (args: list mreg) (s1 s2: node) (info: option bool)
| Ljumptable (arg: mreg) (tbl: list node)
| Lreturn.
@@ -263,11 +263,11 @@ Inductive step: state -> trace -> state -> Prop :=
| exec_Lbranch: forall s f sp pc bb rs m,
step (Block s f sp (Lbranch pc :: bb) rs m)
E0 (State s f sp pc rs m)
- | exec_Lcond: forall s f sp cond args pc1 pc2 bb rs b pc rs' m,
+ | exec_Lcond: forall s f sp cond args pc1 pc2 bb rs b pc rs' m i,
eval_condition cond (reglist rs args) m = Some b ->
pc = (if b then pc1 else pc2) ->
rs' = undef_regs (destroyed_by_cond cond) rs ->
- step (Block s f sp (Lcond cond args pc1 pc2 :: bb) rs m)
+ step (Block s f sp (Lcond cond args pc1 pc2 i :: bb) rs m)
E0 (State s f sp pc rs' m)
| exec_Ljumptable: forall s f sp arg tbl bb rs m n pc rs',
rs (R arg) = Vint n ->
@@ -328,7 +328,7 @@ Fixpoint successors_block (b: bblock) : list node :=
| nil => nil (**r should never happen *)
| Ltailcall _ _ :: _ => nil
| Lbranch s :: _ => s :: nil
- | Lcond _ _ s1 s2 :: _ => s1 :: s2 :: nil
+ | Lcond _ _ s1 s2 _ :: _ => s1 :: s2 :: nil
| Ljumptable _ tbl :: _ => tbl
| Lreturn :: _ => nil
| instr :: b' => successors_block b'
diff --git a/backend/Linearize.v b/backend/Linearize.v
index 4216958c..66b36428 100644
--- a/backend/Linearize.v
+++ b/backend/Linearize.v
@@ -179,7 +179,7 @@ Fixpoint linearize_block (b: LTL.bblock) (k: code) : code :=
Lbuiltin ef args res :: linearize_block b' k
| LTL.Lbranch s :: b' =>
add_branch s k
- | LTL.Lcond cond args s1 s2 :: b' =>
+ | LTL.Lcond cond args s1 s2 _ :: b' =>
if starts_with s1 k then
Lcond (negate_condition cond) args s2 :: add_branch s1 k
else
diff --git a/backend/Linearizeaux.ml b/backend/Linearizeaux.ml
index a813ac96..1381877b 100644
--- a/backend/Linearizeaux.ml
+++ b/backend/Linearizeaux.ml
@@ -13,6 +13,12 @@
open LTL
open Maps
+let debug_flag = ref false
+
+let debug fmt =
+ if !debug_flag then Printf.eprintf fmt
+ else Printf.ifprintf stderr fmt
+
(* Trivial enumeration, in decreasing order of PC *)
(***
@@ -81,7 +87,7 @@ let basic_blocks f joins =
| [] -> assert false
| Lbranch s :: _ -> next_in_block blk minpc s
| Ltailcall (sig0, ros) :: _ -> end_block blk minpc
- | Lcond (cond, args, ifso, ifnot) :: _ ->
+ | Lcond (cond, args, ifso, ifnot, _) :: _ ->
end_block blk minpc; start_block ifso; start_block ifnot
| Ljumptable(arg, tbl) :: _ ->
end_block blk minpc; List.iter start_block tbl
@@ -115,18 +121,11 @@ let enumerate_aux_flat f reach =
flatten_blocks (basic_blocks f (join_points f))
(**
- * Enumeration based on traces as identified by Duplicate.v
- *
- * The Duplicate phase heuristically identifies the most frequented paths. Each
- * Icond is modified so that the preferred condition is a fallthrough (ifnot)
- * rather than a branch (ifso).
+ * Alternate enumeration based on traces as identified by Duplicate.v
*
- * The enumeration below takes advantage of this - preferring to layout nodes
- * following the fallthroughs of the Lcond branches.
- *
- * It is slightly adapted from the work of Petris and Hansen 90 on intraprocedural
- * code positioning - only we do it on a broader grain, since we don't have the exact
- * frequencies (we only know which branch is the preferred one)
+ * This is a slight alteration to the above heuristic, ensuring that any
+ * superblock will be contiguous in memory, while still following the original
+ * heuristic
*)
let get_some = function
@@ -145,16 +144,37 @@ let print_plist l =
| [] -> ()
| n :: l -> Printf.printf "%d, " (P.to_int n); f l
in begin
- Printf.printf "[";
- f l;
- Printf.printf "]"
+ if !debug_flag then begin
+ Printf.printf "[";
+ f l;
+ Printf.printf "]"
+ end
end
+(* adapted from the above join_points function, but with PTree *)
+let get_join_points code entry =
+ let reached = ref (PTree.map (fun n i -> false) code) in
+ let reached_twice = ref (PTree.map (fun n i -> false) code) in
+ let rec traverse pc =
+ if get_some @@ PTree.get pc !reached then begin
+ if not (get_some @@ PTree.get pc !reached_twice) then
+ reached_twice := PTree.set pc true !reached_twice
+ end else begin
+ reached := PTree.set pc true !reached;
+ traverse_succs (successors_block @@ get_some @@ PTree.get pc code)
+ end
+ and traverse_succs = function
+ | [] -> ()
+ | [pc] -> traverse pc
+ | pc :: l -> traverse pc; traverse_succs l
+ in traverse entry; !reached_twice
+
let forward_sequences code entry =
let visited = ref (PTree.map (fun n i -> false) code) in
+ let join_points = get_join_points code entry in
(* returns the list of traversed nodes, and a list of nodes to start traversing next *)
let rec traverse_fallthrough code node =
- (* Printf.printf "Traversing %d..\n" (P.to_int node); *)
+ (* debug "Traversing %d..\n" (P.to_int node); *)
if not (get_some @@ PTree.get node !visited) then begin
visited := PTree.set node true !visited;
match PTree.get node code with
@@ -163,12 +183,19 @@ let forward_sequences code entry =
let ln, rem = match (last_element bb) with
| Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _
| Lbuiltin _ -> assert false
- | Ltailcall _ | Lreturn -> ([], [])
- | Lbranch n -> let ln, rem = traverse_fallthrough code n in (ln, rem)
- | Lcond (_, _, ifso, ifnot) -> let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem)
- | Ljumptable(_, ln) -> match ln with
- | [] -> ([], [])
- | n :: ln -> let lln, rem = traverse_fallthrough code n in (lln, ln @ rem)
+ | Ltailcall _ | Lreturn -> begin (* debug "STOP tailcall/return\n"; *) ([], []) end
+ | Lbranch n ->
+ if get_some @@ PTree.get n join_points then ([], [n])
+ else let ln, rem = traverse_fallthrough code n in (ln, rem)
+ | Lcond (_, _, ifso, ifnot, info) -> (match info with
+ | None -> begin (* debug "STOP Lcond None\n"; *) ([], [ifso; ifnot]) end
+ | Some false ->
+ if get_some @@ PTree.get ifnot join_points then ([], [ifso; ifnot])
+ else let ln, rem = traverse_fallthrough code ifnot in (ln, [ifso] @ rem)
+ | Some true ->
+ let errstr = Printf.sprintf ("Inconsistency detected in node %d: ifnot is not the preferred branch") (P.to_int node) in
+ failwith errstr)
+ | Ljumptable(_, ln) -> begin (* debug "STOP Ljumptable\n"; *) ([], ln) end
in ([node] @ ln, rem)
end
else ([], [])
@@ -179,6 +206,7 @@ let forward_sequences code entry =
in [fs] @ ((f code rem_from_node) @ (f code ln))
in (f code [entry])
+(** Unused code
module PInt = struct
type t = P.t
let compare x y = compare (P.to_int x) (P.to_int y)
@@ -219,7 +247,10 @@ let can_be_merged code s s' =
| Lop _ | Lload _ | Lgetstack _ | Lsetstack _ | Lstore _ | Lcall _
| Lbuiltin _ | Ltailcall _ | Lreturn -> false
| Lbranch n -> n == first_s'
- | Lcond (_, _, ifso, ifnot) -> ifnot == first_s'
+ | Lcond (_, _, ifso, ifnot, info) -> (match info with
+ | None -> false
+ | Some false -> ifnot == first_s'
+ | Some true -> failwith "Inconsistency detected - ifnot is not the preferred branch")
| Ljumptable (_, ln) ->
match ln with
| [] -> false
@@ -256,6 +287,7 @@ let try_merge code (fs: (BinNums.positive list) list) =
end
done;
!seqs
+*)
(** Code adapted from Duplicateaux.get_loop_headers
*
@@ -303,7 +335,7 @@ let get_loop_edges code entry =
| Lbuiltin _ -> assert false
| Ltailcall _ | Lreturn -> []
| Lbranch n -> [n]
- | Lcond (_, _, ifso, ifnot) -> [ifso; ifnot]
+ | Lcond (_, _, ifso, ifnot, _) -> [ifso; ifnot]
| Ljumptable(_, ln) -> ln
) in dfs_visit code (Some node) next_visits;
visited := PTree.set node Visited !visited;
@@ -324,15 +356,19 @@ end
module ISet = Set.Make(Int)
let print_iset s = begin
- Printf.printf "{";
- ISet.iter (fun e -> Printf.printf "%d, " e) s;
- Printf.printf "}"
+ if !debug_flag then begin
+ Printf.printf "{";
+ ISet.iter (fun e -> Printf.printf "%d, " e) s;
+ Printf.printf "}"
+ end
end
let print_depmap dm = begin
- Printf.printf "[|";
- Array.iter (fun s -> print_iset s; Printf.printf ", ") dm;
- Printf.printf "|]\n"
+ if !debug_flag then begin
+ Printf.printf "[|";
+ Array.iter (fun s -> print_iset s; Printf.printf ", ") dm;
+ Printf.printf "|]\n"
+ end
end
let construct_depmap code entry fs =
@@ -350,7 +386,7 @@ let construct_depmap code entry fs =
!index
end
in let check_and_update_depmap from target =
- (* Printf.printf "From %d to %d\n" (P.to_int from) (P.to_int target); *)
+ (* debug "From %d to %d\n" (P.to_int from) (P.to_int target); *)
if not (ppmap_is_true (from, target) is_loop_edge) then
let in_index_fs = find_index_of_node from in
let out_index_fs = find_index_of_node target in
@@ -371,7 +407,7 @@ let construct_depmap code entry fs =
match (last_element bb) with
| Ltailcall _ | Lreturn -> []
| Lbranch n -> (check_and_update_depmap node n; [n])
- | Lcond (_, _, ifso, ifnot) -> begin
+ | Lcond (_, _, ifso, ifnot, _) -> begin
check_and_update_depmap node ifso;
check_and_update_depmap node ifnot;
[ifso; ifnot]
@@ -392,14 +428,18 @@ let construct_depmap code entry fs =
end
let print_sequence s =
- Printf.printf "[";
- List.iter (fun n -> Printf.printf "%d, " (P.to_int n)) s;
- Printf.printf "]\n"
+ if !debug_flag then begin
+ Printf.printf "[";
+ List.iter (fun n -> Printf.printf "%d, " (P.to_int n)) s;
+ Printf.printf "]\n"
+ end
let print_ssequence ofs =
- Printf.printf "[";
- List.iter (fun s -> print_sequence s) ofs;
- Printf.printf "]\n"
+ if !debug_flag then begin
+ Printf.printf "[";
+ List.iter (fun s -> print_sequence s) ofs;
+ Printf.printf "]\n"
+ end
let order_sequences code entry fs =
let fs_a = Array.of_list fs in
@@ -411,40 +451,64 @@ let order_sequences code entry fs =
assert (not fs_evaluated.(s_id));
ordered_fs := fs_a.(s_id) :: !ordered_fs;
fs_evaluated.(s_id) <- true;
+ (* debug "++++++\n";
+ debug "Scheduling %d\n" s_id;
+ debug "Initial depmap: "; print_depmap depmap; *)
Array.iteri (fun i deps ->
depmap.(i) <- ISet.remove s_id deps
- ) depmap
+ ) depmap;
+ (* debug "Final depmap: "; print_depmap depmap; *)
+ end
+ in let choose_best_of candidates =
+ let current_best_id = ref None in
+ let current_best_score = ref None in
+ begin
+ List.iter (fun id ->
+ match !current_best_id with
+ | None -> begin
+ current_best_id := Some id;
+ match fs_a.(id) with
+ | [] -> current_best_score := None
+ | n::l -> current_best_score := Some (P.to_int n)
+ end
+ | Some b -> begin
+ match fs_a.(id) with
+ | [] -> ()
+ | n::l -> let nscore = P.to_int n in
+ match !current_best_score with
+ | None -> (current_best_id := Some id; current_best_score := Some nscore)
+ | Some bs -> if nscore > bs then (current_best_id := Some id; current_best_score := Some nscore)
+ end
+ ) candidates;
+ !current_best_id
end
in let select_next () =
- let selected_id = ref None in
+ let candidates = ref [] in
begin
Array.iteri (fun i deps ->
begin
- (* Printf.printf "Deps: "; print_iset deps; Printf.printf "\n"; *)
- match !selected_id with
- | None -> if (deps == ISet.empty && not fs_evaluated.(i)) then selected_id := Some i
- | Some id -> ()
+ (* debug "Deps of %d: " i; print_iset deps; debug "\n"; *)
+ (* FIXME - if we keep it that way (no dependency check), remove all the unneeded stuff *)
+ if ((* deps == ISet.empty && *) not fs_evaluated.(i)) then
+ candidates := i :: !candidates
end
) depmap;
- match !selected_id with
- | Some id -> id
- | None -> begin
- Array.iteri (fun i deps ->
- match !selected_id with
- | None -> if not fs_evaluated.(i) then selected_id := Some i
- | Some id -> ()
- ) depmap;
- get_some !selected_id
- end
+ if not (List.length !candidates > 0) then begin
+ Array.iteri (fun i deps ->
+ if (not fs_evaluated.(i)) then candidates := i :: !candidates
+ ) depmap;
+ end;
+ get_some (choose_best_of !candidates)
end
in begin
- (* Printf.printf "depmap: "; print_depmap depmap; *)
- (* Printf.printf "forward sequences identified: "; print_ssequence fs; *)
+ debug "-------------------------------\n";
+ debug "depmap: "; print_depmap depmap;
+ debug "forward sequences identified: "; print_ssequence fs;
while List.length !ordered_fs != List.length fs do
let next_id = select_next () in
evaluate next_id
done;
- (* Printf.printf "forward sequences ordered: "; print_ssequence (List.rev (!ordered_fs)); *)
+ debug "forward sequences ordered: "; print_ssequence (List.rev (!ordered_fs));
List.rev (!ordered_fs)
end
diff --git a/backend/Liveness.v b/backend/Liveness.v
index afe11ae6..9652b363 100644
--- a/backend/Liveness.v
+++ b/backend/Liveness.v
@@ -94,7 +94,7 @@ Definition transfer
| Ibuiltin ef args res s =>
reg_list_live (params_of_builtin_args args)
(reg_list_dead (params_of_builtin_res res) after)
- | Icond cond args ifso ifnot =>
+ | Icond cond args ifso ifnot _ =>
reg_list_live args after
| Ijumptable arg tbl =>
reg_live arg after
diff --git a/backend/PrintLTL.ml b/backend/PrintLTL.ml
index b309a9f2..d8f2ac12 100644
--- a/backend/PrintLTL.ml
+++ b/backend/PrintLTL.ml
@@ -83,10 +83,11 @@ let print_instruction pp succ = function
(print_builtin_args loc) args
| Lbranch s ->
print_succ pp s succ
- | Lcond(cond, args, s1, s2) ->
- fprintf pp "if (%a) goto %d else goto %d"
+ | Lcond(cond, args, s1, s2, info) ->
+ fprintf pp "if (%a) goto %d else goto %d (prediction: %s)"
(print_condition mreg) (cond, args)
(P.to_int s1) (P.to_int s2)
+ (match info with None -> "none" | Some true -> "branch" | Some false -> "fallthrough")
| Ljumptable(arg, tbl) ->
let tbl = Array.of_list tbl in
fprintf pp "jumptable (%a)" mreg arg;
diff --git a/backend/PrintRTL.ml b/backend/PrintRTL.ml
index c25773e5..b2ef05ca 100644
--- a/backend/PrintRTL.ml
+++ b/backend/PrintRTL.ml
@@ -75,10 +75,11 @@ let print_instruction pp (pc, i) =
(name_of_external ef)
(print_builtin_args reg) args;
print_succ pp s (pc - 1)
- | Icond(cond, args, s1, s2) ->
- fprintf pp "if (%a) goto %d else goto %d\n"
+ | Icond(cond, args, s1, s2, info) ->
+ fprintf pp "if (%a) goto %d else goto %d (prediction: %s)\n"
(PrintOp.print_condition reg) (cond, args)
(P.to_int s1) (P.to_int s2)
+ (match info with None -> "none" | Some true -> "branch" | Some false -> "fallthrough")
| Ijumptable(arg, tbl) ->
let tbl = Array.of_list tbl in
fprintf pp "jumptable (%a)\n" reg arg;
diff --git a/backend/PrintXTL.ml b/backend/PrintXTL.ml
index 1c7655fb..d1b79623 100644
--- a/backend/PrintXTL.ml
+++ b/backend/PrintXTL.ml
@@ -104,7 +104,7 @@ let print_instruction pp succ = function
(print_builtin_args var) args
| Xbranch s ->
print_succ pp s succ
- | Xcond(cond, args, s1, s2) ->
+ | Xcond(cond, args, s1, s2, _) ->
fprintf pp "if (%a) goto %d else goto %d"
(print_condition var) (cond, args)
(P.to_int s1) (P.to_int s2)
diff --git a/backend/RTL.v b/backend/RTL.v
index 29a49311..dec59ca2 100644
--- a/backend/RTL.v
+++ b/backend/RTL.v
@@ -67,11 +67,12 @@ Inductive instruction: Type :=
(** [Ibuiltin ef args dest succ] calls the built-in function
identified by [ef], giving it the values of [args] as arguments.
It stores the return value in [dest] and branches to [succ]. *)
- | Icond: condition -> list reg -> node -> node -> instruction
- (** [Icond cond args ifso ifnot] evaluates the boolean condition
+ | Icond: condition -> list reg -> node -> node -> option bool -> instruction
+ (** [Icond cond args ifso ifnot info] evaluates the boolean condition
[cond] over the values of registers [args]. If the condition
is true, it transitions to [ifso]. If the condition is false,
- it transitions to [ifnot]. *)
+ it transitions to [ifnot]. [info] is a ghost field there to provide
+ information relative to branch prediction. *)
| Ijumptable: reg -> list node -> instruction
(** [Ijumptable arg tbl] transitions to the node that is the [n]-th
element of the list [tbl], where [n] is the unsigned integer
@@ -262,8 +263,8 @@ Inductive step: state -> trace -> state -> Prop :=
step (State s f sp pc rs m)
t (State s f sp pc' (regmap_setres res vres rs) m')
| exec_Icond:
- forall s f sp pc rs m cond args ifso ifnot b pc',
- (fn_code f)!pc = Some(Icond cond args ifso ifnot) ->
+ forall s f sp pc rs m cond args ifso ifnot b pc' predb,
+ (fn_code f)!pc = Some(Icond cond args ifso ifnot predb) ->
eval_condition cond rs##args m = Some b ->
pc' = (if b then ifso else ifnot) ->
step (State s f sp pc rs m)
@@ -403,7 +404,7 @@ Definition successors_instr (i: instruction) : list node :=
| Icall sig ros args res s => s :: nil
| Itailcall sig ros args => nil
| Ibuiltin ef args res s => s :: nil
- | Icond cond args ifso ifnot => ifso :: ifnot :: nil
+ | Icond cond args ifso ifnot _ => ifso :: ifnot :: nil
| Ijumptable arg tbl => tbl
| Ireturn optarg => nil
end.
@@ -424,7 +425,7 @@ Definition instr_uses (i: instruction) : list reg :=
| Itailcall sig (inl r) args => r :: args
| Itailcall sig (inr id) args => args
| Ibuiltin ef args res s => params_of_builtin_args args
- | Icond cond args ifso ifnot => args
+ | Icond cond args ifso ifnot _ => args
| Ijumptable arg tbl => arg :: nil
| Ireturn None => nil
| Ireturn (Some arg) => arg :: nil
@@ -442,7 +443,7 @@ Definition instr_defs (i: instruction) : option reg :=
| Itailcall sig ros args => None
| Ibuiltin ef args res s =>
match res with BR r => Some r | _ => None end
- | Icond cond args ifso ifnot => None
+ | Icond cond args ifso ifnot _ => None
| Ijumptable arg tbl => None
| Ireturn optarg => None
end.
@@ -485,7 +486,7 @@ Definition max_reg_instr (m: positive) (pc: node) (i: instruction) :=
| Ibuiltin ef args res s =>
fold_left Pos.max (params_of_builtin_args args)
(fold_left Pos.max (params_of_builtin_res res) m)
- | Icond cond args ifso ifnot => fold_left Pos.max args m
+ | Icond cond args ifso ifnot _ => fold_left Pos.max args m
| Ijumptable arg tbl => Pos.max arg m
| Ireturn None => m
| Ireturn (Some arg) => Pos.max arg m
diff --git a/backend/RTLgen.v b/backend/RTLgen.v
index 2c27944a..ac98f3a1 100644
--- a/backend/RTLgen.v
+++ b/backend/RTLgen.v
@@ -479,7 +479,7 @@ with transl_condexpr (map: mapping) (a: condexpr) (ntrue nfalse: node)
match a with
| CEcond c al =>
do rl <- alloc_regs map al;
- do nt <- add_instr (Icond c rl ntrue nfalse);
+ do nt <- add_instr (Icond c rl ntrue nfalse None);
transl_exprlist map al rl nt
| CEcondition a b c =>
do nc <- transl_condexpr map c ntrue nfalse;
diff --git a/backend/RTLgenspec.v b/backend/RTLgenspec.v
index 92b48e2b..30ad7d82 100644
--- a/backend/RTLgenspec.v
+++ b/backend/RTLgenspec.v
@@ -744,9 +744,9 @@ Inductive tr_expr (c: code):
with tr_condition (c: code):
mapping -> list reg -> condexpr -> node -> node -> node -> Prop :=
- | tr_CEcond: forall map pr cond bl ns ntrue nfalse n1 rl,
+ | tr_CEcond: forall map pr cond bl ns ntrue nfalse n1 rl i,
tr_exprlist c map pr bl ns n1 rl ->
- c!n1 = Some (Icond cond rl ntrue nfalse) ->
+ c!n1 = Some (Icond cond rl ntrue nfalse i) ->
tr_condition c map pr (CEcond cond bl) ns ntrue nfalse
| tr_CEcondition: forall map pr a1 a2 a3 ns ntrue nfalse n2 n3,
tr_condition c map pr a1 ns n2 n3 ->
diff --git a/backend/RTLtyping.v b/backend/RTLtyping.v
index 857f2211..15ed6d8a 100644
--- a/backend/RTLtyping.v
+++ b/backend/RTLtyping.v
@@ -139,11 +139,11 @@ Inductive wt_instr : instruction -> Prop :=
valid_successor s ->
wt_instr (Ibuiltin ef args res s)
| wt_Icond:
- forall cond args s1 s2,
+ forall cond args s1 s2 i,
map env args = type_of_condition cond ->
valid_successor s1 ->
valid_successor s2 ->
- wt_instr (Icond cond args s1 s2)
+ wt_instr (Icond cond args s1 s2 i)
| wt_Ijumptable:
forall arg tbl,
env arg = Tint ->
@@ -313,7 +313,7 @@ Definition type_instr (e: S.typenv) (i: instruction) : res S.typenv :=
| _ => type_builtin_args e args sig.(sig_args)
end;
type_builtin_res e1 res (proj_sig_res sig)
- | Icond cond args s1 s2 =>
+ | Icond cond args s1 s2 _ =>
do x1 <- check_successor s1;
do x2 <- check_successor s2;
S.set_list e args (type_of_condition cond)
diff --git a/backend/Regalloc.ml b/backend/Regalloc.ml
index f2658b04..ffe26933 100644
--- a/backend/Regalloc.ml
+++ b/backend/Regalloc.ml
@@ -295,8 +295,8 @@ let block_of_RTL_instr funsig tyenv = function
(Xbuiltin(ef, args2, res2) ::
movelist (params_of_builtin_res res2) (params_of_builtin_res res1)
[Xbranch s])
- | RTL.Icond(cond, args, s1, s2) ->
- [Xcond(cond, vregs tyenv args, s1, s2)]
+ | RTL.Icond(cond, args, s1, s2, i) ->
+ [Xcond(cond, vregs tyenv args, s1, s2, i)]
| RTL.Ijumptable(arg, tbl) ->
[Xjumptable(vreg tyenv arg, tbl)]
| RTL.Ireturn None ->
@@ -380,7 +380,7 @@ let live_before instr after =
vset_addargs args (vset_removeres res after)
| Xbranch s ->
after
- | Xcond(cond, args, s1, s2) ->
+ | Xcond(cond, args, s1, s2, _) ->
List.fold_right VSet.add args after
| Xjumptable(arg, tbl) ->
VSet.add arg after
@@ -575,7 +575,7 @@ let spill_costs f =
charge_list 10 1 (params_of_builtin_res res)
end
| Xbranch _ -> ()
- | Xcond(cond, args, _, _) ->
+ | Xcond(cond, args, _, _, _) ->
charge_list 10 1 args
| Xjumptable(arg, _) ->
charge 10 1 arg
@@ -718,7 +718,7 @@ let add_interfs_instr g instr live =
end
| Xbranch s ->
()
- | Xcond(cond, args, s1, s2) ->
+ | Xcond(cond, args, s1, s2, _) ->
add_interfs_destroyed g live (destroyed_by_cond cond)
| Xjumptable(arg, tbl) ->
add_interfs_destroyed g live destroyed_by_jumptable
@@ -797,7 +797,7 @@ let tospill_instr alloc instr ts =
(addlist_tospill alloc (params_of_builtin_res res) ts)
| Xbranch s ->
ts
- | Xcond(cond, args, s1, s2) ->
+ | Xcond(cond, args, s1, s2, _) ->
addlist_tospill alloc args ts
| Xjumptable(arg, tbl) ->
add_tospill alloc arg ts
@@ -990,9 +990,9 @@ let spill_instr tospill eqs instr =
(c1 @ Xbuiltin(ef, args', res') :: c2, eqs2)
| Xbranch s ->
([instr], eqs)
- | Xcond(cond, args, s1, s2) ->
+ | Xcond(cond, args, s1, s2, i) ->
let (args', c1, eqs1) = reload_vars tospill eqs args in
- (c1 @ [Xcond(cond, args', s1, s2)], eqs1)
+ (c1 @ [Xcond(cond, args', s1, s2, i)], eqs1)
| Xjumptable(arg, tbl) ->
let (arg', c1, eqs1) = reload_var tospill eqs arg in
(c1 @ [Xjumptable(arg', tbl)], eqs1)
@@ -1128,8 +1128,8 @@ let transl_instr alloc instr k =
AST.map_builtin_res (mreg_of alloc) res) :: k
| Xbranch s ->
LTL.Lbranch s :: []
- | Xcond(cond, args, s1, s2) ->
- LTL.Lcond(cond, mregs_of alloc args, s1, s2) :: []
+ | Xcond(cond, args, s1, s2, i) ->
+ LTL.Lcond(cond, mregs_of alloc args, s1, s2, i) :: []
| Xjumptable(arg, tbl) ->
LTL.Ljumptable(mreg_of alloc arg, tbl) :: []
| Xreturn optarg ->
diff --git a/backend/Renumber.v b/backend/Renumber.v
index 7ba16658..2727b979 100644
--- a/backend/Renumber.v
+++ b/backend/Renumber.v
@@ -48,7 +48,7 @@ Definition renum_instr (i: instruction) : instruction :=
| Icall sg ros args res s => Icall sg ros args res (renum_pc s)
| Itailcall sg ros args => i
| Ibuiltin ef args res s => Ibuiltin ef args res (renum_pc s)
- | Icond cond args s1 s2 => Icond cond args (renum_pc s1) (renum_pc s2)
+ | Icond cond args s1 s2 info => Icond cond args (renum_pc s1) (renum_pc s2) info
| Ijumptable arg tbl => Ijumptable arg (List.map renum_pc tbl)
| Ireturn or => i
end.
diff --git a/backend/Splitting.ml b/backend/Splitting.ml
index 78eb66a5..3ca45c3b 100644
--- a/backend/Splitting.ml
+++ b/backend/Splitting.ml
@@ -162,8 +162,8 @@ let ren_instr f maps pc i =
| Ibuiltin(ef, args, res, s) ->
Ibuiltin(ef, List.map (AST.map_builtin_arg (ren_reg before)) args,
AST.map_builtin_res (ren_reg after) res, s)
- | Icond(cond, args, s1, s2) ->
- Icond(cond, ren_regs before args, s1, s2)
+ | Icond(cond, args, s1, s2, i) ->
+ Icond(cond, ren_regs before args, s1, s2, i)
| Ijumptable(arg, tbl) ->
Ijumptable(ren_reg before arg, tbl)
| Ireturn optarg ->
diff --git a/backend/Tunneling.v b/backend/Tunneling.v
index da1ce45a..a4c4a195 100644
--- a/backend/Tunneling.v
+++ b/backend/Tunneling.v
@@ -78,11 +78,11 @@ Definition record_gotos (f: LTL.function) : U.t :=
Definition tunnel_instr (uf: U.t) (i: instruction) : instruction :=
match i with
| Lbranch s => Lbranch (U.repr uf s)
- | Lcond cond args s1 s2 =>
+ | Lcond cond args s1 s2 info =>
let s1' := U.repr uf s1 in let s2' := U.repr uf s2 in
if peq s1' s2'
then Lbranch s1'
- else Lcond cond args s1' s2'
+ else Lcond cond args s1' s2' info
| Ljumptable arg tbl => Ljumptable arg (List.map (U.repr uf) tbl)
| _ => i
end.
diff --git a/backend/Unusedglob.v b/backend/Unusedglob.v
index 1b5f2547..93ca7af4 100644
--- a/backend/Unusedglob.v
+++ b/backend/Unusedglob.v
@@ -53,7 +53,7 @@ Definition ref_instruction (i: instruction) : list ident :=
| Itailcall _ (inl r) _ => nil
| Itailcall _ (inr id) _ => id :: nil
| Ibuiltin _ args _ _ => globals_of_builtin_args args
- | Icond cond _ _ _ => nil
+ | Icond cond _ _ _ _ => nil
| Ijumptable _ _ => nil
| Ireturn _ => nil
end.
diff --git a/backend/ValueAnalysis.v b/backend/ValueAnalysis.v
index e25d2e5f..2e79d1a9 100644
--- a/backend/ValueAnalysis.v
+++ b/backend/ValueAnalysis.v
@@ -156,7 +156,7 @@ Definition transfer (f: function) (rm: romem) (pc: node) (ae: aenv) (am: amem) :
VA.Bot
| Some(Ibuiltin ef args res s) =>
transfer_builtin ae am rm ef args res
- | Some(Icond cond args s1 s2) =>
+ | Some(Icond cond args s1 s2 _) =>
VA.State ae am
| Some(Ijumptable arg tbl) =>
VA.State ae am
diff --git a/backend/XTL.ml b/backend/XTL.ml
index c496fafb..1d8e89c0 100644
--- a/backend/XTL.ml
+++ b/backend/XTL.ml
@@ -36,7 +36,7 @@ type instruction =
| Xtailcall of signature * (var, ident) sum * var list
| Xbuiltin of external_function * var builtin_arg list * var builtin_res
| Xbranch of node
- | Xcond of condition * var list * node * node
+ | Xcond of condition * var list * node * node * bool option
| Xjumptable of var * node list
| Xreturn of var list
@@ -105,7 +105,7 @@ let twin_reg r =
let rec successors_block = function
| Xbranch s :: _ -> [s]
| Xtailcall(sg, vos, args) :: _ -> []
- | Xcond(cond, args, s1, s2) :: _ -> [s1; s2]
+ | Xcond(cond, args, s1, s2, _) :: _ -> [s1; s2]
| Xjumptable(arg, tbl) :: _ -> tbl
| Xreturn _:: _ -> []
| instr :: blk -> successors_block blk
@@ -179,7 +179,7 @@ let type_instr = function
type_builtin_res res (proj_sig_res sg)
| Xbranch s ->
()
- | Xcond(cond, args, s1, s2) ->
+ | Xcond(cond, args, s1, s2, _) ->
set_vars_type args (type_of_condition cond)
| Xjumptable(arg, tbl) ->
set_var_type arg Tint
diff --git a/backend/XTL.mli b/backend/XTL.mli
index b4b77fab..7b7f7186 100644
--- a/backend/XTL.mli
+++ b/backend/XTL.mli
@@ -37,7 +37,7 @@ type instruction =
| Xtailcall of signature * (var, ident) sum * var list
| Xbuiltin of external_function * var builtin_arg list * var builtin_res
| Xbranch of node
- | Xcond of condition * var list * node * node
+ | Xcond of condition * var list * node * node * bool option
| Xjumptable of var * node list
| Xreturn of var list
diff --git a/configure b/configure
index f790281c..cb2f52ba 100755
--- a/configure
+++ b/configure
@@ -568,7 +568,7 @@ missingtools=false
echo "Testing Coq... " | tr -d '\n'
coq_ver=$(${COQBIN}coqc -v 2>/dev/null | sed -n -e 's/The Coq Proof Assistant, version \([^ ]*\).*$/\1/p')
case "$coq_ver" in
- 8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0)
+ 8.9.0|8.9.1|8.10.0|8.10.1|8.10.2|8.11.0|8.11.1)
echo "version $coq_ver -- good!";;
?*)
echo "version $coq_ver -- UNSUPPORTED"