aboutsummaryrefslogtreecommitdiffstats
path: root/src/hls/IfConversion.v
diff options
context:
space:
mode:
Diffstat (limited to 'src/hls/IfConversion.v')
-rw-r--r--src/hls/IfConversion.v17
1 files changed, 8 insertions, 9 deletions
diff --git a/src/hls/IfConversion.v b/src/hls/IfConversion.v
index e893578..b397d43 100644
--- a/src/hls/IfConversion.v
+++ b/src/hls/IfConversion.v
@@ -25,6 +25,7 @@ Require Import compcert.lib.Maps.
Require Import vericert.common.Vericertlib.
Require Import vericert.hls.RTLBlockInstr.
Require Import vericert.hls.RTLBlock.
+Require Import vericert.hls.Predicate.
(*|
=============
@@ -57,10 +58,10 @@ Definition if_convert_block (c: code) (p: predicate) (bb: bblock) : bblock :=
| RBcond cond args n1 n2 =>
match PTree.get n1 c, PTree.get n2 c with
| Some bb1, Some bb2 =>
- let bb1' := List.map (map_if_convert (Pvar p)) bb1.(bb_body) in
- let bb2' := List.map (map_if_convert (Pnot (Pvar p))) bb2.(bb_body) in
- mk_bblock (List.concat (bb.(bb_body) :: ((RBsetpred cond args p) :: bb1') :: bb2' :: nil))
- (RBpred_cf (Pvar p) bb1.(bb_exit) bb2.(bb_exit))
+ let bb1' := List.map (map_if_convert (Plit (true, p))) bb1.(bb_body) in
+ let bb2' := List.map (map_if_convert (Plit (false, p))) bb2.(bb_body) in
+ mk_bblock (List.concat (bb.(bb_body) :: ((RBsetpred None cond args p) :: bb1') :: bb2' :: nil))
+ (RBpred_cf (Plit (true, p)) bb1.(bb_exit) bb2.(bb_exit))
| _, _ => bb
end
| _ => bb
@@ -104,16 +105,14 @@ Definition find_blocks_with_cond (c: code) : list (node * bblock) :=
) (PTree.elements c).
Definition if_convert_code (p: nat * code) (nb: node * bblock) :=
- let (n, bb) := nb in
- let (p', c) := p in
- let nbb := if_convert_block c p' bb in
- (S p', PTree.set n nbb c).
+ let nbb := if_convert_block (snd p) (Pos.of_nat (fst p)) (snd nb) in
+ (S (fst p), PTree.set (fst nb) nbb (snd p)).
Definition transf_function (f: function) : function :=
let (_, c) := List.fold_left if_convert_code
(find_blocks_with_cond f.(fn_code))
(1%nat, f.(fn_code)) in
- mkfunction f.(fn_sig) f.(fn_params) f.(fn_stacksize) c f.(fn_funct_units) f.(fn_entrypoint).
+ mkfunction f.(fn_sig) f.(fn_params) f.(fn_stacksize) c f.(fn_entrypoint).
Definition transf_fundef (fd: fundef) : fundef :=
transf_fundef transf_function fd.