aboutsummaryrefslogtreecommitdiffstats
path: root/src/hls
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2021-12-09 20:06:34 +0000
committerYann Herklotz <git@yannherklotz.com>2021-12-09 20:06:34 +0000
commitb91f6db17ee30efd2068efbeecbf1d2b4c3850ea (patch)
tree01dbe430a4d46d31e38cd6cf0774822b1f3416c4 /src/hls
parenta64ccb64e16175b520a0dc4badde44a1cc46f17d (diff)
downloadvericert-b91f6db17ee30efd2068efbeecbf1d2b4c3850ea.tar.gz
vericert-b91f6db17ee30efd2068efbeecbf1d2b4c3850ea.zip
Add bourdoncle to build
Diffstat (limited to 'src/hls')
-rw-r--r--src/hls/IfConversion.v36
-rw-r--r--src/hls/Pipeline.v3
-rw-r--r--src/hls/PipelineOp.v3
-rw-r--r--src/hls/Schedule.ml5
4 files changed, 40 insertions, 7 deletions
diff --git a/src/hls/IfConversion.v b/src/hls/IfConversion.v
index b397d43..4585770 100644
--- a/src/hls/IfConversion.v
+++ b/src/hls/IfConversion.v
@@ -26,6 +26,9 @@ Require Import vericert.common.Vericertlib.
Require Import vericert.hls.RTLBlockInstr.
Require Import vericert.hls.RTLBlock.
Require Import vericert.hls.Predicate.
+Require Import vericert.bourdoncle.Bourdoncle.
+
+Parameter build_bourdoncle : function -> (bourdoncle * PMap.t N).
(*|
=============
@@ -96,11 +99,34 @@ Definition find_all_backedges (c: code) : list node :=
Definition has_backedge (entry: node) (be: list node) :=
any (fun x => Pos.eqb entry x) be.
-Definition find_blocks_with_cond (c: code) : list (node * bblock) :=
+Fixpoint get_loops (b: bourdoncle): list node :=
+ match b with
+ | L h b' => h::(fold_right (fun a b => get_loops a ++ b) nil b')
+ | I h => nil
+ end.
+
+Definition is_loop (b: list node) (n: node) :=
+ any (fun x => Pos.eqb x n) b.
+
+Definition is_flat_cfi (n: cf_instr) :=
+ match n with
+ | RBcond _ _ _ _ => false
+ | RBjumptable _ _ => false
+ | RBpred_cf _ _ _ => false
+ | _ => true
+ end.
+
+Definition is_flat (c: code) (succ: node) :=
+ match c ! succ with
+ | Some bblock => is_flat_cfi bblock.(bb_exit)
+ | None => false
+ end.
+
+Definition find_blocks_with_cond ep (b: list node) (c: code) : list (node * bblock) :=
let backedges := find_all_backedges c in
List.filter (fun x => is_cond_cfi (snd x).(bb_exit) &&
- negb (has_backedge (fst x) backedges) &&
- all (fun x' => negb (has_backedge x' backedges))
+ (negb (is_loop b (fst x)) || Pos.eqb (fst x) ep) &&
+ all (fun x' => is_flat c x')
(successors_instr (snd x).(bb_exit))
) (PTree.elements c).
@@ -109,8 +135,10 @@ Definition if_convert_code (p: nat * code) (nb: node * bblock) :=
(S (fst p), PTree.set (fst nb) nbb (snd p)).
Definition transf_function (f: function) : function :=
+ let (b, _) := build_bourdoncle f in
+ let b' := get_loops b in
let (_, c) := List.fold_left if_convert_code
- (find_blocks_with_cond f.(fn_code))
+ (find_blocks_with_cond f.(fn_entrypoint) b' f.(fn_code))
(1%nat, f.(fn_code)) in
mkfunction f.(fn_sig) f.(fn_params) f.(fn_stacksize) c f.(fn_entrypoint).
diff --git a/src/hls/Pipeline.v b/src/hls/Pipeline.v
index 7f1485a..67ab1f5 100644
--- a/src/hls/Pipeline.v
+++ b/src/hls/Pipeline.v
@@ -16,7 +16,7 @@
* along with this program. If not, see <https://www.gnu.org/licenses/>.
*)
-Require Import compcert.lib.Maps.
+(*Require Import compcert.lib.Maps.
Require Import compcert.common.AST.
Require Import compcert.backend.RTL.
@@ -26,3 +26,4 @@ Definition transf_fundef := transf_fundef pipeline.
Definition transf_program : program -> program :=
transform_program transf_fundef.
+*)
diff --git a/src/hls/PipelineOp.v b/src/hls/PipelineOp.v
index 793b752..bb01ff9 100644
--- a/src/hls/PipelineOp.v
+++ b/src/hls/PipelineOp.v
@@ -79,7 +79,7 @@ Definition map_at_err {A: Type} (n: nat) (f: A -> A) (l: list A): option (list A
then None
else Some (map_at n f l).
-Definition replace_div' sdiv udiv (d: instr) :=
+(*Definition replace_div' sdiv udiv (d: instr) :=
match d with
| RBop op Odiv args dst => RBpiped op sdiv args
| RBop op Odivu args dst => RBpiped op udiv args
@@ -190,3 +190,4 @@ Definition transf_fundef (fd: fundef) : fundef :=
Definition transf_program (p: program) : program :=
transform_program transf_fundef p.
+*)
diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml
index e557d4b..1969cad 100644
--- a/src/hls/Schedule.ml
+++ b/src/hls/Schedule.ml
@@ -637,7 +637,10 @@ let add_cycle_constr maxf n dfg constr =
- (int_of_float (Float.ceil (Float.div (float_of_int w) (float_of_int maxf))) - 1),
encode_var n (fst v') 0)
) constr (DFG.fold_vertex (fun v l ->
- List.append l (longest_path v |> constrained_paths |> List.map (function (v', w) -> (v, v', - w)))
+ List.append l (longest_path v (*|> (function l -> List.iter (function (a, x) ->
+ printf "c: %d %d\n" (fst a) x) l; l)*) |> constrained_paths (* |> (function l -> List.iter (function (a, x) ->
+ printf "%d %d\n" (fst a) x) l; l)*)
+ |> List.map (function (v', w) -> (v, v', - w)))
) dfg [])
type resource =