aboutsummaryrefslogtreecommitdiffstats
path: root/src/hls/HTLPargen.v
diff options
context:
space:
mode:
authorYann Herklotz <git@yannherklotz.com>2021-10-21 14:03:32 +0100
committerYann Herklotz <git@yannherklotz.com>2021-10-21 14:03:32 +0100
commit7bbedef94189dc9ab094619ee00bc9aaf0fd110a (patch)
tree203053b236c507fa3f0c5d6f8445af625d1bbb14 /src/hls/HTLPargen.v
parent1d86b1c178deb97f3d499f461a417a4fe6846cf8 (diff)
downloadvericert-7bbedef94189dc9ab094619ee00bc9aaf0fd110a.tar.gz
vericert-7bbedef94189dc9ab094619ee00bc9aaf0fd110a.zip
Add work towards decidability of SAT solver
Diffstat (limited to 'src/hls/HTLPargen.v')
-rw-r--r--src/hls/HTLPargen.v11
1 files changed, 6 insertions, 5 deletions
diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v
index 64996c6..e8b18dc 100644
--- a/src/hls/HTLPargen.v
+++ b/src/hls/HTLPargen.v
@@ -28,6 +28,7 @@ Require Import vericert.common.Statemonad.
Require Import vericert.common.Vericertlib.
Require Import vericert.hls.AssocMap.
Require Import vericert.hls.HTL.
+Require Import vericert.hls.Predicate.
Require Import vericert.hls.RTLBlockInstr.
Require Import vericert.hls.RTLPar.
Require Import vericert.hls.ValueInt.
@@ -657,12 +658,12 @@ Definition add_control_instr_force (n : node) (st : stmnt) : mon unit :=
Fixpoint pred_expr (preg: reg) (p: pred_op) :=
match p with
- | Pvar pred =>
- Vrange preg (Vlit (posToValue pred)) (Vlit (posToValue pred))
+ | Pvar (b, pred) =>
+ if b
+ then Vrange preg (Vlit (posToValue pred)) (Vlit (posToValue pred))
+ else Vunop Vnot (Vrange preg (Vlit (posToValue pred)) (Vlit (posToValue pred)))
| Ptrue => Vlit (ZToValue 1)
| Pfalse => Vlit (ZToValue 0)
- | Pnot pred =>
- Vunop Vnot (pred_expr preg pred)
| Pand p1 p2 =>
Vbinop Vand (pred_expr preg p1) (pred_expr preg p2)
| Por p1 p2 =>
@@ -695,7 +696,7 @@ Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr)
translate_predicate a preg p dst (Vvar src)
| RBsetpred c args p =>
do cond <- translate_condition c args;
- ret (a (pred_expr preg (Pvar p)) cond)
+ ret (a (pred_expr preg (Pvar (true, p))) cond)
end.
Lemma create_new_state_state_incr: