aboutsummaryrefslogtreecommitdiffstats
path: root/src/translation/HTLgen.v
diff options
context:
space:
mode:
Diffstat (limited to 'src/translation/HTLgen.v')
-rw-r--r--src/translation/HTLgen.v39
1 files changed, 22 insertions, 17 deletions
diff --git a/src/translation/HTLgen.v b/src/translation/HTLgen.v
index 1c2d786..43a6674 100644
--- a/src/translation/HTLgen.v
+++ b/src/translation/HTLgen.v
@@ -219,7 +219,7 @@ Definition translate_comparison (c : Integers.comparison) (args : list reg) : mo
| Integers.Cgt, r1::r2::nil => ret (bop Vgt r1 r2)
| Integers.Cle, r1::r2::nil => ret (bop Vle r1 r2)
| Integers.Cge, r1::r2::nil => ret (bop Vge r1 r2)
- | _, _ => error (Errors.msg "Veriloggen: comparison instruction not implemented: other")
+ | _, _ => error (Errors.msg "Htlgen: comparison instruction not implemented: other")
end.
Definition translate_comparison_imm (c : Integers.comparison) (args : list reg) (i: Integers.int)
@@ -231,7 +231,7 @@ Definition translate_comparison_imm (c : Integers.comparison) (args : list reg)
| Integers.Cgt, r1::nil => ret (boplit Vgt r1 i)
| Integers.Cle, r1::nil => ret (boplit Vle r1 i)
| Integers.Cge, r1::nil => ret (boplit Vge r1 i)
- | _, _ => error (Errors.msg "Veriloggen: comparison_imm instruction not implemented: other")
+ | _, _ => error (Errors.msg "Htlgen: comparison_imm instruction not implemented: other")
end.
Definition translate_condition (c : Op.condition) (args : list reg) : mon expr :=
@@ -240,9 +240,9 @@ Definition translate_condition (c : Op.condition) (args : list reg) : mon expr :
| Op.Ccompu c, _ => translate_comparison c args
| Op.Ccompimm c i, _ => translate_comparison_imm c args i
| Op.Ccompuimm c i, _ => translate_comparison_imm c args i
- | Op.Cmaskzero n, _ => error (Errors.msg "Veriloggen: condition instruction not implemented: Cmaskzero")
- | Op.Cmasknotzero n, _ => error (Errors.msg "Veriloggen: condition instruction not implemented: Cmasknotzero")
- | _, _ => error (Errors.msg "Veriloggen: condition instruction not implemented: other")
+ | Op.Cmaskzero n, _ => error (Errors.msg "Htlgen: condition instruction not implemented: Cmaskzero")
+ | Op.Cmasknotzero n, _ => error (Errors.msg "Htlgen: condition instruction not implemented: Cmasknotzero")
+ | _, _ => error (Errors.msg "Htlgen: condition instruction not implemented: other")
end.
Definition check_address_parameter (p : Z) : bool :=
@@ -269,7 +269,7 @@ Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon ex
| _, _ => error (Errors.msg "Veriloggen: translate_eff_addressing unsuported addressing")
end.
-(** Translate an instruction to a statement. *)
+(** Translate an instruction to a statement. FIX mulhs mulhu *)
Definition translate_instr (op : Op.operation) (args : list reg) : mon expr :=
match op, args with
| Op.Omove, r::nil => ret (Vvar r)
@@ -278,8 +278,8 @@ Definition translate_instr (op : Op.operation) (args : list reg) : mon expr :=
| Op.Osub, r1::r2::nil => ret (bop Vsub r1 r2)
| Op.Omul, r1::r2::nil => ret (bop Vmul r1 r2)
| Op.Omulimm n, r::nil => ret (boplit Vmul r n)
- | Op.Omulhs, _ => error (Errors.msg "Veriloggen: Instruction not implemented: Omulhs")
- | Op.Omulhu, _ => error (Errors.msg "Veriloggen: Instruction not implemented: Omulhu")
+ | Op.Omulhs, r1::r2::nil => ret (bop Vmul r1 r2)
+ | Op.Omulhu, r1::r2::nil => ret (bop Vmul r1 r2)
| Op.Odiv, r1::r2::nil => ret (bop Vdiv r1 r2)
| Op.Odivu, r1::r2::nil => ret (bop Vdivu r1 r2)
| Op.Omod, r1::r2::nil => ret (bop Vmod r1 r2)
@@ -295,16 +295,21 @@ Definition translate_instr (op : Op.operation) (args : list reg) : mon expr :=
| Op.Oshlimm n, r::nil => ret (boplit Vshl r n)
| Op.Oshr, r1::r2::nil => ret (bop Vshr r1 r2)
| Op.Oshrimm n, r::nil => ret (boplit Vshr r n)
- | Op.Oshrximm n, r::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Oshrximm")
- | Op.Oshru, r1::r2::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Oshru")
- | Op.Oshruimm n, r::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Oshruimm")
- | Op.Ororimm n, r::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Ororimm")
- | Op.Oshldimm n, r::nil => error (Errors.msg "Veriloggen: Instruction not implemented: Oshldimm")
+ | Op.Oshrximm n, r::nil => ret (Vbinop Vdiv (Vvar r)
+ (Vbinop Vshl (Vlit (ZToValue 32 1))
+ (Vlit (intToValue n))))
+ | Op.Oshru, r1::r2::nil => ret (bop Vshru r1 r2)
+ | Op.Oshruimm n, r::nil => ret (boplit Vshru r n)
+ | Op.Ororimm n, r::nil => ret (Vbinop Vor (boplit Vshr r n) (boplit Vshl r (Integers.Int.sub (Integers.Int.repr 32) n)))
+ | Op.Oshldimm n, r::nil => ret (Vbinop Vor (boplit Vshl r n) (boplit Vshr r (Integers.Int.sub (Integers.Int.repr 32) n)))
| Op.Ocmp c, _ => translate_condition c args
+ | Op.Osel c AST.Tint, r1::r2::rl =>
+ do tc <- translate_condition c rl;
+ ret (Vternary tc (Vvar r1) (Vvar r2))
| Op.Olea a, _ => translate_eff_addressing a args
| Op.Oleal a, _ => translate_eff_addressing a args (* FIXME: Need to be careful here; large arrays might fail? *)
| Op.Ocast32signed, r::nil => ret (Vvar r) (* FIXME: Don't need to sign extend for now since everything is 32 bit? *)
- | _, _ => error (Errors.msg "Veriloggen: Instruction not implemented: other")
+ | _, _ => error (Errors.msg "Htlgen: Instruction not implemented: other")
end.
Lemma add_branch_instr_state_incr:
@@ -338,7 +343,7 @@ Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit :=
(AssocMap.set n Vskip (st_datapath s))
(AssocMap.set n (state_cond s.(st_st) e n1 n2) (st_controllogic s)))
(add_branch_instr_state_incr s e n n1 n2 NSTM NTRANS)
- | _, _ => Error (Errors.msg "Veriloggen: add_branch_instr")
+ | _, _ => Error (Errors.msg "Htlgen: add_branch_instr")
end.
Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing)
@@ -349,13 +354,13 @@ Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing)
| Mint32, Op.Ascaled scale offset, r1::nil =>
if (check_address_parameter scale) && (check_address_parameter offset)
then ret (Vvari stack (Vbinop Vadd (boplitz Vmul r1 (scale / 4)) (Vlit (ZToValue 32 (offset / 4)))))
- else error (Errors.msg "Veriloggen: translate_arr_access address misaligned")
+ else error (Errors.msg "Htlgen: translate_arr_access address misaligned")
| Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *)
if (check_address_parameter scale) && (check_address_parameter offset)
then ret (Vvari stack
(Vbinop Vadd (Vbinop Vadd (boplitz Vdiv r1 4) (Vlit (ZToValue 32 (offset / 4))))
(boplitz Vmul r2 (scale / 4))))
- else error (Errors.msg "Veriloggen: translate_arr_access address misaligned")
+ else error (Errors.msg "Htlgen: translate_arr_access address misaligned")
| Mint32, Op.Ainstack a, nil => (* We need to be sure that the base address is aligned *)
let a := Integers.Ptrofs.unsigned a in
if (check_address_parameter a)