From 05347ca5126f335b0479b71a4576b141e082fab5 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 2 Mar 2021 10:27:33 +0000 Subject: Add RAM to HTL --- src/hls/HTLgen.v | 1 + 1 file changed, 1 insertion(+) (limited to 'src/hls/HTLgen.v') diff --git a/src/hls/HTLgen.v b/src/hls/HTLgen.v index f1e6b2a..c071868 100644 --- a/src/hls/HTLgen.v +++ b/src/hls/HTLgen.v @@ -612,6 +612,7 @@ Definition transf_module (f: function) : mon HTL.module := clk current_state.(st_scldecls) current_state.(st_arrdecls) + None (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA))) | _, _ => error (Errors.msg "More than 2^32 states.") end -- cgit From 23a2a2fb2916a7ecb240aa51686bbe049f8418e4 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 6 Apr 2021 23:08:03 +0100 Subject: Finish load and store proof, but broke top-level --- src/hls/HTLgen.v | 48 ++++++++++++++++++++++++------------------------ 1 file changed, 24 insertions(+), 24 deletions(-) (limited to 'src/hls/HTLgen.v') diff --git a/src/hls/HTLgen.v b/src/hls/HTLgen.v index c071868..b55d7a3 100644 --- a/src/hls/HTLgen.v +++ b/src/hls/HTLgen.v @@ -583,7 +583,19 @@ Proof. simplify. transitivity (Z.pos (max_pc_map m)); eauto. Qed. -Definition transf_module (f: function) : mon HTL.module := +Definition decide_order a b c d e f g : {module_ordering a b c d e f g} + {True}. + refine (match bool_dec ((a left _ + | _ => _ + end); auto. + simplify; repeat match goal with + | H: context[(_ apply Pos.ltb_lt in H + end; unfold module_ordering; auto. +Defined. + +Definition transf_module (f: function) : mon HTL.module. + refine ( if stack_correct f.(fn_stacksize) then do fin <- create_reg (Some Voutput) 1; do rtrn <- create_reg (Some Voutput) 32; @@ -595,8 +607,10 @@ Definition transf_module (f: function) : mon HTL.module := do clk <- create_reg (Some Vinput) 1; do current_state <- get; match zle (Z.pos (max_pc_map current_state.(st_datapath))) Integers.Int.max_unsigned, - zle (Z.pos (max_pc_map current_state.(st_controllogic))) Integers.Int.max_unsigned with - | left LEDATA, left LECTRL => + zle (Z.pos (max_pc_map current_state.(st_controllogic))) Integers.Int.max_unsigned, + decide_order (st_st current_state) fin rtrn stack start rst clk + with + | left LEDATA, left LECTRL, left MORD => ret (HTL.mkmodule f.(RTL.fn_params) current_state.(st_datapath) @@ -613,16 +627,19 @@ Definition transf_module (f: function) : mon HTL.module := current_state.(st_scldecls) current_state.(st_arrdecls) None - (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA))) - | _, _ => error (Errors.msg "More than 2^32 states.") + (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA)) + MORD + _) + | _, _, _ => error (Errors.msg "More than 2^32 states.") end - else error (Errors.msg "Stack size misalignment."). + else error (Errors.msg "Stack size misalignment.")); discriminate. +Defined. Definition max_state (f: function) : state := let st := Pos.succ (max_reg_function f) in mkstate st (Pos.succ st) - (Pos.succ (max_pc_function f)) + (Pos.succ (RTL.max_pc_function f)) (AssocMap.set st (None, VScalar 32) (st_scldecls (init_state st))) (st_arrdecls (init_state st)) (st_datapath (init_state st)) @@ -633,23 +650,6 @@ Definition transl_module (f : function) : Errors.res HTL.module := Definition transl_fundef := transf_partial_fundef transl_module. -(* Definition transl_program (p : RTL.program) := transform_partial_program transl_fundef p. *) - -(*Definition transl_main_fundef f : Errors.res HTL.fundef := - match f with - | Internal f => transl_fundef (Internal f) - | External f => Errors.Error (Errors.msg "Could not find internal main function") - end. - -(** Translation of a whole program. *) - -Definition transl_program (p: RTL.program) : Errors.res HTL.program := - transform_partial_program2 (fun i f => if Pos.eqb p.(AST.prog_main) i - then transl_fundef f - else transl_main_fundef f) - (fun i v => Errors.OK v) p. -*) - Definition main_is_internal (p : RTL.program) : bool := let ge := Globalenvs.Genv.globalenv p in match Globalenvs.Genv.find_symbol ge p.(AST.prog_main) with -- cgit From 8573889ca84a84475761b4d75d55547a2995c831 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 7 Apr 2021 01:11:04 +0100 Subject: Basically done with proof --- src/hls/HTLgen.v | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) (limited to 'src/hls/HTLgen.v') diff --git a/src/hls/HTLgen.v b/src/hls/HTLgen.v index b55d7a3..4d60a24 100644 --- a/src/hls/HTLgen.v +++ b/src/hls/HTLgen.v @@ -608,9 +608,10 @@ Definition transf_module (f: function) : mon HTL.module. do current_state <- get; match zle (Z.pos (max_pc_map current_state.(st_datapath))) Integers.Int.max_unsigned, zle (Z.pos (max_pc_map current_state.(st_controllogic))) Integers.Int.max_unsigned, - decide_order (st_st current_state) fin rtrn stack start rst clk + decide_order (st_st current_state) fin rtrn stack start rst clk, + max_list_dec (RTL.fn_params f) (st_st current_state) with - | left LEDATA, left LECTRL, left MORD => + | left LEDATA, left LECTRL, left MORD, left WFPARAMS => ret (HTL.mkmodule f.(RTL.fn_params) current_state.(st_datapath) @@ -629,8 +630,9 @@ Definition transf_module (f: function) : mon HTL.module. None (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA)) MORD - _) - | _, _, _ => error (Errors.msg "More than 2^32 states.") + _ + WFPARAMS) + | _, _, _, _ => error (Errors.msg "More than 2^32 states.") end else error (Errors.msg "Stack size misalignment.")); discriminate. Defined. -- cgit