From 171b326ade18ab77eb155a9d203f2f523708b29b Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 15 Feb 2021 00:06:29 +0000 Subject: Add beginning to scheduling division --- src/hls/HTL.v | 12 +- src/hls/HTLPargen.v | 155 +++++++++++++------------ src/hls/HTLgen.v | 1 + src/hls/HTLgenspec.v | 2 +- src/hls/PrintVerilog.ml | 20 +++- src/hls/Verilog.v | 300 ++++++++++++------------------------------------ src/hls/Veriloggen.v | 1 + 7 files changed, 181 insertions(+), 310 deletions(-) (limited to 'src/hls') diff --git a/src/hls/HTL.v b/src/hls/HTL.v index c8a0041..d91a340 100644 --- a/src/hls/HTL.v +++ b/src/hls/HTL.v @@ -27,16 +27,19 @@ Require compcert.lib.Integers. Require Import compcert.lib.Maps. Require Import vericert.common.Vericertlib. -Require Import vericert.hls.ValueInt. -Require Import vericert.hls.AssocMap. Require Import vericert.hls.Array. +Require Import vericert.hls.AssocMap. +Require Import vericert.hls.FunctionalUnits. +Require Import vericert.hls.ValueInt. Require vericert.hls.Verilog. -(** The purpose of the hardware transfer language (HTL) is to create a more +(*| +The purpose of the hardware transfer language (HTL) is to create a more hardware-like layout that is still similar to the register transfer language (RTL) that it came from. The main change is that function calls become module instantiations and that we now describe a state machine instead of a -control-flow graph. *) +control-flow graph. +|*) Local Open Scope assocmap. @@ -65,6 +68,7 @@ Record module: Type := mod_start : reg; mod_reset : reg; mod_clk : reg; + mod_funct_units: funct_units; mod_scldecls : AssocMap.t (option Verilog.io * Verilog.scl_decl); mod_arrdecls : AssocMap.t (option Verilog.io * Verilog.arr_decl); mod_wf : (map_well_formed mod_controllogic /\ map_well_formed mod_datapath); diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index fcd4441..88fb9b4 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -27,6 +27,7 @@ Require Import compcert.lib.Maps. Require Import vericert.common.Statemonad. Require Import vericert.common.Vericertlib. Require Import vericert.hls.AssocMap. +Require Import vericert.hls.FunctionalUnits. Require Import vericert.hls.HTL. Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLPar. @@ -47,6 +48,7 @@ Record state: Type := mkstate { st_arrdecls: AssocMap.t (option io * arr_decl); st_datapath: datapath; st_controllogic: controllogic; + st_funct_units: funct_units; }. Definition init_state (st : reg) : state := @@ -56,7 +58,8 @@ Definition init_state (st : reg) : state := (AssocMap.empty (option io * scl_decl)) (AssocMap.empty (option io * arr_decl)) (AssocMap.empty stmnt) - (AssocMap.empty stmnt). + (AssocMap.empty stmnt) + initial_funct_units. Module HTLState <: State. @@ -126,7 +129,8 @@ Lemma declare_reg_state_incr : (AssocMap.set r (i, VScalar sz) s.(st_scldecls)) s.(st_arrdecls) s.(st_datapath) - s.(st_controllogic)). + s.(st_controllogic) + s.(st_funct_units)). Proof. auto with htlh. Qed. Definition declare_reg (i : option io) (r : reg) (sz : nat) : mon unit := @@ -137,7 +141,8 @@ Definition declare_reg (i : option io) (r : reg) (sz : nat) : mon unit := (AssocMap.set r (i, VScalar sz) s.(st_scldecls)) s.(st_arrdecls) s.(st_datapath) - s.(st_controllogic)) + s.(st_controllogic) + s.(st_funct_units)) (declare_reg_state_incr i s r sz). Lemma add_instr_state_incr : @@ -151,7 +156,8 @@ Lemma add_instr_state_incr : s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n st s.(st_datapath)) - (AssocMap.set n (state_goto s.(st_st) n') s.(st_controllogic))). + (AssocMap.set n (state_goto s.(st_st) n') s.(st_controllogic)) + s.(st_funct_units)). Proof. constructor; intros; try (simpl; destruct (peq n n0); subst); @@ -169,7 +175,8 @@ Definition add_instr (n : node) (n' : node) (st : stmnt) : mon unit := s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n st s.(st_datapath)) - (AssocMap.set n (state_goto s.(st_st) n') s.(st_controllogic))) + (AssocMap.set n (state_goto s.(st_st) n') s.(st_controllogic)) + s.(st_funct_units)) (add_instr_state_incr s n n' st TRANS) | _ => Error (Errors.msg "HTL.add_instr") end. @@ -185,7 +192,8 @@ Lemma add_instr_skip_state_incr : s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n st s.(st_datapath)) - (AssocMap.set n Vskip s.(st_controllogic))). + (AssocMap.set n Vskip s.(st_controllogic)) + s.(st_funct_units)). Proof. constructor; intros; try (simpl; destruct (peq n n0); subst); @@ -203,7 +211,8 @@ Definition add_instr_skip (n : node) (st : stmnt) : mon unit := s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n st s.(st_datapath)) - (AssocMap.set n Vskip s.(st_controllogic))) + (AssocMap.set n Vskip s.(st_controllogic)) + s.(st_funct_units)) (add_instr_skip_state_incr s n st TRANS) | _ => Error (Errors.msg "HTL.add_instr_skip") end. @@ -219,7 +228,8 @@ Lemma add_node_skip_state_incr : s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n Vskip s.(st_datapath)) - (AssocMap.set n st s.(st_controllogic))). + (AssocMap.set n st s.(st_controllogic)) + s.(st_funct_units)). Proof. constructor; intros; try (simpl; destruct (peq n n0); subst); @@ -237,7 +247,8 @@ Definition add_node_skip (n : node) (st : stmnt) : mon unit := s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n Vskip s.(st_datapath)) - (AssocMap.set n st s.(st_controllogic))) + (AssocMap.set n st s.(st_controllogic)) + s.(st_funct_units)) (add_node_skip_state_incr s n st TRANS) | _ => Error (Errors.msg "HTL.add_node_skip") end. @@ -325,6 +336,32 @@ Definition check_address_parameter_signed (p : Z) : bool := Definition check_address_parameter_unsigned (p : Z) : bool := Z.leb p Integers.Ptrofs.max_unsigned. +Lemma create_reg_state_incr: + forall s sz i, + st_incr s (mkstate + s.(st_st) + (Pos.succ (st_freshreg s)) + (st_freshstate s) + (AssocMap.set s.(st_freshreg) (i, VScalar sz) s.(st_scldecls)) + s.(st_arrdecls) + (st_datapath s) + (st_controllogic s) + s.(st_funct_units)). +Proof. constructor; simpl; auto with htlh. Qed. + +Definition create_reg (i : option io) (sz : nat) : mon reg := + fun s => let r := s.(st_freshreg) in + OK r (mkstate + s.(st_st) + (Pos.succ r) + (st_freshstate s) + (AssocMap.set s.(st_freshreg) (i, VScalar sz) s.(st_scldecls)) + (st_arrdecls s) + (st_datapath s) + (st_controllogic s) + s.(st_funct_units)) + (create_reg_state_incr s sz i). + Definition translate_eff_addressing (a: Op.addressing) (args: list reg) : mon expr := match a, args with (* TODO: We should be more methodical here; what are the possibilities?*) @@ -353,7 +390,7 @@ Definition translate_eff_addressing (a: Op.addressing) (args: list reg) end. (** Translate an instruction to a statement. FIX mulhs mulhu *) -Definition translate_instr (op : Op.operation) (args : list reg) : mon expr := +Definition translate_instr (op : Op.operation) (args : list reg) : mon stmnt := match op, args with | Op.Omove, r::nil => ret (Vvar r) | Op.Ointconst n, _ => ret (Vlit (intToValue n)) @@ -363,7 +400,19 @@ Definition translate_instr (op : Op.operation) (args : list reg) : mon expr := | Op.Omulimm n, r::nil => ret (boplit Vmul r n) | Op.Omulhs, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhs") | Op.Omulhu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhu") - | Op.Odiv, r1::r2::nil => ret (bop Vdiv r1 r2) + | Op.Odiv, r1::r2::nil => + do s <- get; + match sdiv_present s.(st_funct_units) 32%positive with + | SignedDiv s n d q r :: _ => + ret (Vseq (Vnonblock (Vvar n) (Vvar r1)) (Vnonblock (Vvar d) (Vvar r2))) + | _ => + do n <- create_reg None 32; + do d <- create_reg None 32; + do q <- create_reg None 32; + do r <- create_reg None 32; + do _ <- add_funct_unit (SignedDiv 32%positive n d q r); + ret (Vseq (Vnonblock (Vvar n) (Vvar r1)) (Vnonblock (Vvar d) (Vvar r2))) + end | Op.Odivu, r1::r2::nil => ret (bop Vdivu r1 r2) | Op.Omod, r1::r2::nil => ret (bop Vmod r1 r2) | Op.Omodu, r1::r2::nil => ret (bop Vmodu r1 r2) @@ -406,7 +455,8 @@ Lemma add_branch_instr_state_incr: s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n Vskip (st_datapath s)) - (AssocMap.set n (state_cond s.(st_st) e n1 n2) (st_controllogic s))). + (AssocMap.set n (state_cond s.(st_st) e n1 n2) (st_controllogic s)) + s.(st_funct_units)). Proof. intros. apply state_incr_intro; simpl; try (intros; destruct (peq n0 n); subst); @@ -424,7 +474,8 @@ Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit := s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n Vskip (st_datapath s)) - (AssocMap.set n (state_cond s.(st_st) e n1 n2) (st_controllogic s))) + (AssocMap.set n (state_cond s.(st_st) e n1 n2) (st_controllogic s)) + s.(st_funct_units)) (add_branch_instr_state_incr s e n n1 n2 NTRANS) | _ => Error (Errors.msg "Htlgen: add_branch_instr") end. @@ -466,30 +517,6 @@ Definition tbl_to_case_expr (st : reg) (ns : list node) : list (expr * stmnt) := Definition stack_correct (sz : Z) : bool := (0 <=? sz) && (sz let r := s.(st_freshreg) in - OK r (mkstate - s.(st_st) - (Pos.succ r) - (st_freshstate s) - (AssocMap.set s.(st_freshreg) (i, VScalar sz) s.(st_scldecls)) - (st_arrdecls s) - (st_datapath s) - (st_controllogic s)) - (create_reg_state_incr s sz i). - Lemma create_arr_state_incr: forall s sz ln i, st_incr s (mkstate @@ -499,7 +526,8 @@ Lemma create_arr_state_incr: s.(st_scldecls) (AssocMap.set s.(st_freshreg) (i, VArray sz ln) s.(st_arrdecls)) (st_datapath s) - (st_controllogic s)). + (st_controllogic s) + s.(st_funct_units)). Proof. constructor; simpl; auto with htlh. Qed. Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := @@ -511,7 +539,8 @@ Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := s.(st_scldecls) (AssocMap.set s.(st_freshreg) (i, VArray sz ln) s.(st_arrdecls)) (st_datapath s) - (st_controllogic s)) + (st_controllogic s) + s.(st_funct_units)) (create_arr_state_incr s sz ln i). Definition max_pc_map (m : Maps.PTree.t stmnt) := @@ -574,7 +603,8 @@ Lemma add_data_instr_state_incr : s.(st_arrdecls) (AssocMap.set n (Vseq (AssocMapExt.get_default _ Vskip n s.(st_datapath)) st) s.(st_datapath)) - s.(st_controllogic)). + s.(st_controllogic) + s.(st_funct_units)). Proof. constructor; intros; try (simpl; destruct (peq n n0); subst); @@ -590,7 +620,8 @@ Definition add_data_instr (n : node) (st : stmnt) : mon unit := s.(st_scldecls) s.(st_arrdecls) (AssocMap.set n (Vseq (AssocMapExt.get_default _ Vskip n s.(st_datapath)) st) s.(st_datapath)) - s.(st_controllogic)) + s.(st_controllogic) + s.(st_funct_units)) (add_data_instr_state_incr s n st). Lemma add_control_instr_state_incr : @@ -604,7 +635,8 @@ Lemma add_control_instr_state_incr : s.(st_scldecls) s.(st_arrdecls) s.(st_datapath) - (AssocMap.set n st s.(st_controllogic))). + (AssocMap.set n st s.(st_controllogic)) + s.(st_funct_units)). Proof. constructor; intros; try (simpl; destruct (peq n n0); subst); @@ -622,38 +654,13 @@ Definition add_control_instr (n : node) (st : stmnt) : mon unit := s.(st_scldecls) s.(st_arrdecls) s.(st_datapath) - (AssocMap.set n st s.(st_controllogic))) + (AssocMap.set n st s.(st_controllogic)) + s.(st_funct_units)) (add_control_instr_state_incr s n st CTRL) | _ => Error (Errors.msg "HTLPargen.add_control_instr: control logic is not empty") end. -Definition add_control_instr_force_state_incr : - forall s n st, - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - s.(st_datapath) - (AssocMap.set n st s.(st_controllogic))). -Admitted. - -Definition add_control_instr_force (n : node) (st : stmnt) : mon unit := - fun s => - OK tt (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - s.(st_datapath) - (AssocMap.set n st s.(st_controllogic))) - (add_control_instr_force_state_incr s n st). - - Fixpoint pred_expr (preg: reg) (p: pred_op) := match p with | Pvar pred => @@ -707,7 +714,8 @@ Lemma create_new_state_state_incr: s.(st_scldecls) s.(st_arrdecls) s.(st_datapath) - s.(st_controllogic)). + s.(st_controllogic) + s.(st_funct_units)). Admitted. Definition create_new_state (p: node): mon node := @@ -719,7 +727,8 @@ Definition create_new_state (p: node): mon node := s.(st_scldecls) s.(st_arrdecls) s.(st_datapath) - s.(st_controllogic)) + s.(st_controllogic) + s.(st_funct_units)) (create_new_state_state_incr s p). Definition translate_inst_list (fin rtrn stack preg: reg) (ni : node * node * list instr) := @@ -816,6 +825,7 @@ Definition transf_module (f: function) : mon HTL.module := start rst clk + nil current_state.(st_scldecls) current_state.(st_arrdecls) (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA))) @@ -831,7 +841,8 @@ Definition max_state (f: function) : state := (AssocMap.set st (None, VScalar 32) (st_scldecls (init_state st))) (st_arrdecls (init_state st)) (st_datapath (init_state st)) - (st_controllogic (init_state st)). + (st_controllogic (init_state st)) + (st_funct_units (init_state st)). Definition transl_module (f : function) : Errors.res HTL.module := run_mon (max_state f) (transf_module f). diff --git a/src/hls/HTLgen.v b/src/hls/HTLgen.v index f1e6b2a..b323a65 100644 --- a/src/hls/HTLgen.v +++ b/src/hls/HTLgen.v @@ -610,6 +610,7 @@ Definition transf_module (f: function) : mon HTL.module := start rst clk + nil current_state.(st_scldecls) current_state.(st_arrdecls) (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA))) diff --git a/src/hls/HTLgenspec.v b/src/hls/HTLgenspec.v index 845b1d5..ae7917f 100644 --- a/src/hls/HTLgenspec.v +++ b/src/hls/HTLgenspec.v @@ -183,7 +183,7 @@ Inductive tr_module (f : RTL.function) : module -> Prop := data control f.(RTL.fn_entrypoint) - st stk stk_len fin rtrn start rst clk scldecls arrdecls wf) -> + st stk stk_len fin rtrn start rst clk nil scldecls arrdecls wf) -> (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i -> tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> diff --git a/src/hls/PrintVerilog.ml b/src/hls/PrintVerilog.ml index e67b567..e9020ac 100644 --- a/src/hls/PrintVerilog.ml +++ b/src/hls/PrintVerilog.ml @@ -28,6 +28,7 @@ open Clflags open Printf open VericertClflags +open FunctionalUnits module PMap = Map.Make (struct type t = P.t @@ -149,10 +150,10 @@ let declarearr t = " ["; sprintf "%d" (Nat.to_int ln - 1); ":0];\n" ] let print_io = function - | Some Vinput -> "input" - | Some Voutput -> "output reg" - | Some Vinout -> "inout" - | None -> "reg" + | Some Vinput -> "input logic" + | Some Voutput -> "output logic" + | Some Vinout -> "inout logic" + | None -> "logic" let decl i = function | Vdecl (io, r, sz) -> concat [indent i; declare (print_io io) (r, sz)] @@ -175,6 +176,16 @@ let rec intersperse c = function let make_io i io r = concat [indent i; io; " "; register r; ";\n"] +let print_funct_units clk = function + | SignedDiv (stages, numer, denom, quot, rem) -> + sprintf "div_signed #(.stages(%d)) divs(.clk(%s), .clken(1'b1), .numer(%s), .denom(%s), .quotient(%s), .remain(%s))\n" + (P.to_int stages) + (register clk) (register numer) (register denom) (register quot) (register rem) + | UnsignedDiv (stages, numer, denom, quot, rem) -> + sprintf "div_unsigned #(.stages(%d)) divs(.clk(%s), .clken(1'b1), .numer(%s), .denom(%s), .quotient(%s), .remain(%s))\n" + (P.to_int stages) + (register clk) (register numer) (register denom) (register quot) (register rem) + let compose f g x = g x |> f let testbench = "module testbench; @@ -233,6 +244,7 @@ let pprint_module debug i n m = concat [ indent i; "module "; (extern_atom n); "("; concat (intersperse ", " (List.map register (inputs @ outputs))); ");\n"; fold_map (pprint_module_item (i+1)) m.mod_body; + concat (List.map (print_funct_units m.mod_clk) m.mod_funct_units); if !option_initial then print_initial i (Nat.to_int m.mod_stk_len) m.mod_stk else ""; if debug then debug_always i m.mod_clk m.mod_st else ""; indent i; "endmodule\n\n" diff --git a/src/hls/Verilog.v b/src/hls/Verilog.v index a7db3ba..779b05c 100644 --- a/src/hls/Verilog.v +++ b/src/hls/Verilog.v @@ -17,27 +17,32 @@ * along with this program. If not, see . *) -From Coq Require Import - Structures.OrderedTypeEx - FSets.FMapPositive - Program.Basics - PeanoNat - ZArith - Lists.List - Program. +Set Implicit Arguments. -Require Import Lia. +Require Import Coq.Structures.OrderedTypeEx. +Require Import Coq.FSets.FMapPositive. +Require Import Coq.Arith.PeanoNat. +Require Import Coq.ZArith.ZArith. +Require Import Coq.Lists.List. +Require Import Coq.micromega.Lia. + +Require Import compcert.common.Errors. +Require Import compcert.common.Globalenvs. +Require Import compcert.common.Smallstep. +Require Import compcert.lib.Integers. +Require compcert.common.Events. + +Require Import vericert.common.Show. +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.Array. +Require Import vericert.hls.AssocMap. +Require Import vericert.hls.FunctionalUnits. +Require Import vericert.hls.ValueInt. Import ListNotations. -From vericert Require Import Vericertlib Show ValueInt AssocMap Array. -From compcert Require Events. -From compcert Require Import Integers Errors Smallstep Globalenvs. - Local Open Scope assocmap. -Set Implicit Arguments. - Definition reg : Type := positive. Definition node : Type := positive. Definition szreg : Type := reg * nat. @@ -103,23 +108,30 @@ Definition nonblock_reg (r : reg) (asr : reg_associations) (v : value) := Inductive scl_decl : Type := VScalar (sz : nat). Inductive arr_decl : Type := VArray (sz : nat) (ln : nat). -(** * Verilog AST +(*| +Verilog AST +=========== The Verilog AST is defined here, which is the target language of the compilation. -** Value +Value +----- The Verilog [value] is a bitvector containg a size and the actual bitvector. In this case, the [Word.word] type is used as many theorems about that bitvector -have already been proven. *) +have already been proven. +|*) -(** ** Binary Operators +(*| +Binary Operators +---------------- These are the binary operations that can be represented in Verilog. Ideally, multiplication and division would be done by custom modules which could al so be scheduled properly. However, for now every Verilog operator is assumed to take -one clock cycle. *) +one clock cycle. +|*) Inductive binop : Type := | Vadd : binop (** addition (binary [+]) *) @@ -147,13 +159,19 @@ Inductive binop : Type := | Vshru : binop. (** shift right unsigned ([>>]) *) (*| Vror : binop (** shift right unsigned ([>>]) *)*) -(** ** Unary Operators *) +(*| +Unary Operators +--------------- +|*) Inductive unop : Type := | Vneg (** negation ([-]) *) | Vnot. (** not operation [!] *) -(** ** Expressions *) +(*| +Expressions +----------- +|*) Inductive expr : Type := | Vlit : value -> expr @@ -168,7 +186,10 @@ Inductive expr : Type := Definition posToExpr (p : positive) : expr := Vlit (posToValue p). -(** ** Statements *) +(*| +Statements +---------- +|*) Inductive stmnt : Type := | Vskip : stmnt @@ -178,14 +199,17 @@ Inductive stmnt : Type := | Vblock : expr -> expr -> stmnt | Vnonblock : expr -> expr -> stmnt. -(** ** Edges +(*| +Edges +----- These define when an always block should be triggered, for example if the always block should be triggered synchronously at the clock edge, or asynchronously for combinational logic. [edge] is not part of [stmnt] in this formalisation because is closer to the -semantics that are used. *) +semantics that are used. +|*) Inductive edge : Type := | Vposedge : reg -> edge @@ -193,11 +217,14 @@ Inductive edge : Type := | Valledge : edge | Voredge : edge -> edge -> edge. -(** ** Module Items +(*| +Module Items +------------ Module items can either be declarations ([Vdecl]) or always blocks ([Valways]). The declarations are always register declarations as combinational logic can be -done using combinational always block instead of continuous assignments. *) +done using combinational always block instead of continuous assignments. +|*) Inductive io : Type := | Vinput : io @@ -214,7 +241,8 @@ Inductive module_item : Type := | Valways_ff : edge -> stmnt -> module_item | Valways_comb : edge -> stmnt -> module_item. -(** The main module type containing all the important control signals +(*| +The main module type containing all the important control signals - [mod_start] If set, starts the computations in the module. - [mod_reset] If set, reset the state in the module. @@ -223,7 +251,7 @@ Inductive module_item : Type := - [mod_finish] Bit that is set if the result is ready. - [mod_return] The return value that was computed. - [mod_body] The body of the module, including the state machine. -*) +|*) Record module : Type := mkmodule { mod_start : reg; @@ -235,6 +263,7 @@ Record module : Type := mkmodule { mod_stk : reg; mod_stk_len : nat; mod_args : list reg; + mod_funct_units: funct_units; mod_body : list module_item; mod_entrypoint : node; }. @@ -243,15 +272,18 @@ Definition fundef := AST.fundef module. Definition program := AST.program fundef unit. -(** Convert a [positive] to an expression directly, setting it to the right - size. *) +(*| +Convert a [positive] to an expression directly, setting it to the right size. +|*) Definition posToLit (p : positive) : expr := Vlit (posToValue p). Definition fext := unit. Definition fextclk := nat -> fext. -(** ** State +(*| +State +----- The [state] contains the following items: n @@ -270,7 +302,8 @@ retrieved and set appropriately. Verilog, as the Verilog was generated by the RTL, which means that we have to make an assumption about how it looks. In this case, that it contains state which determines which part of the Verilog is executed. This is then the part - of the Verilog that should match with the RTL. *) + of the Verilog that should match with the RTL. +|*) Inductive stackframe : Type := Stackframe : @@ -390,41 +423,9 @@ Definition handle_def {A : Type} (a : A) (val : option A) Local Open Scope error_monad_scope. -(*Definition access_fext (f : fext) (r : reg) : res value := - match AssocMap.get r f with - | Some v => OK v - | _ => OK (ZToValue 0) - end.*) - -(* TODO FIX Vvar case without default *) -(*Fixpoint expr_run (assoc : assocmap) (e : expr) - {struct e} : res value := - match e with - | Vlit v => OK v - | Vvar r => match s with - | State _ assoc _ _ _ => handle_def (ZToValue 32 0) assoc!r - | _ => Error (msg "Verilog: Wrong state") - end - | Vvari _ _ => Error (msg "Verilog: variable indexing not modelled") - | Vinputvar r => access_fext s r - | Vbinop op l r => - let lv := expr_run s l in - let rv := expr_run s r in - let oper := binop_run op in - do lv' <- lv; - do rv' <- rv; - handle_opt (msg "Verilog: sizes are not equal") - (eq_to_opt lv' rv' (oper lv' rv')) - | Vunop op e => - let oper := unop_run op in - do ev <- expr_run s e; - OK (oper ev) - | Vternary c te fe => - do cv <- expr_run s c; - if valueToBool cv then expr_run s te else expr_run s fe - end.*) - -(** Return the location of the lhs of an assignment. *) +(*| +Return the location of the lhs of an assignment. +|*) Inductive location : Type := | LocReg (_ : reg) @@ -436,80 +437,6 @@ Inductive location_is : fext -> assocmap -> assocmap_arr -> expr -> location -> expr_runp f asr asa iexp iv -> location_is f asr asa (Vvari r iexp) (LocArray r (valueToNat iv)). -(* Definition assign_name (f : fext) (asr: assocmap) (asa : assocmap_l) (e : expr) : res reg := *) -(* match e with *) -(* | Vvar r => OK r *) -(* | _ => Error (msg "Verilog: expression not supported on lhs of assignment") *) -(* end. *) - -(*Fixpoint stmnt_height (st : stmnt) {struct st} : nat := - match st with - | Vseq s1 s2 => S (stmnt_height s1 + stmnt_height s2) - | Vcond _ s1 s2 => S (Nat.max (stmnt_height s1) (stmnt_height s2)) - | Vcase _ ls (Some st) => - S (fold_right (fun t acc => Nat.max acc (stmnt_height (snd t))) 0%nat ls) - | _ => 1 - end. - -Fixpoint find_case_stmnt (s : state) (v : value) (cl : list (expr * stmnt)) - {struct cl} : option (res stmnt) := - match cl with - | (e, st)::xs => - match expr_run s e with - | OK v' => - match eq_to_opt v v' (veq v v') with - | Some eq => if valueToBool eq then Some (OK st) else find_case_stmnt s v xs - | None => Some (Error (msg "Verilog: equality check sizes not equal")) - end - | Error msg => Some (Error msg) - end - | _ => None - end. - -Fixpoint stmnt_run' (n : nat) (s : state) (st : stmnt) {struct n} : res state := - match n with - | S n' => - match st with - | Vskip => OK s - | Vseq s1 s2 => - do s' <- stmnt_run' n' s s1; - stmnt_run' n' s' s2 - | Vcond c st sf => - do cv <- expr_run s c; - if valueToBool cv - then stmnt_run' n' s st - else stmnt_run' n' s sf - | Vcase e cl defst => - do v <- expr_run s e; - match find_case_stmnt s v cl with - | Some (OK st') => stmnt_run' n' s st' - | Some (Error msg) => Error msg - | None => do s' <- handle_opt (msg "Verilog: no case was matched") - (option_map (stmnt_run' n' s) defst); s' - end - | Vblock lhs rhs => - match s with - | State m assoc nbassoc f c => - do name <- assign_name lhs; - do rhse <- expr_run s rhs; - OK (State m (PositiveMap.add name rhse assoc) nbassoc f c) - | _ => Error (msg "Verilog: Wrong state") - end - | Vnonblock lhs rhs => - match s with - | State m assoc nbassoc f c => - do name <- assign_name lhs; - do rhse <- expr_run s rhs; - OK (State m assoc (PositiveMap.add name rhse nbassoc) f c) - | _ => Error (msg "Verilog: Wrong state") - end - end - | _ => OK s - end. - -Definition stmnt_run (s : state) (st : stmnt) : res state := - stmnt_run' (stmnt_height st) s st. *) - Inductive stmnt_runp: fext -> reg_associations -> arr_associations -> stmnt -> reg_associations -> arr_associations -> Prop := | stmnt_runp_Vskip: @@ -580,20 +507,6 @@ Inductive stmnt_runp: fext -> reg_associations -> arr_associations -> asr (nonblock_arr r i asa rhsval). Hint Constructors stmnt_runp : verilog. -(*Fixpoint mi_step (s : state) (m : list module_item) {struct m} : res state := - let run := fun st ls => - do s' <- stmnt_run s st; - mi_step s' ls - in - match m with - | (Valways _ st)::ls => run st ls - | (Valways_ff _ st)::ls => run st ls - | (Valways_comb _ st)::ls => run st ls - | (Vdecl _ _)::ls => mi_step s ls - | (Vdeclarr _ _ _)::ls => mi_step s ls - | nil => OK s - end.*) - Inductive mi_stepp : fext -> reg_associations -> arr_associations -> module_item -> reg_associations -> arr_associations -> Prop := | mi_stepp_Valways : @@ -625,80 +538,8 @@ Inductive mis_stepp : fext -> reg_associations -> arr_associations -> mis_stepp f sr sa nil sr sa. Hint Constructors mis_stepp : verilog. -(*Definition mi_step_commit (s : state) (m : list module_item) : res state := - match mi_step s m with - | OK (State m assoc nbassoc f c) => - OK (State m (merge_assocmap nbassoc assoc) empty_assocmap f c) - | Error msg => Error msg - | _ => Error (msg "Verilog: Wrong state") - end. - -Fixpoint mi_run (f : fextclk) (assoc : assocmap) (m : list module_item) (n : nat) - {struct n} : res assocmap := - match n with - | S n' => - do assoc' <- mi_run f assoc m n'; - match mi_step_commit (State assoc' empty_assocmap f (Pos.of_nat n')) m with - | OK (State assoc _ _ _) => OK assoc - | Error m => Error m - end - | O => OK assoc - end.*) - -(** Resets the module into a known state, so that it can be executed. This is -assumed to be the starting state of the module, and may have to be changed if -other arguments to the module are also to be supported. *) - -(*Definition initial_fextclk (m : module) : fextclk := - fun x => match x with - | S O => (AssocMap.set (mod_reset m) (ZToValue 1) empty_assocmap) - | _ => (AssocMap.set (mod_reset m) (ZToValue 0) empty_assocmap) - end.*) - -(*Definition module_run (n : nat) (m : module) : res assocmap := - mi_run (initial_fextclk m) empty_assocmap (mod_body m) n.*) - Local Close Scope error_monad_scope. -(*Theorem value_eq_size_if_eq: - forall lv rv EQ, - vsize lv = vsize rv -> value_eq_size lv rv = left EQ. -Proof. intros. unfold value_eq_size. - -Theorem expr_run_same: - forall assoc e v, expr_run assoc e = OK v <-> expr_runp assoc e v. -Proof. - split; generalize dependent v; generalize dependent assoc. - - induction e. - + intros. inversion H. constructor. - + intros. inversion H. constructor. assumption. - + intros. inversion H. destruct (expr_run assoc e1) eqn:?. destruct (expr_run assoc e2) eqn:?. - unfold eq_to_opt in H1. destruct (value_eq_size v0 v1) eqn:?. inversion H1. - constructor. apply IHe1. assumption. apply IHe2. assumption. reflexivity. discriminate. discriminate. - discriminate. - + intros. inversion H. destruct (expr_run assoc e) eqn:?. unfold option_map in H1. - inversion H1. constructor. apply IHe. assumption. reflexivity. discriminate. - + intros. inversion H. destruct (expr_run assoc e1) eqn:?. destruct (valueToBool v0) eqn:?. - eapply erun_Vternary_true. apply IHe1. eassumption. apply IHe2. eassumption. assumption. - eapply erun_Vternary_false. apply IHe1. eassumption. apply IHe3. eassumption. assumption. - discriminate. - - induction e. - + intros. inversion H. reflexivity. - + intros. inversion H. subst. simpl. assumption. - + intros. inversion H. subst. simpl. apply IHe1 in H4. rewrite H4. - apply IHe2 in H6. rewrite H6. unfold eq_to_opt. assert (vsize lv = vsize rv). - apply EQ. apply (value_eq_size_if_eq lv rv EQ) in H0. rewrite H0. reflexivity. - + intros. inversion H. subst. simpl. destruct (expr_run assoc e) eqn:?. simpl. - apply IHe in H3. rewrite Heqo in H3. - inversion H3. subst. reflexivity. apply IHe in H3. rewrite Heqo in H3. discriminate. - + intros. inversion H. subst. simpl. apply IHe1 in H4. rewrite H4. rewrite H7. - apply IHe2 in H6. rewrite H6. reflexivity. - subst. simpl. apply IHe1 in H4. rewrite H4. rewrite H7. apply IHe3. - assumption. -Qed. - - *) - Fixpoint init_params (vl : list value) (rl : list reg) {struct rl} := match rl, vl with | r :: rl', v :: vl' => AssocMap.set r v (init_params vl' rl') @@ -886,7 +727,8 @@ Lemma semantics_determinate : Proof. intros. constructor; set (ge := Globalenvs.Genv.globalenv p); simplify. - invert H; invert H0; constructor. (* Traces are always empty *) - - invert H; invert H0; crush. assert (f = f0) by (destruct f; destruct f0; auto); subst. + - invert H; invert H0; crush. + assert (f = f0) by (destruct f; destruct f0; auto); subst. pose proof (mis_stepp_determinate H5 H15). crush. - constructor. invert H; crush. diff --git a/src/hls/Veriloggen.v b/src/hls/Veriloggen.v index 80c0669..6ea00e0 100644 --- a/src/hls/Veriloggen.v +++ b/src/hls/Veriloggen.v @@ -61,6 +61,7 @@ Definition transl_module (m : HTL.module) : Verilog.module := m.(HTL.mod_stk) m.(HTL.mod_stk_len) m.(HTL.mod_params) + m.(HTL.mod_funct_units) body m.(HTL.mod_entrypoint). -- cgit From a47cfd17f0e1fc6aca5e10de9362a4be2d4af468 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 21 Feb 2021 21:35:11 +0000 Subject: Correctly add initial scheduling variables --- src/hls/Schedule.ml | 24 ++++++++++++++++++++---- 1 file changed, 20 insertions(+), 4 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index c6c8bf4..9adf17c 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -556,6 +556,18 @@ let gather_bb_constraints debug bb = let encode_var bbn n i = { sv_type = VarType (bbn, n); sv_num = i } let encode_bb bbn i = { sv_type = BBType bbn; sv_num = i } +let add_initial_sv n dfg constr = + let add_initial_sv' (i, instr) g = + let pipes = pipeline_stages instr in + if pipes > 0 then + List.init pipes (fun i' -> i') + |> List.fold_left (fun g i' -> + G.add_edge_e g (encode_var n i i', -1, encode_var n i (i'+1)) + ) g + else g + in + DFG.fold_vertex add_initial_sv' dfg constr + let add_super_nodes n dfg = DFG.fold_vertex (function v1 -> fun g -> (if DFG.in_degree dfg v1 = 0 @@ -563,12 +575,14 @@ let add_super_nodes n dfg = else g) |> (fun g' -> if DFG.out_degree dfg v1 = 0 - then G.add_edge_e g' (encode_var n (fst v1) 0, 0, encode_bb n 1) + then G.add_edge_e g' (encode_var n (fst v1) (pipeline_stages (snd v1)), + 0, encode_bb n 1) else g')) dfg let add_data_deps n = - DFG.fold_edges_e (function ((i1, _), l, (i2, _)) -> fun g -> - G.add_edge_e g (encode_var n i1 0, 0, encode_var n i2 0) + DFG.fold_edges_e (function ((i1, instr1), _, (i2, _)) -> fun g -> + let end_sv = pipeline_stages instr1 in + G.add_edge_e g (encode_var n i1 end_sv, 0, encode_var n i2 0) ) let add_ctrl_deps n succs constr = @@ -596,7 +610,7 @@ let add_cycle_constr max n dfg constr = let negated_dfg = negate_graph dfg in let longest_path v = BFDFG.all_shortest_paths negated_dfg v |> BFDFG.H.to_seq |> List.of_seq in - let constrained_paths = List.filter (function (v, m) -> - m > max) in + let constrained_paths = List.filter (function (_, m) -> - m > max) in List.fold_left (fun g -> function (v, v', w) -> G.add_edge_e g (encode_var n (fst v) 0, - (int_of_float (Float.ceil (Float.div (float_of_int w) (float_of_int max))) - 1), @@ -656,6 +670,7 @@ let gather_cfg_constraints c constr curr = | None -> assert false | Some { bb_exit = ctrl; _ } -> add_super_nodes n dfg constr + |> add_initial_sv n dfg |> add_data_deps n dfg |> add_ctrl_deps n (successors_instr ctrl |> List.map P.to_int @@ -797,5 +812,6 @@ let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = fn_params = f.fn_params; fn_stacksize = f.fn_stacksize; fn_code = List.fold_left (add_to_tree scheduled) PTree.empty reachable; + fn_funct_units = f.fn_funct_units; fn_entrypoint = f.fn_entrypoint } -- cgit From 359194617de51adcc451b45b6c1b0a9332820906 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 21 Feb 2021 21:35:34 +0000 Subject: Add new instructions for pipelines --- src/hls/FunctionalUnits.v | 6 ++++-- src/hls/HTLPargen.v | 26 ++++++++------------------ src/hls/HTLgen.v | 3 ++- src/hls/HTLgenspec.v | 3 ++- src/hls/IfConversion.v | 2 +- src/hls/Partition.ml | 1 + src/hls/PrintVerilog.ml | 4 +++- src/hls/RTLBlockInstr.v | 8 ++++++++ src/hls/RTLPargen.v | 4 ++++ 9 files changed, 33 insertions(+), 24 deletions(-) (limited to 'src/hls') diff --git a/src/hls/FunctionalUnits.v b/src/hls/FunctionalUnits.v index 019cf15..392b1ae 100644 --- a/src/hls/FunctionalUnits.v +++ b/src/hls/FunctionalUnits.v @@ -21,13 +21,15 @@ Require Import compcert.lib.Maps. Require Import vericert.common.Vericertlib. +Definition funct_node := positive. + Inductive funct_unit: Type := | SignedDiv (size: positive) (numer denom quot rem: reg): funct_unit | UnsignedDiv (size: positive) (numer denom quot rem: reg): funct_unit. Record funct_units := mk_avail_funct_units { - avail_sdiv: option positive; - avail_udiv: option positive; + avail_sdiv: option funct_node; + avail_udiv: option funct_node; avail_units: PTree.t funct_unit; }. diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index 01588f6..a5b5f41 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -392,7 +392,7 @@ Definition translate_eff_addressing (a: Op.addressing) (args: list reg) end. (** Translate an instruction to a statement. FIX mulhs mulhu *) -Definition translate_instr (op : Op.operation) (args : list reg) : mon stmnt := +Definition translate_instr (op : Op.operation) (args : list reg) : mon expr := match op, args with | Op.Omove, r::nil => ret (Vvar r) | Op.Ointconst n, _ => ret (Vlit (intToValue n)) @@ -402,22 +402,10 @@ Definition translate_instr (op : Op.operation) (args : list reg) : mon stmnt := | Op.Omulimm n, r::nil => ret (boplit Vmul r n) | Op.Omulhs, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhs") | Op.Omulhu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhu") - | Op.Odiv, r1::r2::nil => - do s <- get; - match sdiv_present s.(st_funct_units) 32%positive with - | SignedDiv s n d q r :: _ => - ret (Vseq (Vnonblock (Vvar n) (Vvar r1)) (Vnonblock (Vvar d) (Vvar r2))) - | _ => - do n <- create_reg None 32; - do d <- create_reg None 32; - do q <- create_reg None 32; - do r <- create_reg None 32; - do _ <- add_funct_unit (SignedDiv 32%positive n d q r); - ret (Vseq (Vnonblock (Vvar n) (Vvar r1)) (Vnonblock (Vvar d) (Vvar r2))) - end - | Op.Odivu, r1::r2::nil => ret (bop Vdivu r1 r2) - | Op.Omod, r1::r2::nil => ret (bop Vmod r1 r2) - | Op.Omodu, r1::r2::nil => ret (bop Vmodu r1 r2) + | Op.Odiv, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: div") + | Op.Odivu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: divu") + | Op.Omod, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mod") + | Op.Omodu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: modu") | Op.Oand, r1::r2::nil => ret (bop Vand r1 r2) | Op.Oandimm n, r::nil => ret (boplit Vand r n) | Op.Oor, r1::r2::nil => ret (bop Vor r1 r2) @@ -702,6 +690,8 @@ Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) | RBsetpred c args p => do cond <- translate_condition c args; ret (a (pred_expr preg (Pvar p)) cond) + | RBpiped p f args => error (Errors.msg "HTLPargen.translate_inst: piped") + | RBassign p f dst => error (Errors.msg "HTLPargen.translate_inst: assign") end. Lemma create_new_state_state_incr: @@ -828,7 +818,7 @@ Definition transf_module (f: function) : mon HTL.module := start rst clk - nil + initial_funct_units current_state.(st_scldecls) current_state.(st_arrdecls) (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA))) diff --git a/src/hls/HTLgen.v b/src/hls/HTLgen.v index b323a65..76616fb 100644 --- a/src/hls/HTLgen.v +++ b/src/hls/HTLgen.v @@ -32,6 +32,7 @@ Require Import vericert.hls.AssocMap. Require Import vericert.hls.HTL. Require Import vericert.hls.ValueInt. Require Import vericert.hls.Verilog. +Require Import vericert.hls.FunctionalUnits. Hint Resolve AssocMap.gempty : htlh. Hint Resolve AssocMap.gso : htlh. @@ -610,7 +611,7 @@ Definition transf_module (f: function) : mon HTL.module := start rst clk - nil + initial_funct_units current_state.(st_scldecls) current_state.(st_arrdecls) (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA))) diff --git a/src/hls/HTLgenspec.v b/src/hls/HTLgenspec.v index ae7917f..556e3cc 100644 --- a/src/hls/HTLgenspec.v +++ b/src/hls/HTLgenspec.v @@ -31,6 +31,7 @@ Require Import vericert.hls.ValueInt. Require Import vericert.hls.HTL. Require Import vericert.hls.HTLgen. Require Import vericert.hls.AssocMap. +Require Import vericert.hls.FunctionalUnits. Hint Resolve Maps.PTree.elements_keys_norepet : htlspec. Hint Resolve Maps.PTree.elements_correct : htlspec. @@ -183,7 +184,7 @@ Inductive tr_module (f : RTL.function) : module -> Prop := data control f.(RTL.fn_entrypoint) - st stk stk_len fin rtrn start rst clk nil scldecls arrdecls wf) -> + st stk stk_len fin rtrn start rst clk initial_funct_units scldecls arrdecls wf) -> (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i -> tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> diff --git a/src/hls/IfConversion.v b/src/hls/IfConversion.v index 39d9fd2..e893578 100644 --- a/src/hls/IfConversion.v +++ b/src/hls/IfConversion.v @@ -113,7 +113,7 @@ 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_entrypoint). + mkfunction f.(fn_sig) f.(fn_params) f.(fn_stacksize) c f.(fn_funct_units) f.(fn_entrypoint). Definition transf_fundef (fd: fundef) : fundef := transf_fundef transf_function fd. diff --git a/src/hls/Partition.ml b/src/hls/Partition.ml index 19c6048..270db14 100644 --- a/src/hls/Partition.ml +++ b/src/hls/Partition.ml @@ -118,6 +118,7 @@ let function_from_RTL f = fn_stacksize = f.RTL.fn_stacksize; fn_params = f.RTL.fn_params; fn_entrypoint = f.RTL.fn_entrypoint; + fn_funct_units = FunctionalUnits.initial_funct_units; fn_code = c } diff --git a/src/hls/PrintVerilog.ml b/src/hls/PrintVerilog.ml index 8c9f20e..824623e 100644 --- a/src/hls/PrintVerilog.ml +++ b/src/hls/PrintVerilog.ml @@ -257,7 +257,9 @@ let pprint_module debug i n m = concat [ indent i; "module "; (extern_atom n); "("; concat (intersperse ", " (List.map register (inputs @ outputs))); ");\n"; fold_map (pprint_module_item (i+1)) m.mod_body; - concat (List.map (print_funct_units m.mod_clk) m.mod_funct_units); + concat (List.map (print_funct_units m.mod_clk) + (Maps.PTree.elements m.mod_funct_units.avail_units + |> List.map snd)); if !option_initial then print_initial i (Nat.to_int m.mod_stk_len) m.mod_stk else ""; if debug then debug_always_verbose i m.mod_clk m.mod_st else ""; indent i; "endmodule\n\n" diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 86f8eba..5ee9702 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -30,6 +30,7 @@ Require Import compcert.verilog.Op. Require Import vericert.common.Vericertlib. Require Import vericert.hls.Sat. +Require Import vericert.hls.FunctionalUnits. Local Open Scope rtl. @@ -250,6 +251,8 @@ Inductive instr : Type := | RBop : option pred_op -> operation -> list reg -> reg -> instr | RBload : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr | RBstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr +| RBpiped : option pred_op -> funct_node -> list reg -> instr +| RBassign : option pred_op -> funct_node -> reg -> instr | RBsetpred : condition -> list reg -> predicate -> instr. Inductive cf_instr : Type := @@ -284,6 +287,10 @@ Definition max_reg_instr (m: positive) (i: instr) := fold_left Pos.max args (Pos.max dst m) | RBstore p chunk addr args src => fold_left Pos.max args (Pos.max src m) + | RBpiped p f args => + fold_left Pos.max args m + | RBassign p f dst => + Pos.max dst m | RBsetpred c args p => fold_left Pos.max args m end. @@ -339,6 +346,7 @@ Section DEFINITION. fn_params: list reg; fn_stacksize: Z; fn_code: code; + fn_funct_units: funct_units; fn_entrypoint: node }. diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index e2e9a90..97ec9aa 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -576,6 +576,8 @@ Definition update (f : forest) (i : instr) : forest := | RBstore p chunk addr rl r => f # Mem <- (Estore (f # Mem) chunk addr (list_translation rl f) (f # (Reg r))) | RBsetpred c addr p => f + | RBpiped p fu args => f + | RBassign p fu dst => f end. (*| @@ -673,6 +675,7 @@ Definition transl_function (f: RTLBlock.function) : Errors.res RTLPar.function : f.(fn_params) f.(fn_stacksize) tfcode + f.(fn_funct_units) f.(fn_entrypoint)) else Errors.Error (Errors.msg "RTLPargen: Could not prove the blocks equivalent."). @@ -683,6 +686,7 @@ Definition transl_function_temp (f: RTLBlock.function) : Errors.res RTLPar.funct f.(fn_params) f.(fn_stacksize) tfcode + f.(fn_funct_units) f.(fn_entrypoint)). Definition transl_fundef := transf_partial_fundef transl_function_temp. -- cgit From a5ae227155f09822207c49d75fdbb6e4b7652936 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 21 Feb 2021 21:35:52 +0000 Subject: Add operator pipelining pass --- src/hls/PipelineOp.v | 67 ++++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 67 insertions(+) create mode 100644 src/hls/PipelineOp.v (limited to 'src/hls') diff --git a/src/hls/PipelineOp.v b/src/hls/PipelineOp.v new file mode 100644 index 0000000..0783e4e --- /dev/null +++ b/src/hls/PipelineOp.v @@ -0,0 +1,67 @@ +(* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2021 Yann Herklotz + * + * This program is free software: you can redistribute it and/or modify + * it under the terms of the GNU General Public License as published by + * the Free Software Foundation, either version 3 of the License, or + * (at your option) any later version. + * + * This program is distributed in the hope that it will be useful, + * but WITHOUT ANY WARRANTY; without even the implied warranty of + * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + * GNU General Public License for more details. + * + * You should have received a copy of the GNU General Public License + * along with this program. If not, see . + *) + +Require Import compcert.common.AST. +Require Import compcert.common.Errors. +Require Import compcert.common.Globalenvs. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. +Require Import compcert.verilog.Op. + +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlockInstr. +Require Import vericert.hls.RTLPar. +Require Import vericert.hls.FunctionalUnits. + +Definition state : Type := code * funct_units. + +Definition div_pos (il: nat * list nat) x := + let (i, l) := il in + match x with + | RBop _ Odiv _ _ => (S i, i :: l) + | RBop _ Odivu _ _ => (S i, i :: l) + | RBop _ Omod _ _ => (S i, i :: l) + | RBop _ Omodu _ _ => (S i, i :: l) + | _ => (S i, l) + end. + +Definition find_divs (bb: bblock) : list nat := + snd (List.fold_left div_pos bb.(bb_body) (1%nat, nil)). + +Definition transf_code (i: state) (bbn: node * bblock) : state := + let (c, fu) := i in + let (n, bb) := bbn in + (PTree.set n bb c, fu). + +Definition transf_function (f: function) : function := + let (c, fu) := List.fold_left transf_code + (PTree.elements f.(fn_code)) + (PTree.empty bblock, f.(fn_funct_units)) in + mkfunction + f.(fn_sig) + f.(fn_params) + f.(fn_stacksize) + c + fu + f.(fn_entrypoint). + +Definition transf_fundef (fd: fundef) : fundef := + transf_fundef transf_function fd. + +Definition transf_program (p: program) : program := + transform_program transf_fundef p. -- cgit From 91b62a7100ccb1f8910422bc4566bc57ade8f7e7 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 22 Feb 2021 18:41:33 +0000 Subject: Add RTLPar printing --- src/hls/printRTLPar.ml | 74 ++++++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 74 insertions(+) create mode 100644 src/hls/printRTLPar.ml (limited to 'src/hls') diff --git a/src/hls/printRTLPar.ml b/src/hls/printRTLPar.ml new file mode 100644 index 0000000..7fac0de --- /dev/null +++ b/src/hls/printRTLPar.ml @@ -0,0 +1,74 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Pretty-printers for RTL code *) + +open Printf +open Camlcoq +open Datatypes +open Maps +open AST +open RTLBlockInstr +open RTLPar +open PrintAST +open PrintRTLBlockInstr + +(* Printing of RTL code *) + +let reg pp r = + fprintf pp "x%d" (P.to_int r) + +let rec regs pp = function + | [] -> () + | [r] -> reg pp r + | r1::rl -> fprintf pp "%a, %a" reg r1 regs rl + +let ros pp = function + | Coq_inl r -> reg pp r + | Coq_inr s -> fprintf pp "\"%s\"" (extern_atom s) + +let print_bblock pp (pc, i) = + fprintf pp "%5d:{\n" pc; + List.iter (fun x -> fprintf pp "["; + List.iter (fun x' -> fprintf pp "["; List.iter (print_bblock_body pp) x'; fprintf pp "]\n") x; + fprintf pp "]\n") i.bb_body; + print_bblock_exit pp i.bb_exit; + fprintf pp "\t}\n\n" + +let print_function pp id f = + fprintf pp "%s(%a) {\n" (extern_atom id) regs f.fn_params; + let instrs = + List.sort + (fun (pc1, _) (pc2, _) -> compare pc2 pc1) + (List.rev_map + (fun (pc, i) -> (P.to_int pc, i)) + (PTree.elements f.fn_code)) in + List.iter (print_bblock pp) instrs; + fprintf pp "}\n\n" + +let print_globdef pp (id, gd) = + match gd with + | Gfun(Internal f) -> print_function pp id f + | _ -> () + +let print_program pp (prog: program) = + List.iter (print_globdef pp) prog.prog_defs + +let destination : string option ref = ref None + +let print_if passno prog = + match !destination with + | None -> () + | Some f -> + let oc = open_out (f ^ "." ^ Z.to_string passno) in + print_program oc prog; + close_out oc -- cgit From 1a7f581dc4b67cbfe17936697aec8f85786c844d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 22 Feb 2021 18:41:47 +0000 Subject: Add operation pipelining --- src/hls/PipelineOp.v | 141 ++++++++++++++++++++++++++++++++++++++++++++++++--- 1 file changed, 133 insertions(+), 8 deletions(-) (limited to 'src/hls') diff --git a/src/hls/PipelineOp.v b/src/hls/PipelineOp.v index 0783e4e..793b752 100644 --- a/src/hls/PipelineOp.v +++ b/src/hls/PipelineOp.v @@ -16,6 +16,9 @@ * along with this program. If not, see . *) +Require Import Coq.Lists.List. + +Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Errors. Require Import compcert.common.Globalenvs. @@ -23,12 +26,15 @@ Require Import compcert.lib.Integers. Require Import compcert.lib.Maps. Require Import compcert.verilog.Op. + Require Import vericert.common.Vericertlib. Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLPar. Require Import vericert.hls.FunctionalUnits. -Definition state : Type := code * funct_units. +Import Option.Notation. + +Local Open Scope positive. Definition div_pos (il: nat * list nat) x := let (i, l) := il in @@ -40,24 +46,143 @@ Definition div_pos (il: nat * list nat) x := | _ => (S i, l) end. -Definition find_divs (bb: bblock) : list nat := - snd (List.fold_left div_pos bb.(bb_body) (1%nat, nil)). +Definition div_pos_in_list (il: nat * list (nat * nat)) bb := + let (i, l) := il in + let dp := snd (List.fold_left div_pos bb (0%nat, nil)) in + (S i, (List.map (fun x => (i, x)) dp) ++ l). + +Definition div_pos_in_block (il: nat * list (nat * nat * nat)) bb := + let (i, l) := il in + let dp := snd (List.fold_left div_pos_in_list bb (0%nat, nil)) in + (S i, (List.map (fun (yx : nat * nat) => let (y, x) := yx in (i, y, x)) dp) ++ l). + +Definition find_divs (bb: bblock) := + snd (List.fold_left div_pos_in_block bb.(bb_body) (0%nat, nil)). + +Fixpoint mapi' {A B: Type} (n: nat) (f: nat -> A -> B) (l: list A): list B := + match l with + | nil => nil + | a :: b => f n a :: mapi' (S n) f b + end. + +Definition mapi {A B: Type} := @mapi' A B 0. + +Definition map_at {A: Type} (n: nat) (f: A -> A) (l: list A): list A := + mapi (fun i il => + if Nat.eqb n i + then f il + else il + ) l. + +Definition map_at_err {A: Type} (n: nat) (f: A -> A) (l: list A): option (list A) := + if (Datatypes.length l <=? n)%nat + then None + else Some (map_at n f l). + +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 + | RBop op Omod args dst => RBpiped op sdiv args + | RBop op Omodu args dst => RBpiped op udiv args + | _ => d + end. + +Definition get_sdiv (fu: funct_unit) : option (reg * reg * reg) := + match fu with + | SignedDiv s a b d _ => Some (a, b, d) + | _ => None + end. -Definition transf_code (i: state) (bbn: node * bblock) : state := +Definition get_udiv (fu: funct_unit) : option (reg * reg * reg) := + match fu with + | UnsignedDiv s a b d _ => Some (a, b, d) + | _ => None + end. + +Definition get_smod (fu: funct_unit) : option (reg * reg * reg) := + match fu with + | SignedDiv s a b _ d => Some (a, b, d) + | _ => None + end. + +Definition get_umod (fu: funct_unit) : option (reg * reg * reg) := + match fu with + | UnsignedDiv s a b _ d => Some (a, b, d) + | _ => None + end. + +Definition bind3 {A B C D: Type} + (f: option (A * B * C)) + (g: A -> B -> C -> option D) : option D := + match f with + | Some (a, b, c) => g a b c + | None => None + end. + +Notation "'do' ( X , Y , Z ) <- A ; B" := (bind3 A (fun X Y Z => B)) + (at level 200, X ident, Y ident, Z ident, A at level 100, B at level 200). + +Definition replace_div_error (fu: funct_units) (bb: bb) (loc: nat * nat * nat) := + match loc with + | (z, y, x) => + do sdiv <- fu.(avail_sdiv); + do udiv <- fu.(avail_udiv); + do sfu <- PTree.get sdiv fu.(avail_units); + do ufu <- PTree.get udiv fu.(avail_units); + do z' <- List.nth_error bb z; + do y' <- List.nth_error z' y; + do x' <- List.nth_error y' x; + let transf := map_at z (map_at y (map_at x (replace_div' sdiv udiv))) bb in + match x' with + | RBop op Odiv args dst => + do (s1, s2, src) <- get_sdiv sfu; + map_at_err (z + 32)%nat (fun x => (RBassign op sdiv src dst :: nil) :: x) transf + | RBop op Odivu args dst => + do (s1, s2, src) <- get_udiv ufu; + map_at_err (z + 32)%nat (fun x => (RBassign op sdiv src dst :: nil) :: x) transf + | RBop op Omod args dst => + do (s1, s2, src) <- get_smod sfu; + map_at_err (z + 32)%nat (fun x => (RBassign op sdiv src dst :: nil) :: x) transf + | RBop op Omodu args dst => + do (s1, s2, src) <- get_umod ufu; + map_at_err (z + 32)%nat (fun x => (RBassign op sdiv src dst :: nil) :: x) transf + | _ => None + end + end. + +Definition replace_div (fu: funct_units) (bb: bb) loc := + match replace_div_error fu bb loc with + | Some bb' => bb' + | _ => bb + end. + +Definition transf_code (i: code * funct_units) (bbn: node * bblock) := let (c, fu) := i in let (n, bb) := bbn in - (PTree.set n bb c, fu). + let divs := find_divs bb in + match divs with + | nil => (PTree.set n bb c, fu) + | _ => + (PTree.set n (mk_bblock (List.fold_left (replace_div fu) divs bb.(bb_body)) bb.(bb_exit)) c, fu) + end. + +Definition create_fu (r: reg) := + let fu := PTree.set 2 (UnsignedDiv 32 (r+5) (r+6) (r+7) (r+8)) + (PTree.set 1 (SignedDiv 32 (r+1) (r+2) (r+3) (r+4)) (PTree.empty _)) in + mk_avail_funct_units (Some 1) (Some 2) fu. Definition transf_function (f: function) : function := - let (c, fu) := List.fold_left transf_code + let fu := create_fu (max_reg_function f) in + let (c, fu') := List.fold_left transf_code (PTree.elements f.(fn_code)) - (PTree.empty bblock, f.(fn_funct_units)) in + (PTree.empty bblock, fu) in mkfunction f.(fn_sig) f.(fn_params) f.(fn_stacksize) c - fu + fu' f.(fn_entrypoint). Definition transf_fundef (fd: fundef) : fundef := -- cgit From 2243443c407a1e951265a9252bac3d3b9b830cbb Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 22 Feb 2021 18:42:11 +0000 Subject: Fix arguments to RBassign and piped --- src/hls/HTLPargen.v | 2 +- src/hls/PrintRTLBlockInstr.ml | 5 +++++ src/hls/RTLBlockInstr.v | 6 +++--- src/hls/RTLPargen.v | 2 +- 4 files changed, 10 insertions(+), 5 deletions(-) (limited to 'src/hls') diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index a5b5f41..e1c7ec8 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -691,7 +691,7 @@ Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) do cond <- translate_condition c args; ret (a (pred_expr preg (Pvar p)) cond) | RBpiped p f args => error (Errors.msg "HTLPargen.translate_inst: piped") - | RBassign p f dst => error (Errors.msg "HTLPargen.translate_inst: assign") + | RBassign p f src dst => error (Errors.msg "HTLPargen.translate_inst: assign") end. Lemma create_new_state_state_incr: diff --git a/src/hls/PrintRTLBlockInstr.ml b/src/hls/PrintRTLBlockInstr.ml index ba7241b..808d342 100644 --- a/src/hls/PrintRTLBlockInstr.ml +++ b/src/hls/PrintRTLBlockInstr.ml @@ -52,6 +52,11 @@ let print_bblock_body pp i = fprintf pp "%a = %a\n" pred p (PrintOp.print_condition reg) (c, args) + | RBpiped (p, fu, args) -> + fprintf pp "%a piped\n" print_pred_option p + | RBassign (p, fu, src, dst) -> + fprintf pp "%a %a = %a" print_pred_option p + reg src reg dst let rec print_bblock_exit pp i = fprintf pp "\t\t"; diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 5ee9702..69cc709 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -252,7 +252,7 @@ Inductive instr : Type := | RBload : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr | RBstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr | RBpiped : option pred_op -> funct_node -> list reg -> instr -| RBassign : option pred_op -> funct_node -> reg -> instr +| RBassign : option pred_op -> funct_node -> reg -> reg -> instr | RBsetpred : condition -> list reg -> predicate -> instr. Inductive cf_instr : Type := @@ -289,8 +289,8 @@ Definition max_reg_instr (m: positive) (i: instr) := fold_left Pos.max args (Pos.max src m) | RBpiped p f args => fold_left Pos.max args m - | RBassign p f dst => - Pos.max dst m + | RBassign p f src dst => + Pos.max src (Pos.max dst m) | RBsetpred c args p => fold_left Pos.max args m end. diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index 97ec9aa..aaabe5d 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -577,7 +577,7 @@ Definition update (f : forest) (i : instr) : forest := f # Mem <- (Estore (f # Mem) chunk addr (list_translation rl f) (f # (Reg r))) | RBsetpred c addr p => f | RBpiped p fu args => f - | RBassign p fu dst => f + | RBassign p fu src dst => f end. (*| -- cgit From b34a08dd656664352e400379d2e890ad95e3afc2 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 22 Feb 2021 18:42:31 +0000 Subject: Fix Scheduling to add missing states --- src/hls/Schedule.ml | 48 ++++++++++++++++++++++++++++++++++-------------- 1 file changed, 34 insertions(+), 14 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index 9adf17c..613236f 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -703,16 +703,21 @@ let print_lp constr = let update_schedule v = function Some l -> Some (v :: l) | None -> Some [ v ] -let parse_soln tree s = +let parse_soln (tree, bbtree) s = let r = Str.regexp "var\\([0-9]+\\)n\\([0-9]+\\)_0[ ]+\\([0-9]+\\)" in - if Str.string_match r s 0 then - IMap.update - (Str.matched_group 1 s |> int_of_string) - (update_schedule - ( Str.matched_group 2 s |> int_of_string, - Str.matched_group 3 s |> int_of_string )) - tree - else tree + let bb = Str.regexp "bb\\([0-9]+\\)_\\([0-9]+\\)[ ]+\\([0-9]+\\)" in + let upd s = IMap.update + (Str.matched_group 1 s |> int_of_string) + (update_schedule + ( Str.matched_group 2 s |> int_of_string, + Str.matched_group 3 s |> int_of_string )) + in + if Str.string_match r s 0 + then (upd s tree, bbtree) + else + (if Str.string_match bb s 0 + then (tree, upd s bbtree) + else (tree, bbtree)) let solve_constraints constr = let oc = open_out "lpsolve.txt" in @@ -721,7 +726,7 @@ let solve_constraints constr = Str.split (Str.regexp_string "\n") (read_process "lp_solve lpsolve.txt") |> drop 3 - |> List.fold_left parse_soln IMap.empty + |> List.fold_left parse_soln (IMap.empty, IMap.empty) let subgraph dfg l = let dfg' = List.fold_left (fun g v -> DFG.add_vertex g v) DFG.empty l in @@ -750,24 +755,39 @@ let combine_bb_schedule schedule s = let i, st = s in IMap.update st (update_schedule i) schedule +let range s e = + List.init (e - s) (fun i -> i) + |> List.map (fun x -> x + s) + (** Should generate the [RTLPar] code based on the input [RTLBlock] description. *) -let transf_rtlpar c c' (schedule : (int * int) list IMap.t) = +let transf_rtlpar c c' schedule = let f i bb : RTLPar.bblock = match bb with | { bb_body = []; bb_exit = c } -> { bb_body = []; bb_exit = c } | { bb_body = bb_body'; bb_exit = ctrl_flow } -> let dfg = match PTree.get i c' with None -> assert false | Some x -> x in - let i_sched = IMap.find (P.to_int i) schedule in + let bb_st_e = + let m = IMap.find (P.to_int i) (snd schedule) in + (List.assq 0 m, List.assq 1 m) in + let i_sched = IMap.find (P.to_int i) (fst schedule) in let i_sched_tree = List.fold_left combine_bb_schedule IMap.empty i_sched in - let body = IMap.to_seq i_sched_tree |> List.of_seq |> List.map snd + (*let body = IMap.to_seq i_sched_tree |> List.of_seq |> List.map snd |> List.map (List.map (fun x -> (x, List.nth bb_body' x))) + in*) + let body2 = List.fold_left (fun x b -> + match IMap.find_opt b i_sched_tree with + | Some i -> i :: x + | None -> [] :: x + ) [] (range (fst bb_st_e) (snd bb_st_e + 1)) + |> List.map (List.map (fun x -> (x, List.nth bb_body' x))) + |> List.rev in (*let final_body = List.map (fun x -> subgraph dfg x |> order_instr) body in*) let final_body2 = List.map (fun x -> subgraph dfg x |> (fun x -> TopoDFG.fold (fun i l -> snd i :: l) x []) - |> List.rev) body + |> List.rev) body2 in { bb_body = List.map (fun x -> [x]) final_body2; bb_exit = ctrl_flow -- cgit From 75641815724c68791cc2754e850b35700e07e586 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 22 Feb 2021 19:04:52 +0000 Subject: Get some Verilog output with dividers --- src/hls/HTLPargen.v | 25 ++++++++++++++++--------- src/hls/PrintVerilog.ml | 8 ++++++-- 2 files changed, 22 insertions(+), 11 deletions(-) (limited to 'src/hls') diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index e1c7ec8..629f53e 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -671,7 +671,7 @@ Definition translate_predicate (a : assignment) ret (a dst (Vternary (pred_expr preg pos) e dst)) end. -Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) +Definition translate_inst a (fu : funct_units) (fin rtrn stack preg : reg) (n : node) (i : instr) : mon stmnt := match i with | RBnop => @@ -690,8 +690,14 @@ Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) | RBsetpred c args p => do cond <- translate_condition c args; ret (a (pred_expr preg (Pvar p)) cond) - | RBpiped p f args => error (Errors.msg "HTLPargen.translate_inst: piped") - | RBassign p f src dst => error (Errors.msg "HTLPargen.translate_inst: assign") + | RBpiped p f args => + match PTree.get f fu.(avail_units), args with + | Some (SignedDiv s n d q _), r1::r2::nil => + ret (Vseq (a (Vvar n) (Vvar r1)) (a (Vvar d) (Vvar r2))) + | _, _ => error (Errors.msg "HTLPargen.translate_inst: not a signed divide.") + end + | RBassign p f src dst => + ret (a (Vvar dst) (Vvar src)) end. Lemma create_new_state_state_incr: @@ -721,12 +727,13 @@ Definition create_new_state (p: node): mon node := s.(st_funct_units)) (create_new_state_state_incr s p). -Definition translate_inst_list (fin rtrn stack preg: reg) (ni : node * node * list (list instr)) := +Definition translate_inst_list (fu: funct_units) + (fin rtrn stack preg: reg) (ni : node * node * list (list instr)) := match ni with | (n, p, li) => do _ <- collectlist (fun l => - do stmnt <- translate_inst Vblock fin rtrn stack preg n l; + do stmnt <- translate_inst Vblock fu fin rtrn stack preg n l; add_data_instr n stmnt) (concat li); do st <- get; add_control_instr n (state_goto st.(st_st) p) @@ -773,11 +780,11 @@ Definition translate_cfi (fin rtrn stack preg: reg) (ni: node * cf_instr) do _ <- add_control_instr n c; add_data_instr n s. -Definition transf_bblock (fin rtrn stack preg: reg) (ni : node * bblock) +Definition transf_bblock (fu: funct_units) (fin rtrn stack preg: reg) (ni : node * bblock) : mon unit := let (n, bb) := ni in do nstate <- create_new_state ((poslength bb.(bb_body)))%positive; - do _ <- collectlist (translate_inst_list fin rtrn stack preg) + do _ <- collectlist (translate_inst_list fu fin rtrn stack preg) (prange n (nstate + poslength bb.(bb_body) - 1)%positive bb.(bb_body)); match bb.(bb_body) with @@ -792,7 +799,7 @@ Definition transf_module (f: function) : mon HTL.module := do (stack, stack_len) <- create_arr None 32 (Z.to_nat (f.(fn_stacksize) / 4)); do preg <- create_reg None 32; - do _ <- collectlist (transf_bblock fin rtrn stack preg) + do _ <- collectlist (transf_bblock f.(fn_funct_units) fin rtrn stack preg) (Maps.PTree.elements f.(fn_code)); do _ <- collectlist (fun r => declare_reg (Some Vinput) r 32) f.(fn_params); @@ -818,7 +825,7 @@ Definition transf_module (f: function) : mon HTL.module := start rst clk - initial_funct_units + f.(fn_funct_units) current_state.(st_scldecls) current_state.(st_arrdecls) (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA))) diff --git a/src/hls/PrintVerilog.ml b/src/hls/PrintVerilog.ml index 824623e..da3bd6e 100644 --- a/src/hls/PrintVerilog.ml +++ b/src/hls/PrintVerilog.ml @@ -178,11 +178,15 @@ let make_io i io r = concat [indent i; io; " "; register r; ";\n"] let print_funct_units clk = function | SignedDiv (stages, numer, denom, quot, rem) -> - sprintf "div_signed #(.stages(%d)) divs(.clk(%s), .clken(1'b1), .numer(%s), .denom(%s), .quotient(%s), .remain(%s))\n" + sprintf ("div_signed #(.stages(%d)) divs(.clk(%s), " ^^ + ".clken(1'b1), .numer(%s), .denom(%s), " ^^ + ".quotient(%s), .remain(%s))\n") (P.to_int stages) (register clk) (register numer) (register denom) (register quot) (register rem) | UnsignedDiv (stages, numer, denom, quot, rem) -> - sprintf "div_unsigned #(.stages(%d)) divs(.clk(%s), .clken(1'b1), .numer(%s), .denom(%s), .quotient(%s), .remain(%s))\n" + sprintf ("div_unsigned #(.stages(%d)) divs(.clk(%s), " ^^ + ".clken(1'b1), .numer(%s), .denom(%s), " ^^ + ".quotient(%s), .remain(%s))\n") (P.to_int stages) (register clk) (register numer) (register denom) (register quot) (register rem) -- cgit From 77d5b29503e1359ac1d61209c843091bb14a5ba4 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:01:11 +0000 Subject: Improve simplification of predicates --- src/hls/Predicate.v | 137 ++++++++++++++++++++++++++++++---------------------- 1 file changed, 80 insertions(+), 57 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Predicate.v b/src/hls/Predicate.v index a2b5400..c393a64 100644 --- a/src/hls/Predicate.v +++ b/src/hls/Predicate.v @@ -358,8 +358,74 @@ Proof. rewrite andb_negb_r. auto. Qed. +Definition simplify' (p: pred_op) := + match p with + | (Plit (b1, a)) ∧ (Plit (b2, b)) as p' => + if Pos.eqb a b then + if negb (xorb b1 b2) then Plit (b1, a) else ⟂ + else p' + | (Plit (b1, a)) ∨ (Plit (b2, b)) as p' => + if Pos.eqb a b then + if negb (xorb b1 b2) then Plit (b1, a) else T + else p' + | A ∧ T => A + | T ∧ A => A + | _ ∧ ⟂ => ⟂ + | ⟂ ∧ _ => ⟂ + | _ ∨ T => T + | T ∨ _ => T + | A ∨ ⟂ => A + | ⟂ ∨ A => A + | A => A + end. + +Lemma pred_op_dec : + forall p1 p2: pred_op, + { p1 = p2 } + { p1 <> p2 }. +Proof. pose proof Pos.eq_dec. repeat decide equality. Qed. + +Fixpoint simplify (p: pred_op) := + match p with + | A ∧ B => + let A' := simplify A in + let B' := simplify B in + simplify' (A' ∧ B') + | A ∨ B => + let A' := simplify A in + let B' := simplify B in + simplify' (A' ∨ B') + | T => T + | ⟂ => ⟂ + | Plit a => Plit a + end. + +Lemma simplify'_correct : + forall h a, + sat_predicate (simplify' h) a = sat_predicate h a. +Proof. + (*destruct h; crush; repeat destruct_match; crush; + solve [rewrite andb_true_r; auto | rewrite orb_false_r; auto]. +Qed.*) Admitted. + +Lemma simplify_correct : + forall h a, + sat_predicate (simplify h) a = sat_predicate h a. +Proof. + Local Opaque simplify'. + induction h; crush. + - replace (sat_predicate h1 a && sat_predicate h2 a) + with (sat_predicate (simplify h1) a && sat_predicate (simplify h2) a) + by crush. + rewrite simplify'_correct. crush. + - replace (sat_predicate h1 a || sat_predicate h2 a) + with (sat_predicate (simplify h1) a || sat_predicate (simplify h2) a) + by crush. + rewrite simplify'_correct. crush. + Local Transparent simplify'. +Qed. + Definition equiv_check p1 p2 := - match sat_pred_simple (p1 ∧ ¬ p2 ∨ p2 ∧ ¬ p1) with + match sat_pred_simple (simplify (p1 ∧ ¬ p2 ∨ p2 ∧ ¬ p1)) with | None => true | _ => false end. @@ -371,8 +437,13 @@ Proof. destruct_match; try discriminate; []. destruct_match. destruct_match. discriminate. eapply sat_equiv3; eauto. + unfold unsat; intros. + rewrite <- simplify_correct. eauto. Qed. +Opaque simplify. +Opaque simplify'. + Lemma equiv_check_correct2 : forall p1 p2, p1 == p2 -> equiv_check p1 p2 = true. Proof. @@ -380,7 +451,9 @@ Proof. destruct_match; auto. destruct_match; try discriminate. destruct_match. simplify. apply sat_equiv4 in H. unfold unsat in H. simplify. - clear Heqs. rewrite H in e. discriminate. + clear Heqs. rewrite simplify_correct in e. + specialize (H (interp_alist a)). simplify. + rewrite H1 in e. rewrite H0 in e. discriminate. Qed. Lemma equiv_check_dec : @@ -444,60 +517,10 @@ Proof. auto. Qed. -Definition simplify' (p: pred_op) := - match p with - | A ∧ T => A - | T ∧ A => A - | _ ∧ ⟂ => ⟂ - | ⟂ ∧ _ => ⟂ - | _ ∨ T => T - | T ∨ _ => T - | A ∨ ⟂ => A - | ⟂ ∨ A => A - | A => A - end. - -Lemma pred_op_dec : - forall p1 p2: pred_op, - { p1 = p2 } + { p1 <> p2 }. -Proof. pose proof Pos.eq_dec. repeat decide equality. Qed. - -Fixpoint simplify (p: pred_op) := - match p with - | A ∧ B => - let A' := simplify A in - let B' := simplify B in - simplify' (A' ∧ B') - | A ∨ B => - let A' := simplify A in - let B' := simplify B in - simplify' (A' ∨ B') - | T => T - | ⟂ => ⟂ - | Plit a => Plit a - end. - -Lemma simplify'_correct : - forall h a, - sat_predicate (simplify' h) a = sat_predicate h a. -Proof. - destruct h; crush; repeat destruct_match; crush; - solve [rewrite andb_true_r; auto | rewrite orb_false_r; auto]. -Qed. - -Lemma simplify_correct : - forall h a, - sat_predicate (simplify h) a = sat_predicate h a. +#[global] +Instance simplifyProper : Proper (equiv ==> equiv) simplify. Proof. - Local Opaque simplify'. - induction h; crush. - - replace (sat_predicate h1 a && sat_predicate h2 a) - with (sat_predicate (simplify h1) a && sat_predicate (simplify h2) a) - by crush. - rewrite simplify'_correct. crush. - - replace (sat_predicate h1 a || sat_predicate h2 a) - with (sat_predicate (simplify h1) a || sat_predicate (simplify h2) a) - by crush. - rewrite simplify'_correct. crush. - Local Transparent simplify'. + unfold Proper, "==>". simplify. unfold "→". + intros. unfold sat_equiv; intros. + rewrite ! simplify_correct; auto. Qed. -- cgit From a3f4b9e52563616f6056a9d67344cc326490f2ff Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:01:26 +0000 Subject: Fix pretty printing of predicates --- src/hls/PrintRTLBlockInstr.ml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) (limited to 'src/hls') diff --git a/src/hls/PrintRTLBlockInstr.ml b/src/hls/PrintRTLBlockInstr.ml index 8e24575..b8e1e2e 100644 --- a/src/hls/PrintRTLBlockInstr.ml +++ b/src/hls/PrintRTLBlockInstr.ml @@ -23,9 +23,9 @@ let ros pp = function | Coq_inr s -> fprintf pp "\"%s\"" (extern_atom s) let rec print_pred_op pp = function - | Plit p -> if fst p then pred pp (snd p) else fprintf pp "~ %a" pred (snd p) - | Pand (p1, p2) -> fprintf pp "(%a & %a)" print_pred_op p1 print_pred_op p2 - | Por (p1, p2) -> fprintf pp "(%a | %a)" print_pred_op p1 print_pred_op p2 + | Plit p -> if fst p then pred pp (snd p) else fprintf pp "~%a" pred (snd p) + | Pand (p1, p2) -> fprintf pp "(%a ∧ %a)" print_pred_op p1 print_pred_op p2 + | Por (p1, p2) -> fprintf pp "(%a ∨ %a)" print_pred_op p1 print_pred_op p2 | Ptrue -> fprintf pp "T" | Pfalse -> fprintf pp "⟂" -- cgit From 1cf2d2ef251f228f2fb2c19dffc8f8dc1d60c519 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:01:38 +0000 Subject: Simplify the RTLPargen update function --- src/hls/RTLPargen.v | 32 +++++++++++++++++++++----------- 1 file changed, 21 insertions(+), 11 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index ae4412b..58b048c 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -132,9 +132,19 @@ Fixpoint NEapp {A} (l m: NE.non_empty A) := | a ::| b => a ::| NEapp b m end. -Definition app_predicated {A: Type} (a b: predicated A) := +Definition app_predicated' {A: Type} (a b: predicated A) := let negation := ¬ (NEfold_left (fun a b => a ∨ (fst b)) b ⟂) in - NEapp (NE.map (fun x => (negation ∧ (fst x), snd x)) a) b. + NEapp (NE.map (fun x => (negation ∧ fst x, snd x)) a) b. + +Definition app_predicated {A: Type} (p: option pred_op) (a b: predicated A) := + match p with + | Some p' => NEapp (NE.map (fun x => (¬ p' ∧ fst x, snd x)) a) + (NE.map (fun x => (p' ∧ fst x, snd x)) b) + | None => b + end. + +Definition pred_ret {A: Type} (a: A) : predicated A := + NE.singleton (T, a). (*| Update Function @@ -156,29 +166,29 @@ Definition update (f : forest) (i : instr) : forest := | RBnop => f | RBop p op rl r => f # (Reg r) <- - (app_predicated + (app_predicated p (f # (Reg r)) - (map_predicated (predicated_from_opt p (Eop op)) (merge (list_translation rl f)))) + (map_predicated (pred_ret (Eop op)) (merge (list_translation rl f)))) | RBload p chunk addr rl r => f # (Reg r) <- - (app_predicated + (app_predicated p (f # (Reg r)) (map_predicated - (map_predicated (predicated_from_opt p (Eload chunk addr)) (merge (list_translation rl f))) + (map_predicated (pred_ret (Eload chunk addr)) (merge (list_translation rl f))) (f # Mem))) | RBstore p chunk addr rl r => f # Mem <- - (app_predicated - (f # (Reg r)) + (app_predicated p + (f # Mem) (map_predicated (map_predicated - (predicated_apply2 (map_predicated (predicated_from_opt p Estore) (f # (Reg r))) chunk addr) + (predicated_apply2 (map_predicated (pred_ret Estore) (f # (Reg r))) chunk addr) (merge (list_translation rl f))) (f # Mem))) | RBsetpred p' c args p => f # (Pred p) <- - (app_predicated + (app_predicated p' (f # (Pred p)) - (map_predicated (predicated_from_opt p' (Esetpred c)) (merge (list_translation args f)))) + (map_predicated (pred_ret (Esetpred c)) (merge (list_translation args f)))) end. (*| -- cgit From 1ca4288064dcf9759c06f30b06b415dc7b4e45a1 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:01:53 +0000 Subject: Small progress on the proof of correctness --- src/hls/RTLPargenproof.v | 6 +++++- 1 file changed, 5 insertions(+), 1 deletion(-) (limited to 'src/hls') diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v index df0615a..588f67f 100644 --- a/src/hls/RTLPargenproof.v +++ b/src/hls/RTLPargenproof.v @@ -742,7 +742,11 @@ Lemma sem_update : match_states st' st''' -> @step_instr A ge sp st''' x st'' -> exists tst, sem (mk_ctx st sp ge) (update f x) tst /\ match_states st'' tst. -Proof. Admitted. +Proof. + intros. destruct x. + - inv H1. econstructor. simplify. eauto. symmetry; auto. + - inv H1. inv H0. econstructor. + Admitted. Lemma rtlblock_trans_correct : forall bb ge sp st st', -- cgit From a281749329194cdfe0e623d343ccc4d55cf3bc76 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:02:05 +0000 Subject: Move lp_solve to the /temp directory --- src/hls/Schedule.ml | 10 ++++++---- 1 file changed, 6 insertions(+), 4 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index c5443c7..84fbcc3 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -715,13 +715,15 @@ let parse_soln tree s = else tree let solve_constraints constr = - let oc = open_out "lpsolve.txt" in + let (fn, oc) = Filename.open_temp_file "vericert_" "_lp_solve" in fprintf oc "%s\n" (print_lp constr); close_out oc; - Str.split (Str.regexp_string "\n") (read_process "lp_solve lpsolve.txt") - |> drop 3 - |> List.fold_left parse_soln IMap.empty + let res = Str.split (Str.regexp_string "\n") (read_process ("lp_solve " ^ fn)) + |> drop 3 + |> List.fold_left parse_soln IMap.empty + in + Sys.remove fn; res let subgraph dfg l = let dfg' = List.fold_left (fun g v -> DFG.add_vertex g v) DFG.empty l in -- cgit From 22322017770c9045657f0d3a43f186ab46b0e127 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:02:21 +0000 Subject: Add commented out pretty printing for Abstr --- src/hls/PrintAbstr.ml | 39 +++++++++++++++++++++++++++++++++++++++ 1 file changed, 39 insertions(+) create mode 100644 src/hls/PrintAbstr.ml (limited to 'src/hls') diff --git a/src/hls/PrintAbstr.ml b/src/hls/PrintAbstr.ml new file mode 100644 index 0000000..c63fa02 --- /dev/null +++ b/src/hls/PrintAbstr.ml @@ -0,0 +1,39 @@ +(**open Camlcoq +open Datatypes +open Maps +open AST +open Abstr +open PrintAST +open Printf + +let rec expr_to_list = function + | Enil -> [] + | Econs (e, el) -> e :: expr_to_list el + +let res pp = function + | Reg r -> fprintf pp "r%d" (P.to_int r) + | Pred r -> fprintf pp "p%d" (P.to_int r) + | Mem -> fprintf pp "M" + +let rec print_expression pp = function + | Ebase r -> fprintf pp "%a'" res r + | Eop (op, el) -> fprintf pp "(%a)" (PrintOp.print_operation print_expression) (op, expr_to_list el) + | Eload (chunk, addr, el, e) -> + fprintf pp "(%s[%a][%a])" + (name_of_chunk chunk) print_expression e + (PrintOp.print_addressing print_expression) (addr, expr_to_list el) + | Estore (e, chunk, addr, el, m) -> + fprintf pp "(%s[%a][%a] = %a)" + (name_of_chunk chunk) print_expression m + (PrintOp.print_addressing print_expression) (addr, expr_to_list el) + print_expression e + | Esetpred (c, el) -> + fprintf pp "(%a)" (PrintOp.print_condition print_expression) (c, expr_to_list el) + +let rec print_predicated pp = function + | NE.Coq_singleton (p, e) -> + fprintf pp "%a %a" PrintRTLBlockInstr.print_pred_option p print_expression e + | NE.Coq_cons ((p, e), pr) -> + fprintf pp "%a %a\n%a" PrintRTLBlockInstr.print_pred_option p print_expression e + print_predicated pr +*) -- cgit From dc9ad1382ee548019e6ff546a24954057cdd8ff0 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 13 Nov 2021 23:02:44 +0000 Subject: Add RTLPar with functional units --- src/hls/HTLParFugen.v | 879 ++++++++++++++++++++++++++++++++++++++++++++++++++ src/hls/RTLParFU.v | 402 +++++++++++++++++++++++ src/hls/RTLParFUgen.v | 21 ++ 3 files changed, 1302 insertions(+) create mode 100644 src/hls/HTLParFugen.v create mode 100644 src/hls/RTLParFU.v create mode 100644 src/hls/RTLParFUgen.v (limited to 'src/hls') diff --git a/src/hls/HTLParFugen.v b/src/hls/HTLParFugen.v new file mode 100644 index 0000000..a002e15 --- /dev/null +++ b/src/hls/HTLParFugen.v @@ -0,0 +1,879 @@ +(* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2020 Yann Herklotz + * + * This program is free software: you can redistribute it and/or modify + * it under the terms of the GNU General Public License as published by + * the Free Software Foundation, either version 3 of the License, or + * (at your option) any later version. + * + * This program is distributed in the hope that it will be useful, + * but WITHOUT ANY WARRANTY; without even the implied warranty of + * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + * GNU General Public License for more details. + * + * You should have received a copy of the GNU General Public License + * along with this program. If not, see . + *) + +Require Import Coq.micromega.Lia. + +Require Import compcert.common.AST. +Require compcert.common.Errors. +Require compcert.common.Globalenvs. +Require compcert.lib.Integers. +Require Import compcert.lib.Maps. + +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.RTLParFU. +Require Import vericert.hls.ValueInt. +Require Import vericert.hls.Verilog. + +#[local] Hint Resolve AssocMap.gempty : htlh. +#[local] Hint Resolve AssocMap.gso : htlh. +#[local] Hint Resolve AssocMap.gss : htlh. +#[local] Hint Resolve Ple_refl : htlh. +#[local] Hint Resolve Ple_succ : htlh. + +Definition assignment : Type := expr -> expr -> stmnt. + +Record state: Type := mkstate { + st_st: reg; + st_freshreg: reg; + st_freshstate: node; + st_scldecls: AssocMap.t (option io * scl_decl); + st_arrdecls: AssocMap.t (option io * arr_decl); + st_datapath: datapath; + st_controllogic: controllogic; +}. + +Definition init_state (st : reg) : state := + mkstate st + 1%positive + 1%positive + (AssocMap.empty (option io * scl_decl)) + (AssocMap.empty (option io * arr_decl)) + (AssocMap.empty stmnt) + (AssocMap.empty stmnt). + +Module HTLState <: State. + + Definition st := state. + + Inductive st_incr: state -> state -> Prop := + state_incr_intro: + forall (s1 s2: state), + st_st s1 = st_st s2 -> + Ple s1.(st_freshreg) s2.(st_freshreg) -> + Ple s1.(st_freshstate) s2.(st_freshstate) -> + (forall n, + s1.(st_controllogic)!n = None + \/ s2.(st_controllogic)!n = s1.(st_controllogic)!n) -> + st_incr s1 s2. + #[local] Hint Constructors st_incr : htlh. + + Definition st_prop := st_incr. + #[local] Hint Unfold st_prop : htlh. + + Lemma st_refl : forall s, st_prop s s. + Proof. auto with htlh. Qed. + + Lemma st_trans : + forall s1 s2 s3, st_prop s1 s2 -> st_prop s2 s3 -> st_prop s1 s3. + Proof. + intros. inv H. inv H0. + apply state_incr_intro; eauto using Ple_trans; intros; try congruence. + destruct H4 with n; destruct H7 with n; intuition congruence. + Qed. + +End HTLState. +Export HTLState. + +Module HTLMonad := Statemonad(HTLState). +Export HTLMonad. + +Module HTLMonadExtra := Monad.MonadExtra(HTLMonad). +Import HTLMonadExtra. +Export MonadNotation. + +Definition state_goto (st : reg) (n : node) : stmnt := + Vnonblock (Vvar st) (Vlit (posToValue n)). + +Definition state_cond (st : reg) (c : expr) (n1 n2 : node) : stmnt := + Vnonblock (Vvar st) (Vternary c (posToExpr n1) (posToExpr n2)). + +Definition check_empty_node_datapath: + forall (s: state) (n: node), { s.(st_datapath)!n = None } + { True }. +Proof. + intros. case (s.(st_datapath)!n); tauto. +Defined. + +Definition check_empty_node_controllogic: + forall (s: state) (n: node), { s.(st_controllogic)!n = None } + { True }. +Proof. + intros. case (s.(st_controllogic)!n); tauto. +Defined. + +Lemma declare_reg_state_incr : + forall i s r sz, + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + s.(st_freshstate) + (AssocMap.set r (i, VScalar sz) s.(st_scldecls)) + s.(st_arrdecls) + s.(st_datapath) + s.(st_controllogic)). +Proof. Admitted. + +Definition declare_reg (i : option io) (r : reg) (sz : nat) : mon unit := + fun s => OK tt (mkstate + s.(st_st) + s.(st_freshreg) + s.(st_freshstate) + (AssocMap.set r (i, VScalar sz) s.(st_scldecls)) + s.(st_arrdecls) + s.(st_datapath) + s.(st_controllogic)) + (declare_reg_state_incr i s r sz). + +Lemma add_instr_state_incr : + forall s n n' st, + (st_controllogic s)!n = None -> + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n st s.(st_datapath)) + (AssocMap.set n (state_goto s.(st_st) n') s.(st_controllogic))). +Proof. + constructor; intros; + try (simpl; destruct (peq n n0); subst); + auto with htlh. +Qed. + +Definition add_instr (n : node) (n' : node) (st : stmnt) : mon unit := + fun s => + match check_empty_node_controllogic s n with + | left TRANS => + OK tt (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n st s.(st_datapath)) + (AssocMap.set n (state_goto s.(st_st) n') s.(st_controllogic))) + (add_instr_state_incr s n n' st TRANS) + | _ => Error (Errors.msg "HTL.add_instr") + end. + +Lemma add_instr_skip_state_incr : + forall s n st, + (st_controllogic s)!n = None -> + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n st s.(st_datapath)) + (AssocMap.set n Vskip s.(st_controllogic))). +Proof. + constructor; intros; + try (simpl; destruct (peq n n0); subst); + auto with htlh. +Qed. + +Definition add_instr_skip (n : node) (st : stmnt) : mon unit := + fun s => + match check_empty_node_controllogic s n with + | left TRANS => + OK tt (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n st s.(st_datapath)) + (AssocMap.set n Vskip s.(st_controllogic))) + (add_instr_skip_state_incr s n st TRANS) + | _ => Error (Errors.msg "HTL.add_instr_skip") + end. + +Lemma add_node_skip_state_incr : + forall s n st, + (st_controllogic s)!n = None -> + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n Vskip s.(st_datapath)) + (AssocMap.set n st s.(st_controllogic))). +Proof. + constructor; intros; + try (simpl; destruct (peq n n0); subst); + auto with htlh. +Qed. + +Definition add_node_skip (n : node) (st : stmnt) : mon unit := + fun s => + match check_empty_node_controllogic s n with + | left TRANS => + OK tt (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n Vskip s.(st_datapath)) + (AssocMap.set n st s.(st_controllogic))) + (add_node_skip_state_incr s n st TRANS) + | _ => Error (Errors.msg "HTL.add_node_skip") + end. + +Definition nonblock (dst : reg) (e : expr) := Vnonblock (Vvar dst) e. +Definition block (dst : reg) (e : expr) := Vblock (Vvar dst) e. + +Definition bop (op : binop) (r1 r2 : reg) : expr := + Vbinop op (Vvar r1) (Vvar r2). + +Definition boplit (op : binop) (r : reg) (l : Integers.int) : expr := + Vbinop op (Vvar r) (Vlit (intToValue l)). + +Definition boplitz (op: binop) (r: reg) (l: Z) : expr := + Vbinop op (Vvar r) (Vlit (ZToValue l)). + +Definition translate_comparison (c : Integers.comparison) (args : list reg) + : mon expr := + match c, args with + | Integers.Ceq, r1::r2::nil => ret (bop Veq r1 r2) + | Integers.Cne, r1::r2::nil => ret (bop Vne r1 r2) + | Integers.Clt, r1::r2::nil => ret (bop Vlt r1 r2) + | 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 + "Htlgen: comparison instruction not implemented: other") + end. + +Definition translate_comparison_imm (c : Integers.comparison) (args : list reg) + (i: Integers.int) : mon expr := + match c, args with + | Integers.Ceq, r1::nil => ret (boplit Veq r1 i) + | Integers.Cne, r1::nil => ret (boplit Vne r1 i) + | Integers.Clt, r1::nil => ret (boplit Vlt r1 i) + | 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 + "Htlgen: comparison_imm instruction not implemented: other") + end. + +Definition translate_comparisonu (c : Integers.comparison) (args : list reg) + : mon expr := + match c, args with + | Integers.Clt, r1::r2::nil => ret (bop Vltu r1 r2) + | Integers.Cgt, r1::r2::nil => ret (bop Vgtu r1 r2) + | Integers.Cle, r1::r2::nil => ret (bop Vleu r1 r2) + | Integers.Cge, r1::r2::nil => ret (bop Vgeu r1 r2) + | _, _ => error (Errors.msg + "Htlgen: comparison instruction not implemented: other") + end. + +Definition translate_comparison_immu (c : Integers.comparison) + (args : list reg) (i: Integers.int) : mon expr := + match c, args with + | Integers.Clt, r1::nil => ret (boplit Vltu r1 i) + | Integers.Cgt, r1::nil => ret (boplit Vgtu r1 i) + | Integers.Cle, r1::nil => ret (boplit Vleu r1 i) + | Integers.Cge, r1::nil => ret (boplit Vgeu r1 i) + | _, _ => error (Errors.msg + "Htlgen: comparison_imm instruction not implemented: other") + end. + +Definition translate_condition (c : Op.condition) (args : list reg) + : mon expr := + match c, args with + | Op.Ccomp c, _ => translate_comparison c args + | Op.Ccompu c, _ => translate_comparisonu c args + | Op.Ccompimm c i, _ => translate_comparison_imm c args i + | Op.Ccompuimm c i, _ => translate_comparison_immu c args i + | 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_signed (p : Z) : bool := + Z.leb Integers.Ptrofs.min_signed p + && Z.leb p Integers.Ptrofs.max_signed. + +Definition check_address_parameter_unsigned (p : Z) : bool := + Z.leb p Integers.Ptrofs.max_unsigned. + +Definition translate_eff_addressing (a: Op.addressing) (args: list reg) + : mon expr := + match a, args with (* TODO: We should be more methodical here; what are the possibilities?*) + | Op.Aindexed off, r1::nil => + if (check_address_parameter_signed off) + then ret (boplitz Vadd r1 off) + else error (Errors.msg ("HTLPargen: translate_eff_addressing (Aindexed): address out of bounds")) + | Op.Ascaled scale offset, r1::nil => + if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) + then ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue offset))) + else error (Errors.msg "HTLPargen: translate_eff_addressing (Ascaled): address out of bounds") + | Op.Aindexed2 offset, r1::r2::nil => + if (check_address_parameter_signed offset) + then ret (Vbinop Vadd (bop Vadd r1 r2) (Vlit (ZToValue offset))) + else error (Errors.msg "HTLPargen: translate_eff_addressing (Aindexed2): address out of bounds") + | Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) + if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) + then ret (Vbinop Vadd (Vvar r1) (Vbinop Vadd (boplitz Vmul r2 scale) (Vlit (ZToValue offset)))) + else error (Errors.msg "HTLPargen: translate_eff_addressing (Aindexed2scaled): address out of bounds") + | 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_unsigned a) + then ret (Vlit (ZToValue a)) + else error (Errors.msg "HTLPargen: translate_eff_addressing (Ainstack): address out of bounds") + | _, _ => error (Errors.msg "HTLPargen: translate_eff_addressing unsuported addressing") + end. + +(** 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) + | Op.Ointconst n, _ => ret (Vlit (intToValue n)) + | Op.Oneg, r::nil => ret (Vunop Vneg (Vvar r)) + | 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, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhs") + | Op.Omulhu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhu") + | 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) + | Op.Omodu, r1::r2::nil => ret (bop Vmodu r1 r2) + | Op.Oand, r1::r2::nil => ret (bop Vand r1 r2) + | Op.Oandimm n, r::nil => ret (boplit Vand r n) + | Op.Oor, r1::r2::nil => ret (bop Vor r1 r2) + | Op.Oorimm n, r::nil => ret (boplit Vor r n) + | Op.Oxor, r1::r2::nil => ret (bop Vxor r1 r2) + | Op.Oxorimm n, r::nil => ret (boplit Vxor r n) + | Op.Onot, r::nil => ret (Vunop Vnot (Vvar r)) + | Op.Oshl, r1::r2::nil => ret (bop Vshl r1 r2) + | 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 => + ret (Vternary (Vbinop Vlt (Vvar r) (Vlit (ZToValue 0))) + (Vunop Vneg (Vbinop Vshru (Vunop Vneg (Vvar r)) (Vlit n))) + (Vbinop Vshru (Vvar r) (Vlit 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 => error (Errors.msg "Htlgen: Instruction not implemented: Ororimm") + (*ret (Vbinop Vor (boplit Vshru r (Integers.Int.modu n (Integers.Int.repr 32))) + (boplit Vshl r (Integers.Int.sub (Integers.Int.repr 32) (Integers.Int.modu n (Integers.Int.repr 32)))))*) + | 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 + | _, _ => error (Errors.msg "Htlgen: Instruction not implemented: other") + end. + +Lemma add_branch_instr_state_incr: + forall s e n n1 n2, + (st_controllogic s) ! n = None -> + st_incr s (mkstate + s.(st_st) + (st_freshreg s) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n Vskip (st_datapath s)) + (AssocMap.set n (state_cond s.(st_st) e n1 n2) (st_controllogic s))). +Proof. + intros. apply state_incr_intro; simpl; + try (intros; destruct (peq n0 n); subst); + auto with htlh. +Qed. + +Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit := + fun s => + match check_empty_node_controllogic s n with + | left NTRANS => + OK tt (mkstate + s.(st_st) + (st_freshreg s) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (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 NTRANS) + | _ => Error (Errors.msg "Htlgen: add_branch_instr") + end. + +Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) + (args : list reg) (stack : reg) : mon expr := + match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) + | Mint32, Op.Aindexed off, r1::nil => + if (check_address_parameter_signed off) + then ret (Vvari stack (Vbinop Vdivu (boplitz Vadd r1 off) (Vlit (ZToValue 4)))) + else error (Errors.msg "HTLgen: translate_arr_access address out of bounds") + | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) + if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) + then ret (Vvari stack + (Vbinop Vdivu + (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) + (Vlit (ZToValue 4)))) + else error (Errors.msg "HTLgen: translate_arr_access address out of bounds") + | 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_unsigned a) + then ret (Vvari stack (Vlit (ZToValue (a / 4)))) + else error (Errors.msg "HTLgen: eff_addressing out of bounds stack offset") + | _, _, _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") + end. + +Fixpoint enumerate (i : nat) (ns : list node) {struct ns} : list (nat * node) := + match ns with + | n :: ns' => (i, n) :: enumerate (i+1) ns' + | nil => nil + end. + +Definition tbl_to_case_expr (st : reg) (ns : list node) : list (expr * stmnt) := + List.map (fun a => match a with + (i, n) => (Vlit (natToValue i), Vnonblock (Vvar st) (Vlit (posToValue n))) + end) + (enumerate 0 ns). + +Definition stack_correct (sz : Z) : bool := + (0 <=? sz) && (sz let r := s.(st_freshreg) in + OK r (mkstate + s.(st_st) + (Pos.succ r) + (st_freshstate s) + (AssocMap.set s.(st_freshreg) (i, VScalar sz) s.(st_scldecls)) + (st_arrdecls s) + (st_datapath s) + (st_controllogic s)) + (create_reg_state_incr s sz i). + +Lemma create_arr_state_incr: + forall s sz ln i, + st_incr s (mkstate + s.(st_st) + (Pos.succ (st_freshreg s)) + (st_freshstate s) + s.(st_scldecls) + (AssocMap.set s.(st_freshreg) (i, VArray sz ln) s.(st_arrdecls)) + (st_datapath s) + (st_controllogic s)). +Proof. constructor; simpl; auto with htlh. Qed. + +Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := + fun s => let r := s.(st_freshreg) in + OK (r, ln) (mkstate + s.(st_st) + (Pos.succ r) + (st_freshstate s) + s.(st_scldecls) + (AssocMap.set s.(st_freshreg) (i, VArray sz ln) s.(st_arrdecls)) + (st_datapath s) + (st_controllogic s)) + (create_arr_state_incr s sz ln i). + +Definition max_pc_map (m : Maps.PTree.t stmnt) := + PTree.fold (fun m pc i => Pos.max m pc) m 1%positive. + +Lemma max_pc_map_sound: + forall m pc i, m!pc = Some i -> Ple pc (max_pc_map m). +Proof. + intros until i. + apply PTree_Properties.fold_rec with (P := fun c m => c!pc = Some i -> Ple pc m). + (* extensionality *) + intros. apply H0. rewrite H; auto. + (* base case *) + rewrite PTree.gempty. congruence. + (* inductive case *) + intros. rewrite PTree.gsspec in H2. destruct (peq pc k). + inv H2. unfold Ple, Plt in *. lia. + apply Ple_trans with a. auto. + unfold Ple, Plt in *. lia. +Qed. + +Lemma max_pc_wf : + forall m, Z.pos (max_pc_map m) <= Integers.Int.max_unsigned -> + map_well_formed m. +Proof. + unfold map_well_formed. intros. + exploit list_in_map_inv. eassumption. intros [x [A B]]. destruct x. + apply Maps.PTree.elements_complete in B. apply max_pc_map_sound in B. + unfold Ple in B. apply Pos2Z.pos_le_pos in B. subst. + simplify. transitivity (Z.pos (max_pc_map m)); eauto. +Qed. + +Definition poslength {A : Type} (l : list A) : positive := + match Zlength l with + | Z.pos p => p + | _ => 1 + end. + +Fixpoint penumerate {A : Type} (p : positive) (l : list A) {struct l} + : list (positive * A) := + match l with + | x :: xs => (p, x) :: penumerate (Pos.pred p) xs + | nil => nil + end. + +Fixpoint prange {A: Type} (p1 p2: positive) (l: list A) {struct l} := + match l with + | x :: xs => (p1, p2, x) :: prange p2 (Pos.pred p2) xs + | nil => nil + end. + +Lemma add_data_instr_state_incr : + forall s n st, + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n (Vseq (AssocMapExt.get_default + _ Vskip n s.(st_datapath)) st) s.(st_datapath)) + s.(st_controllogic)). +Proof. + constructor; intros; + try (simpl; destruct (peq n n0); subst); + auto with htlh. +Qed. + +Definition add_data_instr (n : node) (st : stmnt) : mon unit := + fun s => + OK tt (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + (AssocMap.set n (Vseq (AssocMapExt.get_default _ Vskip n s.(st_datapath)) st) s.(st_datapath)) + s.(st_controllogic)) + (add_data_instr_state_incr s n st). + +Lemma add_control_instr_state_incr : + forall s n st, + (st_controllogic s) ! n = None -> + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + s.(st_datapath) + (AssocMap.set n st s.(st_controllogic))). +Proof. + constructor; intros; + try (simpl; destruct (peq n n0); subst); + auto with htlh. +Qed. + +Definition add_control_instr (n : node) (st : stmnt) : mon unit := + fun s => + match check_empty_node_controllogic s n with + | left CTRL => + OK tt (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + s.(st_datapath) + (AssocMap.set n st s.(st_controllogic))) + (add_control_instr_state_incr s n st CTRL) + | _ => + Error (Errors.msg "HTLPargen.add_control_instr: control logic is not empty") + end. + +Definition add_control_instr_force_state_incr : + forall s n st, + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + s.(st_datapath) + (AssocMap.set n st s.(st_controllogic))). +Admitted. + +Definition add_control_instr_force (n : node) (st : stmnt) : mon unit := + fun s => + OK tt (mkstate + s.(st_st) + s.(st_freshreg) + (st_freshstate s) + s.(st_scldecls) + s.(st_arrdecls) + s.(st_datapath) + (AssocMap.set n st s.(st_controllogic))) + (add_control_instr_force_state_incr s n st). + +Fixpoint pred_expr (preg: reg) (p: pred_op) := + match p with + | Plit (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) + | Pand p1 p2 => + Vbinop Vand (pred_expr preg p1) (pred_expr preg p2) + | Por p1 p2 => + Vbinop Vor (pred_expr preg p1) (pred_expr preg p2) + end. + +Definition translate_predicate (a : assignment) + (preg: reg) (p: option pred_op) (dst e: expr) := + match p with + | None => ret (a dst e) + | Some pos => + ret (a dst (Vternary (pred_expr preg pos) e dst)) + end. + +Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) + : mon stmnt := + match i with + | FUnop => + ret Vskip + | FUop p op args dst => + do instr <- translate_instr op args; + do _ <- declare_reg None dst 32; + translate_predicate a preg p (Vvar dst) instr + | FUload p chunk addr args dst => + do src <- translate_arr_access chunk addr args stack; + do _ <- declare_reg None dst 32; + translate_predicate a preg p (Vvar dst) src + | FUstore p chunk addr args src => + do dst <- translate_arr_access chunk addr args stack; + translate_predicate a preg p dst (Vvar src) + | FUsetpred _ c args p => + do cond <- translate_condition c args; + ret (a (pred_expr preg (Plit (true, p))) cond) + end. + +Lemma create_new_state_state_incr: + forall s p, + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + (s.(st_freshstate) + p)%positive + s.(st_scldecls) + s.(st_arrdecls) + s.(st_datapath) + s.(st_controllogic)). +Admitted. + +Definition create_new_state (p: node): mon node := + fun s => OK s.(st_freshstate) + (mkstate + s.(st_st) + s.(st_freshreg) + (s.(st_freshstate) + p)%positive + s.(st_scldecls) + s.(st_arrdecls) + s.(st_datapath) + s.(st_controllogic)) + (create_new_state_state_incr s p). + +Definition translate_inst_list (fin rtrn stack preg: reg) (ni : node * node * list (list instr)) := + match ni with + | (n, p, li) => + do _ <- collectlist + (fun l => + do stmnt <- translate_inst Vblock fin rtrn stack preg n l; + add_data_instr n stmnt) (concat li); + do st <- get; + add_control_instr n (state_goto st.(st_st) p) + end. + +Fixpoint translate_cfi' (fin rtrn stack preg: reg) (cfi: cf_instr) + : mon (stmnt * stmnt) := + match cfi with + | FUgoto n' => + do st <- get; + ret (Vskip, state_goto st.(st_st) n') + | FUcond c args n1 n2 => + do st <- get; + do e <- translate_condition c args; + ret (Vskip, state_cond st.(st_st) e n1 n2) + | FUreturn r => + match r with + | Some r' => + ret ((Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vvar r'))), + Vskip) + | None => + ret ((Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vlit (ZToValue 0%Z)))), + Vskip) + end + | FUpred_cf p c1 c2 => + do (tc1s, tc1c) <- translate_cfi' fin rtrn stack preg c1; + do (tc2s, tc2c) <- translate_cfi' fin rtrn stack preg c2; + ret ((Vcond (pred_expr preg p) tc1s tc2s), (Vcond (pred_expr preg p) tc1c tc2c)) + | FUjumptable r tbl => + do s <- get; + ret (Vskip, Vcase (Vvar r) (list_to_stmnt (tbl_to_case_expr s.(st_st) tbl)) (Some Vskip)) + | FUcall sig ri rl r n => + error (Errors.msg "HTLPargen: RPcall not supported.") + | FUtailcall sig ri lr => + error (Errors.msg "HTLPargen: RPtailcall not supported.") + | FUbuiltin e lb b n => + error (Errors.msg "HTLPargen: RPbuildin not supported.") + end. + +Definition translate_cfi (fin rtrn stack preg: reg) (ni: node * cf_instr) + : mon unit := + let (n, cfi) := ni in + do (s, c) <- translate_cfi' fin rtrn stack preg cfi; + do _ <- add_control_instr n c; + add_data_instr n s. + +Definition transf_bblock (fin rtrn stack preg: reg) (ni : node * bblock) + : mon unit := + let (n, bb) := ni in + do nstate <- create_new_state ((poslength bb.(bb_body)))%positive; + do _ <- collectlist (translate_inst_list fin rtrn stack preg) + (prange n (nstate + poslength bb.(bb_body) - 1)%positive + bb.(bb_body)); + match bb.(bb_body) with + | nil => translate_cfi fin rtrn stack preg (n, bb.(bb_exit)) + | _ => translate_cfi fin rtrn stack preg (nstate, bb.(bb_exit)) + end. + +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; + do (stack, stack_len) <- create_arr None 32 + (Z.to_nat (f.(fn_stacksize) / 4)); + do preg <- create_reg None 32; + do _ <- collectlist (transf_bblock fin rtrn stack preg) + (Maps.PTree.elements f.(fn_code)); + do _ <- collectlist (fun r => declare_reg (Some Vinput) r 32) + f.(fn_params); + do start <- create_reg (Some Vinput) 1; + do rst <- create_reg (Some Vinput) 1; + 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, + decide_order (st_st current_state) fin rtrn stack start rst clk, + max_list_dec (fn_params f) (st_st current_state) + with + | left LEDATA, left LECTRL, left MORD, left WFPARAMS => + ret (HTL.mkmodule + f.(fn_params) + current_state.(st_datapath) + current_state.(st_controllogic) + f.(fn_entrypoint) + current_state.(st_st) + stack + stack_len + fin + rtrn + start + rst + clk + current_state.(st_scldecls) + current_state.(st_arrdecls) + None + (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA)) + MORD + _ + WFPARAMS) + | _, _, _, _ => error (Errors.msg "More than 2^32 states.") + end + 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)) + (AssocMap.set st (None, VScalar 32) (st_scldecls (init_state st))) + (st_arrdecls (init_state st)) + (st_datapath (init_state st)) + (st_controllogic (init_state st)). + +Definition transl_module (f : function) : Errors.res HTL.module := + run_mon (max_state f) (transf_module f). + +Definition transl_fundef := transf_partial_fundef transl_module. + +Definition main_is_internal (p : RTLParFU.program) : bool := + let ge := Globalenvs.Genv.globalenv p in + match Globalenvs.Genv.find_symbol ge p.(AST.prog_main) with + | Some b => + match Globalenvs.Genv.find_funct_ptr ge b with + | Some (AST.Internal _) => true + | _ => false + end + | _ => false + end. + +Definition transl_program (p : RTLParFU.program) : Errors.res HTL.program := + if main_is_internal p + then transform_partial_program transl_fundef p + else Errors.Error (Errors.msg "Main function is not Internal."). diff --git a/src/hls/RTLParFU.v b/src/hls/RTLParFU.v new file mode 100644 index 0000000..3442ff0 --- /dev/null +++ b/src/hls/RTLParFU.v @@ -0,0 +1,402 @@ +(* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2020-2021 Yann Herklotz + * + * This program is free software: you can redistribute it and/or modify + * it under the terms of the GNU General Public License as published by + * the Free Software Foundation, either version 3 of the License, or + * (at your option) any later version. + * + * This program is distributed in the hope that it will be useful, + * but WITHOUT ANY WARRANTY; without even the implied warranty of + * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + * GNU General Public License for more details. + * + * You should have received a copy of the GNU General Public License + * along with this program. If not, see . + *) + +Require Import compcert.backend.Registers. +Require Import compcert.common.AST. +Require Import compcert.common.Events. +Require Import compcert.common.Globalenvs. +Require Import compcert.common.Memory. +Require Import compcert.common.Smallstep. +Require Import compcert.common.Values. +Require Import compcert.lib.Coqlib. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. +Require Import compcert.verilog.Op. + +Require Import vericert.hls.FunctionalUnits. +Require Import Predicate. +Require Import Vericertlib. + +Definition node := positive. + +Inductive instr : Type := +| FUnop : instr +| FUop : option pred_op -> operation -> list reg -> reg -> instr +| FUload : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr +| FUstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr +| FUsetpred : option pred_op -> condition -> list reg -> predicate -> instr. + +Inductive cf_instr : Type := +| FUcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr +| FUtailcall : signature -> reg + ident -> list reg -> cf_instr +| FUbuiltin : external_function -> list (builtin_arg reg) -> + builtin_res reg -> node -> cf_instr +| FUcond : condition -> list reg -> node -> node -> cf_instr +| FUjumptable : reg -> list node -> cf_instr +| FUreturn : option reg -> cf_instr +| FUgoto : node -> cf_instr +| FUpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. + +Fixpoint successors_instr (i : cf_instr) : list node := + match i with + | FUcall sig ros args res s => s :: nil + | FUtailcall sig ros args => nil + | FUbuiltin ef args res s => s :: nil + | FUcond cond args ifso ifnot => ifso :: ifnot :: nil + | FUjumptable arg tbl => tbl + | FUreturn optarg => nil + | FUgoto n => n :: nil + | FUpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil) + end. + +Definition max_reg_instr (m: positive) (i: instr) := + match i with + | FUnop => m + | FUop p op args res => + fold_left Pos.max args (Pos.max res m) + | FUload p chunk addr args dst => + fold_left Pos.max args (Pos.max dst m) + | FUstore p chunk addr args src => + fold_left Pos.max args (Pos.max src m) + | FUsetpred p' c args p => + fold_left Pos.max args m + end. + +Fixpoint max_reg_cfi (m : positive) (i : cf_instr) := + match i with + | FUcall sig (inl r) args res s => + fold_left Pos.max args (Pos.max r (Pos.max res m)) + | FUcall sig (inr id) args res s => + fold_left Pos.max args (Pos.max res m) + | FUtailcall sig (inl r) args => + fold_left Pos.max args (Pos.max r m) + | FUtailcall sig (inr id) args => + fold_left Pos.max args m + | FUbuiltin ef args res s => + fold_left Pos.max (params_of_builtin_args args) + (fold_left Pos.max (params_of_builtin_res res) m) + | FUcond cond args ifso ifnot => fold_left Pos.max args m + | FUjumptable arg tbl => Pos.max arg m + | FUreturn None => m + | FUreturn (Some arg) => Pos.max arg m + | FUgoto n => m + | FUpred_cf p c1 c2 => Pos.max (max_reg_cfi m c1) (max_reg_cfi m c2) + end. + +Definition regset := Regmap.t val. +Definition predset := PMap.t bool. + +Definition eval_predf (pr: predset) (p: pred_op) := + sat_predicate p (fun x => pr !! (Pos.of_nat x)). + +#[global] + Instance eval_predf_Proper : Proper (eq ==> equiv ==> eq) eval_predf. +Proof. + unfold Proper. simplify. unfold "==>". + intros. + unfold sat_equiv in *. intros. unfold eval_predf. subst. apply H0. +Qed. + +#[local] Open Scope pred_op. + +Lemma eval_predf_Pand : + forall ps p p', + eval_predf ps (p ∧ p') = eval_predf ps p && eval_predf ps p'. +Proof. unfold eval_predf; split; simplify; auto with bool. Qed. + +Lemma eval_predf_Por : + forall ps p p', + eval_predf ps (p ∨ p') = eval_predf ps p || eval_predf ps p'. +Proof. unfold eval_predf; split; simplify; auto with bool. Qed. + +Lemma eval_predf_pr_equiv : + forall p ps ps', + (forall x, ps !! x = ps' !! x) -> + eval_predf ps p = eval_predf ps' p. +Proof. + induction p; simplify; auto; + try (unfold eval_predf; simplify; repeat (destruct_match; []); inv Heqp0; rewrite <- H; auto); + [repeat rewrite eval_predf_Pand|repeat rewrite eval_predf_Por]; + erewrite IHp1; try eassumption; erewrite IHp2; eauto. +Qed. + +Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset := + match rl, vl with + | r1 :: rs, v1 :: vs => Regmap.set r1 v1 (init_regs vs rs) + | _, _ => Regmap.init Vundef + end. + +Definition bblock_body := list (list (list instr)). + +Record bblock : Type := + mk_bblock { + bb_body: bblock_body; + bb_exit: cf_instr + }. + +Definition code: Type := PTree.t bblock. + +Record function: Type := + mkfunction { + fn_sig: signature; + fn_params: list reg; + fn_stacksize: Z; + fn_code: code; + fn_funct_units: funct_unit; + fn_entrypoint: node; + }. + +Definition fundef := AST.fundef function. +Definition program := AST.program fundef unit. + +Definition funsig (fd: fundef) := + match fd with + | Internal f => fn_sig f + | External ef => ef_sig ef + end. + +Inductive stackframe : Type := +| Stackframe: + forall (res: reg) (**r where to store the result *) + (f: function) (**r calling function *) + (sp: val) (**r stack pointer in calling function *) + (pc: node) (**r program point in calling function *) + (rs: regset) (**r register state in calling function *) + (pr: predset), (**r predicate state of the calling function *) + stackframe. + +Inductive state : Type := +| State: + forall (stack: list stackframe) (**r call stack *) + (f: function) (**r current function *) + (sp: val) (**r stack pointer *) + (pc: node) (**r current program point in [c] *) + (rs: regset) (**r register state *) + (pr: predset) (**r predicate register state *) + (m: mem), (**r memory state *) + state +| Callstate: + forall (stack: list stackframe) (**r call stack *) + (f: fundef) (**r function to call *) + (args: list val) (**r arguments to the call *) + (m: mem), (**r memory state *) + state +| Returnstate: + forall (stack: list stackframe) (**r call stack *) + (v: val) (**r return value for the call *) + (m: mem), (**r memory state *) + state. + +Record instr_state := mk_instr_state { + is_rs: regset; + is_ps: predset; + is_mem: mem; + }. + +Definition genv := Genv.t fundef unit. + +Section RELSEM. + + Context (ge: genv). + + Definition find_function + (ros: reg + ident) (rs: regset) : option fundef := + match ros with + | inl r => Genv.find_funct ge rs#r + | inr symb => + match Genv.find_symbol ge symb with + | None => None + | Some b => Genv.find_funct_ptr ge b + end + end. + + Inductive eval_pred: option pred_op -> instr_state -> instr_state -> instr_state -> Prop := + | eval_pred_true: + forall i i' p, + eval_predf (is_ps i) p = true -> + eval_pred (Some p) i i' i' + | eval_pred_false: + forall i i' p, + eval_predf (is_ps i) p = false -> + eval_pred (Some p) i i' i + | eval_pred_none: + forall i i', eval_pred None i i' i. + + Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := + | exec_FUnop: + forall sp ist, + step_instr sp ist FUnop ist + | exec_FUop: + forall op v res args rs m sp p ist pr, + eval_operation ge sp op rs##args m = Some v -> + eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#res <- v) pr m) ist -> + step_instr sp (mk_instr_state rs pr m) (FUop p op args res) ist + | exec_FUload: + forall addr rs args a chunk m v dst sp p pr ist, + eval_addressing ge sp addr rs##args = Some a -> + Mem.loadv chunk m a = Some v -> + eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#dst <- v) pr m) ist -> + step_instr sp (mk_instr_state rs pr m) (FUload p chunk addr args dst) ist + | exec_FUstore: + forall addr rs args a chunk m src m' sp p pr ist, + eval_addressing ge sp addr rs##args = Some a -> + Mem.storev chunk m a rs#src = Some m' -> + eval_pred p (mk_instr_state rs pr m) (mk_instr_state rs pr m') ist -> + step_instr sp (mk_instr_state rs pr m) (FUstore p chunk addr args src) ist + | exec_FUsetpred: + forall sp rs pr m p c b args p' ist, + Op.eval_condition c rs##args m = Some b -> + eval_pred p' (mk_instr_state rs pr m) (mk_instr_state rs (pr#p <- b) m) ist -> + step_instr sp (mk_instr_state rs pr m) (FUsetpred p' c args p) ist. + + Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop := + | exec_RBcons: + forall state i state' state'' instrs sp, + step_instr sp state i state' -> + step_instr_list sp state' instrs state'' -> + step_instr_list sp state (i :: instrs) state'' + | exec_RBnil: + forall state sp, + step_instr_list sp state nil state. + + Inductive step_instr_seq (sp : val) + : instr_state -> list (list instr) -> instr_state -> Prop := + | exec_instr_seq_cons: + forall state i state' state'' instrs, + step_instr_list sp state i state' -> + step_instr_seq sp state' instrs state'' -> + step_instr_seq sp state (i :: instrs) state'' + | exec_instr_seq_nil: + forall state, + step_instr_seq sp state nil state. + + Inductive step_instr_block (sp : val) + : instr_state -> bblock_body -> instr_state -> Prop := + | exec_instr_block_cons: + forall state i state' state'' instrs, + step_instr_seq sp state i state' -> + step_instr_block sp state' instrs state'' -> + step_instr_block sp state (i :: instrs) state'' + | exec_instr_block_nil: + forall state, + step_instr_block sp state nil state. + + Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop := + | exec_FUcall: + forall s f sp rs m res fd ros sig args pc pc' pr, + find_function ros rs = Some fd -> + funsig fd = sig -> + step_cf_instr (State s f sp pc rs pr m) (FUcall sig ros args res pc') + E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m) + | exec_FUtailcall: + forall s f stk rs m sig ros args fd m' pc pr, + find_function ros rs = Some fd -> + funsig fd = sig -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (FUtailcall sig ros args) + E0 (Callstate s fd rs##args m') + | exec_FUbuiltin: + forall s f sp rs m ef args res pc' vargs t vres m' pc pr, + eval_builtin_args ge (fun r => rs#r) sp m args vargs -> + external_call ef ge vargs m t vres m' -> + step_cf_instr (State s f sp pc rs pr m) (FUbuiltin ef args res pc') + t (State s f sp pc' (regmap_setres res vres rs) pr m') + | exec_FUcond: + forall s f sp rs m cond args ifso ifnot b pc pc' pr, + eval_condition cond rs##args m = Some b -> + pc' = (if b then ifso else ifnot) -> + step_cf_instr (State s f sp pc rs pr m) (FUcond cond args ifso ifnot) + E0 (State s f sp pc' rs pr m) + | exec_FUjumptable: + forall s f sp rs m arg tbl n pc pc' pr, + rs#arg = Vint n -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> + step_cf_instr (State s f sp pc rs pr m) (FUjumptable arg tbl) + E0 (State s f sp pc' rs pr m) + | exec_FUreturn: + forall s f stk rs m or pc m' pr, + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (FUreturn or) + E0 (Returnstate s (regmap_optget or Vundef rs) m') + | exec_FUgoto: + forall s f sp pc rs pr m pc', + step_cf_instr (State s f sp pc rs pr m) (FUgoto pc') E0 (State s f sp pc' rs pr m) + | exec_FUpred_cf: + forall s f sp pc rs pr m cf1 cf2 st' p t, + step_cf_instr (State s f sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> + step_cf_instr (State s f sp pc rs pr m) (FUpred_cf p cf1 cf2) t st'. + + Inductive step: state -> trace -> state -> Prop := + | exec_bblock: + forall s f sp pc rs rs' m m' t s' bb pr pr', + f.(fn_code)!pc = Some bb -> + step_instr_block sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') -> + step_cf_instr (State s f sp pc rs' pr' m') bb.(bb_exit) t s' -> + step (State s f sp pc rs pr m) t s' + | exec_function_internal: + forall s f args m m' stk, + Mem.alloc m 0 f.(fn_stacksize) = (m', stk) -> + step (Callstate s (Internal f) args m) + E0 (State s + f + (Vptr stk Ptrofs.zero) + f.(fn_entrypoint) + (init_regs args f.(fn_params)) + (PMap.init false) m') + | exec_function_external: + forall s ef args res t m m', + external_call ef ge args m t res m' -> + step (Callstate s (External ef) args m) + t (Returnstate s res m') + | exec_return: + forall res f sp pc rs s vres m pr, + step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m) + E0 (State s f sp pc (rs#res <- vres) pr m). + +End RELSEM. + +Inductive initial_state (p: program): state -> Prop := +| initial_state_intro: forall b f m0, + let ge := Genv.globalenv p in + Genv.init_mem p = Some m0 -> + Genv.find_symbol ge p.(prog_main) = Some b -> + Genv.find_funct_ptr ge b = Some f -> + funsig f = signature_main -> + initial_state p (Callstate nil f nil m0). + +Inductive final_state: state -> int -> Prop := +| final_state_intro: forall r m, + final_state (Returnstate nil (Vint r) m) r. + +Definition semantics (p: program) := + Semantics step (initial_state p) final_state (Genv.globalenv p). + +Definition max_reg_bblock (m : positive) (pc : node) (bb : bblock) := + let max_body := fold_left (fun x l => fold_left (fun x' l' => fold_left max_reg_instr l' x') l x) bb.(bb_body) m in + max_reg_cfi max_body bb.(bb_exit). + +Definition max_reg_function (f: function) := + Pos.max + (PTree.fold max_reg_bblock f.(fn_code) 1%positive) + (fold_left Pos.max f.(fn_params) 1%positive). + +Definition max_pc_function (f: function) : positive := + PTree.fold (fun m pc i => (Pos.max m + (pc + match Zlength i.(bb_body) + with Z.pos p => p | _ => 1 end))%positive) + f.(fn_code) 1%positive. diff --git a/src/hls/RTLParFUgen.v b/src/hls/RTLParFUgen.v new file mode 100644 index 0000000..389d76c --- /dev/null +++ b/src/hls/RTLParFUgen.v @@ -0,0 +1,21 @@ +Require Import Coq.micromega.Lia. + +Require Import compcert.common.AST. +Require compcert.common.Errors. +Require compcert.common.Globalenvs. +Require compcert.lib.Integers. +Require Import compcert.lib.Maps. + +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.RTLParFU. +Require Import vericert.hls.ValueInt. +Require Import vericert.hls.Verilog. +Require Import vericert.hls.RTLBlockInstr. +Require Import vericert.hls.RTLPar. + +Definition transl_program (p : RTLParFU.program) : Errors.res HTL.program := + transform_partial_program transl_fundef p. -- cgit From 9afc00b337eb4eb704bf95d642073135f3345dca Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 14 Nov 2021 15:58:08 +0000 Subject: Add the tseytin transformation instead --- src/hls/Abstr.v | 9 - src/hls/Predicate.v | 569 +++++++++++++++++++++++++++++++++------------------- 2 files changed, 359 insertions(+), 219 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v index abc4181..2ab79cf 100644 --- a/src/hls/Abstr.v +++ b/src/hls/Abstr.v @@ -630,15 +630,6 @@ Definition encode_expression max pe h := (Pand (Por (Pnot p) (Pvar p')) p'', h'') end.*) -Fixpoint max_predicate (p: pred_op) : positive := - match p with - | Plit (b, p) => p - | Ptrue => 1 - | Pfalse => 1 - | Pand a b => Pos.max (max_predicate a) (max_predicate b) - | Por a b => Pos.max (max_predicate a) (max_predicate b) - end. - Fixpoint max_pred_expr (pe: pred_expr) : positive := match pe with | NE.singleton (p, e) => max_predicate p diff --git a/src/hls/Predicate.v b/src/hls/Predicate.v index c393a64..9f9ec65 100644 --- a/src/hls/Predicate.v +++ b/src/hls/Predicate.v @@ -24,6 +24,8 @@ Notation "A ∨ B" := (Por A B) (at level 25) : pred_op. Notation "⟂" := (Pfalse) : pred_op. Notation "'T'" := (Ptrue) : pred_op. +#[local] Open Scope pred_op. + Fixpoint sat_predicate (p: pred_op) (a: asgn) : bool := match p with | Plit (b, p') => if b then a (Pos.to_nat p') else negb (a (Pos.to_nat p')) @@ -33,160 +35,57 @@ Fixpoint sat_predicate (p: pred_op) (a: asgn) : bool := | Por p1 p2 => sat_predicate p1 a || sat_predicate p2 a end. -Fixpoint mult {A: Type} (a b: list (list A)) : list (list A) := - match a with - | nil => nil - | l :: ls => mult ls b ++ (List.map (fun x => l ++ x) b) - end. - -Lemma satFormula_concat: - forall a b agn, - satFormula a agn -> - satFormula b agn -> - satFormula (a ++ b) agn. -Proof. induction a; crush. Qed. +Definition sat_equiv p1 p2 := forall c, sat_predicate p1 c = sat_predicate p2 c. -Lemma satFormula_concat2: - forall a b agn, - satFormula (a ++ b) agn -> - satFormula a agn /\ satFormula b agn. -Proof. - induction a; simplify; - try apply IHa in H1; crush. -Qed. +Lemma equiv_symm : forall a b, sat_equiv a b -> sat_equiv b a. +Proof. crush. Qed. -Lemma satClause_concat: - forall a a1 a0, - satClause a a1 -> - satClause (a0 ++ a) a1. -Proof. induction a0; crush. Qed. +Lemma equiv_trans : forall a b c, sat_equiv a b -> sat_equiv b c -> sat_equiv a c. +Proof. crush. Qed. -Lemma satClause_concat2: - forall a a1 a0, - satClause a0 a1 -> - satClause (a0 ++ a) a1. -Proof. - induction a0; crush. - inv H; crush. -Qed. +Lemma equiv_refl : forall a, sat_equiv a a. +Proof. crush. Qed. -Lemma satClause_concat3: - forall a b c, - satClause (a ++ b) c -> - satClause a c \/ satClause b c. -Proof. - induction a; crush. - inv H; crush. - apply IHa in H0; crush. - inv H0; crush. -Qed. +#[global] +Instance Equivalence_SAT : Equivalence sat_equiv := + { Equivalence_Reflexive := equiv_refl ; + Equivalence_Symmetric := equiv_symm ; + Equivalence_Transitive := equiv_trans ; + }. -Lemma satFormula_mult': - forall p2 a a0, - satFormula p2 a0 \/ satClause a a0 -> - satFormula (map (fun x : list lit => a ++ x) p2) a0. -Proof. - induction p2; crush. - - inv H. inv H0. apply satClause_concat. auto. - apply satClause_concat2; auto. - - apply IHp2. - inv H; crush; inv H0; crush. -Qed. +#[global] +Instance SATSetoid : Setoid pred_op := + { equiv := sat_equiv; }. -Lemma satFormula_mult2': - forall p2 a a0, - satFormula (map (fun x : list lit => a ++ x) p2) a0 -> - satClause a a0 \/ satFormula p2 a0. +#[global] +Instance PandProper : Proper (equiv ==> equiv ==> equiv) Pand. Proof. - induction p2; crush. - apply IHp2 in H1. inv H1; crush. - apply satClause_concat3 in H0. - inv H0; crush. + unfold Proper. simplify. unfold "==>". + intros. + unfold sat_equiv in *. intros. + simplify. rewrite H0. rewrite H. + auto. Qed. -Lemma satFormula_mult: - forall p1 p2 a, - satFormula p1 a \/ satFormula p2 a -> - satFormula (mult p1 p2) a. +#[global] +Instance PorProper : Proper (equiv ==> equiv ==> equiv) Por. Proof. - induction p1; crush. - apply satFormula_concat; crush. - inv H. inv H0. - apply IHp1. auto. - apply IHp1. auto. - apply satFormula_mult'; - inv H; crush. + unfold Proper, "==>". simplify. + intros. + unfold sat_equiv in *. intros. + simplify. rewrite H0. rewrite H. + auto. Qed. -Lemma satFormula_mult2: - forall p1 p2 a, - satFormula (mult p1 p2) a -> - satFormula p1 a \/ satFormula p2 a. +#[global] +Instance sat_predicate_Proper : Proper (equiv ==> eq ==> eq) sat_predicate. Proof. - induction p1; crush. - apply satFormula_concat2 in H; crush. - apply IHp1 in H0. - inv H0; crush. - apply satFormula_mult2' in H1. inv H1; crush. + unfold Proper, "==>". simplify. + intros. + unfold sat_equiv in *. subst. + apply H. Qed. -Fixpoint trans_pred (p: pred_op) : - {fm: formula | forall a, - sat_predicate p a = true <-> satFormula fm a}. - refine - (match p with - | Plit (b, p') => exist _ (((b, Pos.to_nat p') :: nil) :: nil) _ - | Ptrue => exist _ nil _ - | Pfalse => exist _ (nil::nil) _ - | Pand p1 p2 => - match trans_pred p1, trans_pred p2 with - | exist _ p1' _, exist _ p2' _ => exist _ (p1' ++ p2') _ - end - | Por p1 p2 => - match trans_pred p1, trans_pred p2 with - | exist _ p1' _, exist _ p2' _ => exist _ (mult p1' p2') _ - end - end); split; intros; simpl in *; auto; try solve [crush]. - - destruct b; auto. apply negb_true_iff in H. auto. - - destruct b. inv H. inv H0; auto. apply negb_true_iff. inv H. inv H0; eauto. contradiction. - - apply satFormula_concat. - apply andb_prop in H. inv H. apply i in H0. auto. - apply andb_prop in H. inv H. apply i0 in H1. auto. - - apply satFormula_concat2 in H. simplify. apply andb_true_intro. - split. apply i in H0. auto. - apply i0 in H1. auto. - - apply orb_prop in H. inv H; apply satFormula_mult. apply i in H0. auto. - apply i0 in H0. auto. - - apply orb_true_intro. - apply satFormula_mult2 in H. inv H. apply i in H0. auto. - apply i0 in H0. auto. -Defined. - -Definition sat_pred (p: pred_op) : - ({al : alist | sat_predicate p (interp_alist al) = true} - + {forall a : asgn, sat_predicate p a = false}). - refine - ( match trans_pred p with - | exist _ fm _ => - match satSolve fm with - | inleft (exist _ a _) => inleft (exist _ a _) - | inright _ => inright _ - end - end ). - - apply i in s0. auto. - - intros. specialize (n a). specialize (i a). - destruct (sat_predicate p a). exfalso. - apply n. apply i. auto. auto. -Defined. - -Definition sat_pred_simple (p: pred_op) := - match sat_pred p with - | inleft (exist _ al _) => Some al - | inright _ => None - end. - -#[local] Open Scope pred_op. - Fixpoint negate (p: pred_op) := match p with | Plit (b, pr) => Plit (negb b, pr) @@ -243,66 +142,6 @@ Qed. Lemma unsat_commut a b: unsat (a ∧ b) -> unsat (b ∧ a). Proof. unfold unsat; simplify; eauto with bool. Qed. -Lemma sat_dec a: {sat a} + {unsat a}. -Proof. - unfold sat, unsat. - destruct (sat_pred a). - intros. left. destruct s. - exists (Sat.interp_alist x). auto. - intros. tauto. -Qed. - -Definition sat_equiv p1 p2 := forall c, sat_predicate p1 c = sat_predicate p2 c. - -Lemma equiv_symm : forall a b, sat_equiv a b -> sat_equiv b a. -Proof. crush. Qed. - -Lemma equiv_trans : forall a b c, sat_equiv a b -> sat_equiv b c -> sat_equiv a c. -Proof. crush. Qed. - -Lemma equiv_refl : forall a, sat_equiv a a. -Proof. crush. Qed. - -#[global] -Instance Equivalence_SAT : Equivalence sat_equiv := - { Equivalence_Reflexive := equiv_refl ; - Equivalence_Symmetric := equiv_symm ; - Equivalence_Transitive := equiv_trans ; - }. - -#[global] -Instance SATSetoid : Setoid pred_op := - { equiv := sat_equiv; }. - -#[global] -Instance PandProper : Proper (equiv ==> equiv ==> equiv) Pand. -Proof. - unfold Proper. simplify. unfold "==>". - intros. - unfold sat_equiv in *. intros. - simplify. rewrite H0. rewrite H. - auto. -Qed. - -#[global] -Instance PorProper : Proper (equiv ==> equiv ==> equiv) Por. -Proof. - unfold Proper, "==>". simplify. - intros. - unfold sat_equiv in *. intros. - simplify. rewrite H0. rewrite H. - auto. -Qed. - -#[global] -Instance sat_predicate_Proper : Proper (equiv ==> eq ==> eq) sat_predicate. -Proof. - unfold Proper, "==>". simplify. - intros. - unfold sat_equiv in *. subst. - apply H. -Qed. - Lemma sat_imp_equiv : forall a b, unsat (a ∧ ¬ b ∨ ¬ a ∧ b) -> a == b. @@ -360,14 +199,6 @@ Qed. Definition simplify' (p: pred_op) := match p with - | (Plit (b1, a)) ∧ (Plit (b2, b)) as p' => - if Pos.eqb a b then - if negb (xorb b1 b2) then Plit (b1, a) else ⟂ - else p' - | (Plit (b1, a)) ∨ (Plit (b2, b)) as p' => - if Pos.eqb a b then - if negb (xorb b1 b2) then Plit (b1, a) else T - else p' | A ∧ T => A | T ∧ A => A | _ ∧ ⟂ => ⟂ @@ -403,9 +234,9 @@ Lemma simplify'_correct : forall h a, sat_predicate (simplify' h) a = sat_predicate h a. Proof. - (*destruct h; crush; repeat destruct_match; crush; + destruct h; crush; repeat destruct_match; crush; solve [rewrite andb_true_r; auto | rewrite orb_false_r; auto]. -Qed.*) Admitted. +Qed. Lemma simplify_correct : forall h a, @@ -424,12 +255,329 @@ Proof. Local Transparent simplify'. Qed. +Fixpoint mult {A: Type} (a b: list (list A)) : list (list A) := + match a with + | nil => nil + | l :: ls => mult ls b ++ (List.map (fun x => l ++ x) b) + end. + +Lemma satFormula_concat: + forall a b agn, + satFormula a agn -> + satFormula b agn -> + satFormula (a ++ b) agn. +Proof. induction a; crush. Qed. + +Lemma satFormula_concat2: + forall a b agn, + satFormula (a ++ b) agn -> + satFormula a agn /\ satFormula b agn. +Proof. + induction a; simplify; + try apply IHa in H1; crush. +Qed. + +Lemma satClause_concat: + forall a a1 a0, + satClause a a1 -> + satClause (a0 ++ a) a1. +Proof. induction a0; crush. Qed. + +Lemma satClause_concat2: + forall a a1 a0, + satClause a0 a1 -> + satClause (a0 ++ a) a1. +Proof. + induction a0; crush. + inv H; crush. +Qed. + +Lemma satClause_concat3: + forall a b c, + satClause (a ++ b) c -> + satClause a c \/ satClause b c. +Proof. + induction a; crush. + inv H; crush. + apply IHa in H0; crush. + inv H0; crush. +Qed. + +Lemma satFormula_mult': + forall p2 a a0, + satFormula p2 a0 \/ satClause a a0 -> + satFormula (map (fun x : list lit => a ++ x) p2) a0. +Proof. + induction p2; crush. + - inv H. inv H0. apply satClause_concat. auto. + apply satClause_concat2; auto. + - apply IHp2. + inv H; crush; inv H0; crush. +Qed. + +Lemma satFormula_mult2': + forall p2 a a0, + satFormula (map (fun x : list lit => a ++ x) p2) a0 -> + satClause a a0 \/ satFormula p2 a0. +Proof. + induction p2; crush. + apply IHp2 in H1. inv H1; crush. + apply satClause_concat3 in H0. + inv H0; crush. +Qed. + +Lemma satFormula_mult: + forall p1 p2 a, + satFormula p1 a \/ satFormula p2 a -> + satFormula (mult p1 p2) a. +Proof. + induction p1; crush. + apply satFormula_concat; crush. + inv H. inv H0. + apply IHp1. auto. + apply IHp1. auto. + apply satFormula_mult'; + inv H; crush. +Qed. + +Lemma satFormula_mult2: + forall p1 p2 a, + satFormula (mult p1 p2) a -> + satFormula p1 a \/ satFormula p2 a. +Proof. + induction p1; crush. + apply satFormula_concat2 in H; crush. + apply IHp1 in H0. + inv H0; crush. + apply satFormula_mult2' in H1. inv H1; crush. +Qed. + +Fixpoint trans_pred (p: pred_op) : + {fm: formula | forall a, + sat_predicate p a = true <-> satFormula fm a}. + refine + (match p with + | Plit (b, p') => exist _ (((b, Pos.to_nat p') :: nil) :: nil) _ + | Ptrue => exist _ nil _ + | Pfalse => exist _ (nil::nil) _ + | Pand p1 p2 => + match trans_pred p1, trans_pred p2 with + | exist _ p1' _, exist _ p2' _ => exist _ (p1' ++ p2') _ + end + | Por p1 p2 => + match trans_pred p1, trans_pred p2 with + | exist _ p1' _, exist _ p2' _ => exist _ (mult p1' p2') _ + end + end); split; intros; simpl in *; auto; try solve [crush]. + - destruct b; auto. apply negb_true_iff in H. auto. + - destruct b. inv H. inv H0; auto. apply negb_true_iff. inv H. inv H0; eauto. contradiction. + - apply satFormula_concat. + apply andb_prop in H. inv H. apply i in H0. auto. + apply andb_prop in H. inv H. apply i0 in H1. auto. + - apply satFormula_concat2 in H. simplify. apply andb_true_intro. + split. apply i in H0. auto. + apply i0 in H1. auto. + - apply orb_prop in H. inv H; apply satFormula_mult. apply i in H0. auto. + apply i0 in H0. auto. + - apply orb_true_intro. + apply satFormula_mult2 in H. inv H. apply i in H0. auto. + apply i0 in H0. auto. +Defined. + +Definition bar (p1: lit): lit := (negb (fst p1), (snd p1)). + +Definition stseytin_or (cur p1 p2: lit) : formula := + (bar cur :: p1 :: p2 :: nil) + :: (cur :: bar p1 :: nil) + :: (cur :: bar p2 :: nil) :: nil. + +Definition stseytin_and (cur p1 p2: lit) : formula := + (cur :: bar p1 :: bar p2 :: nil) + :: (bar cur :: p1 :: nil) + :: (bar cur :: p2 :: nil) :: nil. + +Fixpoint xtseytin (next: nat) (p: pred_op) {struct p} : (nat * lit * formula) := + match p with + | Plit (b, p') => (next, (b, Pos.to_nat p'), nil) + | Ptrue => + ((next+1)%nat, (true, next), ((true, next)::nil)::nil) + | Pfalse => + ((next+1)%nat, (true, next), ((false, next)::nil)::nil) + | Por p1 p2 => + let '(m1, n1, f1) := xtseytin next p1 in + let '(m2, n2, f2) := xtseytin m1 p2 in + ((m2+1)%nat, (true, m2), stseytin_or (true, m2) n1 n2 ++ f1 ++ f2) + | Pand p1 p2 => + let '(m1, n1, f1) := xtseytin next p1 in + let '(m2, n2, f2) := xtseytin m1 p2 in + ((m2+1)%nat, (true, m2), stseytin_and (true, m2) n1 n2 ++ f1 ++ f2) + end. + +Lemma stseytin_and_correct : + forall cur p1 p2 fm c, + stseytin_and cur p1 p2 = fm -> + satLit cur c -> + satLit p1 c /\ satLit p2 c -> + satFormula fm c. +Proof. + intros. + unfold stseytin_and in *. rewrite <- H. + unfold satLit in *. destruct p1. destruct p2. destruct cur. + simpl in *|-. cbn. unfold satLit. cbn. crush. +Qed. + +Lemma stseytin_and_correct2 : + forall cur p1 p2 fm c, + stseytin_and cur p1 p2 = fm -> + satFormula fm c -> + satLit cur c <-> satLit p1 c /\ satLit p2 c. +Proof. + intros. split. intros. inv H1. unfold stseytin_and in *. + inv H0; try contradiction. Admitted. + +Lemma stseytin_or_correct : + forall cur p1 p2 fm c, + stseytin_or cur p1 p2 = fm -> + satLit cur c -> + satLit p1 c \/ satLit p2 c -> + satFormula fm c. +Proof. + intros. + unfold stseytin_or in *. rewrite <- H. inv H1. + unfold satLit in *. destruct p1. destruct p2. destruct cur. + simpl in *|-. cbn. unfold satLit. cbn. crush. + unfold satLit in *. destruct p1. destruct p2. destruct cur. + simpl in *|-. cbn. unfold satLit. cbn. crush. +Qed. + +Lemma stseytin_or_correct2 : + forall cur p1 p2 fm c, + stseytin_or cur p1 p2 = fm -> + satFormula fm c -> + satLit cur c <-> satLit p1 c \/ satLit p2 c. +Proof. Admitted. + +Lemma xtseytin_correct : + forall p next l n fm c, + xtseytin next p = (n, l, fm) -> + sat_predicate p c = true <-> satFormula ((l::nil)::fm) c. +Proof. + induction p. + - intros. simplify. destruct p. + inv H. split. + intros. destruct b. split; crush. + apply negb_true_iff in H. + split; crush. + intros. inv H. inv H0; try contradiction. + inv H. simplify. rewrite <- H0. + destruct b. + rewrite -> H0; auto. + rewrite -> H0; auto. + - admit. + - admit. + - intros. split. intros. simpl in H0. + apply andb_prop in H0. inv H0. + cbn in H. + repeat destruct_match; try discriminate; []. inv H. eapply IHp1 in Heqp. + eapply IHp2 in Heqp1. apply Heqp1 in H2. + apply Heqp in H1. inv H1. inv H2. + assert + (satFormula + (((true, n1) :: bar l0 :: bar l1 :: nil) + :: (bar (true, n1) :: l0 :: nil) + :: (bar (true, n1) :: l1 :: nil) :: nil) c). + eapply stseytin_and_correct. unfold stseytin_and. eauto. + unfold satLit. simpl. admit. + inv H; try contradiction. inv H1; try contradiction. eauto. +Admitted. + +Fixpoint max_predicate (p: pred_op) : positive := + match p with + | Plit (b, p) => p + | Ptrue => 1 + | Pfalse => 1 + | Pand a b => Pos.max (max_predicate a) (max_predicate b) + | Por a b => Pos.max (max_predicate a) (max_predicate b) + end. + +Definition tseytin (p: pred_op) : + {fm: formula | forall a, + sat_predicate p a = true <-> satFormula fm a}. + refine ( + (match xtseytin (Pos.to_nat (max_predicate p + 1)) p as X + return xtseytin (Pos.to_nat (max_predicate p + 1)) p = X -> + {fm: formula | forall a, sat_predicate p a = true <-> satFormula fm a} + with (m, n, fm) => fun H => exist _ ((n::nil) :: fm) _ + end) (eq_refl (xtseytin (Pos.to_nat (max_predicate p + 1)) p))). + intros. eapply xtseytin_correct; eauto. Defined. + +Definition tseytin_simple (p: pred_op) : formula := + let m := Pos.to_nat (max_predicate p + 1) in + let '(m, n, fm) := xtseytin m p in + (n::nil) :: fm. + +Definition sat_pred_tseytin (p: pred_op) : + ({al : alist | sat_predicate p (interp_alist al) = true} + + {forall a : asgn, sat_predicate p a = false}). + refine + ( match tseytin p with + | exist _ fm _ => + match satSolve fm with + | inleft (exist _ a _) => inleft (exist _ a _) + | inright _ => inright _ + end + end ). + - apply i in s0. auto. + - intros. specialize (n a). specialize (i a). + destruct (sat_predicate p a). exfalso. + apply n. apply i. auto. auto. +Defined. + +Definition sat_pred_simple (p: pred_op) : option alist := + match sat_pred_tseytin p with + | inleft (exist _ a _) => Some a + | inright _ => None + end. + +Definition sat_pred (p: pred_op) : + ({al : alist | sat_predicate p (interp_alist al) = true} + + {forall a : asgn, sat_predicate p a = false}). + refine + ( match trans_pred p with + | exist _ fm _ => + match satSolve fm with + | inleft (exist _ a _) => inleft (exist _ a _) + | inright _ => inright _ + end + end ). + - apply i in s0. auto. + - intros. specialize (n a). specialize (i a). + destruct (sat_predicate p a). exfalso. + apply n. apply i. auto. auto. +Defined. + +#[local] Open Scope positive. + +Compute tseytin_simple (Por (negate (Pand (Por (Plit (true, 1)) (Plit (true, 2))) (Plit (true, 3)))) (Plit (false, 4))). +Compute sat_pred_simple (Por Pfalse (Pand (Plit (true, 1)) (Plit (false, 1)))). + +Lemma sat_dec a: {sat a} + {unsat a}. +Proof. + unfold sat, unsat. + destruct (sat_pred a). + intros. left. destruct s. + exists (Sat.interp_alist x). auto. + intros. tauto. +Qed. + Definition equiv_check p1 p2 := match sat_pred_simple (simplify (p1 ∧ ¬ p2 ∨ p2 ∧ ¬ p1)) with | None => true | _ => false end. +Compute equiv_check Pfalse (Pand (Plit (true, 1%positive)) (Plit (false, 1%positive))). + Lemma equiv_check_correct : forall p1 p2, equiv_check p1 p2 = true -> p1 == p2. Proof. @@ -448,7 +596,8 @@ Lemma equiv_check_correct2 : forall p1 p2, p1 == p2 -> equiv_check p1 p2 = true. Proof. unfold equiv_check, equiv, sat_pred_simple. intros. - destruct_match; auto. destruct_match; try discriminate. destruct_match. + destruct_match; auto. destruct_match; try discriminate. + destruct_match. simplify. apply sat_equiv4 in H. unfold unsat in H. simplify. clear Heqs. rewrite simplify_correct in e. -- cgit From dcd9104c1b9de5c856e1dfb95788dc514ec7bc5f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 14 Nov 2021 22:22:51 +0000 Subject: Add helper functions to FunctionalUnits --- src/hls/FunctionalUnits.v | 74 +++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 74 insertions(+) (limited to 'src/hls') diff --git a/src/hls/FunctionalUnits.v b/src/hls/FunctionalUnits.v index e94b8e8..4a9bf5b 100644 --- a/src/hls/FunctionalUnits.v +++ b/src/hls/FunctionalUnits.v @@ -77,6 +77,80 @@ Record resources := mk_resources { res_arch: arch; }. +Definition index_div {b:bool} r (d: divider b) := + match r with + | 1 => div_numer d + | 2 => div_denom d + | 3 => div_quot d + | _ => div_rem d + end. + +Definition index_ram r (d: ram) := + match r with + | 1 => ram_mem d + | 2 => ram_en d + | 3 => ram_u_en d + | 4 => ram_addr d + | 5 => ram_wr_en d + | 6 => ram_d_in d + | _ => ram_d_out d + end. + +Definition index_res u r res := + match PTree.get u res with + | Some (SignedDiv d) => Some (index_div r d) + | Some (UnsignedDiv d) => Some (index_div r d) + | Some (Ram d) => Some (index_ram r d) + | None => None + end. + +Definition get_ram n res: option (positive * ram) := + match nth_error (arch_ram (res_arch res)) n with + | Some ri => + match PTree.get ri (res_funct_units res) with + | Some (Ram r) => Some (ri, r) + | _ => None + end + | None => None + end. + +Definition get_div n res := + match nth_error (arch_div (res_arch res)) n with + | Some ri => + match PTree.get ri (res_funct_units res) with + | Some (UnsignedDiv d) => Some (ri, d) + | _ => None + end + | None => None + end. + +Definition get_sdiv n res := + match nth_error (arch_sdiv (res_arch res)) n with + | Some ri => + match PTree.get ri (res_funct_units res) with + | Some (SignedDiv d) => Some (ri, d) + | _ => None + end + | None => None + end. + +Definition set_res fu res := + let max := fold_left Pos.max ((arch_sdiv (res_arch res)) + ++ (arch_div (res_arch res)) + ++ (arch_ram (res_arch res))) 1 in + let nt := PTree.set (max + 1)%positive fu (res_funct_units res) in + match fu with + | UnsignedDiv _ => mk_resources nt (mk_arch (max :: arch_div (res_arch res)) + (arch_sdiv (res_arch res)) + (arch_ram (res_arch res))) + | SignedDiv _ => mk_resources nt (mk_arch (arch_div (res_arch res)) + (max :: arch_sdiv (res_arch res)) + (arch_ram (res_arch res))) + | Ram _ => mk_resources nt (mk_arch (arch_div (res_arch res)) + (arch_sdiv (res_arch res)) + (max :: arch_ram (res_arch res))) + end. + Definition initial_funct_units: funct_units := PTree.empty _. Definition initial_arch := mk_arch nil nil nil. -- cgit From d772e22704ffe806b9962507c9faf05ce0159133 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 14 Nov 2021 22:23:06 +0000 Subject: Add proper functional unit generation --- src/hls/RTLParFU.v | 24 +++---------- src/hls/RTLParFUgen.v | 97 ++++++++++++++++++++++++++++++++++++++++++++++++--- 2 files changed, 98 insertions(+), 23 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLParFU.v b/src/hls/RTLParFU.v index 3442ff0..b18aef6 100644 --- a/src/hls/RTLParFU.v +++ b/src/hls/RTLParFU.v @@ -37,8 +37,8 @@ Definition node := positive. Inductive instr : Type := | FUnop : instr | FUop : option pred_op -> operation -> list reg -> reg -> instr -| FUload : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr -| FUstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr +| FUread : positive -> positive -> reg -> instr +| FUwrite : positive -> positive -> reg -> instr | FUsetpred : option pred_op -> condition -> list reg -> predicate -> instr. Inductive cf_instr : Type := @@ -69,10 +69,8 @@ Definition max_reg_instr (m: positive) (i: instr) := | FUnop => m | FUop p op args res => fold_left Pos.max args (Pos.max res m) - | FUload p chunk addr args dst => - fold_left Pos.max args (Pos.max dst m) - | FUstore p chunk addr args src => - fold_left Pos.max args (Pos.max src m) + | FUread p1 p2 r => Pos.max m r + | FUwrite p1 p2 r => Pos.max m r | FUsetpred p' c args p => fold_left Pos.max args m end. @@ -157,7 +155,7 @@ Record function: Type := fn_params: list reg; fn_stacksize: Z; fn_code: code; - fn_funct_units: funct_unit; + fn_funct_units: resources; fn_entrypoint: node; }. @@ -246,18 +244,6 @@ Section RELSEM. eval_operation ge sp op rs##args m = Some v -> eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#res <- v) pr m) ist -> step_instr sp (mk_instr_state rs pr m) (FUop p op args res) ist - | exec_FUload: - forall addr rs args a chunk m v dst sp p pr ist, - eval_addressing ge sp addr rs##args = Some a -> - Mem.loadv chunk m a = Some v -> - eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#dst <- v) pr m) ist -> - step_instr sp (mk_instr_state rs pr m) (FUload p chunk addr args dst) ist - | exec_FUstore: - forall addr rs args a chunk m src m' sp p pr ist, - eval_addressing ge sp addr rs##args = Some a -> - Mem.storev chunk m a rs#src = Some m' -> - eval_pred p (mk_instr_state rs pr m) (mk_instr_state rs pr m') ist -> - step_instr sp (mk_instr_state rs pr m) (FUstore p chunk addr args src) ist | exec_FUsetpred: forall sp rs pr m p c b args p' ist, Op.eval_condition c rs##args m = Some b -> diff --git a/src/hls/RTLParFUgen.v b/src/hls/RTLParFUgen.v index 389d76c..0098f57 100644 --- a/src/hls/RTLParFUgen.v +++ b/src/hls/RTLParFUgen.v @@ -1,21 +1,110 @@ +(* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2021 Yann Herklotz + * + * This program is free software: you can redistribute it and/or modify + * it under the terms of the GNU General Public License as published by + * the Free Software Foundation, either version 3 of the License, or + * (at your option) any later version. + * + * This program is distributed in the hope that it will be useful, + * but WITHOUT ANY WARRANTY; without even the implied warranty of + * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + * GNU General Public License for more details. + * + * You should have received a copy of the GNU General Public License + * along with this program. If not, see . + *) + Require Import Coq.micromega.Lia. Require Import compcert.common.AST. -Require compcert.common.Errors. +Require Import compcert.common.Errors. Require compcert.common.Globalenvs. -Require compcert.lib.Integers. +Require Import compcert.lib.Integers. Require Import compcert.lib.Maps. 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.RTLParFU. Require Import vericert.hls.ValueInt. Require Import vericert.hls.Verilog. Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLPar. +Require Import vericert.hls.FunctionalUnits. + +#[local] Open Scope error_monad_scope. + +Definition transl_instr (res: resources) (i: RTLBlockInstr.instr): Errors.res (list (Z * RTLParFU.instr)) := + match i with + | RBnop => Errors.OK ((0, FUnop)::nil) + | RBop po op args d => Errors.OK ((0, FUop po op args d)::nil) + | RBload po chunk addr args d => + match get_ram 0 res with + | Some (ri, r) => + Errors.OK ((0, FUop po Op.Onot (ram_u_en r::nil) (ram_u_en r)) + :: (0, FUop po (Op.Ointconst (Int.repr 0)) nil (ram_wr_en r)) + :: (0, FUop po (Op.Olea addr) args (ram_addr r)) + :: (1, FUop po Op.Omove (ram_d_out r::nil) d) + :: nil) + | _ => Errors.Error (Errors.msg "Could not find RAM") + end + | RBstore po chunk addr args d => + match get_ram 0 res with + | Some (ri, r) => + Errors.OK ((0, FUop po Op.Onot (ram_u_en r::nil) (ram_u_en r)) + :: (0, FUop po (Op.Ointconst (Int.repr 1)) nil (ram_wr_en r)) + :: (0, FUop po Op.Omove (d::nil) (ram_d_in r)) + :: (0, FUop po (Op.Olea addr) args (ram_addr r)) + :: nil) + | _ => Errors.Error (Errors.msg "Could not find RAM") + end + | RBsetpred op c args p => Errors.OK ((0, FUsetpred op c args p)::nil) + end. + +Fixpoint transl_cf_instr (i: RTLBlockInstr.cf_instr): RTLParFU.cf_instr := + match i with + | RBcall sig r args d n => FUcall sig r args d n + | RBtailcall sig r args => FUtailcall sig r args + | RBbuiltin ef args r n => FUbuiltin ef args r n + | RBcond c args n1 n2 => FUcond c args n1 n2 + | RBjumptable r ns => FUjumptable r ns + | RBreturn r => FUreturn r + | RBgoto n => FUgoto n + | RBpred_cf po c1 c2 => FUpred_cf po (transl_cf_instr c1) (transl_cf_instr c2) + end. + +Definition list_split {A:Type} (l: list (Z * A)) : (list (Z * A)) * (list (Z * A)) := + (filter (fun x => Z.eqb 0 (fst x)) l, + map (fun x => (Z.pred (fst x), snd x)) (filter (fun x => negb (Z.eqb 0 (fst x))) l)). + +Fixpoint map_error {A B : Type} (f : A -> res B) (l : list A) {struct l} : res (list B) := + match l with + | nil => OK nil + | x::xs => + do y <- f x ; + do ys <- map_error f xs ; + OK (y::ys) + end. + +Definition transl_bb (res: resources) (bb: RTLPar.bb): RTLParFU.bblock_body := + map_error (map_error (fold_right + (fun c n => transl_instr res))) bb. + +Definition transl_bblock (bb: RTLPar.bblock): RTLParFU.bblock := + RTLParFU.mk_bblock (transl_bb (bb_body bb)) (transl_cf_instr (bb_exit bb)). + +Definition transl_code (c: RTLPar.code): RTLParFU.code := + PTree.map (fun _ => transl_bblock) c. + +Definition transl_function (f: RTLPar.function): Errors.res RTLParFU.function := + Errors.OK (RTLParFU.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (transl_code (fn_code f)) + initial_resources (fn_entrypoint f)). + +Definition transl_fundef p := + transf_partial_fundef transl_function p. -Definition transl_program (p : RTLParFU.program) : Errors.res HTL.program := +Definition transl_program (p : RTLPar.program) : Errors.res RTLParFU.program := transform_partial_program transl_fundef p. -- cgit From b8616c44c16bf3edecd7d4569afcf8ff0f7992ef Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 14 Nov 2021 22:23:16 +0000 Subject: Improve simplification of predicates --- src/hls/Predicate.v | 12 ++++++++++-- 1 file changed, 10 insertions(+), 2 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Predicate.v b/src/hls/Predicate.v index 9f9ec65..b19ae98 100644 --- a/src/hls/Predicate.v +++ b/src/hls/Predicate.v @@ -199,6 +199,14 @@ Qed. Definition simplify' (p: pred_op) := match p with + | (Plit (b1, a)) ∧ (Plit (b2, b)) as p' => + if Pos.eqb a b then + if negb (xorb b1 b2) then Plit (b1, a) else ⟂ + else p' + | (Plit (b1, a)) ∨ (Plit (b2, b)) as p' => + if Pos.eqb a b then + if negb (xorb b1 b2) then Plit (b1, a) else T + else p' | A ∧ T => A | T ∧ A => A | _ ∧ ⟂ => ⟂ @@ -234,9 +242,9 @@ Lemma simplify'_correct : forall h a, sat_predicate (simplify' h) a = sat_predicate h a. Proof. - destruct h; crush; repeat destruct_match; crush; + (*destruct h; crush; repeat destruct_match; crush; solve [rewrite andb_true_r; auto | rewrite orb_false_r; auto]. -Qed. +Qed.*) Admitted. Lemma simplify_correct : forall h a, -- cgit From 5367a016299be89083254fcb972e4a5911a645e9 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 14 Nov 2021 22:23:26 +0000 Subject: Replace HTLPargen by HTLParFUgen --- src/hls/HTLParFugen.v | 879 -------------------------------------------------- src/hls/HTLPargen.v | 36 ++- 2 files changed, 19 insertions(+), 896 deletions(-) delete mode 100644 src/hls/HTLParFugen.v (limited to 'src/hls') diff --git a/src/hls/HTLParFugen.v b/src/hls/HTLParFugen.v deleted file mode 100644 index a002e15..0000000 --- a/src/hls/HTLParFugen.v +++ /dev/null @@ -1,879 +0,0 @@ -(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 2020 Yann Herklotz - * - * This program is free software: you can redistribute it and/or modify - * it under the terms of the GNU General Public License as published by - * the Free Software Foundation, either version 3 of the License, or - * (at your option) any later version. - * - * This program is distributed in the hope that it will be useful, - * but WITHOUT ANY WARRANTY; without even the implied warranty of - * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the - * GNU General Public License for more details. - * - * You should have received a copy of the GNU General Public License - * along with this program. If not, see . - *) - -Require Import Coq.micromega.Lia. - -Require Import compcert.common.AST. -Require compcert.common.Errors. -Require compcert.common.Globalenvs. -Require compcert.lib.Integers. -Require Import compcert.lib.Maps. - -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.RTLParFU. -Require Import vericert.hls.ValueInt. -Require Import vericert.hls.Verilog. - -#[local] Hint Resolve AssocMap.gempty : htlh. -#[local] Hint Resolve AssocMap.gso : htlh. -#[local] Hint Resolve AssocMap.gss : htlh. -#[local] Hint Resolve Ple_refl : htlh. -#[local] Hint Resolve Ple_succ : htlh. - -Definition assignment : Type := expr -> expr -> stmnt. - -Record state: Type := mkstate { - st_st: reg; - st_freshreg: reg; - st_freshstate: node; - st_scldecls: AssocMap.t (option io * scl_decl); - st_arrdecls: AssocMap.t (option io * arr_decl); - st_datapath: datapath; - st_controllogic: controllogic; -}. - -Definition init_state (st : reg) : state := - mkstate st - 1%positive - 1%positive - (AssocMap.empty (option io * scl_decl)) - (AssocMap.empty (option io * arr_decl)) - (AssocMap.empty stmnt) - (AssocMap.empty stmnt). - -Module HTLState <: State. - - Definition st := state. - - Inductive st_incr: state -> state -> Prop := - state_incr_intro: - forall (s1 s2: state), - st_st s1 = st_st s2 -> - Ple s1.(st_freshreg) s2.(st_freshreg) -> - Ple s1.(st_freshstate) s2.(st_freshstate) -> - (forall n, - s1.(st_controllogic)!n = None - \/ s2.(st_controllogic)!n = s1.(st_controllogic)!n) -> - st_incr s1 s2. - #[local] Hint Constructors st_incr : htlh. - - Definition st_prop := st_incr. - #[local] Hint Unfold st_prop : htlh. - - Lemma st_refl : forall s, st_prop s s. - Proof. auto with htlh. Qed. - - Lemma st_trans : - forall s1 s2 s3, st_prop s1 s2 -> st_prop s2 s3 -> st_prop s1 s3. - Proof. - intros. inv H. inv H0. - apply state_incr_intro; eauto using Ple_trans; intros; try congruence. - destruct H4 with n; destruct H7 with n; intuition congruence. - Qed. - -End HTLState. -Export HTLState. - -Module HTLMonad := Statemonad(HTLState). -Export HTLMonad. - -Module HTLMonadExtra := Monad.MonadExtra(HTLMonad). -Import HTLMonadExtra. -Export MonadNotation. - -Definition state_goto (st : reg) (n : node) : stmnt := - Vnonblock (Vvar st) (Vlit (posToValue n)). - -Definition state_cond (st : reg) (c : expr) (n1 n2 : node) : stmnt := - Vnonblock (Vvar st) (Vternary c (posToExpr n1) (posToExpr n2)). - -Definition check_empty_node_datapath: - forall (s: state) (n: node), { s.(st_datapath)!n = None } + { True }. -Proof. - intros. case (s.(st_datapath)!n); tauto. -Defined. - -Definition check_empty_node_controllogic: - forall (s: state) (n: node), { s.(st_controllogic)!n = None } + { True }. -Proof. - intros. case (s.(st_controllogic)!n); tauto. -Defined. - -Lemma declare_reg_state_incr : - forall i s r sz, - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - s.(st_freshstate) - (AssocMap.set r (i, VScalar sz) s.(st_scldecls)) - s.(st_arrdecls) - s.(st_datapath) - s.(st_controllogic)). -Proof. Admitted. - -Definition declare_reg (i : option io) (r : reg) (sz : nat) : mon unit := - fun s => OK tt (mkstate - s.(st_st) - s.(st_freshreg) - s.(st_freshstate) - (AssocMap.set r (i, VScalar sz) s.(st_scldecls)) - s.(st_arrdecls) - s.(st_datapath) - s.(st_controllogic)) - (declare_reg_state_incr i s r sz). - -Lemma add_instr_state_incr : - forall s n n' st, - (st_controllogic s)!n = None -> - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n st s.(st_datapath)) - (AssocMap.set n (state_goto s.(st_st) n') s.(st_controllogic))). -Proof. - constructor; intros; - try (simpl; destruct (peq n n0); subst); - auto with htlh. -Qed. - -Definition add_instr (n : node) (n' : node) (st : stmnt) : mon unit := - fun s => - match check_empty_node_controllogic s n with - | left TRANS => - OK tt (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n st s.(st_datapath)) - (AssocMap.set n (state_goto s.(st_st) n') s.(st_controllogic))) - (add_instr_state_incr s n n' st TRANS) - | _ => Error (Errors.msg "HTL.add_instr") - end. - -Lemma add_instr_skip_state_incr : - forall s n st, - (st_controllogic s)!n = None -> - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n st s.(st_datapath)) - (AssocMap.set n Vskip s.(st_controllogic))). -Proof. - constructor; intros; - try (simpl; destruct (peq n n0); subst); - auto with htlh. -Qed. - -Definition add_instr_skip (n : node) (st : stmnt) : mon unit := - fun s => - match check_empty_node_controllogic s n with - | left TRANS => - OK tt (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n st s.(st_datapath)) - (AssocMap.set n Vskip s.(st_controllogic))) - (add_instr_skip_state_incr s n st TRANS) - | _ => Error (Errors.msg "HTL.add_instr_skip") - end. - -Lemma add_node_skip_state_incr : - forall s n st, - (st_controllogic s)!n = None -> - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n Vskip s.(st_datapath)) - (AssocMap.set n st s.(st_controllogic))). -Proof. - constructor; intros; - try (simpl; destruct (peq n n0); subst); - auto with htlh. -Qed. - -Definition add_node_skip (n : node) (st : stmnt) : mon unit := - fun s => - match check_empty_node_controllogic s n with - | left TRANS => - OK tt (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n Vskip s.(st_datapath)) - (AssocMap.set n st s.(st_controllogic))) - (add_node_skip_state_incr s n st TRANS) - | _ => Error (Errors.msg "HTL.add_node_skip") - end. - -Definition nonblock (dst : reg) (e : expr) := Vnonblock (Vvar dst) e. -Definition block (dst : reg) (e : expr) := Vblock (Vvar dst) e. - -Definition bop (op : binop) (r1 r2 : reg) : expr := - Vbinop op (Vvar r1) (Vvar r2). - -Definition boplit (op : binop) (r : reg) (l : Integers.int) : expr := - Vbinop op (Vvar r) (Vlit (intToValue l)). - -Definition boplitz (op: binop) (r: reg) (l: Z) : expr := - Vbinop op (Vvar r) (Vlit (ZToValue l)). - -Definition translate_comparison (c : Integers.comparison) (args : list reg) - : mon expr := - match c, args with - | Integers.Ceq, r1::r2::nil => ret (bop Veq r1 r2) - | Integers.Cne, r1::r2::nil => ret (bop Vne r1 r2) - | Integers.Clt, r1::r2::nil => ret (bop Vlt r1 r2) - | 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 - "Htlgen: comparison instruction not implemented: other") - end. - -Definition translate_comparison_imm (c : Integers.comparison) (args : list reg) - (i: Integers.int) : mon expr := - match c, args with - | Integers.Ceq, r1::nil => ret (boplit Veq r1 i) - | Integers.Cne, r1::nil => ret (boplit Vne r1 i) - | Integers.Clt, r1::nil => ret (boplit Vlt r1 i) - | 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 - "Htlgen: comparison_imm instruction not implemented: other") - end. - -Definition translate_comparisonu (c : Integers.comparison) (args : list reg) - : mon expr := - match c, args with - | Integers.Clt, r1::r2::nil => ret (bop Vltu r1 r2) - | Integers.Cgt, r1::r2::nil => ret (bop Vgtu r1 r2) - | Integers.Cle, r1::r2::nil => ret (bop Vleu r1 r2) - | Integers.Cge, r1::r2::nil => ret (bop Vgeu r1 r2) - | _, _ => error (Errors.msg - "Htlgen: comparison instruction not implemented: other") - end. - -Definition translate_comparison_immu (c : Integers.comparison) - (args : list reg) (i: Integers.int) : mon expr := - match c, args with - | Integers.Clt, r1::nil => ret (boplit Vltu r1 i) - | Integers.Cgt, r1::nil => ret (boplit Vgtu r1 i) - | Integers.Cle, r1::nil => ret (boplit Vleu r1 i) - | Integers.Cge, r1::nil => ret (boplit Vgeu r1 i) - | _, _ => error (Errors.msg - "Htlgen: comparison_imm instruction not implemented: other") - end. - -Definition translate_condition (c : Op.condition) (args : list reg) - : mon expr := - match c, args with - | Op.Ccomp c, _ => translate_comparison c args - | Op.Ccompu c, _ => translate_comparisonu c args - | Op.Ccompimm c i, _ => translate_comparison_imm c args i - | Op.Ccompuimm c i, _ => translate_comparison_immu c args i - | 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_signed (p : Z) : bool := - Z.leb Integers.Ptrofs.min_signed p - && Z.leb p Integers.Ptrofs.max_signed. - -Definition check_address_parameter_unsigned (p : Z) : bool := - Z.leb p Integers.Ptrofs.max_unsigned. - -Definition translate_eff_addressing (a: Op.addressing) (args: list reg) - : mon expr := - match a, args with (* TODO: We should be more methodical here; what are the possibilities?*) - | Op.Aindexed off, r1::nil => - if (check_address_parameter_signed off) - then ret (boplitz Vadd r1 off) - else error (Errors.msg ("HTLPargen: translate_eff_addressing (Aindexed): address out of bounds")) - | Op.Ascaled scale offset, r1::nil => - if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) - then ret (Vbinop Vadd (boplitz Vmul r1 scale) (Vlit (ZToValue offset))) - else error (Errors.msg "HTLPargen: translate_eff_addressing (Ascaled): address out of bounds") - | Op.Aindexed2 offset, r1::r2::nil => - if (check_address_parameter_signed offset) - then ret (Vbinop Vadd (bop Vadd r1 r2) (Vlit (ZToValue offset))) - else error (Errors.msg "HTLPargen: translate_eff_addressing (Aindexed2): address out of bounds") - | Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) - if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) - then ret (Vbinop Vadd (Vvar r1) (Vbinop Vadd (boplitz Vmul r2 scale) (Vlit (ZToValue offset)))) - else error (Errors.msg "HTLPargen: translate_eff_addressing (Aindexed2scaled): address out of bounds") - | 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_unsigned a) - then ret (Vlit (ZToValue a)) - else error (Errors.msg "HTLPargen: translate_eff_addressing (Ainstack): address out of bounds") - | _, _ => error (Errors.msg "HTLPargen: translate_eff_addressing unsuported addressing") - end. - -(** 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) - | Op.Ointconst n, _ => ret (Vlit (intToValue n)) - | Op.Oneg, r::nil => ret (Vunop Vneg (Vvar r)) - | 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, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhs") - | Op.Omulhu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhu") - | 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) - | Op.Omodu, r1::r2::nil => ret (bop Vmodu r1 r2) - | Op.Oand, r1::r2::nil => ret (bop Vand r1 r2) - | Op.Oandimm n, r::nil => ret (boplit Vand r n) - | Op.Oor, r1::r2::nil => ret (bop Vor r1 r2) - | Op.Oorimm n, r::nil => ret (boplit Vor r n) - | Op.Oxor, r1::r2::nil => ret (bop Vxor r1 r2) - | Op.Oxorimm n, r::nil => ret (boplit Vxor r n) - | Op.Onot, r::nil => ret (Vunop Vnot (Vvar r)) - | Op.Oshl, r1::r2::nil => ret (bop Vshl r1 r2) - | 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 => - ret (Vternary (Vbinop Vlt (Vvar r) (Vlit (ZToValue 0))) - (Vunop Vneg (Vbinop Vshru (Vunop Vneg (Vvar r)) (Vlit n))) - (Vbinop Vshru (Vvar r) (Vlit 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 => error (Errors.msg "Htlgen: Instruction not implemented: Ororimm") - (*ret (Vbinop Vor (boplit Vshru r (Integers.Int.modu n (Integers.Int.repr 32))) - (boplit Vshl r (Integers.Int.sub (Integers.Int.repr 32) (Integers.Int.modu n (Integers.Int.repr 32)))))*) - | 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 - | _, _ => error (Errors.msg "Htlgen: Instruction not implemented: other") - end. - -Lemma add_branch_instr_state_incr: - forall s e n n1 n2, - (st_controllogic s) ! n = None -> - st_incr s (mkstate - s.(st_st) - (st_freshreg s) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n Vskip (st_datapath s)) - (AssocMap.set n (state_cond s.(st_st) e n1 n2) (st_controllogic s))). -Proof. - intros. apply state_incr_intro; simpl; - try (intros; destruct (peq n0 n); subst); - auto with htlh. -Qed. - -Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit := - fun s => - match check_empty_node_controllogic s n with - | left NTRANS => - OK tt (mkstate - s.(st_st) - (st_freshreg s) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (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 NTRANS) - | _ => Error (Errors.msg "Htlgen: add_branch_instr") - end. - -Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) - (args : list reg) (stack : reg) : mon expr := - match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) - | Mint32, Op.Aindexed off, r1::nil => - if (check_address_parameter_signed off) - then ret (Vvari stack (Vbinop Vdivu (boplitz Vadd r1 off) (Vlit (ZToValue 4)))) - else error (Errors.msg "HTLgen: translate_arr_access address out of bounds") - | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) - if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) - then ret (Vvari stack - (Vbinop Vdivu - (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) - (Vlit (ZToValue 4)))) - else error (Errors.msg "HTLgen: translate_arr_access address out of bounds") - | 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_unsigned a) - then ret (Vvari stack (Vlit (ZToValue (a / 4)))) - else error (Errors.msg "HTLgen: eff_addressing out of bounds stack offset") - | _, _, _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") - end. - -Fixpoint enumerate (i : nat) (ns : list node) {struct ns} : list (nat * node) := - match ns with - | n :: ns' => (i, n) :: enumerate (i+1) ns' - | nil => nil - end. - -Definition tbl_to_case_expr (st : reg) (ns : list node) : list (expr * stmnt) := - List.map (fun a => match a with - (i, n) => (Vlit (natToValue i), Vnonblock (Vvar st) (Vlit (posToValue n))) - end) - (enumerate 0 ns). - -Definition stack_correct (sz : Z) : bool := - (0 <=? sz) && (sz let r := s.(st_freshreg) in - OK r (mkstate - s.(st_st) - (Pos.succ r) - (st_freshstate s) - (AssocMap.set s.(st_freshreg) (i, VScalar sz) s.(st_scldecls)) - (st_arrdecls s) - (st_datapath s) - (st_controllogic s)) - (create_reg_state_incr s sz i). - -Lemma create_arr_state_incr: - forall s sz ln i, - st_incr s (mkstate - s.(st_st) - (Pos.succ (st_freshreg s)) - (st_freshstate s) - s.(st_scldecls) - (AssocMap.set s.(st_freshreg) (i, VArray sz ln) s.(st_arrdecls)) - (st_datapath s) - (st_controllogic s)). -Proof. constructor; simpl; auto with htlh. Qed. - -Definition create_arr (i : option io) (sz : nat) (ln : nat) : mon (reg * nat) := - fun s => let r := s.(st_freshreg) in - OK (r, ln) (mkstate - s.(st_st) - (Pos.succ r) - (st_freshstate s) - s.(st_scldecls) - (AssocMap.set s.(st_freshreg) (i, VArray sz ln) s.(st_arrdecls)) - (st_datapath s) - (st_controllogic s)) - (create_arr_state_incr s sz ln i). - -Definition max_pc_map (m : Maps.PTree.t stmnt) := - PTree.fold (fun m pc i => Pos.max m pc) m 1%positive. - -Lemma max_pc_map_sound: - forall m pc i, m!pc = Some i -> Ple pc (max_pc_map m). -Proof. - intros until i. - apply PTree_Properties.fold_rec with (P := fun c m => c!pc = Some i -> Ple pc m). - (* extensionality *) - intros. apply H0. rewrite H; auto. - (* base case *) - rewrite PTree.gempty. congruence. - (* inductive case *) - intros. rewrite PTree.gsspec in H2. destruct (peq pc k). - inv H2. unfold Ple, Plt in *. lia. - apply Ple_trans with a. auto. - unfold Ple, Plt in *. lia. -Qed. - -Lemma max_pc_wf : - forall m, Z.pos (max_pc_map m) <= Integers.Int.max_unsigned -> - map_well_formed m. -Proof. - unfold map_well_formed. intros. - exploit list_in_map_inv. eassumption. intros [x [A B]]. destruct x. - apply Maps.PTree.elements_complete in B. apply max_pc_map_sound in B. - unfold Ple in B. apply Pos2Z.pos_le_pos in B. subst. - simplify. transitivity (Z.pos (max_pc_map m)); eauto. -Qed. - -Definition poslength {A : Type} (l : list A) : positive := - match Zlength l with - | Z.pos p => p - | _ => 1 - end. - -Fixpoint penumerate {A : Type} (p : positive) (l : list A) {struct l} - : list (positive * A) := - match l with - | x :: xs => (p, x) :: penumerate (Pos.pred p) xs - | nil => nil - end. - -Fixpoint prange {A: Type} (p1 p2: positive) (l: list A) {struct l} := - match l with - | x :: xs => (p1, p2, x) :: prange p2 (Pos.pred p2) xs - | nil => nil - end. - -Lemma add_data_instr_state_incr : - forall s n st, - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n (Vseq (AssocMapExt.get_default - _ Vskip n s.(st_datapath)) st) s.(st_datapath)) - s.(st_controllogic)). -Proof. - constructor; intros; - try (simpl; destruct (peq n n0); subst); - auto with htlh. -Qed. - -Definition add_data_instr (n : node) (st : stmnt) : mon unit := - fun s => - OK tt (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - (AssocMap.set n (Vseq (AssocMapExt.get_default _ Vskip n s.(st_datapath)) st) s.(st_datapath)) - s.(st_controllogic)) - (add_data_instr_state_incr s n st). - -Lemma add_control_instr_state_incr : - forall s n st, - (st_controllogic s) ! n = None -> - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - s.(st_datapath) - (AssocMap.set n st s.(st_controllogic))). -Proof. - constructor; intros; - try (simpl; destruct (peq n n0); subst); - auto with htlh. -Qed. - -Definition add_control_instr (n : node) (st : stmnt) : mon unit := - fun s => - match check_empty_node_controllogic s n with - | left CTRL => - OK tt (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - s.(st_datapath) - (AssocMap.set n st s.(st_controllogic))) - (add_control_instr_state_incr s n st CTRL) - | _ => - Error (Errors.msg "HTLPargen.add_control_instr: control logic is not empty") - end. - -Definition add_control_instr_force_state_incr : - forall s n st, - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - s.(st_datapath) - (AssocMap.set n st s.(st_controllogic))). -Admitted. - -Definition add_control_instr_force (n : node) (st : stmnt) : mon unit := - fun s => - OK tt (mkstate - s.(st_st) - s.(st_freshreg) - (st_freshstate s) - s.(st_scldecls) - s.(st_arrdecls) - s.(st_datapath) - (AssocMap.set n st s.(st_controllogic))) - (add_control_instr_force_state_incr s n st). - -Fixpoint pred_expr (preg: reg) (p: pred_op) := - match p with - | Plit (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) - | Pand p1 p2 => - Vbinop Vand (pred_expr preg p1) (pred_expr preg p2) - | Por p1 p2 => - Vbinop Vor (pred_expr preg p1) (pred_expr preg p2) - end. - -Definition translate_predicate (a : assignment) - (preg: reg) (p: option pred_op) (dst e: expr) := - match p with - | None => ret (a dst e) - | Some pos => - ret (a dst (Vternary (pred_expr preg pos) e dst)) - end. - -Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) - : mon stmnt := - match i with - | FUnop => - ret Vskip - | FUop p op args dst => - do instr <- translate_instr op args; - do _ <- declare_reg None dst 32; - translate_predicate a preg p (Vvar dst) instr - | FUload p chunk addr args dst => - do src <- translate_arr_access chunk addr args stack; - do _ <- declare_reg None dst 32; - translate_predicate a preg p (Vvar dst) src - | FUstore p chunk addr args src => - do dst <- translate_arr_access chunk addr args stack; - translate_predicate a preg p dst (Vvar src) - | FUsetpred _ c args p => - do cond <- translate_condition c args; - ret (a (pred_expr preg (Plit (true, p))) cond) - end. - -Lemma create_new_state_state_incr: - forall s p, - st_incr s - (mkstate - s.(st_st) - s.(st_freshreg) - (s.(st_freshstate) + p)%positive - s.(st_scldecls) - s.(st_arrdecls) - s.(st_datapath) - s.(st_controllogic)). -Admitted. - -Definition create_new_state (p: node): mon node := - fun s => OK s.(st_freshstate) - (mkstate - s.(st_st) - s.(st_freshreg) - (s.(st_freshstate) + p)%positive - s.(st_scldecls) - s.(st_arrdecls) - s.(st_datapath) - s.(st_controllogic)) - (create_new_state_state_incr s p). - -Definition translate_inst_list (fin rtrn stack preg: reg) (ni : node * node * list (list instr)) := - match ni with - | (n, p, li) => - do _ <- collectlist - (fun l => - do stmnt <- translate_inst Vblock fin rtrn stack preg n l; - add_data_instr n stmnt) (concat li); - do st <- get; - add_control_instr n (state_goto st.(st_st) p) - end. - -Fixpoint translate_cfi' (fin rtrn stack preg: reg) (cfi: cf_instr) - : mon (stmnt * stmnt) := - match cfi with - | FUgoto n' => - do st <- get; - ret (Vskip, state_goto st.(st_st) n') - | FUcond c args n1 n2 => - do st <- get; - do e <- translate_condition c args; - ret (Vskip, state_cond st.(st_st) e n1 n2) - | FUreturn r => - match r with - | Some r' => - ret ((Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vvar r'))), - Vskip) - | None => - ret ((Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vlit (ZToValue 0%Z)))), - Vskip) - end - | FUpred_cf p c1 c2 => - do (tc1s, tc1c) <- translate_cfi' fin rtrn stack preg c1; - do (tc2s, tc2c) <- translate_cfi' fin rtrn stack preg c2; - ret ((Vcond (pred_expr preg p) tc1s tc2s), (Vcond (pred_expr preg p) tc1c tc2c)) - | FUjumptable r tbl => - do s <- get; - ret (Vskip, Vcase (Vvar r) (list_to_stmnt (tbl_to_case_expr s.(st_st) tbl)) (Some Vskip)) - | FUcall sig ri rl r n => - error (Errors.msg "HTLPargen: RPcall not supported.") - | FUtailcall sig ri lr => - error (Errors.msg "HTLPargen: RPtailcall not supported.") - | FUbuiltin e lb b n => - error (Errors.msg "HTLPargen: RPbuildin not supported.") - end. - -Definition translate_cfi (fin rtrn stack preg: reg) (ni: node * cf_instr) - : mon unit := - let (n, cfi) := ni in - do (s, c) <- translate_cfi' fin rtrn stack preg cfi; - do _ <- add_control_instr n c; - add_data_instr n s. - -Definition transf_bblock (fin rtrn stack preg: reg) (ni : node * bblock) - : mon unit := - let (n, bb) := ni in - do nstate <- create_new_state ((poslength bb.(bb_body)))%positive; - do _ <- collectlist (translate_inst_list fin rtrn stack preg) - (prange n (nstate + poslength bb.(bb_body) - 1)%positive - bb.(bb_body)); - match bb.(bb_body) with - | nil => translate_cfi fin rtrn stack preg (n, bb.(bb_exit)) - | _ => translate_cfi fin rtrn stack preg (nstate, bb.(bb_exit)) - end. - -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; - do (stack, stack_len) <- create_arr None 32 - (Z.to_nat (f.(fn_stacksize) / 4)); - do preg <- create_reg None 32; - do _ <- collectlist (transf_bblock fin rtrn stack preg) - (Maps.PTree.elements f.(fn_code)); - do _ <- collectlist (fun r => declare_reg (Some Vinput) r 32) - f.(fn_params); - do start <- create_reg (Some Vinput) 1; - do rst <- create_reg (Some Vinput) 1; - 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, - decide_order (st_st current_state) fin rtrn stack start rst clk, - max_list_dec (fn_params f) (st_st current_state) - with - | left LEDATA, left LECTRL, left MORD, left WFPARAMS => - ret (HTL.mkmodule - f.(fn_params) - current_state.(st_datapath) - current_state.(st_controllogic) - f.(fn_entrypoint) - current_state.(st_st) - stack - stack_len - fin - rtrn - start - rst - clk - current_state.(st_scldecls) - current_state.(st_arrdecls) - None - (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA)) - MORD - _ - WFPARAMS) - | _, _, _, _ => error (Errors.msg "More than 2^32 states.") - end - 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)) - (AssocMap.set st (None, VScalar 32) (st_scldecls (init_state st))) - (st_arrdecls (init_state st)) - (st_datapath (init_state st)) - (st_controllogic (init_state st)). - -Definition transl_module (f : function) : Errors.res HTL.module := - run_mon (max_state f) (transf_module f). - -Definition transl_fundef := transf_partial_fundef transl_module. - -Definition main_is_internal (p : RTLParFU.program) : bool := - let ge := Globalenvs.Genv.globalenv p in - match Globalenvs.Genv.find_symbol ge p.(AST.prog_main) with - | Some b => - match Globalenvs.Genv.find_funct_ptr ge b with - | Some (AST.Internal _) => true - | _ => false - end - | _ => false - end. - -Definition transl_program (p : RTLParFU.program) : Errors.res HTL.program := - if main_is_internal p - then transform_partial_program transl_fundef p - else Errors.Error (Errors.msg "Main function is not Internal."). diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index b0815b7..dd0944f 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2020 Yann Herklotz + * Copyright (C) 2021 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -30,7 +30,7 @@ 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.RTLParFU. Require Import vericert.hls.ValueInt. Require Import vericert.hls.Verilog. @@ -681,20 +681,22 @@ Definition translate_predicate (a : assignment) Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) : mon stmnt := match i with - | RBnop => + | FUnop => ret Vskip - | RBop p op args dst => + | FUop p op args dst => do instr <- translate_instr op args; do _ <- declare_reg None dst 32; translate_predicate a preg p (Vvar dst) instr - | RBload p chunk addr args dst => + | FUload p chunk addr args dst => do src <- translate_arr_access chunk addr args stack; do _ <- declare_reg None dst 32; translate_predicate a preg p (Vvar dst) src - | RBstore p chunk addr args src => + | FUstore p chunk addr args src => do dst <- translate_arr_access chunk addr args stack; translate_predicate a preg p dst (Vvar src) - | RBsetpred _ c args p => + | FUread p1 p2 r => ret Vskip + | FUwrite p1 p2 r => ret Vskip + | FUsetpred _ c args p => do cond <- translate_condition c args; ret (a (pred_expr preg (Plit (true, p))) cond) end. @@ -738,14 +740,14 @@ Definition translate_inst_list (fin rtrn stack preg: reg) (ni : node * node * li Fixpoint translate_cfi' (fin rtrn stack preg: reg) (cfi: cf_instr) : mon (stmnt * stmnt) := match cfi with - | RBgoto n' => + | FUgoto n' => do st <- get; ret (Vskip, state_goto st.(st_st) n') - | RBcond c args n1 n2 => + | FUcond c args n1 n2 => do st <- get; do e <- translate_condition c args; ret (Vskip, state_cond st.(st_st) e n1 n2) - | RBreturn r => + | FUreturn r => match r with | Some r' => ret ((Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vvar r'))), @@ -754,18 +756,18 @@ Fixpoint translate_cfi' (fin rtrn stack preg: reg) (cfi: cf_instr) ret ((Vseq (block fin (Vlit (ZToValue 1%Z))) (block rtrn (Vlit (ZToValue 0%Z)))), Vskip) end - | RBpred_cf p c1 c2 => + | FUpred_cf p c1 c2 => do (tc1s, tc1c) <- translate_cfi' fin rtrn stack preg c1; do (tc2s, tc2c) <- translate_cfi' fin rtrn stack preg c2; ret ((Vcond (pred_expr preg p) tc1s tc2s), (Vcond (pred_expr preg p) tc1c tc2c)) - | RBjumptable r tbl => + | FUjumptable r tbl => do s <- get; ret (Vskip, Vcase (Vvar r) (list_to_stmnt (tbl_to_case_expr s.(st_st) tbl)) (Some Vskip)) - | RBcall sig ri rl r n => + | FUcall sig ri rl r n => error (Errors.msg "HTLPargen: RPcall not supported.") - | RBtailcall sig ri lr => + | FUtailcall sig ri lr => error (Errors.msg "HTLPargen: RPtailcall not supported.") - | RBbuiltin e lb b n => + | FUbuiltin e lb b n => error (Errors.msg "HTLPargen: RPbuildin not supported.") end. @@ -863,7 +865,7 @@ Definition transl_module (f : function) : Errors.res HTL.module := Definition transl_fundef := transf_partial_fundef transl_module. -Definition main_is_internal (p : RTLPar.program) : bool := +Definition main_is_internal (p : RTLParFU.program) : bool := let ge := Globalenvs.Genv.globalenv p in match Globalenvs.Genv.find_symbol ge p.(AST.prog_main) with | Some b => @@ -874,7 +876,7 @@ Definition main_is_internal (p : RTLPar.program) : bool := | _ => false end. -Definition transl_program (p : RTLBlockInstr.program) : Errors.res HTL.program := +Definition transl_program (p : RTLParFU.program) : Errors.res HTL.program := if main_is_internal p then transform_partial_program transl_fundef p else Errors.Error (Errors.msg "Main function is not Internal."). -- cgit From 39200b933b230579b90089539216077f016bacc5 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 15 Nov 2021 21:01:15 +0000 Subject: Add helper functions to FunctionalUnits --- src/hls/FunctionalUnits.v | 25 ++++++++++++++++++++++--- 1 file changed, 22 insertions(+), 3 deletions(-) (limited to 'src/hls') diff --git a/src/hls/FunctionalUnits.v b/src/hls/FunctionalUnits.v index 4a9bf5b..9504bb1 100644 --- a/src/hls/FunctionalUnits.v +++ b/src/hls/FunctionalUnits.v @@ -135,10 +135,10 @@ Definition get_sdiv n res := end. Definition set_res fu res := - let max := fold_left Pos.max ((arch_sdiv (res_arch res)) + let max := ((fold_left Pos.max ((arch_sdiv (res_arch res)) ++ (arch_div (res_arch res)) - ++ (arch_ram (res_arch res))) 1 in - let nt := PTree.set (max + 1)%positive fu (res_funct_units res) in + ++ (arch_ram (res_arch res))) 1) + 1)%positive in + let nt := PTree.set max fu (res_funct_units res) in match fu with | UnsignedDiv _ => mk_resources nt (mk_arch (max :: arch_div (res_arch res)) (arch_sdiv (res_arch res)) @@ -164,3 +164,22 @@ Definition funct_unit_stages (f: funct_unit) : positive := | UnsignedDiv d => div_stages d | _ => 1 end. + +Definition max_reg_ram r := + fold_right Pos.max 1 (ram_mem r::ram_en r::ram_u_en r::ram_addr r + ::ram_wr_en r::ram_d_in r::ram_d_out r::nil). + +Definition max_reg_divider {b: bool} (d: divider b) := + fold_right Pos.max 1 (div_numer d::div_denom d::div_quot d::div_rem d::nil). + +Definition max_reg_fu fu := + match fu with + | SignedDiv d | UnsignedDiv d => max_reg_divider d + | Ram r => max_reg_ram r + end. + +Definition max_reg_funct_units r := + PTree.fold (fun m _ a => Pos.max m (max_reg_fu a)) r 1. + +Definition max_reg_resources r := + max_reg_funct_units r.(res_funct_units). -- cgit From c54576bce7bccb71a807f8331b9e9d7fbdf1eec3 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 15 Nov 2021 21:01:26 +0000 Subject: Remove unnecessary RAM --- src/hls/HTL.v | 17 +---------------- 1 file changed, 1 insertion(+), 16 deletions(-) (limited to 'src/hls') diff --git a/src/hls/HTL.v b/src/hls/HTL.v index 8cebbfd..d5a99fc 100644 --- a/src/hls/HTL.v +++ b/src/hls/HTL.v @@ -31,6 +31,7 @@ Require Import vericert.common.Vericertlib. Require Import vericert.hls.ValueInt. Require Import vericert.hls.AssocMap. Require Import vericert.hls.Array. +Require Import vericert.hls.FunctionalUnits. Require vericert.hls.Verilog. Local Open Scope positive. @@ -54,22 +55,6 @@ Definition map_well_formed {A : Type} (m : PTree.t A) : Prop := In p0 (map fst (Maps.PTree.elements m)) -> (Z.pos p0 <= Integers.Int.max_unsigned)%Z. -Record ram := mk_ram { - ram_size: nat; - ram_mem: reg; - ram_en: reg; - ram_u_en: reg; - ram_addr: reg; - ram_wr_en: reg; - ram_d_in: reg; - ram_d_out: reg; - ram_ordering: (ram_addr < ram_en - /\ ram_en < ram_d_in - /\ ram_d_in < ram_d_out - /\ ram_d_out < ram_wr_en - /\ ram_wr_en < ram_u_en) -}. - Definition module_ordering a b c d e f g := a < b < c /\ c < d < e /\ e < f < g. Record module: Type := -- cgit From 2f424a8b7b1acd7bbf81c68f64ba2e26bb46b315 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 15 Nov 2021 21:01:36 +0000 Subject: Use new RAM defined in FunctionalUnits.v --- src/hls/Memorygen.v | 5 +++-- src/hls/Veriloggen.v | 1 + 2 files changed, 4 insertions(+), 2 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Memorygen.v b/src/hls/Memorygen.v index 96c11c3..4ff4a19 100644 --- a/src/hls/Memorygen.v +++ b/src/hls/Memorygen.v @@ -35,6 +35,7 @@ Require Import vericert.hls.Verilog. Require Import vericert.hls.HTL. Require Import vericert.hls.AssocMap. Require Import vericert.hls.Array. +Require Import vericert.hls.FunctionalUnits. Local Open Scope positive. Local Open Scope assocmap. @@ -459,8 +460,8 @@ Lemma forall_ram_lt : forall_ram (fun x => x < max_reg_module m + 1) r. Proof. assert (X: forall a b c, a < b + 1 -> a < Pos.max c b + 1) by lia. - unfold forall_ram; simplify; unfold max_reg_module; repeat apply X; - unfold max_reg_ram; rewrite H; try lia. + unfold forall_ram; simplify; unfold HTL.max_reg_module; repeat apply X; + unfold HTL.max_reg_ram; rewrite H; try lia. Qed. #[local] Hint Resolve forall_ram_lt : mgen. diff --git a/src/hls/Veriloggen.v b/src/hls/Veriloggen.v index aba2293..035e7a4 100644 --- a/src/hls/Veriloggen.v +++ b/src/hls/Veriloggen.v @@ -25,6 +25,7 @@ Require Import vericert.hls.AssocMap. Require Import vericert.hls.HTL. Require Import vericert.hls.ValueInt. Require Import vericert.hls.Verilog. +Require Import vericert.hls.FunctionalUnits. Definition transl_list_fun (a : node * Verilog.stmnt) := let (n, stmnt) := a in -- cgit From c7cc9a3eb82e36c23caf494e144a3391a31e6260 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 15 Nov 2021 21:01:50 +0000 Subject: Fix HTL generation from RTLParFU --- src/hls/HTLPargen.v | 45 ++++++++++++++++++++++++++++++++++----------- 1 file changed, 34 insertions(+), 11 deletions(-) (limited to 'src/hls') diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index dd0944f..373f704 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -31,6 +31,7 @@ Require Import vericert.hls.HTL. Require Import vericert.hls.Predicate. Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLParFU. +Require Import vericert.hls.FunctionalUnits. Require Import vericert.hls.ValueInt. Require Import vericert.hls.Verilog. @@ -687,13 +688,6 @@ Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) do instr <- translate_instr op args; do _ <- declare_reg None dst 32; translate_predicate a preg p (Vvar dst) instr - | FUload p chunk addr args dst => - do src <- translate_arr_access chunk addr args stack; - do _ <- declare_reg None dst 32; - translate_predicate a preg p (Vvar dst) src - | FUstore p chunk addr args src => - do dst <- translate_arr_access chunk addr args stack; - translate_predicate a preg p dst (Vvar src) | FUread p1 p2 r => ret Vskip | FUwrite p1 p2 r => ret Vskip | FUsetpred _ c args p => @@ -801,6 +795,11 @@ Definition decide_order a b c d e f g : {module_ordering a b c d e f g} + {True} end; unfold module_ordering; auto. Defined. +Lemma clk_greater : + forall ram clk r', + Some ram = Some r' -> (clk < ram_addr r')%positive. +Proof. Admitted. + Definition transf_module (f: function) : mon HTL.module. refine ( if stack_correct f.(fn_stacksize) then @@ -822,9 +821,31 @@ Definition transf_module (f: function) : mon HTL.module. 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, - max_list_dec (fn_params f) (st_st current_state) + max_list_dec (fn_params f) (st_st current_state), + get_ram 0 (fn_funct_units f) with - | left LEDATA, left LECTRL, left MORD, left WFPARAMS => + | left LEDATA, left LECTRL, left MORD, left WFPARAMS, Some (i, ram) => + ret (HTL.mkmodule + f.(fn_params) + current_state.(st_datapath) + current_state.(st_controllogic) + f.(fn_entrypoint) + current_state.(st_st) + stack + stack_len + fin + rtrn + start + rst + clk + current_state.(st_scldecls) + current_state.(st_arrdecls) + (Some ram) + (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA)) + MORD + _ + WFPARAMS) + | left LEDATA, left LECTRL, left MORD, left WFPARAMS, _ => ret (HTL.mkmodule f.(fn_params) current_state.(st_datapath) @@ -845,9 +866,11 @@ Definition transf_module (f: function) : mon HTL.module. MORD _ WFPARAMS) - | _, _, _, _ => error (Errors.msg "More than 2^32 states.") + | _, _, _, _, _ => error (Errors.msg "More than 2^32 states.") end - else error (Errors.msg "Stack size misalignment.")); discriminate. + else error (Errors.msg "Stack size misalignment.")). + apply clk_greater. + discriminate. Defined. Definition max_state (f: function) : state := -- cgit From 894789af7aa5b3f748b3d3be40fe8a4f227b693f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 15 Nov 2021 21:01:59 +0000 Subject: Fix max funtion in RTLParFU --- src/hls/RTLParFU.v | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) (limited to 'src/hls') diff --git a/src/hls/RTLParFU.v b/src/hls/RTLParFU.v index b18aef6..f0ceafd 100644 --- a/src/hls/RTLParFU.v +++ b/src/hls/RTLParFU.v @@ -379,7 +379,8 @@ Definition max_reg_bblock (m : positive) (pc : node) (bb : bblock) := Definition max_reg_function (f: function) := Pos.max (PTree.fold max_reg_bblock f.(fn_code) 1%positive) - (fold_left Pos.max f.(fn_params) 1%positive). + (Pos.max (fold_left Pos.max f.(fn_params) 1%positive) + (max_reg_resources f.(fn_funct_units))). Definition max_pc_function (f: function) : positive := PTree.fold (fun m pc i => (Pos.max m -- cgit From 665945e7512a19aa600c51d164651ad6a00e5713 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 15 Nov 2021 21:02:09 +0000 Subject: Finish generation of RTLParFU with RAM insertion --- src/hls/RTLParFUgen.v | 116 +++++++++++++++++++++++++++++++++++++++----------- 1 file changed, 92 insertions(+), 24 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLParFUgen.v b/src/hls/RTLParFUgen.v index 0098f57..55fe4e7 100644 --- a/src/hls/RTLParFUgen.v +++ b/src/hls/RTLParFUgen.v @@ -28,40 +28,52 @@ Require Import vericert.common.Statemonad. Require Import vericert.common.Vericertlib. Require Import vericert.hls.AssocMap. Require Import vericert.hls.Predicate. -Require Import vericert.hls.RTLParFU. Require Import vericert.hls.ValueInt. Require Import vericert.hls.Verilog. Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLPar. +Require Import vericert.hls.RTLParFU. Require Import vericert.hls.FunctionalUnits. #[local] Open Scope error_monad_scope. -Definition transl_instr (res: resources) (i: RTLBlockInstr.instr): Errors.res (list (Z * RTLParFU.instr)) := +Definition update {A: Type} (i: positive) (f: option A -> A) (pt: PTree.t A) := + PTree.set i (f (pt ! i)) pt. + +Definition add_instr (instr_: instr) x := + match x with Some i => instr_ :: i | None => instr_ :: nil end. + +Definition transl_instr (res: resources) (cycle: positive) (i: RTLBlockInstr.instr) + (li: Errors.res (list instr * PTree.t (list instr))): + Errors.res (list instr * PTree.t (list instr)) := + do (instr_list, d_tree) <- li; match i with - | RBnop => Errors.OK ((0, FUnop)::nil) - | RBop po op args d => Errors.OK ((0, FUop po op args d)::nil) + | RBnop => Errors.OK (FUnop :: instr_list, d_tree) + | RBop po op args d => Errors.OK (FUop po op args d :: instr_list, d_tree) | RBload po chunk addr args d => match get_ram 0 res with | Some (ri, r) => - Errors.OK ((0, FUop po Op.Onot (ram_u_en r::nil) (ram_u_en r)) - :: (0, FUop po (Op.Ointconst (Int.repr 0)) nil (ram_wr_en r)) - :: (0, FUop po (Op.Olea addr) args (ram_addr r)) - :: (1, FUop po Op.Omove (ram_d_out r::nil) d) - :: nil) + Errors.OK (FUop po Op.Onot (ram_u_en r::nil) (ram_u_en r) + :: FUop po (Op.Ointconst (Int.repr 0)) nil (ram_wr_en r) + :: FUop po (Op.Olea addr) args (ram_addr r) + :: FUop po (Op.Oshruimm (Int.repr 2)) ((ram_addr r)::nil) (ram_addr r) + :: instr_list, update (cycle+1) + (add_instr (FUop po Op.Omove (ram_d_out r::nil) d)) + d_tree) | _ => Errors.Error (Errors.msg "Could not find RAM") end | RBstore po chunk addr args d => match get_ram 0 res with | Some (ri, r) => - Errors.OK ((0, FUop po Op.Onot (ram_u_en r::nil) (ram_u_en r)) - :: (0, FUop po (Op.Ointconst (Int.repr 1)) nil (ram_wr_en r)) - :: (0, FUop po Op.Omove (d::nil) (ram_d_in r)) - :: (0, FUop po (Op.Olea addr) args (ram_addr r)) - :: nil) + Errors.OK (FUop po Op.Onot (ram_u_en r::nil) (ram_u_en r) + :: FUop po (Op.Ointconst (Int.repr 1)) nil (ram_wr_en r) + :: FUop po Op.Omove (d::nil) (ram_d_in r) + :: FUop po (Op.Olea addr) args (ram_addr r) + :: FUop po (Op.Oshruimm (Int.repr 2)) ((ram_addr r)::nil) (ram_addr r) + :: instr_list, d_tree) | _ => Errors.Error (Errors.msg "Could not find RAM") end - | RBsetpred op c args p => Errors.OK ((0, FUsetpred op c args p)::nil) + | RBsetpred op c args p => Errors.OK (FUsetpred op c args p :: instr_list, d_tree) end. Fixpoint transl_cf_instr (i: RTLBlockInstr.cf_instr): RTLParFU.cf_instr := @@ -89,19 +101,75 @@ Fixpoint map_error {A B : Type} (f : A -> res B) (l : list A) {struct l} : res ( OK (y::ys) end. -Definition transl_bb (res: resources) (bb: RTLPar.bb): RTLParFU.bblock_body := - map_error (map_error (fold_right - (fun c n => transl_instr res))) bb. +Definition transl_op_chain_block (res: resources) (cycle: positive) (instrs: list RTLBlockInstr.instr) + (state: Errors.res (list (list instr) * PTree.t (list instr))) + : Errors.res (list (list instr) * PTree.t (list instr)) := + do (li, tr) <- state; + do (li', tr') <- fold_right (transl_instr res cycle) (OK (nil, tr)) instrs; + OK (li' :: li, tr'). + +(*Compute transl_op_chain_block initial_resources 10%nat (RBop None (Op.Ointconst (Int.repr 1)) nil 1%positive::RBnop::RBnop::RBnop::nil) (OK (nil, PTree.empty _)).*) + +Definition transl_par_block (res: resources) (cycle: positive) (instrs: list (list RTLBlockInstr.instr)) + (state: Errors.res (list (list (list instr)) * PTree.t (list instr))) + : Errors.res (list (list (list instr)) * PTree.t (list instr)) := + do (li, tr) <- state; + do (li', tr') <- fold_right (transl_op_chain_block res cycle) (OK (nil, tr)) instrs; + OK (li' :: li, tr'). + +(*Compute transl_par_block initial_resources 10%nat ((RBop None (Op.Ointconst (Int.repr 1)) nil 1%positive::RBnop::nil)::(RBop None (Op.Ointconst (Int.repr 2)) nil 2%positive::RBnop::nil)::nil) (OK (((FUnop::nil)::nil)::nil, PTree.empty _)).*) + +Definition transl_seq_block (res: resources) (b: list (list RTLBlockInstr.instr)) + (a: Errors.res (list (list (list instr)) * PTree.t (list instr) * positive)) := + do (litr, n) <- a; + let (li, tr) := litr in + do (li', tr') <- transl_par_block res n b (OK (li, tr)); + OK (li', tr', (n+1)%positive). + +Definition insert_extra (pt: PTree.t (list instr)) (curr: list (list instr)) + (cycle_bb: (positive * list (list (list instr)))) := + let (cycle, bb) := cycle_bb in + match pt ! cycle with + | Some instrs => ((cycle + 1)%positive, (curr ++ (map (fun x => x :: nil) instrs)) :: bb) + | None => ((cycle + 1)%positive, curr :: bb) + end. + +Definition transl_bb (res: resources) (bb: RTLPar.bb): Errors.res RTLParFU.bblock_body := + do (litr, n) <- fold_right (transl_seq_block res) (OK (nil, PTree.empty _, 1%positive)) bb; + let (li, tr) := litr in + OK (snd (fold_right (insert_extra tr) (1%positive, nil) li)). + +Definition transl_bblock (res: resources) (bb: RTLPar.bblock): Errors.res bblock := + do bb' <- transl_bb res (RTLBlockInstr.bb_body bb); + OK (mk_bblock bb' (transl_cf_instr (RTLBlockInstr.bb_exit bb))). -Definition transl_bblock (bb: RTLPar.bblock): RTLParFU.bblock := - RTLParFU.mk_bblock (transl_bb (bb_body bb)) (transl_cf_instr (bb_exit bb)). +Definition error_map_ptree {A B: Type} (f: positive -> A -> res B) (pt: PTree.t A) := + do ptl' <- map_error (fun x => do x' <- uncurry f x; OK (fst x, x')) (PTree.elements pt); + OK (PTree_Properties.of_list ptl'). -Definition transl_code (c: RTLPar.code): RTLParFU.code := - PTree.map (fun _ => transl_bblock) c. +Definition transl_code (fu: resources) (c: RTLPar.code): res code := + error_map_ptree (fun _ => transl_bblock fu) c. Definition transl_function (f: RTLPar.function): Errors.res RTLParFU.function := - Errors.OK (RTLParFU.mkfunction (fn_sig f) (fn_params f) (fn_stacksize f) (transl_code (fn_code f)) - initial_resources (fn_entrypoint f)). + let max := RTLPar.max_reg_function f in + let fu := set_res (Ram (mk_ram + (Z.to_nat (RTLBlockInstr.fn_stacksize f)) + (max+1)%positive + (max+3)%positive + (max+7)%positive + (max+2)%positive + (max+6)%positive + (max+4)%positive + (max+5)%positive + ltac:(lia) + )) initial_resources in + do c' <- transl_code fu (RTLBlockInstr.fn_code f); + Errors.OK (mkfunction (RTLBlockInstr.fn_sig f) + (RTLBlockInstr.fn_params f) + (RTLBlockInstr.fn_stacksize f) + c' + fu + (RTLBlockInstr.fn_entrypoint f)). Definition transl_fundef p := transf_partial_fundef transl_function p. -- cgit From fea7ee4d30aa7597ff5b8e2a2954ed452a1a7a57 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 17 Nov 2021 12:11:12 +0000 Subject: Fix generation of RTLParFU --- src/hls/HTLPargen.v | 58 +++++++++++++++------------------------------------ src/hls/RTLParFUgen.v | 2 +- src/hls/Schedule.ml | 9 ++++---- 3 files changed, 23 insertions(+), 46 deletions(-) (limited to 'src/hls') diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index b66a704..0695f07 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -392,10 +392,10 @@ Definition translate_instr (op : Op.operation) (args : list reg) : mon expr := | Op.Omulimm n, r::nil => ret (boplit Vmul r n) | Op.Omulhs, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhs") | Op.Omulhu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mulhu") - | Op.Odiv, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: div") - | Op.Odivu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: divu") - | Op.Omod, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: mod") - | Op.Omodu, r1::r2::nil => error (Errors.msg "Htlgen: Instruction not implemented: modu") + | 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) + | Op.Omodu, r1::r2::nil => ret (bop Vmodu r1 r2) | Op.Oand, r1::r2::nil => ret (bop Vand r1 r2) | Op.Oandimm n, r::nil => ret (boplit Vand r n) | Op.Oor, r1::r2::nil => ret (bop Vor r1 r2) @@ -458,28 +458,6 @@ Definition add_branch_instr (e: expr) (n n1 n2: node) : mon unit := | _ => Error (Errors.msg "Htlgen: add_branch_instr") end. -Definition translate_arr_access (mem : AST.memory_chunk) (addr : Op.addressing) - (args : list reg) (stack : reg) : mon expr := - match mem, addr, args with (* TODO: We should be more methodical here; what are the possibilities?*) - | Mint32, Op.Aindexed off, r1::nil => - if (check_address_parameter_signed off) - then ret (Vvari stack (Vbinop Vdivu (boplitz Vadd r1 off) (Vlit (ZToValue 4)))) - else error (Errors.msg "HTLgen: translate_arr_access address out of bounds") - | Mint32, Op.Aindexed2scaled scale offset, r1::r2::nil => (* Typical for dynamic array addressing *) - if (check_address_parameter_signed scale) && (check_address_parameter_signed offset) - then ret (Vvari stack - (Vbinop Vdivu - (Vbinop Vadd (boplitz Vadd r1 offset) (boplitz Vmul r2 scale)) - (Vlit (ZToValue 4)))) - else error (Errors.msg "HTLgen: translate_arr_access address out of bounds") - | 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_unsigned a) - then ret (Vvari stack (Vlit (ZToValue (a / 4)))) - else error (Errors.msg "HTLgen: eff_addressing out of bounds stack offset") - | _, _, _ => error (Errors.msg "HTLgen: translate_arr_access unsuported addressing") - end. - Fixpoint enumerate (i : nat) (ns : list node) {struct ns} : list (nat * node) := match ns with | n :: ns' => (i, n) :: enumerate (i+1) ns' @@ -680,7 +658,7 @@ Definition translate_predicate (a : assignment) ret (a dst (Vternary (pred_expr preg pos) e dst)) end. -Definition translate_inst a (fin rtrn stack preg : reg) (n : node) (i : instr) +Definition translate_inst a (fin rtrn preg : reg) (n : node) (i : instr) : mon stmnt := match i with | FUnop => @@ -721,18 +699,18 @@ Definition create_new_state (p: node): mon node := s.(st_controllogic)) (create_new_state_state_incr s p). -Definition translate_inst_list (fin rtrn stack preg: reg) (ni : node * node * list (list instr)) := +Definition translate_inst_list (fin rtrn preg: reg) (ni : node * node * list (list instr)) := match ni with | (n, p, li) => do _ <- collectlist (fun l => - do stmnt <- translate_inst Vblock fin rtrn stack preg n l; + do stmnt <- translate_inst Vblock fin rtrn preg n l; add_data_instr n stmnt) (concat li); do st <- get; add_control_instr n (state_goto st.(st_st) p) end. -Fixpoint translate_cfi' (fin rtrn stack preg: reg) (cfi: cf_instr) +Fixpoint translate_cfi' (fin rtrn preg: reg) (cfi: cf_instr) : mon (stmnt * stmnt) := match cfi with | FUgoto n' => @@ -752,8 +730,8 @@ Fixpoint translate_cfi' (fin rtrn stack preg: reg) (cfi: cf_instr) Vskip) end | FUpred_cf p c1 c2 => - do (tc1s, tc1c) <- translate_cfi' fin rtrn stack preg c1; - do (tc2s, tc2c) <- translate_cfi' fin rtrn stack preg c2; + do (tc1s, tc1c) <- translate_cfi' fin rtrn preg c1; + do (tc2s, tc2c) <- translate_cfi' fin rtrn preg c2; ret ((Vcond (pred_expr preg p) tc1s tc2s), (Vcond (pred_expr preg p) tc1c tc2c)) | FUjumptable r tbl => do s <- get; @@ -766,23 +744,23 @@ Fixpoint translate_cfi' (fin rtrn stack preg: reg) (cfi: cf_instr) error (Errors.msg "HTLPargen: RPbuildin not supported.") end. -Definition translate_cfi (fin rtrn stack preg: reg) (ni: node * cf_instr) +Definition translate_cfi (fin rtrn preg: reg) (ni: node * cf_instr) : mon unit := let (n, cfi) := ni in - do (s, c) <- translate_cfi' fin rtrn stack preg cfi; + do (s, c) <- translate_cfi' fin rtrn preg cfi; do _ <- add_control_instr n c; add_data_instr n s. -Definition transf_bblock (fin rtrn stack preg: reg) (ni : node * bblock) +Definition transf_bblock (fin rtrn preg: reg) (ni : node * bblock) : mon unit := let (n, bb) := ni in do nstate <- create_new_state ((poslength bb.(bb_body)))%positive; - do _ <- collectlist (translate_inst_list fin rtrn stack preg) + do _ <- collectlist (translate_inst_list fin rtrn preg) (prange n (nstate + poslength bb.(bb_body) - 1)%positive bb.(bb_body)); match bb.(bb_body) with - | nil => translate_cfi fin rtrn stack preg (n, bb.(bb_exit)) - | _ => translate_cfi fin rtrn stack preg (nstate, bb.(bb_exit)) + | nil => translate_cfi fin rtrn preg (n, bb.(bb_exit)) + | _ => translate_cfi fin rtrn preg (nstate, bb.(bb_exit)) end. Definition decide_order a b c d e f g : {module_ordering a b c d e f g} + {True}. @@ -806,10 +784,8 @@ Definition transf_module (f: function) : mon HTL.module. if stack_correct f.(fn_stacksize) then do fin <- create_reg (Some Voutput) 1; do rtrn <- create_reg (Some Voutput) 32; - do (stack, stack_len) <- create_arr None 32 - (Z.to_nat (f.(fn_stacksize) / 4)); do preg <- create_reg None 32; - do _ <- collectlist (transf_bblock fin rtrn stack preg) + do _ <- collectlist (transf_bblock fin rtrn preg) (Maps.PTree.elements f.(fn_code)); do _ <- collectlist (fun r => declare_reg (Some Vinput) r 32) f.(fn_params); diff --git a/src/hls/RTLParFUgen.v b/src/hls/RTLParFUgen.v index 55fe4e7..65ddf74 100644 --- a/src/hls/RTLParFUgen.v +++ b/src/hls/RTLParFUgen.v @@ -57,7 +57,7 @@ Definition transl_instr (res: resources) (cycle: positive) (i: RTLBlockInstr.ins :: FUop po (Op.Ointconst (Int.repr 0)) nil (ram_wr_en r) :: FUop po (Op.Olea addr) args (ram_addr r) :: FUop po (Op.Oshruimm (Int.repr 2)) ((ram_addr r)::nil) (ram_addr r) - :: instr_list, update (cycle+1) + :: instr_list, update (Pos.pred cycle) (add_instr (FUop po Op.Omove (ram_d_out r::nil) d)) d_tree) | _ => Errors.Error (Errors.msg "Could not find RAM") diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index 94225fa..b75465b 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -289,6 +289,7 @@ let comb_delay = function | _ -> 1 let pipeline_stages = function + | RBload _ -> 2 | RBop (_, op, _, _) -> (match op with | Odiv -> 32 @@ -743,7 +744,7 @@ let solve_constraints constr = |> drop 3 |> List.fold_left parse_soln (IMap.empty, IMap.empty) in - Sys.remove fn; res + (*Sys.remove fn;*) res let subgraph dfg l = let dfg' = List.fold_left (fun g v -> DFG.add_vertex g v) DFG.empty l in @@ -806,14 +807,14 @@ let transf_rtlpar c c' schedule = let body = IMap.to_seq i_sched_tree |> List.of_seq |> List.map snd |> List.map (List.map (fun x -> (x, List.nth bb_body' x))) in - (*let body2 = List.fold_left (fun x b -> + let body2 = List.fold_left (fun x b -> match IMap.find_opt b i_sched_tree with | Some i -> i :: x | None -> [] :: x ) [] (range (fst bb_st_e) (snd bb_st_e + 1)) |> List.map (List.map (fun x -> (x, List.nth bb_body' x))) |> List.rev - in*) + in (*let final_body = List.map (fun x -> subgraph dfg x |> order_instr) body in*) let final_body2 = List.map (fun x -> subgraph dfg x |> (fun x -> @@ -821,7 +822,7 @@ let transf_rtlpar c c' schedule = |> List.map (subgraph x) |> List.map (fun y -> TopoDFG.fold (fun i l -> snd i :: l) y [] - |> List.rev))) body + |> List.rev))) body2 (*|> (fun x -> TopoDFG.fold (fun i l -> snd i :: l) x []) |> List.rev) body2*) in -- cgit From 3c5bd88f22f744e4908afbc5a56e202dfa469360 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:14:30 +0000 Subject: Fix compilation with new HTL language --- src/hls/FunctionalUnits.v | 8 ++- src/hls/HTL.v | 72 +++++++++++++++------------ src/hls/HTLPargen.v | 122 ++++++++++++++++++++++++---------------------- src/hls/HTLgen.v | 23 +++++---- src/hls/HTLgenproof.v | 15 +++--- src/hls/HTLgenspec.v | 14 ++++-- src/hls/Memorygen.v | 3 +- src/hls/Veriloggen.v | 31 ++---------- src/hls/Veriloggenproof.v | 12 ++--- 9 files changed, 151 insertions(+), 149 deletions(-) (limited to 'src/hls') diff --git a/src/hls/FunctionalUnits.v b/src/hls/FunctionalUnits.v index 9504bb1..dcbe22f 100644 --- a/src/hls/FunctionalUnits.v +++ b/src/hls/FunctionalUnits.v @@ -52,13 +52,11 @@ Record ram := mk_ram { ram_wr_en: reg; ram_d_in: reg; ram_d_out: reg; - ram_ordering: (ram_addr < ram_en - /\ ram_en < ram_d_in - /\ ram_d_in < ram_d_out - /\ ram_d_out < ram_wr_en - /\ ram_wr_en < ram_u_en) }. +Definition all_ram_regs r := + ram_mem r::ram_en r::ram_u_en r::ram_addr r::ram_wr_en r::ram_d_in r::ram_d_out r::nil. + Inductive funct_unit: Type := | SignedDiv: divider true -> funct_unit | UnsignedDiv: divider false -> funct_unit diff --git a/src/hls/HTL.v b/src/hls/HTL.v index f4552a5..551c66e 100644 --- a/src/hls/HTL.v +++ b/src/hls/HTL.v @@ -57,8 +57,6 @@ Definition map_well_formed {A : Type} (m : PTree.t A) : Prop := In p0 (map fst (Maps.PTree.elements m)) -> (Z.pos p0 <= Integers.Int.max_unsigned)%Z. -Definition module_ordering a b c d e f g := a < b < c /\ c < d < e /\ e < f < g. - Record module: Type := mkmodule { mod_params : list reg; @@ -66,8 +64,6 @@ Record module: Type := mod_controllogic : controllogic; mod_entrypoint : node; mod_st : reg; - mod_stk : reg; - mod_stk_len : nat; mod_finish : reg; mod_return : reg; mod_start : reg; @@ -75,11 +71,7 @@ Record module: Type := mod_clk : reg; mod_scldecls : AssocMap.t (option Verilog.io * Verilog.scl_decl); mod_arrdecls : AssocMap.t (option Verilog.io * Verilog.arr_decl); - mod_ram : option ram; - mod_wf : map_well_formed mod_controllogic /\ map_well_formed mod_datapath; - mod_ordering_wf : module_ordering mod_st mod_finish mod_return mod_stk mod_start mod_reset mod_clk; - mod_ram_wf : forall r', mod_ram = Some r' -> mod_clk < ram_addr r'; - mod_params_wf : Forall (Pos.gt mod_st) mod_params; + mod_ram : ram; }. Definition fundef := AST.fundef module. @@ -93,7 +85,7 @@ Fixpoint init_regs (vl : list value) (rl : list reg) {struct rl} := end. Definition empty_stack (m : module) : Verilog.assocmap_arr := - (AssocMap.set m.(mod_stk) (Array.arr_repeat None m.(mod_stk_len)) (AssocMap.empty Verilog.arr)). + (AssocMap.set m.(mod_ram).(ram_mem) (Array.arr_repeat None m.(mod_ram).(ram_size)) (AssocMap.empty Verilog.arr)). (** * Operational Semantics *) @@ -124,13 +116,13 @@ Inductive state : Type := (args : list value), state. Inductive exec_ram: - Verilog.reg_associations -> Verilog.arr_associations -> option ram -> + Verilog.reg_associations -> Verilog.arr_associations -> ram -> Verilog.reg_associations -> Verilog.arr_associations -> Prop := | exec_ram_Some_idle: forall ra ar r, Int.eq (Verilog.assoc_blocking ra)#(ram_en r, 32) (Verilog.assoc_blocking ra)#(ram_u_en r, 32) = true -> - exec_ram ra ar (Some r) ra ar + exec_ram ra ar r ra ar | exec_ram_Some_write: forall ra ar r d_in addr en wr_en u_en, Int.eq en u_en = false -> @@ -140,7 +132,7 @@ Inductive exec_ram: (Verilog.assoc_blocking ra)!(ram_wr_en r) = Some wr_en -> (Verilog.assoc_blocking ra)!(ram_d_in r) = Some d_in -> (Verilog.assoc_blocking ra)!(ram_addr r) = Some addr -> - exec_ram ra ar (Some r) (Verilog.nonblock_reg (ram_en r) ra u_en) + exec_ram ra ar r (Verilog.nonblock_reg (ram_en r) ra u_en) (Verilog.nonblock_arr (ram_mem r) (valueToNat addr) ar d_in) | exec_ram_Some_read: forall ra ar r addr v_d_out en u_en, @@ -151,11 +143,8 @@ Inductive exec_ram: (Verilog.assoc_blocking ra)!(ram_addr r) = Some addr -> Verilog.arr_assocmap_lookup (Verilog.assoc_blocking ar) (ram_mem r) (valueToNat addr) = Some v_d_out -> - exec_ram ra ar (Some r) (Verilog.nonblock_reg (ram_en r) - (Verilog.nonblock_reg (ram_d_out r) ra v_d_out) u_en) ar -| exec_ram_None: - forall r a, - exec_ram r a None r a. + exec_ram ra ar r (Verilog.nonblock_reg (ram_en r) + (Verilog.nonblock_reg (ram_d_out r) ra v_d_out) u_en) ar. Inductive step : genv -> state -> Events.trace -> state -> Prop := | step_module : @@ -233,6 +222,10 @@ Definition semantics (m : program) := Smallstep.Semantics step (initial_state m) final_state (Globalenvs.Genv.globalenv m). +Definition all_module_regs m := + all_ram_regs (mod_ram m) ++ + (mod_st m::mod_finish m::mod_return m::mod_start m::mod_reset m::mod_clk m::nil). + Definition max_pc_function (m: module) := List.fold_left Pos.max (List.map fst (PTree.elements m.(mod_controllogic))) 1. @@ -253,17 +246,20 @@ Definition max_reg_ram r := (Pos.max (ram_d_out ram) (ram_u_en ram))))))) end. -Definition max_reg_module m := +Definition max_reg_body m := Pos.max (max_list (mod_params m)) - (Pos.max (max_stmnt_tree (mod_datapath m)) - (Pos.max (max_stmnt_tree (mod_controllogic m)) - (Pos.max (mod_st m) - (Pos.max (mod_stk m) - (Pos.max (mod_finish m) - (Pos.max (mod_return m) - (Pos.max (mod_start m) - (Pos.max (mod_reset m) - (Pos.max (mod_clk m) (max_reg_ram (mod_ram m))))))))))). + (Pos.max (max_stmnt_tree (mod_datapath m)) + (max_stmnt_tree (mod_controllogic m))). + +Definition max_reg_module m := + Pos.max (max_reg_body m) (max_list (all_module_regs m)). + +Record wf_htl_module m := + mk_wf_htl_module { + mod_wf : map_well_formed (mod_controllogic m) /\ map_well_formed (mod_datapath m); + mod_ordering_wf : list_norepet (all_module_regs m); + mod_gt : Forall (Pos.lt (max_reg_body m)) (all_module_regs m); + }. Lemma max_fold_lt : forall m l n, m <= n -> m <= (fold_left Pos.max l n). @@ -320,12 +316,16 @@ Proof. intros. apply max_reg_stmnt_le_stmnt_tree in H; lia. Qed. Lemma max_stmnt_lt_module : forall m d i, + wf_htl_module m -> (mod_controllogic m) ! i = Some d \/ (mod_datapath m) ! i = Some d -> Verilog.max_reg_stmnt d < max_reg_module m + 1. Proof. - inversion 1 as [ Hv | Hv ]; unfold max_reg_module; - apply max_reg_stmnt_le_stmnt_tree in Hv; lia. -Qed. + intros. apply mod_gt in H. + unfold Pos.lt, max_reg_body, max_reg_module, max_list, all_module_regs, all_ram_regs in *. + simplify. + repeat match goal with H: Forall _ _ |- _ => inv H end. + inversion H0 as [Hv | Hv]; apply max_reg_stmnt_le_stmnt_tree in Hv. + Admitted. Lemma max_list_correct l st : st > max_list l -> Forall (Pos.gt st) l. Proof. induction l; crush; constructor; [|apply IHl]; lia. Qed. @@ -339,3 +339,13 @@ Definition max_list_dec (l: list reg) (st: reg) : {Forall (Pos.gt st) l} + {True ); auto. apply max_list_correct. apply Pos.ltb_lt in e. lia. Qed. + +Variant wf_htl_fundef: fundef -> Prop := + | wf_htl_fundef_external: forall ef, + wf_htl_fundef (AST.External ef) + | wf_htl_function_internal: forall f, + wf_htl_module f -> + wf_htl_fundef (AST.Internal f). + +Definition wf_htl_program (p: program) : Prop := + forall f id, In (id, AST.Gfun f) (AST.prog_defs p) -> wf_htl_fundef f. diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index 0695f07..8c85701 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -145,6 +145,30 @@ Definition declare_reg (i : option io) (r : reg) (sz : nat) : mon unit := s.(st_controllogic)) (declare_reg_state_incr i s r sz). +Lemma declare_arr_state_incr : + forall i s r sz ln, + st_incr s + (mkstate + s.(st_st) + s.(st_freshreg) + s.(st_freshstate) + s.(st_scldecls) + (AssocMap.set r (i, VArray sz ln) s.(st_arrdecls)) + s.(st_datapath) + s.(st_controllogic)). +Proof. Admitted. + +Definition declare_arr (i : option io) (r : reg) (sz : nat) (ln : nat) : mon unit := + fun s => OK tt (mkstate + s.(st_st) + s.(st_freshreg) + s.(st_freshstate) + s.(st_scldecls) + (AssocMap.set r (i, VArray sz ln) s.(st_arrdecls)) + s.(st_datapath) + s.(st_controllogic)) + (declare_arr_state_incr i s r sz ln). + Lemma add_instr_state_incr : forall s n n' st, (st_controllogic s)!n = None -> @@ -763,7 +787,7 @@ Definition transf_bblock (fin rtrn preg: reg) (ni : node * bblock) | _ => translate_cfi fin rtrn preg (nstate, bb.(bb_exit)) end. -Definition decide_order a b c d e f g : {module_ordering a b c d e f g} + {True}. +(*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 _ @@ -772,15 +796,24 @@ Definition decide_order a b c d e f g : {module_ordering a b c d e f g} + {True} simplify; repeat match goal with | H: context[(_ apply Pos.ltb_lt in H end; unfold module_ordering; auto. -Defined. +Defined.*) Lemma clk_greater : forall ram clk r', Some ram = Some r' -> (clk < ram_addr r')%positive. Proof. Admitted. -Definition transf_module (f: function) : mon HTL.module. - refine ( +Definition declare_ram (r: ram) : mon unit := + do _ <- declare_reg None r.(ram_en) 1; + do _ <- declare_reg None r.(ram_u_en) 1; + do _ <- declare_reg None r.(ram_wr_en) 1; + do _ <- declare_reg None r.(ram_addr) 32; + do _ <- declare_reg None r.(ram_d_in) 32; + do _ <- declare_reg None r.(ram_d_out) 32; + do _ <- declare_arr None r.(ram_mem) 32 r.(ram_size); + ret tt. + +Definition transf_module (f: function) : mon HTL.module := if stack_correct f.(fn_stacksize) then do fin <- create_reg (Some Voutput) 1; do rtrn <- create_reg (Some Voutput) 32; @@ -792,63 +825,36 @@ Definition transf_module (f: function) : mon HTL.module. do start <- create_reg (Some Vinput) 1; do rst <- create_reg (Some Vinput) 1; 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, + match get_ram 0 (fn_funct_units f) with + | Some (_, r) => + do _ <- declare_ram r; + 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, - max_list_dec (fn_params f) (st_st current_state), - get_ram 0 (fn_funct_units f) - with - | left LEDATA, left LECTRL, left MORD, left WFPARAMS, Some (i, ram) => - ret (HTL.mkmodule - f.(fn_params) - current_state.(st_datapath) - current_state.(st_controllogic) - f.(fn_entrypoint) - current_state.(st_st) - stack - stack_len - fin - rtrn - start - rst - clk - current_state.(st_scldecls) - current_state.(st_arrdecls) - (Some ram) - (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA)) - MORD - _ - WFPARAMS) - | left LEDATA, left LECTRL, left MORD, left WFPARAMS, _ => - ret (HTL.mkmodule - f.(fn_params) - current_state.(st_datapath) - current_state.(st_controllogic) - f.(fn_entrypoint) - current_state.(st_st) - stack - stack_len - fin - rtrn - start - rst - clk - current_state.(st_scldecls) - current_state.(st_arrdecls) - None - (conj (max_pc_wf _ LECTRL) (max_pc_wf _ LEDATA)) - MORD - _ - WFPARAMS) - | _, _, _, _, _ => error (Errors.msg "More than 2^32 states.") + max_list_dec (fn_params f) (st_st current_state) + with + | left LEDATA, left LECTRL, left WFPARAMS => + ret (HTL.mkmodule + f.(fn_params) + current_state.(st_datapath) + current_state.(st_controllogic) + f.(fn_entrypoint) + current_state.(st_st) + fin + rtrn + start + rst + clk + current_state.(st_scldecls) + current_state.(st_arrdecls) + r) + | _, _, _=> error (Errors.msg "More than 2^32 states.") + end + | _ => error (Errors.msg "Stack RAM not found.") end - else error (Errors.msg "Stack size misalignment.")). - apply clk_greater. - discriminate. -Defined. + else error (Errors.msg "Stack size misalignment."). Definition max_state (f: function) : state := let st := Pos.succ (max_reg_function f) in diff --git a/src/hls/HTLgen.v b/src/hls/HTLgen.v index b879c8d..c793b1b 100644 --- a/src/hls/HTLgen.v +++ b/src/hls/HTLgen.v @@ -584,7 +584,7 @@ Proof. simplify. transitivity (Z.pos (max_pc_map m)); eauto. Qed. -Definition decide_order a b c d e f g : {module_ordering a b c d e f g} + {True}. +(*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 _ @@ -593,7 +593,7 @@ Definition decide_order a b c d e f g : {module_ordering a b c d e f g} + {True} simplify; repeat match goal with | H: context[(_ apply Pos.ltb_lt in H end; unfold module_ordering; auto. -Defined. +Defined.*) Definition transf_module (f: function) : mon HTL.module. refine ( @@ -606,21 +606,24 @@ Definition transf_module (f: function) : mon HTL.module. do start <- create_reg (Some Vinput) 1; do rst <- create_reg (Some Vinput) 1; do clk <- create_reg (Some Vinput) 1; + do r_en <- create_reg None 1; + do r_u_en <- create_reg None 1; + do r_addr <- create_reg None 32; + do r_wr_en <- create_reg None 1; + do r_d_in <- create_reg None 32; + do r_d_out <- create_reg None 32; 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, max_list_dec (RTL.fn_params f) (st_st current_state) with - | left LEDATA, left LECTRL, left MORD, left WFPARAMS => + | left LEDATA, left LECTRL, left WFPARAMS => ret (HTL.mkmodule f.(RTL.fn_params) current_state.(st_datapath) current_state.(st_controllogic) f.(fn_entrypoint) current_state.(st_st) - stack - stack_len fin rtrn start @@ -628,12 +631,8 @@ 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)) - MORD - _ - WFPARAMS) - | _, _, _, _ => error (Errors.msg "More than 2^32 states.") + (mk_ram stack_len stack r_en r_u_en r_addr r_wr_en r_d_in r_d_out)) + | _, _, _ => error (Errors.msg "More than 2^32 states.") end else error (Errors.msg "Stack size misalignment.")); discriminate. Defined. diff --git a/src/hls/HTLgenproof.v b/src/hls/HTLgenproof.v index fc7af6b..9930f4d 100644 --- a/src/hls/HTLgenproof.v +++ b/src/hls/HTLgenproof.v @@ -34,6 +34,7 @@ Require vericert.hls.HTL. Require Import vericert.hls.HTLgen. Require Import vericert.hls.HTLgenspec. Require Import vericert.hls.ValueInt. +Require Import vericert.hls.FunctionalUnits. Require vericert.hls.Verilog. Require Import Lia. @@ -62,10 +63,10 @@ Definition state_st_wf (m : HTL.module) (s : HTL.state) := Inductive match_arrs (m : HTL.module) (f : RTL.function) (sp : Values.val) (mem : mem) : Verilog.assocmap_arr -> Prop := | match_arr : forall asa stack, - asa ! (m.(HTL.mod_stk)) = Some stack /\ + asa ! (m.(HTL.mod_ram).(ram_mem)) = Some stack /\ stack.(arr_length) = Z.to_nat (f.(RTL.fn_stacksize) / 4) /\ (forall ptr, - 0 <= ptr < Z.of_nat m.(HTL.mod_stk_len) -> + 0 <= ptr < Z.of_nat m.(HTL.mod_ram).(ram_size) -> opt_val_value_lessdef (Mem.loadv AST.Mint32 mem (Values.Val.offset_ptr sp (Integers.Ptrofs.repr (4 * ptr)))) (Option.default (NToValue 0) @@ -1022,7 +1023,7 @@ Section CORRECTNESS. Definition transl_instr_prop (instr : RTL.instruction) : Prop := forall m asr asa fin rtrn st stmt trans res, - tr_instr fin rtrn st (m.(HTL.mod_stk)) instr stmt trans -> + tr_instr fin rtrn st (m.(HTL.mod_ram).(ram_mem)) instr stmt trans -> exists asr' asa', HTL.step tge (HTL.State res m st asr asa) Events.E0 (HTL.State res m st asr' asa'). @@ -1098,9 +1099,9 @@ Section CORRECTNESS. econstructor. econstructor. - all: invert MARR; big_tac. + all: invert MARR; big_tac. Abort. - inv CONST; constructor; simplify; rewrite AssocMap.gso; auto; lia. +(* inv CONST; constructor; simplify; rewrite AssocMap.gso; auto; lia. Unshelve. exact tt. Qed. @@ -2862,13 +2863,13 @@ Section CORRECTNESS. Proof. induction 1; eauto with htlproof; (intros; inv_state). Qed. - #[local] Hint Resolve transl_step_correct : htlproof. + #[local] Hint Resolve transl_step_correct : htlproof.*) Theorem transf_program_correct: Smallstep.forward_simulation (RTL.semantics prog) (HTL.semantics tprog). Proof. eapply Smallstep.forward_simulation_plus; eauto with htlproof. apply senv_preserved. - Qed. + (*Qed.*)Admitted. End CORRECTNESS. diff --git a/src/hls/HTLgenspec.v b/src/hls/HTLgenspec.v index 75d5321..c7dbe72 100644 --- a/src/hls/HTLgenspec.v +++ b/src/hls/HTLgenspec.v @@ -179,12 +179,13 @@ Inductive tr_code (c : RTL.code) (pc : RTL.node) (i : RTL.instruction) (stmnts t Inductive tr_module (f : RTL.function) : module -> Prop := tr_module_intro : - forall data control fin rtrn st stk stk_len m start rst clk scldecls arrdecls wf1 wf2 wf3 wf4, + forall data control fin rtrn st stk stk_len m start rst clk scldecls arrdecls r_en r_u_en r_addr r_wr_en r_d_in r_d_out, m = (mkmodule f.(RTL.fn_params) data control f.(RTL.fn_entrypoint) - st stk stk_len fin rtrn start rst clk scldecls arrdecls None wf1 wf2 wf3 wf4) -> + st fin rtrn start rst clk scldecls arrdecls + (mk_ram stk_len stk r_en r_u_en r_addr r_wr_en r_d_in r_d_out)) -> (forall pc i, Maps.PTree.get pc f.(RTL.fn_code) = Some i -> tr_code f.(RTL.fn_code) pc i data control fin rtrn st stk) -> stk_len = Z.to_nat (f.(RTL.fn_stacksize) / 4) -> @@ -197,6 +198,12 @@ Inductive tr_module (f : RTL.function) : module -> Prop := start = ((RTL.max_reg_function f) + 5)%positive -> rst = ((RTL.max_reg_function f) + 6)%positive -> clk = ((RTL.max_reg_function f) + 7)%positive -> + r_en = ((RTL.max_reg_function f) + 8)%positive -> + r_u_en = ((RTL.max_reg_function f) + 9)%positive -> + r_addr = ((RTL.max_reg_function f) + 10)%positive -> + r_wr_en = ((RTL.max_reg_function f) + 11)%positive -> + r_d_in = ((RTL.max_reg_function f) + 12)%positive -> + r_d_out = ((RTL.max_reg_function f) + 13)%positive -> tr_module f m. #[local] Hint Constructors tr_module : htlspec. @@ -645,7 +652,7 @@ Proof. assert (EQ3D := EQ3). apply collect_declare_datapath_trans in EQ3. apply collect_declare_controllogic_trans in EQ3D. - replace (st_controllogic s10) with (st_controllogic s3) by congruence. + (*replace (st_controllogic s10) with (st_controllogic s3) by congruence. replace (st_datapath s10) with (st_datapath s3) by congruence. replace (st_st s10) with (st_st s3) by congruence. eapply iter_expand_instr_spec; eauto with htlspec. @@ -655,3 +662,4 @@ Proof. erewrite <- collect_declare_freshreg_trans; try eassumption. lia. Qed. +*)Admitted. diff --git a/src/hls/Memorygen.v b/src/hls/Memorygen.v index 4ff4a19..e28bbc7 100644 --- a/src/hls/Memorygen.v +++ b/src/hls/Memorygen.v @@ -44,7 +44,7 @@ Local Open Scope assocmap. #[local] Hint Resolve max_reg_stmnt_lt_stmnt_tree : mgen. #[local] Hint Resolve max_stmnt_lt_module : mgen. -Lemma int_eq_not_false : forall x, Int.eq x (Int.not x) = false. +(*Lemma int_eq_not_false : forall x, Int.eq x (Int.not x) = false. Proof. intros. unfold Int.eq. rewrite Int.unsigned_not. @@ -3202,3 +3202,4 @@ Section CORRECTNESS. Qed. End CORRECTNESS. +*) diff --git a/src/hls/Veriloggen.v b/src/hls/Veriloggen.v index 035e7a4..c07e18d 100644 --- a/src/hls/Veriloggen.v +++ b/src/hls/Veriloggen.v @@ -57,9 +57,8 @@ Definition inst_ram clk ram := Definition transl_module (m : HTL.module) : Verilog.module := let case_el_ctrl := list_to_stmnt (transl_list (PTree.elements m.(mod_controllogic))) in let case_el_data := list_to_stmnt (transl_list (PTree.elements m.(mod_datapath))) in - match m.(HTL.mod_ram) with - | Some ram => - let body := + let ram := m.(HTL.mod_ram) in + let body := Valways (Vposedge m.(HTL.mod_clk)) (Vcond (Vbinop Veq (Vvar m.(HTL.mod_reset)) (Vlit (ZToValue 1))) (Vnonblock (Vvar m.(HTL.mod_st)) (Vlit (posToValue m.(HTL.mod_entrypoint)))) (Vcase (Vvar m.(HTL.mod_st)) case_el_ctrl (Some Vskip))) @@ -73,31 +72,11 @@ Definition transl_module (m : HTL.module) : Verilog.module := m.(HTL.mod_finish) m.(HTL.mod_return) m.(HTL.mod_st) - m.(HTL.mod_stk) - m.(HTL.mod_stk_len) + m.(HTL.mod_ram).(ram_mem) + m.(HTL.mod_ram).(ram_size) m.(HTL.mod_params) body - m.(HTL.mod_entrypoint) - | None => - let body := - Valways (Vposedge m.(HTL.mod_clk)) (Vcond (Vbinop Veq (Vvar m.(HTL.mod_reset)) (Vlit (ZToValue 1))) - (Vnonblock (Vvar m.(HTL.mod_st)) (Vlit (posToValue m.(HTL.mod_entrypoint)))) - (Vcase (Vvar m.(HTL.mod_st)) case_el_ctrl (Some Vskip))) - :: Valways (Vposedge m.(HTL.mod_clk)) (Vcase (Vvar m.(HTL.mod_st)) case_el_data (Some Vskip)) - :: List.map Vdeclaration (arr_to_Vdeclarr (AssocMap.elements m.(mod_arrdecls)) - ++ scl_to_Vdecl (AssocMap.elements m.(mod_scldecls))) in - Verilog.mkmodule m.(HTL.mod_start) - m.(HTL.mod_reset) - m.(HTL.mod_clk) - m.(HTL.mod_finish) - m.(HTL.mod_return) - m.(HTL.mod_st) - m.(HTL.mod_stk) - m.(HTL.mod_stk_len) - m.(HTL.mod_params) - body - m.(HTL.mod_entrypoint) - end. + m.(HTL.mod_entrypoint). Definition transl_fundef := transf_fundef transl_module. diff --git a/src/hls/Veriloggenproof.v b/src/hls/Veriloggenproof.v index d1494ec..90cf4cb 100644 --- a/src/hls/Veriloggenproof.v +++ b/src/hls/Veriloggenproof.v @@ -241,7 +241,7 @@ Proof. unfold transl_module; intros; destruct (HTL.mod_ram m) eqn:?; crush. Qed. Lemma mod_st_equiv m : mod_st (transl_module m) = HTL.mod_st m. Proof. unfold transl_module; intros; destruct (HTL.mod_ram m) eqn:?; crush. Qed. -Lemma mod_stk_equiv m : mod_stk (transl_module m) = HTL.mod_stk m. +(*Lemma mod_stk_equiv m : mod_stk (transl_module m) = HTL.mod_stk m. Proof. unfold transl_module; intros; destruct (HTL.mod_ram m) eqn:?; crush. Qed. Lemma mod_stk_len_equiv m : mod_stk_len (transl_module m) = HTL.mod_stk_len m. @@ -335,7 +335,7 @@ Proof. Qed. -Section CORRECTNESS. +*)Section CORRECTNESS. Variable prog: HTL.program. Variable tprog: program. @@ -345,7 +345,7 @@ Section CORRECTNESS. Let ge : HTL.genv := Globalenvs.Genv.globalenv prog. Let tge : genv := Globalenvs.Genv.globalenv tprog. - Lemma symbols_preserved: +(* Lemma symbols_preserved: forall (s: AST.ident), Genv.find_symbol tge s = Genv.find_symbol ge s. Proof. intros. eapply (Genv.find_symbol_match TRANSL). Qed. #[local] Hint Resolve symbols_preserved : verilogproof. @@ -532,13 +532,13 @@ Section CORRECTNESS. Proof. intros. inv H0. inv H. inv H3. constructor. reflexivity. Qed. - #[local] Hint Resolve transl_final_states : verilogproof. + #[local] Hint Resolve transl_final_states : verilogproof.*) Theorem transf_program_correct: forward_simulation (HTL.semantics prog) (Verilog.semantics tprog). Proof. eapply Smallstep.forward_simulation_plus; eauto with verilogproof. - apply senv_preserved. - Qed. + (*apply senv_preserved. + Qed.*) Admitted. End CORRECTNESS. -- cgit From 1c70273a90f4917f8e2385f80ceff7e2a9864209 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:15:15 +0000 Subject: Remove unnecessary proof from RTLParFUgen --- src/hls/RTLParFUgen.v | 1 - 1 file changed, 1 deletion(-) (limited to 'src/hls') diff --git a/src/hls/RTLParFUgen.v b/src/hls/RTLParFUgen.v index 65ddf74..49ee6e7 100644 --- a/src/hls/RTLParFUgen.v +++ b/src/hls/RTLParFUgen.v @@ -161,7 +161,6 @@ Definition transl_function (f: RTLPar.function): Errors.res RTLParFU.function := (max+6)%positive (max+4)%positive (max+5)%positive - ltac:(lia) )) initial_resources in do c' <- transl_code fu (RTLBlockInstr.fn_code f); Errors.OK (mkfunction (RTLBlockInstr.fn_sig f) -- cgit From f88d08c460b08f990e221e6aaba5186330b72ef0 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:15:28 +0000 Subject: Fix operation chaining in scheduler It now respects the delays properly. --- src/hls/Schedule.ml | 20 ++++++++++++-------- 1 file changed, 12 insertions(+), 8 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index b75465b..e557d4b 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -305,7 +305,7 @@ let add_dep map i tree dfg curr = | None -> dfg | Some ip -> let ipv = (List.nth map ip) in - DFG.add_edge_e dfg (ipv, comb_delay (snd ipv), List.nth map i) + DFG.add_edge_e dfg (ipv, comb_delay (snd (List.nth map i)), List.nth map i) (** This function calculates the dependencies of each instruction. The nodes correspond to previous registers that were allocated and show which instruction caused it. @@ -408,7 +408,7 @@ let accumulate_WAW_mem_deps instrs dfg curri = | RBstore (_, _, _, _, _) -> ( match next_store 0 (take i instrs |> List.rev) with | None -> dfg - | Some d -> DFG.add_edge dfg (gen_vertex instrs (i - d - 1)) (gen_vertex instrs i) ) + | Some d -> DFG.add_edge dfg (gen_vertex instrs (i - d - 1)) (gen_vertex instrs i)) | _ -> dfg (** Predicate dependencies. *) @@ -622,18 +622,22 @@ let negate_graph constr = DFG.add_edge_e g (v1, -e, v2) ) constr DFG.empty -let add_cycle_constr max n dfg constr = +let add_cycle_constr maxf n dfg constr = let negated_dfg = negate_graph dfg in + let max_initial_del = DFG.fold_vertex (fun v1 g -> + if DFG.in_degree dfg v1 = 0 + then max g (comb_delay (snd v1)) + else g) dfg 0 in let longest_path v = BFDFG.all_shortest_paths negated_dfg v - |> BFDFG.H.to_seq |> List.of_seq in - let constrained_paths = List.filter (function (_, m) -> - m > max) in + |> BFDFG.H.to_seq |> List.of_seq + |> List.map (function (x, y) -> (x, y - max_initial_del)) in + let constrained_paths = List.filter (function (_, m) -> - m > maxf) in List.fold_left (fun g -> function (v, v', w) -> G.add_edge_e g (encode_var n (fst v) 0, - - (int_of_float (Float.ceil (Float.div (float_of_int w) (float_of_int max))) - 1), + - (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 |> constrained_paths |> List.map (function (v', w) -> (v, v', - w))) ) dfg []) type resource = -- cgit From 47acd7870b8a577349016502779799c77fc91734 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 18 Nov 2021 22:16:06 +0000 Subject: Add a printer for RTLParFU --- src/hls/PrintRTLParFU.ml | 120 +++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 120 insertions(+) create mode 100644 src/hls/PrintRTLParFU.ml (limited to 'src/hls') diff --git a/src/hls/PrintRTLParFU.ml b/src/hls/PrintRTLParFU.ml new file mode 100644 index 0000000..ec4d851 --- /dev/null +++ b/src/hls/PrintRTLParFU.ml @@ -0,0 +1,120 @@ +(* *********************************************************************) +(* *) +(* The Compcert verified compiler *) +(* *) +(* Xavier Leroy, INRIA Paris-Rocquencourt *) +(* *) +(* Copyright Institut National de Recherche en Informatique et en *) +(* Automatique. All rights reserved. This file is distributed *) +(* under the terms of the INRIA Non-Commercial License Agreement. *) +(* *) +(* *********************************************************************) + +(** Pretty-printers for RTL code *) + +open Printf +open Camlcoq +open Datatypes +open Maps +open AST +open PrintRTLBlockInstr +open RTLParFU +open PrintAST + +(* Printing of RTL code *) + +let reg pp r = + fprintf pp "x%d" (P.to_int r) + +let rec regs pp = function + | [] -> () + | [r] -> reg pp r + | r1::rl -> fprintf pp "%a, %a" reg r1 regs rl + +let ros pp = function + | Coq_inl r -> reg pp r + | Coq_inr s -> fprintf pp "\"%s\"" (extern_atom s) + +let print_bblock_body pp i = + fprintf pp "\t\t"; + match i with + | FUnop -> fprintf pp "nop\n" + | FUop(p, op, ls, dst) -> + fprintf pp "%a %a = %a\n" + print_pred_option p reg dst (PrintOp.print_operation reg) (op, ls) + | FUsetpred (p', c, args, p) -> + fprintf pp "%a %a = %a\n" + print_pred_option p' + pred p + (PrintOp.print_condition reg) (c, args) + | _ -> assert false + +let rec print_bblock_exit pp i = + fprintf pp "\t\t"; + match i with + | FUcall(_, fn, args, res, _) -> + fprintf pp "%a = %a(%a)\n" + reg res ros fn regs args; + | FUtailcall(_, fn, args) -> + fprintf pp "tailcall %a(%a)\n" + ros fn regs args + | FUbuiltin(ef, args, res, _) -> + fprintf pp "%a = %s(%a)\n" + (print_builtin_res reg) res + (name_of_external ef) + (print_builtin_args reg) args + | FUcond(cond, args, s1, s2) -> + fprintf pp "if (%a) goto %d else goto %d\n" + (PrintOp.print_condition reg) (cond, args) + (P.to_int s1) (P.to_int s2) + | FUjumptable(arg, tbl) -> + let tbl = Array.of_list tbl in + fprintf pp "jumptable (%a)\n" reg arg; + for i = 0 to Array.length tbl - 1 do + fprintf pp "\tcase %d: goto %d\n" i (P.to_int tbl.(i)) + done + | FUreturn None -> + fprintf pp "return\n" + | FUreturn (Some arg) -> + fprintf pp "return %a\n" reg arg + | FUgoto n -> + fprintf pp "goto %d\n" (P.to_int n) + | FUpred_cf (p, c1, c2) -> + fprintf pp "if %a then (%a) else (%a)\n" print_pred_op p print_bblock_exit c1 print_bblock_exit c2 + +let print_bblock pp (pc, i) = + fprintf pp "%5d:{\n" pc; + List.iter (fun x -> fprintf pp "{\n"; + List.iter (fun x -> fprintf pp "( "; List.iter (print_bblock_body pp) x; fprintf pp " )") x; + fprintf pp "}\n") i.bb_body; + print_bblock_exit pp i.bb_exit; + fprintf pp "\t}\n\n" + +let print_function pp id f = + fprintf pp "%s(%a) {\n" (extern_atom id) regs f.fn_params; + let instrs = + List.sort + (fun (pc1, _) (pc2, _) -> compare pc2 pc1) + (List.rev_map + (fun (pc, i) -> (P.to_int pc, i)) + (PTree.elements f.fn_code)) in + List.iter (print_bblock pp) instrs; + fprintf pp "}\n\n" + +let print_globdef pp (id, gd) = + match gd with + | Gfun(Internal f) -> print_function pp id f + | _ -> () + +let print_program pp (prog: program) = + List.iter (print_globdef pp) prog.prog_defs + +let destination : string option ref = ref None + +let print_if passno prog = + match !destination with + | None -> () + | Some f -> + let oc = open_out (f ^ "." ^ Z.to_string passno) in + print_program oc prog; + close_out oc -- cgit From b91f6db17ee30efd2068efbeecbf1d2b4c3850ea Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 9 Dec 2021 20:06:34 +0000 Subject: Add bourdoncle to build --- src/hls/IfConversion.v | 36 ++++++++++++++++++++++++++++++++---- src/hls/Pipeline.v | 3 ++- src/hls/PipelineOp.v | 3 ++- src/hls/Schedule.ml | 5 ++++- 4 files changed, 40 insertions(+), 7 deletions(-) (limited to 'src/hls') 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 . *) -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 = -- cgit From 8d71333042d9ed87a80cffd4005daa0a0acc1810 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 25 Feb 2022 13:56:59 +0000 Subject: Start converting comments --- src/hls/IfConversion.v | 8 ++------ src/hls/RTLPargen.v | 20 +++++--------------- 2 files changed, 7 insertions(+), 21 deletions(-) (limited to 'src/hls') diff --git a/src/hls/IfConversion.v b/src/hls/IfConversion.v index 4585770..9b8f4d4 100644 --- a/src/hls/IfConversion.v +++ b/src/hls/IfConversion.v @@ -30,14 +30,10 @@ Require Import vericert.bourdoncle.Bourdoncle. Parameter build_bourdoncle : function -> (bourdoncle * PMap.t N). -(*| -============= -If conversion -============= +(** * If conversion This conversion is a verified conversion from RTLBlock back to itself, which performs if-conversion -on basic blocks to make basic blocks larger. -|*) +on basic blocks to make basic blocks larger. *) Definition combine_pred (p: pred_op) (optp: option pred_op) := match optp with diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index 58b048c..82f3f5e 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -34,24 +34,17 @@ Require Import vericert.hls.Predicate. Require Import vericert.hls.Abstr. Import NE.NonEmptyNotation. -(*| -================= -RTLPar Generation -================= -|*) +(** * RTLPar Generation *) #[local] Open Scope positive. #[local] Open Scope forest. #[local] Open Scope pred_op. -(*| -Abstract Computations -===================== +(** ** Abstract Computations Define the abstract computation using the ``update`` function, which will set each register to its symbolic value. First we need to define a few helper functions to correctly translate the -predicates. -|*) +predicates. *) Fixpoint list_translation (l : list reg) (f : forest) {struct l} : list pred_expr := match l with @@ -146,9 +139,7 @@ Definition app_predicated {A: Type} (p: option pred_op) (a b: predicated A) := Definition pred_ret {A: Type} (a: A) : predicated A := NE.singleton (T, a). -(*| -Update Function ---------------- +(** *** Update Function The ``update`` function will generate a new forest given an existing forest and a new instruction, so that it can evaluate a symbolic expression by folding over a list of instructions. The main @@ -158,8 +149,7 @@ problem is that predicates need to be merged as well, so that: 2. The expression assigned to the register should still be correct. This is done by multiplying the predicates together, and assigning the negation of the expression to -the other predicates. -|*) +the other predicates. *) Definition update (f : forest) (i : instr) : forest := match i with -- cgit From 5f34267c4bccb471c71fd5698ec49adc99940850 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 25 Feb 2022 18:32:35 +0000 Subject: Fix up some more documentation --- src/hls/Abstr.v | 39 ++++++++--------------- src/hls/HTL.v | 11 +++---- src/hls/HTLgenspec.v | 2 +- src/hls/RTLBlockInstr.v | 45 ++++++++------------------- src/hls/RTLPargen.v | 30 +++++++----------- src/hls/RTLPargenproof.v | 7 ++--- src/hls/Verilog.v | 80 +++++++++++++++--------------------------------- 7 files changed, 68 insertions(+), 146 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v index 2ab79cf..01f2fd9 100644 --- a/src/hls/Abstr.v +++ b/src/hls/Abstr.v @@ -38,14 +38,11 @@ Require Import vericert.hls.Predicate. #[local] Open Scope positive. #[local] Open Scope pred_op. -(*| -Schedule Oracle -=============== +(** ** Schedule Oracle This oracle determines if a schedule was valid by performing symbolic execution on the input and output and showing that these behave the same. This acts on each basic block separately, as the -rest of the functions should be equivalent. -|*) +rest of the functions should be equivalent. *) Definition reg := positive. @@ -54,10 +51,8 @@ Inductive resource : Set := | Pred : reg -> resource | Mem : resource. -(*| -The following defines quite a few equality comparisons automatically, however, these can be -optimised heavily if written manually, as their proofs are not needed. -|*) +(** The following defines quite a few equality comparisons automatically, however, these can be +optimised heavily if written manually, as their proofs are not needed. *) Lemma resource_eq : forall (r1 r2 : resource), {r1 = r2} + {r1 <> r2}. Proof. @@ -179,11 +174,9 @@ Proof. repeat decide equality. Defined. -(*| -We then create equality lemmas for a resource and a module to index resources uniquely. The +(** We then create equality lemmas for a resource and a module to index resources uniquely. The indexing is done by setting Mem to 1, whereas all other infinitely many registers will all be -shifted right by 1. This means that they will never overlap. -|*) +shifted right by 1. This means that they will never overlap. *) Module R_indexed. Definition t := resource. @@ -200,8 +193,7 @@ Module R_indexed. Definition eq := resource_eq. End R_indexed. -(*| -We can then create expressions that mimic the expressions defined in RTLBlock and RTLPar, which use +(** We can then create expressions that mimic the expressions defined in RTLBlock and RTLPar, which use expressions instead of registers as their inputs and outputs. This means that we can accumulate all the results of the operations as general expressions that will be present in those registers. @@ -211,8 +203,7 @@ the results of the operations as general expressions that will be present in tho - Estore: A store from a register to a memory location. Then, to make recursion over expressions easier, expression_list is also defined in the datatype, as -that enables mutual recursive definitions over the datatypes. -|*) +that enables mutual recursive definitions over the datatypes. *) Inductive expression : Type := | Ebase : resource -> expression @@ -321,10 +312,8 @@ Definition predicated A := NE.non_empty (pred_op * A). Definition pred_expr := predicated expression. -(*| -Using IMap we can create a map from resources to any other type, as resources can be uniquely -identified as positive numbers. -|*) +(** Using IMap we can create a map from resources to any other type, as resources can be uniquely +identified as positive numbers. *) Module Rtree := ITree(R_indexed). @@ -356,10 +345,8 @@ Definition get_m i := match i with mk_instr_state a b c => c end. Definition eval_predf_opt pr p := match p with Some p' => eval_predf pr p' | None => true end. -(*| -Finally we want to define the semantics of execution for the expressions with symbolic values, so -the result of executing the expressions will be an expressions. -|*) +(** Finally we want to define the semantics of execution for the expressions with symbolic values, +so the result of executing the expressions will be an expressions. *) Section SEMANTICS. @@ -682,7 +669,7 @@ Inductive similar {A B} : @ctx A -> @ctx B -> Prop := Definition beq_pred_expr_once (pe1 pe2: pred_expr) : bool := match pe1, pe2 with - (*| PEsingleton None e1, PEsingleton None e2 => beq_expression e1 e2 + (* PEsingleton None e1, PEsingleton None e2 => beq_expression e1 e2 | PEsingleton (Some p1) e1, PEsingleton (Some p2) e2 => if beq_expression e1 e2 then match sat_pred_simple bound (Por (Pand p1 (Pnot p2)) (Pand p2 (Pnot p1))) with diff --git a/src/hls/HTL.v b/src/hls/HTL.v index 551c66e..47f2092 100644 --- a/src/hls/HTL.v +++ b/src/hls/HTL.v @@ -36,13 +36,10 @@ Require Import ValueInt. Local Open Scope positive. -(*| -The purpose of the hardware transfer language (HTL) is to create a more -hardware-like layout that is still similar to the register transfer language -(RTL) that it came from. The main change is that function calls become module -instantiations and that we now describe a state machine instead of a -control-flow graph. -|*) +(** The purpose of the hardware transfer language (HTL) is to create a more hardware-like layout +that is still similar to the register transfer language (RTL) that it came from. The main change is +that function calls become module instantiations and that we now describe a state machine instead of +a control-flow graph. *) Local Open Scope assocmap. diff --git a/src/hls/HTLgenspec.v b/src/hls/HTLgenspec.v index c7dbe72..265b8f7 100644 --- a/src/hls/HTLgenspec.v +++ b/src/hls/HTLgenspec.v @@ -160,7 +160,7 @@ Inductive tr_instr (fin rtrn st stk : reg) : RTL.instruction -> stmnt -> stmnt - translate_arr_access mem addr args stk s = OK c s' i -> tr_instr fin rtrn st stk (RTL.Istore mem addr args src n) (Vnonblock c (Vvar src)) (state_goto st n). -(*| tr_instr_Ijumptable : +(* tr_instr_Ijumptable : forall cexpr tbl r, cexpr = tbl_to_case_expr st tbl -> tr_instr fin rtrn st stk (RTL.Ijumptable r tbl) (Vskip) (Vcase (Vvar r) cexpr (Some Vskip)).*) diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index d9f3e74..7eabcf7 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -31,22 +31,18 @@ Require Import compcert.verilog.Op. Require Import Predicate. Require Import Vericertlib. -(*| -===================== -RTLBlock Instructions -===================== +(** * RTLBlock Instructions These instructions are used for ``RTLBlock`` and ``RTLPar``, so that they have consistent instructions, which greatly simplifies the proofs, as they will by default have the same instruction syntax and semantics. The only changes are therefore at the top-level of the instructions. -Instruction Definition -====================== +** Instruction Definition First, we define the instructions that can be placed into a basic block, meaning they won't branch. The main changes to how instructions are defined in ``RTL``, is that these instructions don't have a next node, as they will be in a basic block, and they also have an optional predicate (``pred_op``). -|*) +*) Definition node := positive. @@ -57,13 +53,10 @@ Inductive instr : Type := | RBstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr | RBsetpred : option pred_op -> condition -> list reg -> predicate -> instr. -(*| -Control-Flow Instruction Definition -=================================== +(** ** Control-Flow Instruction Definition These are the instructions that count as control-flow, and will be placed at the end of the basic -blocks. -|*) +blocks. *) Inductive cf_instr : Type := | RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr @@ -76,10 +69,7 @@ Inductive cf_instr : Type := | RBgoto : node -> cf_instr | RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. -(*| -Helper functions -================ -|*) +(** ** Helper functions *) Fixpoint successors_instr (i : cf_instr) : list node := match i with @@ -170,17 +160,14 @@ Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset := | _, _ => Regmap.init Vundef end. -(*| -Instruction State ------------------ +(** *** Instruction State Definition of the instruction state, which contains the following: -:is_rs: This is the current state of the registers. -:is_ps: This is the current state of the predicate registers, which is in a separate namespace and - area compared to the standard registers in ``is_rs``. -:is_mem: The current state of the memory. -|*) +- [is_rs] This is the current state of the registers. +- [is_ps] This is the current state of the predicate registers, which is in a separate namespace and + area compared to the standard registers in [is_rs]. +- [is_mem] The current state of the memory. *) Record instr_state := mk_instr_state { is_rs: regset; @@ -188,10 +175,7 @@ Record instr_state := mk_instr_state { is_mem: mem; }. -(*| -Top-Level Type Definitions -========================== -|*) +(** ** Top-Level Type Definitions *) Section DEFINITION. @@ -256,10 +240,7 @@ Section DEFINITION. End DEFINITION. -(*| -Semantics -========= -|*) +(** ** Semantics *) Section RELSEM. diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index 82f3f5e..ab4c0da 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -42,7 +42,7 @@ Import NE.NonEmptyNotation. (** ** Abstract Computations -Define the abstract computation using the ``update`` function, which will set each register to its +Define the abstract computation using the [update] function, which will set each register to its symbolic value. First we need to define a few helper functions to correctly translate the predicates. *) @@ -141,7 +141,7 @@ Definition pred_ret {A: Type} (a: A) : predicated A := (** *** Update Function -The ``update`` function will generate a new forest given an existing forest and a new instruction, +The [update] function will generate a new forest given an existing forest and a new instruction, so that it can evaluate a symbolic expression by folding over a list of instructions. The main problem is that predicates need to be merged as well, so that: @@ -181,12 +181,10 @@ Definition update (f : forest) (i : instr) : forest := (map_predicated (pred_ret (Esetpred c)) (merge (list_translation args f)))) end. -(*| -Implementing which are necessary to show the correctness of the translation validation by showing -that there aren't any more effects in the resultant RTLPar code than in the RTLBlock code. +(** Implementing which are necessary to show the correctness of the translation validation by +showing that there aren't any more effects in the resultant RTLPar code than in the RTLBlock code. -Get a sequence from the basic block. -|*) +Get a sequence from the basic block. *) Fixpoint abstract_sequence (f : forest) (b : list instr) : forest := match b with @@ -194,18 +192,15 @@ Fixpoint abstract_sequence (f : forest) (b : list instr) : forest := | i :: l => abstract_sequence (update f i) l end. -(*| -Check equivalence of control flow instructions. As none of the basic blocks should have been moved, -none of the labels should be different, meaning the control-flow instructions should match exactly. -|*) +(** Check equivalence of control flow instructions. As none of the basic blocks should have been +moved, none of the labels should be different, meaning the control-flow instructions should match +exactly. *) Definition check_control_flow_instr (c1 c2: cf_instr) : bool := if cf_instr_eq c1 c2 then true else false. -(*| -We define the top-level oracle that will check if two basic blocks are equivalent after a scheduling -transformation. -|*) +(** We define the top-level oracle that will check if two basic blocks are equivalent after a +scheduling transformation. *) Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := match bb with @@ -246,10 +241,7 @@ Lemma check_scheduled_trees_correct2: PTree.get x f2 = None. Proof. solve_scheduled_trees_correct. Qed. -(*| -Top-level Functions -=================== -|*) +(** ** Top-level Functions *) Parameter schedule : RTLBlock.function -> RTLPar.function. diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v index 588f67f..0023edc 100644 --- a/src/hls/RTLPargenproof.v +++ b/src/hls/RTLPargenproof.v @@ -46,12 +46,9 @@ Inductive state_lessdef : instr_state -> instr_state -> Prop := (forall x, rs1 !! x = rs2 !! x) -> state_lessdef (mk_instr_state rs1 m1) (mk_instr_state rs2 m1). -(*| -RTLBlock to abstract translation --------------------------------- +(** *** RTLBlock to abstract translation -Correctness of translation from RTLBlock to the abstract interpretation language. -|*) +Correctness of translation from RTLBlock to the abstract interpretation language. *) Ltac inv_simp := repeat match goal with diff --git a/src/hls/Verilog.v b/src/hls/Verilog.v index 3a2c81d..7561f77 100644 --- a/src/hls/Verilog.v +++ b/src/hls/Verilog.v @@ -142,30 +142,23 @@ Definition nonblock_reg (r : reg) (asr : reg_associations) (v : value) := Inductive scl_decl : Type := VScalar (sz : nat). Inductive arr_decl : Type := VArray (sz : nat) (ln : nat). -(*| -Verilog AST -=========== +(** ** Verilog AST The Verilog AST is defined here, which is the target language of the compilation. -Value ------ +*** Value The Verilog [value] is a bitvector containg a size and the actual bitvector. In this case, the [Word.word] type is used as many theorems about that bitvector have already been proven. -|*) -(*| -Binary Operators ----------------- +*** Binary Operators These are the binary operations that can be represented in Verilog. Ideally, multiplication and division would be done by custom modules which could al so be scheduled properly. However, for now every Verilog operator is assumed to take -one clock cycle. -|*) +one clock cycle. *) Inductive binop : Type := | Vadd : binop (** addition (binary [+]) *) @@ -191,21 +184,15 @@ Inductive binop : Type := | Vshl : binop (** shift left ([<<]) *) | Vshr : binop (** shift right ([>>>]) *) | Vshru : binop. (** shift right unsigned ([>>]) *) -(*| Vror : binop (** shift right unsigned ([>>]) *)*) +(* Vror : binop (** shift right unsigned ([>>]) *)*) -(*| -Unary Operators ---------------- -|*) +(** *** Unary Operators *) Inductive unop : Type := | Vneg (** negation ([-]) *) | Vnot. (** not operation [!] *) -(*| -Expressions ------------ -|*) +(** *** Expressions *) Inductive expr : Type := | Vlit : value -> expr @@ -220,10 +207,7 @@ Inductive expr : Type := Definition posToExpr (p : positive) : expr := Vlit (posToValue p). -(*| -Statements ----------- -|*) +(** *** Statements *) Inductive stmnt : Type := | Vskip : stmnt @@ -236,17 +220,13 @@ with stmnt_expr_list : Type := | Stmntnil : stmnt_expr_list | Stmntcons : expr -> stmnt -> stmnt_expr_list -> stmnt_expr_list. -(*| -Edges ------ +(** *** Edges -These define when an always block should be triggered, for example if the always -block should be triggered synchronously at the clock edge, or asynchronously for -combinational logic. +These define when an always block should be triggered, for example if the always block should be +triggered synchronously at the clock edge, or asynchronously for combinational logic. -[edge] is not part of [stmnt] in this formalisation because is closer to the -semantics that are used. -|*) +[edge] is not part of [stmnt] in this formalisation because is closer to the semantics that are +used. *) Inductive edge : Type := | Vposedge : reg -> edge @@ -254,14 +234,11 @@ Inductive edge : Type := | Valledge : edge | Voredge : edge -> edge -> edge. -(*| -Module Items ------------- +(** *** Module Items -Module items can either be declarations ([Vdecl]) or always blocks ([Valways]). -The declarations are always register declarations as combinational logic can be -done using combinational always block instead of continuous assignments. -|*) +Module items can either be declarations ([Vdecl]) or always blocks ([Valways]). The declarations +are always register declarations as combinational logic can be done using combinational always block +instead of continuous assignments. *) Inductive io : Type := | Vinput : io @@ -278,8 +255,7 @@ Inductive module_item : Type := | Valways_ff : edge -> stmnt -> module_item | Valways_comb : edge -> stmnt -> module_item. -(*| -The main module type containing all the important control signals +(** The main module type containing all the important control signals - [mod_start] If set, starts the computations in the module. - [mod_reset] If set, reset the state in the module. @@ -287,8 +263,7 @@ The main module type containing all the important control signals - [mod_args] The arguments to the module. - [mod_finish] Bit that is set if the result is ready. - [mod_return] The return value that was computed. -- [mod_body] The body of the module, including the state machine. -|*) +- [mod_body] The body of the module, including the state machine. *) Record module : Type := mkmodule { mod_start : reg; @@ -308,21 +283,17 @@ Definition fundef := AST.fundef module. Definition program := AST.program fundef unit. -(*| -Convert a [positive] to an expression directly, setting it to the right size. -|*) +(** Convert a [positive] to an expression directly, setting it to the right size. *) Definition posToLit (p : positive) : expr := Vlit (posToValue p). Definition fext := unit. Definition fextclk := nat -> fext. -(*| -State ------ +(** *** State The [state] contains the following items: -n + - Current [module] that is being worked on, so that the state variable can be retrieved and set appropriately. - Current [module_item] which is being worked on. @@ -338,8 +309,7 @@ retrieved and set appropriately. Verilog, as the Verilog was generated by the RTL, which means that we have to make an assumption about how it looks. In this case, that it contains state which determines which part of the Verilog is executed. This is then the part - of the Verilog that should match with the RTL. -|*) + of the Verilog that should match with the RTL. *) Inductive stackframe : Type := Stackframe : @@ -459,9 +429,7 @@ Definition handle_def {A : Type} (a : A) (val : option A) Local Open Scope error_monad_scope. -(*| -Return the location of the lhs of an assignment. -|*) +(** Return the location of the lhs of an assignment. *) Inductive location : Type := | LocReg (_ : reg) -- cgit From 314e1178ccede8ed42cbfc14b68352a51dcd014b Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Fri, 25 Feb 2022 19:15:40 +0000 Subject: Final updates to the current documentation --- src/hls/HTLgenproof.v | 12 ++++++------ src/hls/RTLBlockInstr.v | 6 +++--- 2 files changed, 9 insertions(+), 9 deletions(-) (limited to 'src/hls') diff --git a/src/hls/HTLgenproof.v b/src/hls/HTLgenproof.v index 9930f4d..bf4b057 100644 --- a/src/hls/HTLgenproof.v +++ b/src/hls/HTLgenproof.v @@ -1008,12 +1008,12 @@ Section CORRECTNESS. is a simulation argument based on diagrams of the following form: << match_states - code st rs ---------------- State m st assoc - || | - || | - || | - \/ v - code st rs' --------------- State m st assoc' + code st rs ------------------------- State m st assoc + || | + || | + || | + \/ v + code st rs' ------------------------ State m st assoc' match_states >> where [tr_code c data control fin rtrn st] is assumed to hold. diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 7eabcf7..52259a0 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -33,15 +33,15 @@ Require Import Vericertlib. (** * RTLBlock Instructions -These instructions are used for ``RTLBlock`` and ``RTLPar``, so that they have consistent +These instructions are used for [RTLBlock] and [RTLPar], so that they have consistent instructions, which greatly simplifies the proofs, as they will by default have the same instruction syntax and semantics. The only changes are therefore at the top-level of the instructions. ** Instruction Definition First, we define the instructions that can be placed into a basic block, meaning they won't branch. -The main changes to how instructions are defined in ``RTL``, is that these instructions don't have a -next node, as they will be in a basic block, and they also have an optional predicate (``pred_op``). +The main changes to how instructions are defined in [RTL], is that these instructions don't have a +next node, as they will be in a basic block, and they also have an optional predicate ([pred_op]). *) Definition node := positive. -- cgit From c2ba535604cef5bc86369512e6f3c2833753a55a Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 2 Mar 2022 18:06:18 +0000 Subject: Update Coq version to 8.14.1 --- src/hls/AssocMap.v | 86 +++++++++++++++++++++++++++++++----------------------- 1 file changed, 49 insertions(+), 37 deletions(-) (limited to 'src/hls') diff --git a/src/hls/AssocMap.v b/src/hls/AssocMap.v index 8dbc6b2..10bd8eb 100644 --- a/src/hls/AssocMap.v +++ b/src/hls/AssocMap.v @@ -42,63 +42,75 @@ Module AssocMapExt. | Some v => v end. - Fixpoint merge (m1 m2 : t A) : t A := - match m1, m2 with - | Node l1 (Some a) r1, Node l2 _ r2 => Node (merge l1 l2) (Some a) (merge r1 r2) - | Node l1 None r1, Node l2 o r2 => Node (merge l1 l2) o (merge r1 r2) - | Leaf, _ => m2 - | _, _ => m1 + Definition merge_atom (new : option A) (old : option A) : option A := + match new, old with + | Some _, _ => new + | _, _ => old end. + Definition merge (m1 m2 : t A) : t A := + PTree.combine merge_atom m1 m2. + Lemma merge_base_1 : - forall am, - merge (empty A) am = am. - Proof. auto. Qed. + forall am x, + (merge (empty A) am) ! x = am ! x. + Proof using. + unfold merge; intros. + rewrite PTree.gcombine; eauto. + Qed. Lemma merge_base_2 : - forall am, - merge am (empty A) = am. - Proof. - unfold merge. - destruct am; trivial. - destruct o; trivial. + forall am x, + (merge am (empty A)) ! x = am ! x. + Proof using. + unfold merge; intros. + rewrite PTree.gcombine; eauto. + cbv [merge_atom]. destruct_match; crush. Qed. Lemma merge_add_assoc : - forall r am am' v, - merge (set r v am) am' = set r v (merge am am'). - Proof. - induction r; intros; destruct am; destruct am'; try (destruct o); simpl; - try rewrite IHr; try reflexivity. + forall r am am' v x, + (merge (set r v am) am') ! x = (set r v (merge am am')) ! x. + Proof using. + unfold merge. intros; rewrite PTree.gcombine by auto. + unfold merge_atom. destruct_match. + destruct (peq r x); subst. + rewrite PTree.gss in Heqo. inv Heqo. rewrite PTree.gss. auto. + rewrite PTree.gso in Heqo by auto. rewrite PTree.gso by auto. + rewrite PTree.gcombine by auto. rewrite Heqo. auto. + destruct (peq r x); subst. + rewrite PTree.gss in Heqo. inv Heqo. + rewrite PTree.gso in Heqo by auto. rewrite PTree.gso by auto. + rewrite PTree.gcombine by auto. rewrite Heqo. auto. Qed. Lemma merge_correct_1 : forall am bm k v, - get k am = Some v -> - get k (merge am bm) = Some v. - Proof. - induction am; intros; destruct k; destruct bm; try (destruct o); simpl; - try rewrite gempty in H; try discriminate; try assumption; auto. + am ! k = Some v -> + (merge am bm) ! k = Some v. + Proof using. + unfold merge; intros. rewrite PTree.gcombine by auto. + rewrite H; auto. Qed. Lemma merge_correct_2 : forall am bm k v, - get k am = None -> - get k bm = Some v -> - get k (merge am bm) = Some v. - Proof. - induction am; intros; destruct k; destruct bm; try (destruct o); simpl; - try rewrite gempty in H; try discriminate; try assumption; auto. + am ! k = None -> + bm ! k = Some v -> + (merge am bm) ! k = Some v. + Proof using. + unfold merge; intros. rewrite PTree.gcombine by auto. + rewrite H. rewrite H0. auto. Qed. Lemma merge_correct_3 : forall am bm k, - get k am = None -> - get k bm = None -> - get k (merge am bm) = None. - Proof. - induction am; intros; destruct k; destruct bm; try (destruct o); simpl; - try rewrite gempty in H; try discriminate; try assumption; auto. + am ! k = None -> + bm ! k = None -> + (merge am bm) ! k = None. + Proof using. + unfold merge; intros. rewrite PTree.gcombine by auto. + rewrite H. rewrite H0. auto. Qed. Definition merge_fold (am bm : t A) : t A := -- cgit From 2e1cdc46f0e5275fadba200a54f4862972086f59 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 3 Mar 2022 14:39:08 +0000 Subject: Fix Abstr --- src/hls/Abstr.v | 47 +++++++++++++++++------------------------------ 1 file changed, 17 insertions(+), 30 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v index 01f2fd9..9ded7fe 100644 --- a/src/hls/Abstr.v +++ b/src/hls/Abstr.v @@ -1364,19 +1364,25 @@ Proof. end; discriminate. Qed. -Fixpoint beq2 {A B : Type} (beqA : A -> B -> bool) (m1 : PTree.t A) (m2 : PTree.t B) {struct m1} : bool := +Fixpoint beq2' {A B : Type} (beqA : A -> B -> bool) (m1: PTree.tree' A) (m2: PTree.tree' B) {struct m1} : bool := match m1, m2 with - | PTree.Leaf, _ => PTree.bempty m2 - | _, PTree.Leaf => PTree.bempty m1 - | PTree.Node l1 o1 r1, PTree.Node l2 o2 r2 => - match o1, o2 with - | None, None => true - | Some y1, Some y2 => beqA y1 y2 - | _, _ => false - end - && beq2 beqA l1 l2 && beq2 beqA r1 r2 + | PTree.Node001 r1, PTree.Node001 r2 => beq2' beqA r1 r2 + | PTree.Node010 x1, PTree.Node010 x2 => beqA x1 x2 + | PTree.Node011 x1 r1, PTree.Node011 x2 r2 => beqA x1 x2 && beq2' beqA r1 r2 + | PTree.Node100 l1, PTree.Node100 l2 => beq2' beqA l1 l2 + | PTree.Node101 l1 r1, PTree.Node101 l2 r2 => beq2' beqA l1 l2 && beq2' beqA r1 r2 + | PTree.Node110 l1 x1, PTree.Node110 l2 x2 => beqA x1 x2 && beq2' beqA l1 l2 + | PTree.Node111 l1 x1 r1, PTree.Node111 l2 x2 r2 => beqA x1 x2 && beq2' beqA l1 l2 && beq2' beqA r1 r2 + | _, _ => false end. +Definition beq2 {A B : Type} (beqA: A -> B -> bool) (m1: PTree.t A) (m2 : PTree.t B) : bool := + match m1, m2 with + | PTree.Empty, PTree.Empty => true + | PTree.Nodes m1', PTree.Nodes m2' => beq2' beqA m1' m2' + | _, _ => false + end. + Lemma beq2_correct: forall A B beqA m1 m2, @beq2 A B beqA m1 m2 = true <-> @@ -1386,26 +1392,7 @@ Lemma beq2_correct: | Some y1, Some y2 => beqA y1 y2 = true | _, _ => False end). -Proof. - induction m1; intros. - - simpl. rewrite PTree.bempty_correct. split; intros. - rewrite PTree.gleaf. rewrite H. auto. - generalize (H x). rewrite PTree.gleaf. destruct (PTree.get x m2); tauto. - - destruct m2. - + unfold beq2. rewrite PTree.bempty_correct. split; intros. - rewrite H. rewrite PTree.gleaf. auto. - generalize (H x). rewrite PTree.gleaf. - destruct (PTree.get x (PTree.Node m1_1 o m1_2)); tauto. - + simpl. split; intros. - * destruct (andb_prop _ _ H). destruct (andb_prop _ _ H0). - rewrite IHm1_1 in H3. rewrite IHm1_2 in H1. - destruct x; simpl. apply H1. apply H3. - destruct o; destruct o0; auto || congruence. - * apply andb_true_intro. split. apply andb_true_intro. split. - generalize (H xH); simpl. destruct o; destruct o0; tauto. - apply IHm1_1. intros; apply (H (xO x)). - apply IHm1_2. intros; apply (H (xI x)). -Qed. +Proof. Admitted. Lemma map1: forall w dst dst', -- cgit From b4d024615688e7d2ee59581e482afee303e79779 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 3 Mar 2022 17:51:48 +0000 Subject: Add back proof of beq2_correct --- src/hls/Abstr.v | 99 +++++++++++++++++++++++++++++++++++++++++++-------------- 1 file changed, 76 insertions(+), 23 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v index 9ded7fe..8c77636 100644 --- a/src/hls/Abstr.v +++ b/src/hls/Abstr.v @@ -1364,35 +1364,88 @@ Proof. end; discriminate. Qed. -Fixpoint beq2' {A B : Type} (beqA : A -> B -> bool) (m1: PTree.tree' A) (m2: PTree.tree' B) {struct m1} : bool := - match m1, m2 with - | PTree.Node001 r1, PTree.Node001 r2 => beq2' beqA r1 r2 - | PTree.Node010 x1, PTree.Node010 x2 => beqA x1 x2 - | PTree.Node011 x1 r1, PTree.Node011 x2 r2 => beqA x1 x2 && beq2' beqA r1 r2 - | PTree.Node100 l1, PTree.Node100 l2 => beq2' beqA l1 l2 - | PTree.Node101 l1 r1, PTree.Node101 l2 r2 => beq2' beqA l1 l2 && beq2' beqA r1 r2 - | PTree.Node110 l1 x1, PTree.Node110 l2 x2 => beqA x1 x2 && beq2' beqA l1 l2 - | PTree.Node111 l1 x1 r1, PTree.Node111 l2 x2 r2 => beqA x1 x2 && beq2' beqA l1 l2 && beq2' beqA r1 r2 - | _, _ => false - end. +Section BOOLEAN_EQUALITY. + + Context {A B: Type}. + Context (beqA: A -> B -> bool). -Definition beq2 {A B : Type} (beqA: A -> B -> bool) (m1: PTree.t A) (m2 : PTree.t B) : bool := + Fixpoint beq2' (m1: PTree.tree' A) (m2: PTree.tree' B) {struct m1} : bool := + match m1, m2 with + | PTree.Node001 r1, PTree.Node001 r2 => beq2' r1 r2 + | PTree.Node010 x1, PTree.Node010 x2 => beqA x1 x2 + | PTree.Node011 x1 r1, PTree.Node011 x2 r2 => beqA x1 x2 && beq2' r1 r2 + | PTree.Node100 l1, PTree.Node100 l2 => beq2' l1 l2 + | PTree.Node101 l1 r1, PTree.Node101 l2 r2 => beq2' l1 l2 && beq2' r1 r2 + | PTree.Node110 l1 x1, PTree.Node110 l2 x2 => beqA x1 x2 && beq2' l1 l2 + | PTree.Node111 l1 x1 r1, PTree.Node111 l2 x2 r2 => beqA x1 x2 && beq2' l1 l2 && beq2' r1 r2 + | _, _ => false + end. + + Definition beq2 (m1: PTree.t A) (m2 : PTree.t B) : bool := match m1, m2 with | PTree.Empty, PTree.Empty => true - | PTree.Nodes m1', PTree.Nodes m2' => beq2' beqA m1' m2' + | PTree.Nodes m1', PTree.Nodes m2' => beq2' m1' m2' | _, _ => false end. -Lemma beq2_correct: - forall A B beqA m1 m2, - @beq2 A B beqA m1 m2 = true <-> - (forall (x: PTree.elt), - match PTree.get x m1, PTree.get x m2 with - | None, None => True - | Some y1, Some y2 => beqA y1 y2 = true - | _, _ => False - end). -Proof. Admitted. + Let beq2_optA (o1: option A) (o2: option B) : bool := + match o1, o2 with + | None, None => true + | Some a1, Some a2 => beqA a1 a2 + | _, _ => false + end. + + Lemma beq_correct_bool: + forall m1 m2, + beq2 m1 m2 = true <-> (forall x, beq2_optA (m1 ! x) (m2 ! x) = true). + Proof. + Local Transparent PTree.Node. + assert (beq_NN: forall l1 o1 r1 l2 o2 r2, + PTree.not_trivially_empty l1 o1 r1 -> + PTree.not_trivially_empty l2 o2 r2 -> + beq2 (PTree.Node l1 o1 r1) (PTree.Node l2 o2 r2) = + beq2 l1 l2 && beq2_optA o1 o2 && beq2 r1 r2). + { intros. + destruct l1, o1, r1; try contradiction; destruct l2, o2, r2; try contradiction; + simpl; rewrite ? andb_true_r, ? andb_false_r; auto. + rewrite andb_comm; auto. + f_equal; rewrite andb_comm; auto. } + induction m1 using PTree.tree_ind; [|induction m2 using PTree.tree_ind]. + - intros. simpl; split; intros. + + destruct m2; try discriminate. simpl; auto. + + replace m2 with (@PTree.Empty B); auto. apply PTree.extensionality; intros x. + specialize (H x). destruct (m2 ! x); simpl; easy. + - split; intros. + + destruct (PTree.Node l o r); try discriminate. simpl; auto. + + replace (PTree.Node l o r) with (@PTree.Empty A); auto. apply PTree.extensionality; intros x. + specialize (H0 x). destruct ((PTree.Node l o r) ! x); simpl in *; easy. + - rewrite beq_NN by auto. split; intros. + + InvBooleans. rewrite ! PTree.gNode. destruct x. + * apply IHm0; auto. + * apply IHm1; auto. + * auto. + + apply andb_true_intro; split; [apply andb_true_intro; split|]. + * apply IHm1. intros. specialize (H1 (xO x)); rewrite ! PTree.gNode in H1; auto. + * specialize (H1 xH); rewrite ! PTree.gNode in H1; auto. + * apply IHm0. intros. specialize (H1 (xI x)); rewrite ! PTree.gNode in H1; auto. + Qed. + + Theorem beq2_correct: + forall m1 m2, + beq2 m1 m2 = true <-> + (forall (x: PTree.elt), + match m1 ! x, m2 ! x with + | None, None => True + | Some y1, Some y2 => beqA y1 y2 = true + | _, _ => False + end). + Proof. + intros. rewrite beq_correct_bool. unfold beq2_optA. split; intros. + - specialize (H x). destruct (m1 ! x), (m2 ! x); intuition congruence. + - specialize (H x). destruct (m1 ! x), (m2 ! x); intuition auto. + Qed. + +End BOOLEAN_EQUALITY. Lemma map1: forall w dst dst', -- cgit From d3f4c2ef8afa87619af1a92c23e3e97711835015 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sun, 6 Mar 2022 17:52:06 +0000 Subject: Fix RAW dependency calculation for predicates --- src/hls/Schedule.ml | 3 +++ 1 file changed, 3 insertions(+) (limited to 'src/hls') diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index 1969cad..11c7308 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -328,6 +328,7 @@ let accumulate_RAW_deps map dfg curr = match curr with | RBop (op, _, rs, dst) -> acc_dep_instruction rs dst | RBload (op, _mem, _addr, rs, dst) -> acc_dep_instruction rs dst + | RBsetpred (_op, _mem, rs, _p) -> acc_dep_instruction_nodst rs | RBstore (op, _mem, _addr, rs, src) -> acc_dep_instruction_nodst (src :: rs) | _ -> (i + 1, dst_map, graph) @@ -425,6 +426,7 @@ let get_predicate = function | RBop (p, _, _, _) -> p | RBload (p, _, _, _, _) -> p | RBstore (p, _, _, _, _) -> p + | RBsetpred (p, _, _, _) -> p | _ -> None let rec next_setpred p i = function @@ -448,6 +450,7 @@ let rec next_preduse p i instr= | RBload (Some p', _, _, _, _) :: rst -> next p' rst | RBstore (Some p', _, _, _, _) :: rst -> next p' rst | RBop (Some p', _, _, _) :: rst -> next p' rst + | RBsetpred (Some p', _, _, _) :: rst -> next p' rst | _ :: rst -> next_load (i + 1) rst let accumulate_RAW_pred_deps instrs dfg curri = -- cgit From 3117d20ad51f98ec7af9dba9969d2ee998667306 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 7 Mar 2022 16:41:48 +0000 Subject: Start work on reverse if-conversion --- src/hls/IfConversion.v | 2 +- src/hls/RICtransf.v | 83 ++++++++++++++++++++++++++++++++++++++++++++++++++ 2 files changed, 84 insertions(+), 1 deletion(-) create mode 100644 src/hls/RICtransf.v (limited to 'src/hls') diff --git a/src/hls/IfConversion.v b/src/hls/IfConversion.v index 9b8f4d4..1879205 100644 --- a/src/hls/IfConversion.v +++ b/src/hls/IfConversion.v @@ -30,7 +30,7 @@ Require Import vericert.bourdoncle.Bourdoncle. Parameter build_bourdoncle : function -> (bourdoncle * PMap.t N). -(** * If conversion +(** * If-Conversion This conversion is a verified conversion from RTLBlock back to itself, which performs if-conversion on basic blocks to make basic blocks larger. *) diff --git a/src/hls/RICtransf.v b/src/hls/RICtransf.v new file mode 100644 index 0000000..d6e926b --- /dev/null +++ b/src/hls/RICtransf.v @@ -0,0 +1,83 @@ +(* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2022 Yann Herklotz + * + * This program is free software: you can redistribute it and/or modify + * it under the terms of the GNU General Public License as published by + * the Free Software Foundation, either version 3 of the License, or + * (at your option) any later version. + * + * This program is distributed in the hope that it will be useful, + * but WITHOUT ANY WARRANTY; without even the implied warranty of + * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + * GNU General Public License for more details. + * + * You should have received a copy of the GNU General Public License + * along with this program. If not, see . + *) + +Require Import compcert.common.AST. +Require Import compcert.common.Errors. +Require Import compcert.common.Globalenvs. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Maps. + +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlockInstr. +Require Import vericert.hls.RTLPar. +Require Import vericert.hls.Predicate. + +(** * Reverse If-Conversion + +This transformation takes a scheduling RTLPar block and removes any predicates from it. It can then +be trivially transformed into linear code again. It works by iteratively selecting a predicate, +then creating a branch based on it and placing subsequent instructions in one branch or the other. +*) + +Fixpoint existsb_val {A B} (f: A -> option B) (l : list A) : option B := + match l with + | nil => None + | a :: l0 => + match f a, existsb_val f l0 with + | _, Some v => Some v + | Some v, _ => Some v + | _, _ => None + end + end. + +Definition includes_setpred (bb: list (list instr)) := + existsb_val (existsb_val (fun x => match x with RBsetpred a _ _ _ => Some a | _ => None end)) bb. + +Definition add_bb (should_split: bool) (i: list (list instr)) + (bbc: list (list (list instr)) * list (list (list instr))) := + match bbc with + | (a, b) => if should_split then (a, i::b) else (i::a, b) + end. + +Fixpoint first_split (saw_pred: bool) (bb: bb) := + match bb with + | a :: b => + match includes_setpred a with + | Some v => + let (_, res) := first_split true b in + (Some v, add_bb saw_pred a res) + | _ => + let (v, res) := first_split saw_pred b in + (v, add_bb saw_pred a res) + end + | nil => (None, (nil, nil)) + end. + +Definition split_bb (fresh: node) (b: bb) : bb * bb * bb := + match first_split false b with + | (Some p, (bb1, bb2)) => (bb1 ++ bb2, nil, nil) + | (None, (bb1, bb2)) => (bb1 ++ bb2, nil, nil) + end. + +Definition transf_function (f: function) : function := f. + +Definition transf_fundef (fd: fundef) : fundef := + transf_fundef transf_function fd. + +Definition transf_program (p: program) : program := + transform_program transf_fundef p. -- cgit From c897f9ab9c515501f88199003bd3e696bf7b9692 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 17 Mar 2022 14:57:07 +0000 Subject: Update copyright notice --- src/hls/AssocMap.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) (limited to 'src/hls') diff --git a/src/hls/AssocMap.v b/src/hls/AssocMap.v index 10bd8eb..ee39e8e 100644 --- a/src/hls/AssocMap.v +++ b/src/hls/AssocMap.v @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2020 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * 2020 James Pollard * * This program is free software: you can redistribute it and/or modify -- cgit From 27fea2fe14a81f4e73e0e3e53ec5ac5db07a5d82 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 19 Mar 2022 12:05:44 +0000 Subject: Delete extra data files and scripts --- src/hls/main | Bin 0 -> 54136 bytes 1 file changed, 0 insertions(+), 0 deletions(-) create mode 100755 src/hls/main (limited to 'src/hls') diff --git a/src/hls/main b/src/hls/main new file mode 100755 index 0000000..98b4e1b Binary files /dev/null and b/src/hls/main differ -- cgit From 8681d039abc5b2beac21ee17d7d944b00552aef6 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 16:41:37 +0000 Subject: Add RTLBlockgenproof --- src/hls/RTLBlockgenproof.v | 52 ++++++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 52 insertions(+) create mode 100644 src/hls/RTLBlockgenproof.v (limited to 'src/hls') diff --git a/src/hls/RTLBlockgenproof.v b/src/hls/RTLBlockgenproof.v new file mode 100644 index 0000000..b65abca --- /dev/null +++ b/src/hls/RTLBlockgenproof.v @@ -0,0 +1,52 @@ +(* + * Vericert: Verified high-level synthesis. + * Copyright (C) 2020-2022 Yann Herklotz + * + * This program is free software: you can redistribute it and/or modify + * it under the terms of the GNU General Public License as published by + * the Free Software Foundation, either version 3 of the License, or + * (at your option) any later version. + * + * This program is distributed in the hope that it will be useful, + * but WITHOUT ANY WARRANTY; without even the implied warranty of + * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + * GNU General Public License for more details. + * + * You should have received a copy of the GNU General Public License + * along with this program. If not, see . + *) + +(* [[file:../../lit/scheduling.org::rtlblockgenproof-imports][rtlblockgenproof-imports]] *) +Require compcert.backend.RTL. +Require Import compcert.common.AST. +Require Import compcert.lib.Maps. + +Require Import vericert.hls.RTLBlock. +Require Import vericert.hls.RTLBlockgen. +(* rtlblockgenproof-imports ends here *) + +(* [[file:../../lit/scheduling.org::rtlblockgenproof-match-states][rtlblockgenproof-match-states]] *) +Inductive match_states : RTL.state -> RTLBlock.state -> Prop := +| match_state : + forall stk f tf sp pc rs m + (TF: transl_function f = OK tf), + match_states (RTL.State stk f sp pc rs m) + (RTLBlock.State stk tf sp (find_block max n i) rs m). +(* rtlblockgenproof-match-states ends here *) + +(* [[file:../../lit/scheduling.org::rtlblockgenproof-correctness][rtlblockgenproof-correctness]] *) +Section CORRECTNESS. + + Context (prog : RTL.program). + Context (tprog : RTLBlock.program). + + Context (TRANSL : match_prog prog tprog). + + Theorem transf_program_correct: + Smallstep.forward_simulation (RTL.semantics prog) (RTLBlock.semantics tprog). + Proof. + eapply Smallstep.forward_simulation_plus; eauto with htlproof. + apply senv_preserved. + +End CORRECTNESS. +(* rtlblockgenproof-correctness ends here *) -- cgit From 125348f2a99c947b1a141866e0e8f33f33d7134a Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 16:42:15 +0000 Subject: Remove main binary --- src/hls/main | Bin 54136 -> 0 bytes 1 file changed, 0 insertions(+), 0 deletions(-) delete mode 100755 src/hls/main (limited to 'src/hls') diff --git a/src/hls/main b/src/hls/main deleted file mode 100755 index 98b4e1b..0000000 Binary files a/src/hls/main and /dev/null differ -- cgit From 5bb3e077854e33a7bd51d38f97970b08da171130 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 16:42:39 +0000 Subject: Fix proofs in Sat.v --- src/hls/Sat.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) (limited to 'src/hls') diff --git a/src/hls/Sat.v b/src/hls/Sat.v index b7596f6..428c176 100644 --- a/src/hls/Sat.v +++ b/src/hls/Sat.v @@ -119,7 +119,6 @@ Lemma satLit_contra : forall s l a cl, -> satClause cl a. unfold satLit, upd; simpl; intros. destruct (eq_nat_dec (snd l) (snd l)); intuition. - assert False; intuition. Qed. #[local] Hint Resolve satLit_contra : core. @@ -287,6 +286,7 @@ Lemma unitClauseTrue : forall l a fm, induction fm; intuition. inversion H. inversion H; subst; simpl in H0; intuition. + apply IHfm; eauto. inv H0. eauto. Qed. #[local] Hint Resolve unitClauseTrue : core. -- cgit From 2926dabad750aa862c7108f53a195c7973f5d619 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 16:42:55 +0000 Subject: Move forall_ptree into common --- src/hls/Abstr.v | 42 ------------------------------------------ 1 file changed, 42 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v index 8c77636..43d7783 100644 --- a/src/hls/Abstr.v +++ b/src/hls/Abstr.v @@ -692,48 +692,6 @@ Definition beq_pred_expr_once (pe1 pe2: pred_expr) : bool := equiv_check p1 p2 end. -Definition forall_ptree {A:Type} (f:positive->A->bool) (m:Maps.PTree.t A) : bool := - Maps.PTree.fold (fun (res: bool) (i: positive) (x: A) => res && f i x) m true. - -Ltac boolInv := - match goal with - | [ H: _ && _ = true |- _ ] => - destruct (andb_prop _ _ H); clear H; boolInv - | [ H: proj_sumbool _ = true |- _ ] => - generalize (proj_sumbool_true _ H); clear H; - let EQ := fresh in (intro EQ; try subst; boolInv) - | _ => - idtac - end. - -Remark ptree_forall: - forall (A: Type) (f: positive -> A -> bool) (m: Maps.PTree.t A), - Maps.PTree.fold (fun (res: bool) (i: positive) (x: A) => res && f i x) m true = true -> - forall i x, Maps.PTree.get i m = Some x -> f i x = true. -Proof. - intros. - set (f' := fun (res: bool) (i: positive) (x: A) => res && f i x). - set (P := fun (m: Maps.PTree.t A) (res: bool) => - res = true -> Maps.PTree.get i m = Some x -> f i x = true). - assert (P m true). - rewrite <- H. fold f'. apply Maps.PTree_Properties.fold_rec. - unfold P; intros. rewrite <- H1 in H4. auto. - red; intros. now rewrite Maps.PTree.gempty in H2. - unfold P; intros. unfold f' in H4. boolInv. - rewrite Maps.PTree.gsspec in H5. destruct (peq i k). - subst. inv H5. auto. - apply H3; auto. - red in H1. auto. -Qed. - -Lemma forall_ptree_true: - forall (A: Type) (f: positive -> A -> bool) (m: Maps.PTree.t A), - forall_ptree f m = true -> - forall i x, Maps.PTree.get i m = Some x -> f i x = true. -Proof. - apply ptree_forall. -Qed. - Definition tree_equiv_check_el (np2: PTree.t pred_op) (n: positive) (p: pred_op): bool := match np2 ! n with | Some p' => equiv_check p p' -- cgit From 2c33145e2de887964b68466d1691d0455d63334b Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 16:43:22 +0000 Subject: Add nops to partition generation --- src/hls/Partition.ml | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Partition.ml b/src/hls/Partition.ml index 19c6048..d7972e5 100644 --- a/src/hls/Partition.ml +++ b/src/hls/Partition.ml @@ -28,7 +28,7 @@ open Op open RTLBlockInstr open RTLBlock -(* Assuming that the nodes of the CFG [code] are numbered in reverse postorder (cf. pass +(** Assuming that the nodes of the CFG [code] are numbered in reverse postorder (cf. pass [Renumber]), an edge from [n] to [s] is a normal edge if [s < n] and a back-edge otherwise. *) let find_edge i n = let succ = RTL.successors_instr i in @@ -65,14 +65,14 @@ let rec next_bblock_from_RTL is_start e (c : RTL.code) s i = match trans_inst, succ with | (None, Some i'), _ -> if List.exists (fun x -> x = s) (snd e) && not is_start then - Errors.OK { bb_body = []; bb_exit = RBgoto s } + Errors.OK { bb_body = [RBnop]; bb_exit = RBgoto s } else - Errors.OK { bb_body = []; bb_exit = i' } + Errors.OK { bb_body = [RBnop]; bb_exit = i' } | (Some i', None), (s', Some i_n)::[] -> if List.exists (fun x -> x = s) (fst e) then Errors.OK { bb_body = [i']; bb_exit = RBgoto s' } else if List.exists (fun x -> x = s) (snd e) && not is_start then - Errors.OK { bb_body = []; bb_exit = RBgoto s } + Errors.OK { bb_body = [RBnop]; bb_exit = RBgoto s } else begin match next_bblock_from_RTL false e c s' i_n with | Errors.OK bb -> -- cgit From 2d647ce5fdf5343a7d9961a63d66b5191706aeaf Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 16:43:39 +0000 Subject: Add comments to allow for literate detangling --- src/hls/RTLBlock.v | 94 +++++++++++++++-- src/hls/RTLBlockInstr.v | 6 +- src/hls/RTLBlockgen.v | 264 ++++++++++++++++++++++++++++++++++++++++++++++- src/hls/RTLPar.v | 4 +- src/hls/RTLPargen.v | 4 +- src/hls/RTLPargenproof.v | 4 +- src/hls/Schedule.ml | 4 +- 7 files changed, 362 insertions(+), 18 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v index bf5c37a..4fae701 100644 --- a/src/hls/RTLBlock.v +++ b/src/hls/RTLBlock.v @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2020-2021 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -16,6 +16,7 @@ * along with this program. If not, see . *) +(* [[file:../../lit/scheduling.org::rtlblock-main][rtlblock-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Events. @@ -39,9 +40,32 @@ Definition fundef := @fundef bb. Definition program := @program bb. Definition funsig := @funsig bb. Definition stackframe := @stackframe bb. -Definition state := @state bb. + Definition genv := @genv bb. +Inductive state : Type := +| State: + forall (stack: list stackframe) (**r call stack *) + (f: function) (**r current function *) + (b: bb) (**r current block being executed *) + (sp: val) (**r stack pointer *) + (pc: node) (**r current program point in [c] *) + (rs: regset) (**r register state *) + (pr: predset) (**r predicate register state *) + (m: mem), (**r memory state *) + state +| Callstate: + forall (stack: list stackframe) (**r call stack *) + (f: fundef) (**r function to call *) + (args: list val) (**r arguments to the call *) + (m: mem), (**r memory state *) + state +| Returnstate: + forall (stack: list stackframe) (**r call stack *) + (v: val) (**r return value for the call *) + (m: mem), (**r memory state *) + state. + Section RELSEM. Context (ge: genv). @@ -56,19 +80,68 @@ Section RELSEM. forall state sp, step_instr_list sp state nil state. + Definition find_function + (ros: reg + ident) (rs: regset) : option fundef := + match ros with + | inl r => Genv.find_funct ge rs#r + | inr symb => + match Genv.find_symbol ge symb with + | None => None + | Some b => Genv.find_funct_ptr ge b + end + end. + + Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop := + | exec_RBcall: + forall s f b sp rs m res fd ros sig args pc pc' pr, + find_function ros rs = Some fd -> + funsig fd = sig -> + step_cf_instr (State s f b sp pc rs pr m) (RBcall sig ros args res pc') + E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m) + | exec_RBtailcall: + forall s f b stk rs m sig ros args fd m' pc pr, + find_function ros rs = Some fd -> + funsig fd = sig -> + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args) + E0 (Callstate s fd rs##args m') + | exec_RBbuiltin: + forall s f b sp rs m ef args res pc' vargs t vres m' pc pr, + eval_builtin_args ge (fun r => rs#r) sp m args vargs -> + external_call ef ge vargs m t vres m' -> + step_cf_instr (State s f b sp pc rs pr m) (RBbuiltin ef args res pc') + t (State s f b sp pc' (regmap_setres res vres rs) pr m') + | exec_RBcond: + forall s f block sp rs m cond args ifso ifnot b pc pc' pr, + eval_condition cond rs##args m = Some b -> + pc' = (if b then ifso else ifnot) -> + step_cf_instr (State s f block sp pc rs pr m) (RBcond cond args ifso ifnot) + E0 (State s f block sp pc' rs pr m) + | exec_RBjumptable: + forall s f b sp rs m arg tbl n pc pc' pr, + rs#arg = Vint n -> + list_nth_z tbl (Int.unsigned n) = Some pc' -> + step_cf_instr (State s f b sp pc rs pr m) (RBjumptable arg tbl) + E0 (State s f b sp pc' rs pr m) + | exec_RBreturn: + forall s f b stk rs m or pc m' pr, + Mem.free m stk 0 f.(fn_stacksize) = Some m' -> + step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or) + E0 (Returnstate s (regmap_optget or Vundef rs) m') + | exec_RBgoto: + forall s f b sp pc rs pr m pc', + step_cf_instr (State s f b sp pc rs pr m) (RBgoto pc') E0 (State s f b sp pc' rs pr m) + | exec_RBpred_cf: + forall s f b sp pc rs pr m cf1 cf2 st' p t, + step_cf_instr (State s f b sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> + step_cf_instr (State s f b sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'. + Inductive step: state -> trace -> state -> Prop := - | exec_bblock: - forall s f sp pc rs rs' m m' t s' bb pr pr', - f.(fn_code)!pc = Some bb -> - step_instr_list sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') -> - step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' -> - step (State s f sp pc rs pr m) t s' | exec_function_internal: forall s f args m m' stk, Mem.alloc m 0 f.(fn_stacksize) = (m', stk) -> step (Callstate s (Internal f) args m) - E0 (State s - f + E0 (State s f (Vptr stk Ptrofs.zero) f.(fn_entrypoint) (init_regs args f.(fn_params)) @@ -101,3 +174,4 @@ Inductive final_state: state -> int -> Prop := Definition semantics (p: program) := Semantics step (initial_state p) final_state (Genv.globalenv p). +(* rtlblock-main ends here *) diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 52259a0..7391f97 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2021 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -16,6 +16,7 @@ * along with this program. If not, see . *) +(* [[file:../../lit/scheduling.org::rtlblockinstr-main][rtlblockinstr-main]] *) Require Import Coq.micromega.Lia. Require Import compcert.backend.Registers. @@ -37,7 +38,7 @@ These instructions are used for [RTLBlock] and [RTLPar], so that they have consi instructions, which greatly simplifies the proofs, as they will by default have the same instruction syntax and semantics. The only changes are therefore at the top-level of the instructions. -** Instruction Definition + ** Instruction Definition First, we define the instructions that can be placed into a basic block, meaning they won't branch. The main changes to how instructions are defined in [RTL], is that these instructions don't have a @@ -346,3 +347,4 @@ Section RELSEM. step_cf_instr (State s f sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'. End RELSEM. +(* rtlblockinstr-main ends here *) diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index 889e104..af2c5af 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2020 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -16,15 +16,275 @@ * along with this program. If not, see . *) +(* [[file:../../lit/scheduling.org::rtlblockgen-main][rtlblockgen-main]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. +Require Import compcert.lib.Integers. +Require Import compcert.lib.Floats. +Require Import vericert.common.Vericertlib. +Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLBlock. +#[local] Open Scope positive. + Parameter partition : RTL.function -> Errors.res function. -Definition transl_fundef := transf_partial_fundef partition. +(** [find_block max nodes index]: Does not need to be sorted, because we use filter and the max fold + function to find the desired element. *) +Definition find_block (max: positive) (nodes: list positive) (index: positive) : positive := + List.fold_right Pos.min max (List.filter (fun x => (index <=? x)) nodes). + +(*Compute find_block (2::94::28::40::19::nil) 40.*) + +(* [[[[file:~/projects/vericert/lit/scheduling.org::rtlblockgen-equalities][rtlblockgen-equalities]]][rtlblockgen-equalities]] *) +Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}. +Proof. + decide equality. +Defined. + +Lemma condition_eq: forall (x y : Op.condition), {x = y} + {x <> y}. +Proof. + generalize comparison_eq; intro. + generalize Int.eq_dec; intro. + generalize Int64.eq_dec; intro. + decide equality. +Defined. + +Lemma addressing_eq : forall (x y : Op.addressing), {x = y} + {x <> y}. +Proof. + generalize Int.eq_dec; intro. + generalize AST.ident_eq; intro. + generalize Z.eq_dec; intro. + generalize Ptrofs.eq_dec; intro. + decide equality. +Defined. + +Lemma typ_eq : forall (x y : AST.typ), {x = y} + {x <> y}. +Proof. + decide equality. +Defined. + +Lemma operation_eq: forall (x y : Op.operation), {x = y} + {x <> y}. +Proof. + generalize Int.eq_dec; intro. + generalize Int64.eq_dec; intro. + generalize Float.eq_dec; intro. + generalize Float32.eq_dec; intro. + generalize AST.ident_eq; intro. + generalize condition_eq; intro. + generalize addressing_eq; intro. + generalize typ_eq; intro. + decide equality. +Defined. + +Lemma memory_chunk_eq : forall (x y : AST.memory_chunk), {x = y} + {x <> y}. +Proof. + decide equality. +Defined. + +Lemma list_typ_eq: forall (x y : list AST.typ), {x = y} + {x <> y}. +Proof. + generalize typ_eq; intro. + decide equality. +Defined. + +Lemma option_typ_eq : forall (x y : option AST.typ), {x = y} + {x <> y}. +Proof. + generalize typ_eq; intro. + decide equality. +Defined. + +Lemma signature_eq: forall (x y : AST.signature), {x = y} + {x <> y}. +Proof. + repeat decide equality. +Defined. + +Lemma list_operation_eq : forall (x y : list Op.operation), {x = y} + {x <> y}. +Proof. + generalize operation_eq; intro. + decide equality. +Defined. + +Lemma list_pos_eq : forall (x y : list positive), {x = y} + {x <> y}. +Proof. + generalize Pos.eq_dec; intros. + decide equality. +Defined. + +Lemma sig_eq : forall (x y : AST.signature), {x = y} + {x <> y}. +Proof. + repeat decide equality. +Defined. + +Lemma instr_eq: forall (x y : instr), {x = y} + {x <> y}. +Proof. + generalize Pos.eq_dec; intro. + generalize typ_eq; intro. + generalize Int.eq_dec; intro. + generalize memory_chunk_eq; intro. + generalize addressing_eq; intro. + generalize operation_eq; intro. + generalize condition_eq; intro. + generalize signature_eq; intro. + generalize list_operation_eq; intro. + generalize list_pos_eq; intro. + generalize AST.ident_eq; intro. + repeat decide equality. +Defined. + +Lemma cf_instr_eq: forall (x y : cf_instr), {x = y} + {x <> y}. +Proof. + generalize Pos.eq_dec; intro. + generalize typ_eq; intro. + generalize Int.eq_dec; intro. + generalize Int64.eq_dec; intro. + generalize Float.eq_dec; intro. + generalize Float32.eq_dec; intro. + generalize Ptrofs.eq_dec; intro. + generalize memory_chunk_eq; intro. + generalize addressing_eq; intro. + generalize operation_eq; intro. + generalize condition_eq; intro. + generalize signature_eq; intro. + generalize list_operation_eq; intro. + generalize list_pos_eq; intro. + generalize AST.ident_eq; intro. + repeat decide equality. +Defined. + +Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := + if eqd a b then true else false. +(* rtlblockgen-equalities ends here *) + +Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) := + match istr, istr' with + | RTL.Inop n', RBnop => (n' + 1 =? n) + | RTL.Iop op args dst n', RBop None op' args' dst' => + ceq operation_eq op op' && + ceq list_pos_eq args args' && + ceq peq dst dst' && (n' + 1 =? n) + | RTL.Iload chunk addr args dst n', RBload None chunk' addr' args' dst' => + ceq memory_chunk_eq chunk chunk' && + ceq addressing_eq addr addr' && + ceq list_pos_eq args args' && + ceq peq dst dst' && + (n' + 1 =? n) + | RTL.Istore chunk addr args src n', RBstore None chunk' addr' args' src' => + ceq memory_chunk_eq chunk chunk' && + ceq addressing_eq addr addr' && + ceq list_pos_eq args args' && + ceq peq src src' && + (n' + 1 =? n) + | _, _ => false + end. + +Definition check_cf_instr_body (istr: RTL.instruction) (istr': instr): bool := + match istr, istr' with + | RTL.Iop op args dst _, RBop None op' args' dst' => + ceq operation_eq op op' && + ceq list_pos_eq args args' && + ceq peq dst dst' + | RTL.Iload chunk addr args dst _, RBload None chunk' addr' args' dst' => + ceq memory_chunk_eq chunk chunk' && + ceq addressing_eq addr addr' && + ceq list_pos_eq args args' && + ceq peq dst dst' + | RTL.Istore chunk addr args src _, RBstore None chunk' addr' args' src' => + ceq memory_chunk_eq chunk chunk' && + ceq addressing_eq addr addr' && + ceq list_pos_eq args args' && + ceq peq src src' + | RTL.Inop _, RBnop + | RTL.Icall _ _ _ _ _, RBnop + | RTL.Itailcall _ _ _, RBnop + | RTL.Ibuiltin _ _ _ _, RBnop + | RTL.Icond _ _ _ _, RBnop + | RTL.Ijumptable _ _, RBnop + | RTL.Ireturn _, RBnop => true + | _, _ => false + end. + +Definition check_cf_instr (istr: RTL.instruction) (istr': cf_instr) := + match istr, istr' with + | RTL.Inop n, RBgoto n' => (n =? n') + | RTL.Iop _ _ _ n, RBgoto n' => (n =? n') + | RTL.Iload _ _ _ _ n, RBgoto n' => (n =? n') + | RTL.Istore _ _ _ _ n, RBgoto n' => (n =? n') + | RTL.Icall sig (inl r) args dst n, RBcall sig' (inl r') args' dst' n' => + ceq signature_eq sig sig' && + ceq peq r r' && + ceq list_pos_eq args args' && + ceq peq dst dst' && + (n =? n') + | RTL.Icall sig (inr i) args dst n, RBcall sig' (inr i') args' dst' n' => + ceq signature_eq sig sig' && + ceq peq i i' && + ceq list_pos_eq args args' && + ceq peq dst dst' && + (n =? n') + | RTL.Itailcall sig (inl r) args, RBtailcall sig' (inl r') args' => + ceq signature_eq sig sig' && + ceq peq r r' && + ceq list_pos_eq args args' + | RTL.Itailcall sig (inr r) args, RBtailcall sig' (inr r') args' => + ceq signature_eq sig sig' && + ceq peq r r' && + ceq list_pos_eq args args' + | RTL.Icond cond args n1 n2, RBcond cond' args' n1' n2' => + ceq condition_eq cond cond' && + ceq list_pos_eq args args' && + ceq peq n1 n1' && ceq peq n2 n2' + | RTL.Ijumptable r ns, RBjumptable r' ns' => + ceq peq r r' && ceq list_pos_eq ns ns' + | RTL.Ireturn (Some r), RBreturn (Some r') => + ceq peq r r' + | RTL.Ireturn None, RBreturn None => true + | _, _ => false + end. + +Definition is_cf_instr (n: positive) (i: RTL.instruction) := + match i with + | RTL.Inop n' => negb (n' + 1 =? n) + | RTL.Iop _ _ _ n' => negb (n' + 1 =? n) + | RTL.Iload _ _ _ _ n' => negb (n' + 1 =? n) + | RTL.Istore _ _ _ _ n' => negb (n' + 1 =? n) + | RTL.Icall _ _ _ _ _ => true + | RTL.Itailcall _ _ _ => true + | RTL.Ibuiltin _ _ _ _ => true + | RTL.Icond _ _ _ _ => true + | RTL.Ijumptable _ _ => true + | RTL.Ireturn _ => true + end. + +Definition check_present_blocks (c: code) (n: list positive) (max: positive) (i: positive) (istr: RTL.instruction) := + let blockn := find_block max n i in + match c ! blockn with + | Some istrs => + match List.nth_error istrs.(bb_body) (Pos.to_nat blockn - Pos.to_nat i)%nat with + | Some istr' => + if is_cf_instr i istr + then check_cf_instr istr istrs.(bb_exit) && check_cf_instr_body istr istr' + else check_instr i istr istr' + | None => false + end + | None => false + end. + +Definition transl_function (f: RTL.function) := + match partition f with + | Errors.OK f' => + let blockids := map fst (PTree.elements f'.(fn_code)) in + if forall_ptree (check_present_blocks f'.(fn_code) blockids (fold_right Pos.max 1 blockids)) + f.(RTL.fn_code) then + Errors.OK f' + else Errors.Error (Errors.msg "check_present_blocks failed") + | Errors.Error msg => Errors.Error msg + end. + +Definition transl_fundef := transf_partial_fundef transl_function. Definition transl_program : RTL.program -> Errors.res program := transform_partial_program transl_fundef. +(* rtlblockgen-main ends here *) diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v index 4986cff..bcb51c6 100644 --- a/src/hls/RTLPar.v +++ b/src/hls/RTLPar.v @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2020-2021 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -16,6 +16,7 @@ * along with this program. If not, see . *) +(* [[file:../../lit/scheduling.org::rtlpar-main][rtlpar-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Events. @@ -138,3 +139,4 @@ Definition max_pc_function (f: function) : positive := (pc + match Zlength i.(bb_body) with Z.pos p => p | _ => 1 end))%positive) f.(fn_code) 1%positive. +(* rtlpar-main ends here *) diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index ab4c0da..f3d13f5 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2020-2021 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -16,6 +16,7 @@ * along with this program. If not, see . *) +(* [[file:../../lit/scheduling.org::rtlpargen-main][rtlpargen-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Globalenvs. @@ -260,3 +261,4 @@ Definition transl_fundef := transf_partial_fundef transl_function. Definition transl_program (p : RTLBlock.program) : Errors.res RTLPar.program := transform_partial_program transl_fundef p. +(* rtlpargen-main ends here *) diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v index 0023edc..63a294e 100644 --- a/src/hls/RTLPargenproof.v +++ b/src/hls/RTLPargenproof.v @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2021 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -16,6 +16,7 @@ * along with this program. If not, see . *) +(* [[file:../../lit/scheduling.org::rtlpargenproof-main][rtlpargenproof-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Errors. @@ -1135,3 +1136,4 @@ Proof. induction 2; try rewrite H; eauto with barg. Qed. Qed. End CORRECTNESS. +(* rtlpargenproof-main ends here *) diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index 11c7308..88aa6fd 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -1,6 +1,6 @@ (* * Vericert: Verified high-level synthesis. - * Copyright (C) 2020 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -16,6 +16,7 @@ * along with this program. If not, see . *) +(* [[file:../../lit/scheduling.org::scheduler-main][scheduler-main]] *) open Printf open Clflags open Camlcoq @@ -881,3 +882,4 @@ let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = fn_code = scheduled (*List.fold_left (add_to_tree scheduled) PTree.empty reachable*); fn_entrypoint = f.fn_entrypoint } +(* scheduler-main ends here *) -- cgit From 79ebc1c11ac3daccbc13b56043bdc89b14b23c60 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 17:14:08 +0000 Subject: Add literate Coq file --- src/hls/RTLBlockgen.v | 2 -- src/hls/RTLPargenproof.v | 8 +++----- 2 files changed, 3 insertions(+), 7 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index af2c5af..f45151d 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -38,7 +38,6 @@ Definition find_block (max: positive) (nodes: list positive) (index: positive) : (*Compute find_block (2::94::28::40::19::nil) 40.*) -(* [[[[file:~/projects/vericert/lit/scheduling.org::rtlblockgen-equalities][rtlblockgen-equalities]]][rtlblockgen-equalities]] *) Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}. Proof. decide equality. @@ -156,7 +155,6 @@ Defined. Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := if eqd a b then true else false. -(* rtlblockgen-equalities ends here *) Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) := match istr, istr' with diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v index 63a294e..7b262db 100644 --- a/src/hls/RTLPargenproof.v +++ b/src/hls/RTLPargenproof.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlpargenproof-main][rtlpargenproof-main]] *) +(* [[file:../../lit/scheduling.org::rtlpargenproof-imports][rtlpargenproof-imports]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Errors. @@ -37,7 +37,9 @@ Require Import vericert.hls.Abstr. #[local] Open Scope positive. #[local] Open Scope forest. #[local] Open Scope pred_op. +(* rtlpargenproof-imports ends here *) +(* [[file:../../lit/scheduling.org::rtlpargenproof-main][rtlpargenproof-main]] *) (*Definition is_regs i := match i with mk_instr_state rs _ => rs end. Definition is_mem i := match i with mk_instr_state _ m => m end. @@ -47,10 +49,6 @@ Inductive state_lessdef : instr_state -> instr_state -> Prop := (forall x, rs1 !! x = rs2 !! x) -> state_lessdef (mk_instr_state rs1 m1) (mk_instr_state rs2 m1). -(** *** RTLBlock to abstract translation - -Correctness of translation from RTLBlock to the abstract interpretation language. *) - Ltac inv_simp := repeat match goal with | H: exists _, _ |- _ => inv H -- cgit From 0f223a52dc59646ce2731ebd9b2141ce1ad82ba1 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 18:04:18 +0000 Subject: Add first descriptions to org file --- src/hls/RTLBlockInstr.v | 49 ++++++++++++++----------------------------------- src/hls/RTLBlockgen.v | 26 ++++++++++++++++---------- 2 files changed, 30 insertions(+), 45 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 7391f97..770f377 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlblockinstr-main][rtlblockinstr-main]] *) +(* [[file:../../lit/scheduling.org::rtlblockinstr-imports][rtlblockinstr-imports]] *) Require Import Coq.micromega.Lia. Require Import compcert.backend.Registers. @@ -31,20 +31,9 @@ Require Import compcert.verilog.Op. Require Import Predicate. Require Import Vericertlib. +(* rtlblockinstr-imports ends here *) -(** * RTLBlock Instructions - -These instructions are used for [RTLBlock] and [RTLPar], so that they have consistent -instructions, which greatly simplifies the proofs, as they will by default have the same instruction -syntax and semantics. The only changes are therefore at the top-level of the instructions. - - ** Instruction Definition - -First, we define the instructions that can be placed into a basic block, meaning they won't branch. -The main changes to how instructions are defined in [RTL], is that these instructions don't have a -next node, as they will be in a basic block, and they also have an optional predicate ([pred_op]). -*) - +(* [[file:../../lit/scheduling.org::rtlblockinstr-instr-def][rtlblockinstr-instr-def]] *) Definition node := positive. Inductive instr : Type := @@ -53,12 +42,9 @@ Inductive instr : Type := | RBload : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr | RBstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr | RBsetpred : option pred_op -> condition -> list reg -> predicate -> instr. +(* rtlblockinstr-instr-def ends here *) -(** ** Control-Flow Instruction Definition - -These are the instructions that count as control-flow, and will be placed at the end of the basic -blocks. *) - +(* [[file:../../lit/scheduling.org::rtlblockinstr-cf-instr-def][rtlblockinstr-cf-instr-def]] *) Inductive cf_instr : Type := | RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr | RBtailcall : signature -> reg + ident -> list reg -> cf_instr @@ -69,9 +55,9 @@ Inductive cf_instr : Type := | RBreturn : option reg -> cf_instr | RBgoto : node -> cf_instr | RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. +(* rtlblockinstr-cf-instr-def ends here *) -(** ** Helper functions *) - +(* [[file:../../lit/scheduling.org::rtlblockinstr-helpers][rtlblockinstr-helpers]] *) Fixpoint successors_instr (i : cf_instr) : list node := match i with | RBcall sig ros args res s => s :: nil @@ -160,24 +146,17 @@ Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset := | r1 :: rs, v1 :: vs => Regmap.set r1 v1 (init_regs vs rs) | _, _ => Regmap.init Vundef end. +(* rtlblockinstr-helpers ends here *) -(** *** Instruction State - -Definition of the instruction state, which contains the following: - -- [is_rs] This is the current state of the registers. -- [is_ps] This is the current state of the predicate registers, which is in a separate namespace and - area compared to the standard registers in [is_rs]. -- [is_mem] The current state of the memory. *) - +(* [[file:../../lit/scheduling.org::rtlblockinstr-instr-state][rtlblockinstr-instr-state]] *) Record instr_state := mk_instr_state { is_rs: regset; is_ps: predset; is_mem: mem; }. +(* rtlblockinstr-instr-state ends here *) -(** ** Top-Level Type Definitions *) - +(* [[file:../../lit/scheduling.org::rtlblockinstr-type-def][rtlblockinstr-type-def]] *) Section DEFINITION. Context {bblock_body: Type}. @@ -240,9 +219,9 @@ Section DEFINITION. state. End DEFINITION. +(* rtlblockinstr-type-def ends here *) -(** ** Semantics *) - +(* [[file:../../lit/scheduling.org::rtlblockinstr-semantics][rtlblockinstr-semantics]] *) Section RELSEM. Context {bblock_body : Type}. @@ -347,4 +326,4 @@ Section RELSEM. step_cf_instr (State s f sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'. End RELSEM. -(* rtlblockinstr-main ends here *) +(* rtlblockinstr-semantics ends here *) diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index f45151d..706cb09 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlblockgen-main][rtlblockgen-main]] *) +(* [[file:../../lit/scheduling.org::rtlblockgen-imports][rtlblockgen-imports]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. @@ -28,16 +28,10 @@ Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLBlock. #[local] Open Scope positive. +(* rtlblockgen-imports ends here *) -Parameter partition : RTL.function -> Errors.res function. - -(** [find_block max nodes index]: Does not need to be sorted, because we use filter and the max fold - function to find the desired element. *) -Definition find_block (max: positive) (nodes: list positive) (index: positive) : positive := - List.fold_right Pos.min max (List.filter (fun x => (index <=? x)) nodes). - -(*Compute find_block (2::94::28::40::19::nil) 40.*) - +(* [[file:../../lit/scheduling.org::rtlblockgen-equalities-insert][rtlblockgen-equalities-insert]] *) +(* [[[[file:~/projects/vericert/lit/scheduling.org::rtlblockgen-equalities][rtlblockgen-equalities]]][rtlblockgen-equalities]] *) Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}. Proof. decide equality. @@ -155,6 +149,18 @@ Defined. Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := if eqd a b then true else false. +(* rtlblockgen-equalities ends here *) +(* rtlblockgen-equalities-insert ends here *) + +(* [[file:../../lit/scheduling.org::rtlblockgen-main][rtlblockgen-main]] *) +Parameter partition : RTL.function -> Errors.res function. + +(** [find_block max nodes index]: Does not need to be sorted, because we use filter and the max fold + function to find the desired element. *) +Definition find_block (max: positive) (nodes: list positive) (index: positive) : positive := + List.fold_right Pos.min max (List.filter (fun x => (index <=? x)) nodes). + +(*Compute find_block (2::94::28::40::19::nil) 40.*) Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) := match istr, istr' with -- cgit From 115f5b18952bdeea150ce54b156cb9a96c8d3e96 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Tue, 22 Mar 2022 19:01:10 +0000 Subject: Add back equalities --- src/hls/RTLBlockgen.v | 2 -- 1 file changed, 2 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index 706cb09..fa462e7 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -31,7 +31,6 @@ Require Import vericert.hls.RTLBlock. (* rtlblockgen-imports ends here *) (* [[file:../../lit/scheduling.org::rtlblockgen-equalities-insert][rtlblockgen-equalities-insert]] *) -(* [[[[file:~/projects/vericert/lit/scheduling.org::rtlblockgen-equalities][rtlblockgen-equalities]]][rtlblockgen-equalities]] *) Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}. Proof. decide equality. @@ -149,7 +148,6 @@ Defined. Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := if eqd a b then true else false. -(* rtlblockgen-equalities ends here *) (* rtlblockgen-equalities-insert ends here *) (* [[file:../../lit/scheduling.org::rtlblockgen-main][rtlblockgen-main]] *) -- cgit From f00a195ac17fe47047fafc183663a96ec4125f0d Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Wed, 23 Mar 2022 09:47:56 +0000 Subject: Change origin of tangled files --- src/hls/RTLBlock.v | 2 +- src/hls/RTLBlockInstr.v | 14 +++++++------- src/hls/RTLBlockgen.v | 6 ++---- src/hls/RTLBlockgenproof.v | 6 +++--- src/hls/RTLPar.v | 2 +- src/hls/RTLPargen.v | 2 +- src/hls/RTLPargenproof.v | 6 +++--- src/hls/Schedule.ml | 2 +- 8 files changed, 19 insertions(+), 21 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v index 4fae701..539d5bc 100644 --- a/src/hls/RTLBlock.v +++ b/src/hls/RTLBlock.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlblock-main][rtlblock-main]] *) +(* [[file:../../lit/scheduler-languages.org::rtlblock-main][rtlblock-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Events. diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 770f377..19b3928 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlblockinstr-imports][rtlblockinstr-imports]] *) +(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-imports][rtlblockinstr-imports]] *) Require Import Coq.micromega.Lia. Require Import compcert.backend.Registers. @@ -33,7 +33,7 @@ Require Import Predicate. Require Import Vericertlib. (* rtlblockinstr-imports ends here *) -(* [[file:../../lit/scheduling.org::rtlblockinstr-instr-def][rtlblockinstr-instr-def]] *) +(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-instr-def][rtlblockinstr-instr-def]] *) Definition node := positive. Inductive instr : Type := @@ -44,7 +44,7 @@ Inductive instr : Type := | RBsetpred : option pred_op -> condition -> list reg -> predicate -> instr. (* rtlblockinstr-instr-def ends here *) -(* [[file:../../lit/scheduling.org::rtlblockinstr-cf-instr-def][rtlblockinstr-cf-instr-def]] *) +(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-cf-instr-def][rtlblockinstr-cf-instr-def]] *) Inductive cf_instr : Type := | RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr | RBtailcall : signature -> reg + ident -> list reg -> cf_instr @@ -57,7 +57,7 @@ Inductive cf_instr : Type := | RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. (* rtlblockinstr-cf-instr-def ends here *) -(* [[file:../../lit/scheduling.org::rtlblockinstr-helpers][rtlblockinstr-helpers]] *) +(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-helpers][rtlblockinstr-helpers]] *) Fixpoint successors_instr (i : cf_instr) : list node := match i with | RBcall sig ros args res s => s :: nil @@ -148,7 +148,7 @@ Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset := end. (* rtlblockinstr-helpers ends here *) -(* [[file:../../lit/scheduling.org::rtlblockinstr-instr-state][rtlblockinstr-instr-state]] *) +(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-instr-state][rtlblockinstr-instr-state]] *) Record instr_state := mk_instr_state { is_rs: regset; is_ps: predset; @@ -156,7 +156,7 @@ Record instr_state := mk_instr_state { }. (* rtlblockinstr-instr-state ends here *) -(* [[file:../../lit/scheduling.org::rtlblockinstr-type-def][rtlblockinstr-type-def]] *) +(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-type-def][rtlblockinstr-type-def]] *) Section DEFINITION. Context {bblock_body: Type}. @@ -221,7 +221,7 @@ Section DEFINITION. End DEFINITION. (* rtlblockinstr-type-def ends here *) -(* [[file:../../lit/scheduling.org::rtlblockinstr-semantics][rtlblockinstr-semantics]] *) +(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-semantics][rtlblockinstr-semantics]] *) Section RELSEM. Context {bblock_body : Type}. diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index fa462e7..6d38e4f 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlblockgen-imports][rtlblockgen-imports]] *) +(* [[file:../../lit/basic-block-generation.org::rtlblockgen-imports][rtlblockgen-imports]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. @@ -30,7 +30,6 @@ Require Import vericert.hls.RTLBlock. #[local] Open Scope positive. (* rtlblockgen-imports ends here *) -(* [[file:../../lit/scheduling.org::rtlblockgen-equalities-insert][rtlblockgen-equalities-insert]] *) Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}. Proof. decide equality. @@ -148,9 +147,8 @@ Defined. Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := if eqd a b then true else false. -(* rtlblockgen-equalities-insert ends here *) -(* [[file:../../lit/scheduling.org::rtlblockgen-main][rtlblockgen-main]] *) +(* [[file:../../lit/basic-block-generation.org::rtlblockgen-main][rtlblockgen-main]] *) Parameter partition : RTL.function -> Errors.res function. (** [find_block max nodes index]: Does not need to be sorted, because we use filter and the max fold diff --git a/src/hls/RTLBlockgenproof.v b/src/hls/RTLBlockgenproof.v index b65abca..4433d52 100644 --- a/src/hls/RTLBlockgenproof.v +++ b/src/hls/RTLBlockgenproof.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlblockgenproof-imports][rtlblockgenproof-imports]] *) +(* [[file:../../lit/basic-block-generation.org::rtlblockgenproof-imports][rtlblockgenproof-imports]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. @@ -25,7 +25,7 @@ Require Import vericert.hls.RTLBlock. Require Import vericert.hls.RTLBlockgen. (* rtlblockgenproof-imports ends here *) -(* [[file:../../lit/scheduling.org::rtlblockgenproof-match-states][rtlblockgenproof-match-states]] *) +(* [[file:../../lit/basic-block-generation.org::rtlblockgenproof-match-states][rtlblockgenproof-match-states]] *) Inductive match_states : RTL.state -> RTLBlock.state -> Prop := | match_state : forall stk f tf sp pc rs m @@ -34,7 +34,7 @@ Inductive match_states : RTL.state -> RTLBlock.state -> Prop := (RTLBlock.State stk tf sp (find_block max n i) rs m). (* rtlblockgenproof-match-states ends here *) -(* [[file:../../lit/scheduling.org::rtlblockgenproof-correctness][rtlblockgenproof-correctness]] *) +(* [[file:../../lit/basic-block-generation.org::rtlblockgenproof-correctness][rtlblockgenproof-correctness]] *) Section CORRECTNESS. Context (prog : RTL.program). diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v index bcb51c6..160f71c 100644 --- a/src/hls/RTLPar.v +++ b/src/hls/RTLPar.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlpar-main][rtlpar-main]] *) +(* [[file:../../lit/scheduler-languages.org::rtlpar-main][rtlpar-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Events. diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index f3d13f5..ac9ac2c 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlpargen-main][rtlpargen-main]] *) +(* [[file:../../lit/scheduler.org::rtlpargen-main][rtlpargen-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Globalenvs. diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v index 7b262db..5188709 100644 --- a/src/hls/RTLPargenproof.v +++ b/src/hls/RTLPargenproof.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::rtlpargenproof-imports][rtlpargenproof-imports]] *) +(* [[file:../../lit/scheduler.org::rtlpargenproof-imports][rtlpargenproof-imports]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Errors. @@ -39,7 +39,7 @@ Require Import vericert.hls.Abstr. #[local] Open Scope pred_op. (* rtlpargenproof-imports ends here *) -(* [[file:../../lit/scheduling.org::rtlpargenproof-main][rtlpargenproof-main]] *) +(* [[file:../../lit/scheduler.org::rtlpargenproof-main][rtlpargenproof-main]] *) (*Definition is_regs i := match i with mk_instr_state rs _ => rs end. Definition is_mem i := match i with mk_instr_state _ m => m end. @@ -1092,7 +1092,7 @@ Proof. induction 2; try rewrite H; eauto with barg. Qed. eauto. eauto. simplify. eauto. eauto. } { unfold bind in *. inv TRANSL0. clear Learn. inv H0. destruct_match; crush. inv H2. unfold transl_function in Heqr. destruct_match; crush. - inv Heqr. + inv Heqr. repeat econstructor; eauto. unfold bind in *. destruct_match; crush. } { inv TRANSL0. repeat econstructor; eauto using Events.external_call_symbols_preserved, symbols_preserved, senv_preserved, Events.E0_right. } diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index 88aa6fd..36dbc2b 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduling.org::scheduler-main][scheduler-main]] *) +(* [[file:../../lit/scheduler.org::scheduler-main][scheduler-main]] *) open Printf open Clflags open Camlcoq -- cgit From 4b012187df7c66bef2300252058f27ac79337325 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 24 Mar 2022 10:04:47 +0000 Subject: Rename lit directory --- src/hls/Partition.ml | 6 ++++-- src/hls/RTLBlock.v | 2 +- src/hls/RTLBlockInstr.v | 14 +++++++------- src/hls/RTLBlockgen.v | 4 ++-- src/hls/RTLBlockgenproof.v | 6 +++--- src/hls/RTLPar.v | 2 +- 6 files changed, 18 insertions(+), 16 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Partition.ml b/src/hls/Partition.ml index d7972e5..545277b 100644 --- a/src/hls/Partition.ml +++ b/src/hls/Partition.ml @@ -1,6 +1,6 @@ - (* +(* * Vericert: Verified high-level synthesis. - * Copyright (C) 2020 Yann Herklotz + * Copyright (C) 2020-2022 Yann Herklotz * * This program is free software: you can redistribute it and/or modify * it under the terms of the GNU General Public License as published by @@ -16,6 +16,7 @@ * along with this program. If not, see . *) +(* [[file:../../docs/basic-block-generation.org::partition-main][partition-main]] *) open Printf open Clflags open Camlcoq @@ -122,3 +123,4 @@ let function_from_RTL f = } let partition = function_from_RTL +(* partition-main ends here *) diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v index 539d5bc..50bff90 100644 --- a/src/hls/RTLBlock.v +++ b/src/hls/RTLBlock.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduler-languages.org::rtlblock-main][rtlblock-main]] *) +(* [[file:../../docs/scheduler-languages.org::rtlblock-main][rtlblock-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Events. diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 19b3928..801a5ea 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-imports][rtlblockinstr-imports]] *) +(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-imports][rtlblockinstr-imports]] *) Require Import Coq.micromega.Lia. Require Import compcert.backend.Registers. @@ -33,7 +33,7 @@ Require Import Predicate. Require Import Vericertlib. (* rtlblockinstr-imports ends here *) -(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-instr-def][rtlblockinstr-instr-def]] *) +(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-instr-def][rtlblockinstr-instr-def]] *) Definition node := positive. Inductive instr : Type := @@ -44,7 +44,7 @@ Inductive instr : Type := | RBsetpred : option pred_op -> condition -> list reg -> predicate -> instr. (* rtlblockinstr-instr-def ends here *) -(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-cf-instr-def][rtlblockinstr-cf-instr-def]] *) +(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-cf-instr-def][rtlblockinstr-cf-instr-def]] *) Inductive cf_instr : Type := | RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr | RBtailcall : signature -> reg + ident -> list reg -> cf_instr @@ -57,7 +57,7 @@ Inductive cf_instr : Type := | RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. (* rtlblockinstr-cf-instr-def ends here *) -(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-helpers][rtlblockinstr-helpers]] *) +(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-helpers][rtlblockinstr-helpers]] *) Fixpoint successors_instr (i : cf_instr) : list node := match i with | RBcall sig ros args res s => s :: nil @@ -148,7 +148,7 @@ Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset := end. (* rtlblockinstr-helpers ends here *) -(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-instr-state][rtlblockinstr-instr-state]] *) +(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-instr-state][rtlblockinstr-instr-state]] *) Record instr_state := mk_instr_state { is_rs: regset; is_ps: predset; @@ -156,7 +156,7 @@ Record instr_state := mk_instr_state { }. (* rtlblockinstr-instr-state ends here *) -(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-type-def][rtlblockinstr-type-def]] *) +(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-type-def][rtlblockinstr-type-def]] *) Section DEFINITION. Context {bblock_body: Type}. @@ -221,7 +221,7 @@ Section DEFINITION. End DEFINITION. (* rtlblockinstr-type-def ends here *) -(* [[file:../../lit/scheduler-languages.org::rtlblockinstr-semantics][rtlblockinstr-semantics]] *) +(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-semantics][rtlblockinstr-semantics]] *) Section RELSEM. Context {bblock_body : Type}. diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index 6d38e4f..beca0ea 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/basic-block-generation.org::rtlblockgen-imports][rtlblockgen-imports]] *) +(* [[file:../../docs/basic-block-generation.org::rtlblockgen-imports][rtlblockgen-imports]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. @@ -148,7 +148,7 @@ Defined. Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := if eqd a b then true else false. -(* [[file:../../lit/basic-block-generation.org::rtlblockgen-main][rtlblockgen-main]] *) +(* [[file:../../docs/basic-block-generation.org::rtlblockgen-main][rtlblockgen-main]] *) Parameter partition : RTL.function -> Errors.res function. (** [find_block max nodes index]: Does not need to be sorted, because we use filter and the max fold diff --git a/src/hls/RTLBlockgenproof.v b/src/hls/RTLBlockgenproof.v index 4433d52..d51e5d4 100644 --- a/src/hls/RTLBlockgenproof.v +++ b/src/hls/RTLBlockgenproof.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/basic-block-generation.org::rtlblockgenproof-imports][rtlblockgenproof-imports]] *) +(* [[file:../../docs/basic-block-generation.org::rtlblockgenproof-imports][rtlblockgenproof-imports]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. @@ -25,7 +25,7 @@ Require Import vericert.hls.RTLBlock. Require Import vericert.hls.RTLBlockgen. (* rtlblockgenproof-imports ends here *) -(* [[file:../../lit/basic-block-generation.org::rtlblockgenproof-match-states][rtlblockgenproof-match-states]] *) +(* [[file:../../docs/basic-block-generation.org::rtlblockgenproof-match-states][rtlblockgenproof-match-states]] *) Inductive match_states : RTL.state -> RTLBlock.state -> Prop := | match_state : forall stk f tf sp pc rs m @@ -34,7 +34,7 @@ Inductive match_states : RTL.state -> RTLBlock.state -> Prop := (RTLBlock.State stk tf sp (find_block max n i) rs m). (* rtlblockgenproof-match-states ends here *) -(* [[file:../../lit/basic-block-generation.org::rtlblockgenproof-correctness][rtlblockgenproof-correctness]] *) +(* [[file:../../docs/basic-block-generation.org::rtlblockgenproof-correctness][rtlblockgenproof-correctness]] *) Section CORRECTNESS. Context (prog : RTL.program). diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v index 160f71c..f380d19 100644 --- a/src/hls/RTLPar.v +++ b/src/hls/RTLPar.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduler-languages.org::rtlpar-main][rtlpar-main]] *) +(* [[file:../../docs/scheduler-languages.org::rtlpar-main][rtlpar-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Events. -- cgit From 30baa719fb15c45b13cb869056e51ec7446c0207 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 24 Mar 2022 11:46:15 +0000 Subject: Add more documentation --- src/hls/RTLBlock.v | 150 ++++++++++++++--------------------------------- src/hls/RTLBlockInstr.v | 16 ++++- src/hls/RTLPargen.v | 50 +++++----------- src/hls/RTLPargenproof.v | 4 +- src/hls/Schedule.ml | 2 +- 5 files changed, 73 insertions(+), 149 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v index 50bff90..98e032b 100644 --- a/src/hls/RTLBlock.v +++ b/src/hls/RTLBlock.v @@ -40,138 +40,74 @@ Definition fundef := @fundef bb. Definition program := @program bb. Definition funsig := @funsig bb. Definition stackframe := @stackframe bb. +Definition state := @state bb. Definition genv := @genv bb. +(* rtlblock-main ends here *) -Inductive state : Type := -| State: - forall (stack: list stackframe) (**r call stack *) - (f: function) (**r current function *) - (b: bb) (**r current block being executed *) - (sp: val) (**r stack pointer *) - (pc: node) (**r current program point in [c] *) - (rs: regset) (**r register state *) - (pr: predset) (**r predicate register state *) - (m: mem), (**r memory state *) - state -| Callstate: - forall (stack: list stackframe) (**r call stack *) - (f: fundef) (**r function to call *) - (args: list val) (**r arguments to the call *) - (m: mem), (**r memory state *) - state -| Returnstate: - forall (stack: list stackframe) (**r call stack *) - (v: val) (**r return value for the call *) - (m: mem), (**r memory state *) - state. - +(* [[file:../../docs/scheduler-languages.org::rtlblock-semantics][rtlblock-semantics]] *) Section RELSEM. Context (ge: genv). +(* rtlblock-semantics ends here *) +(* [[file:../../docs/scheduler-languages.org::#step_instr_list][rtlblock-step_instr_list]] *) Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop := - | exec_RBcons: - forall state i state' state'' instrs sp, - step_instr ge sp state i state' -> - step_instr_list sp state' instrs state'' -> - step_instr_list sp state (i :: instrs) state'' - | exec_RBnil: - forall state sp, - step_instr_list sp state nil state. - - Definition find_function - (ros: reg + ident) (rs: regset) : option fundef := - match ros with - | inl r => Genv.find_funct ge rs#r - | inr symb => - match Genv.find_symbol ge symb with - | None => None - | Some b => Genv.find_funct_ptr ge b - end - end. - - Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop := - | exec_RBcall: - forall s f b sp rs m res fd ros sig args pc pc' pr, - find_function ros rs = Some fd -> - funsig fd = sig -> - step_cf_instr (State s f b sp pc rs pr m) (RBcall sig ros args res pc') - E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m) - | exec_RBtailcall: - forall s f b stk rs m sig ros args fd m' pc pr, - find_function ros rs = Some fd -> - funsig fd = sig -> - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args) - E0 (Callstate s fd rs##args m') - | exec_RBbuiltin: - forall s f b sp rs m ef args res pc' vargs t vres m' pc pr, - eval_builtin_args ge (fun r => rs#r) sp m args vargs -> - external_call ef ge vargs m t vres m' -> - step_cf_instr (State s f b sp pc rs pr m) (RBbuiltin ef args res pc') - t (State s f b sp pc' (regmap_setres res vres rs) pr m') - | exec_RBcond: - forall s f block sp rs m cond args ifso ifnot b pc pc' pr, - eval_condition cond rs##args m = Some b -> - pc' = (if b then ifso else ifnot) -> - step_cf_instr (State s f block sp pc rs pr m) (RBcond cond args ifso ifnot) - E0 (State s f block sp pc' rs pr m) - | exec_RBjumptable: - forall s f b sp rs m arg tbl n pc pc' pr, - rs#arg = Vint n -> - list_nth_z tbl (Int.unsigned n) = Some pc' -> - step_cf_instr (State s f b sp pc rs pr m) (RBjumptable arg tbl) - E0 (State s f b sp pc' rs pr m) - | exec_RBreturn: - forall s f b stk rs m or pc m' pr, - Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step_cf_instr (State s f b (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or) - E0 (Returnstate s (regmap_optget or Vundef rs) m') - | exec_RBgoto: - forall s f b sp pc rs pr m pc', - step_cf_instr (State s f b sp pc rs pr m) (RBgoto pc') E0 (State s f b sp pc' rs pr m) - | exec_RBpred_cf: - forall s f b sp pc rs pr m cf1 cf2 st' p t, - step_cf_instr (State s f b sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> - step_cf_instr (State s f b sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'. - - Inductive step: state -> trace -> state -> Prop := + | exec_RBcons: + forall state i state' state'' instrs sp, + step_instr ge sp state i state' -> + step_instr_list sp state' instrs state'' -> + step_instr_list sp state (i :: instrs) state'' + | exec_RBnil: + forall state sp, + step_instr_list sp state nil state. +(* rtlblock-step_instr_list ends here *) + +(* [[file:../../docs/scheduler-languages.org::#rtlblock-step][rtlblock-step]] *) + Variant step: state -> trace -> state -> Prop := + | exec_bblock: + forall s f sp pc rs rs' m m' t s' bb pr pr', + f.(fn_code)!pc = Some bb -> + step_instr_list sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') -> + step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' -> + step (State s f sp pc rs pr m) t s' | exec_function_internal: forall s f args m m' stk, Mem.alloc m 0 f.(fn_stacksize) = (m', stk) -> step (Callstate s (Internal f) args m) - E0 (State s f - (Vptr stk Ptrofs.zero) - f.(fn_entrypoint) - (init_regs args f.(fn_params)) - (PMap.init false) - m') + E0 (State s f + (Vptr stk Ptrofs.zero) + f.(fn_entrypoint) + (init_regs args f.(fn_params)) + (PMap.init false) + m') | exec_function_external: forall s ef args res t m m', external_call ef ge args m t res m' -> step (Callstate s (External ef) args m) - t (Returnstate s res m') + t (Returnstate s res m') | exec_return: forall res f sp pc rs s vres m pr, step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m) - E0 (State s f sp pc (rs#res <- vres) pr m). + E0 (State s f sp pc (rs#res <- vres) pr m). +(* rtlblock-step ends here *) +(* [[file:../../docs/scheduler-languages.org::#rtlblock-step][rtlblock-rest]] *) End RELSEM. Inductive initial_state (p: program): state -> Prop := - | initial_state_intro: forall b f m0, - let ge := Genv.globalenv p in - Genv.init_mem p = Some m0 -> - Genv.find_symbol ge p.(prog_main) = Some b -> - Genv.find_funct_ptr ge b = Some f -> - funsig f = signature_main -> - initial_state p (Callstate nil f nil m0). +| initial_state_intro: forall b f m0, + let ge := Genv.globalenv p in + Genv.init_mem p = Some m0 -> + Genv.find_symbol ge p.(prog_main) = Some b -> + Genv.find_funct_ptr ge b = Some f -> + funsig f = signature_main -> + initial_state p (Callstate nil f nil m0). Inductive final_state: state -> int -> Prop := - | final_state_intro: forall r m, - final_state (Returnstate nil (Vint r) m) r. +| final_state_intro: forall r m, + final_state (Returnstate nil (Vint r) m) r. Definition semantics (p: program) := Semantics step (initial_state p) final_state (Genv.globalenv p). -(* rtlblock-main ends here *) +(* rtlblock-rest ends here *) diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 801a5ea..bd40516 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -195,8 +195,10 @@ Section DEFINITION. (rs: regset) (**r register state in calling function *) (pr: predset), (**r predicate state of the calling function *) stackframe. +(* rtlblockinstr-type-def ends here *) - Inductive state : Type := +(* [[file:../../docs/scheduler-languages.org::#state][rtlblockinstr-state]] *) + Variant state : Type := | State: forall (stack: list stackframe) (**r call stack *) (f: function) (**r current function *) @@ -217,9 +219,11 @@ Section DEFINITION. (v: val) (**r return value for the call *) (m: mem), (**r memory state *) state. +(* rtlblockinstr-state ends here *) +(* [[file:../../docs/scheduler-languages.org::#state][rtlblockinstr-state]] *) End DEFINITION. -(* rtlblockinstr-type-def ends here *) +(* rtlblockinstr-state ends here *) (* [[file:../../docs/scheduler-languages.org::rtlblockinstr-semantics][rtlblockinstr-semantics]] *) Section RELSEM. @@ -252,7 +256,9 @@ Section RELSEM. eval_pred (Some p) i i' i | eval_pred_none: forall i i', eval_pred None i i' i. +(* rtlblockinstr-semantics ends here *) +(* [[file:../../docs/scheduler-languages.org::#step_instr][rtlblockinstr-step_instr]] *) Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := | exec_RBnop: forall sp ist, @@ -279,7 +285,9 @@ Section RELSEM. Op.eval_condition c rs##args m = Some b -> eval_pred p' (mk_instr_state rs pr m) (mk_instr_state rs (pr#p <- b) m) ist -> step_instr sp (mk_instr_state rs pr m) (RBsetpred p' c args p) ist. +(* rtlblockinstr-step_instr ends here *) +(* [[file:../../docs/scheduler-languages.org::#step_cf_instr][rtlblockinstr-step_cf_instr]] *) Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop := | exec_RBcall: forall s f sp rs m res fd ros sig args pc pc' pr, @@ -324,6 +332,8 @@ Section RELSEM. forall s f sp pc rs pr m cf1 cf2 st' p t, step_cf_instr (State s f sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> step_cf_instr (State s f sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'. +(* rtlblockinstr-step_cf_instr ends here *) +(* [[file:../../docs/scheduler-languages.org::#step_cf_instr][rtlblockinstr-end_RELSEM]] *) End RELSEM. -(* rtlblockinstr-semantics ends here *) +(* rtlblockinstr-end_RELSEM ends here *) diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index ac9ac2c..31ea51f 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduler.org::rtlpargen-main][rtlpargen-main]] *) +(* [[file:../../docs/scheduler.org::rtlpargen-main][rtlpargen-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Globalenvs. @@ -35,18 +35,12 @@ Require Import vericert.hls.Predicate. Require Import vericert.hls.Abstr. Import NE.NonEmptyNotation. -(** * RTLPar Generation *) - #[local] Open Scope positive. #[local] Open Scope forest. #[local] Open Scope pred_op. +(* rtlpargen-main ends here *) -(** ** Abstract Computations - -Define the abstract computation using the [update] function, which will set each register to its -symbolic value. First we need to define a few helper functions to correctly translate the -predicates. *) - +(* [[file:../../docs/scheduler.org::rtlpargen-generation][rtlpargen-generation]] *) Fixpoint list_translation (l : list reg) (f : forest) {struct l} : list pred_expr := match l with | nil => nil @@ -139,19 +133,9 @@ Definition app_predicated {A: Type} (p: option pred_op) (a b: predicated A) := Definition pred_ret {A: Type} (a: A) : predicated A := NE.singleton (T, a). +(* rtlpargen-generation ends here *) -(** *** Update Function - -The [update] function will generate a new forest given an existing forest and a new instruction, -so that it can evaluate a symbolic expression by folding over a list of instructions. The main -problem is that predicates need to be merged as well, so that: - -1. The predicates are *independent*. -2. The expression assigned to the register should still be correct. - -This is done by multiplying the predicates together, and assigning the negation of the expression to -the other predicates. *) - +(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-update-function]] *) Definition update (f : forest) (i : instr) : forest := match i with | RBnop => f @@ -181,28 +165,22 @@ Definition update (f : forest) (i : instr) : forest := (f # (Pred p)) (map_predicated (pred_ret (Esetpred c)) (merge (list_translation args f)))) end. +(* rtlpargen-update-function ends here *) -(** Implementing which are necessary to show the correctness of the translation validation by -showing that there aren't any more effects in the resultant RTLPar code than in the RTLBlock code. - -Get a sequence from the basic block. *) - +(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-abstract-seq]] *) Fixpoint abstract_sequence (f : forest) (b : list instr) : forest := match b with | nil => f | i :: l => abstract_sequence (update f i) l end. +(* rtlpargen-abstract-seq ends here *) -(** Check equivalence of control flow instructions. As none of the basic blocks should have been -moved, none of the labels should be different, meaning the control-flow instructions should match -exactly. *) - +(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-check-functions]] *) Definition check_control_flow_instr (c1 c2: cf_instr) : bool := if cf_instr_eq c1 c2 then true else false. +(* rtlpargen-check-functions ends here *) -(** We define the top-level oracle that will check if two basic blocks are equivalent after a -scheduling transformation. *) - +(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-top-check-functions]] *) Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := match bb with | nil => @@ -241,9 +219,9 @@ Lemma check_scheduled_trees_correct2: PTree.get x f1 = None -> PTree.get x f2 = None. Proof. solve_scheduled_trees_correct. Qed. +(* rtlpargen-top-check-functions ends here *) -(** ** Top-level Functions *) - +(* [[file:../../docs/scheduler.org::rtlpargen-top-level-functions][rtlpargen-top-level-functions]] *) Parameter schedule : RTLBlock.function -> RTLPar.function. Definition transl_function (f: RTLBlock.function) : Errors.res RTLPar.function := @@ -261,4 +239,4 @@ Definition transl_fundef := transf_partial_fundef transl_function. Definition transl_program (p : RTLBlock.program) : Errors.res RTLPar.program := transform_partial_program transl_fundef p. -(* rtlpargen-main ends here *) +(* rtlpargen-top-level-functions ends here *) diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v index 5188709..6d61572 100644 --- a/src/hls/RTLPargenproof.v +++ b/src/hls/RTLPargenproof.v @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduler.org::rtlpargenproof-imports][rtlpargenproof-imports]] *) +(* [[file:../../docs/scheduler.org::rtlpargenproof-imports][rtlpargenproof-imports]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Errors. @@ -39,7 +39,7 @@ Require Import vericert.hls.Abstr. #[local] Open Scope pred_op. (* rtlpargenproof-imports ends here *) -(* [[file:../../lit/scheduler.org::rtlpargenproof-main][rtlpargenproof-main]] *) +(* [[file:../../docs/scheduler.org::rtlpargenproof-main][rtlpargenproof-main]] *) (*Definition is_regs i := match i with mk_instr_state rs _ => rs end. Definition is_mem i := match i with mk_instr_state _ m => m end. diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index 36dbc2b..3a5351e 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -16,7 +16,7 @@ * along with this program. If not, see . *) -(* [[file:../../lit/scheduler.org::scheduler-main][scheduler-main]] *) +(* [[file:../../docs/scheduler.org::scheduler-main][scheduler-main]] *) open Printf open Clflags open Camlcoq -- cgit From dd8d4ae9c320668ac5fd70f72ea76b768edf8165 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 26 Mar 2022 15:48:47 +0000 Subject: Remove literal files again --- src/hls/Abstr.v | 54 +++++++++++------ src/hls/Array.v | 5 +- src/hls/HTL.v | 16 +++-- src/hls/HTLPargen.v | 5 +- src/hls/HTLgenproof.v | 39 +++++++------ src/hls/HTLgenspec.v | 10 +++- src/hls/IfConversion.v | 12 ++-- src/hls/RICtransf.v | 14 +++-- src/hls/RTLBlock.v | 47 +++++++++++---- src/hls/RTLBlockInstr.v | 142 +++++++++++++++++++++++++++++++-------------- src/hls/RTLBlockgen.v | 15 ++--- src/hls/RTLBlockgenproof.v | 38 +++++++++--- src/hls/RTLPar.v | 40 +++++++------ src/hls/RTLPargen.v | 104 ++++++++++++++++++++++++--------- src/hls/RTLPargenproof.v | 15 +++-- src/hls/Schedule.ml | 2 - src/hls/ValueInt.v | 42 +++++++++----- src/hls/Verilog.v | 119 +++++++++++++++++++++++-------------- 18 files changed, 487 insertions(+), 232 deletions(-) (limited to 'src/hls') diff --git a/src/hls/Abstr.v b/src/hls/Abstr.v index 43d7783..9d310fb 100644 --- a/src/hls/Abstr.v +++ b/src/hls/Abstr.v @@ -38,11 +38,14 @@ Require Import vericert.hls.Predicate. #[local] Open Scope positive. #[local] Open Scope pred_op. -(** ** Schedule Oracle +(*| +Schedule Oracle +=============== -This oracle determines if a schedule was valid by performing symbolic execution on the input and -output and showing that these behave the same. This acts on each basic block separately, as the -rest of the functions should be equivalent. *) +This oracle determines if a schedule was valid by performing symbolic execution +on the input and output and showing that these behave the same. This acts on +each basic block separately, as the rest of the functions should be equivalent. +|*) Definition reg := positive. @@ -51,8 +54,11 @@ Inductive resource : Set := | Pred : reg -> resource | Mem : resource. -(** The following defines quite a few equality comparisons automatically, however, these can be -optimised heavily if written manually, as their proofs are not needed. *) +(*| +The following defines quite a few equality comparisons automatically, however, +these can be optimised heavily if written manually, as their proofs are not +needed. +|*) Lemma resource_eq : forall (r1 r2 : resource), {r1 = r2} + {r1 <> r2}. Proof. @@ -174,9 +180,12 @@ Proof. repeat decide equality. Defined. -(** We then create equality lemmas for a resource and a module to index resources uniquely. The -indexing is done by setting Mem to 1, whereas all other infinitely many registers will all be -shifted right by 1. This means that they will never overlap. *) +(*| +We then create equality lemmas for a resource and a module to index resources +uniquely. The indexing is done by setting Mem to 1, whereas all other +infinitely many registers will all be shifted right by 1. This means that they +will never overlap. +|*) Module R_indexed. Definition t := resource. @@ -193,17 +202,21 @@ Module R_indexed. Definition eq := resource_eq. End R_indexed. -(** We can then create expressions that mimic the expressions defined in RTLBlock and RTLPar, which use -expressions instead of registers as their inputs and outputs. This means that we can accumulate all -the results of the operations as general expressions that will be present in those registers. +(*| +We can then create expressions that mimic the expressions defined in RTLBlock +and RTLPar, which use expressions instead of registers as their inputs and +outputs. This means that we can accumulate all the results of the operations as +general expressions that will be present in those registers. - Ebase: the starting value of the register. - Eop: Some arithmetic operation on a number of registers. - Eload: A load from a memory location into a register. - Estore: A store from a register to a memory location. -Then, to make recursion over expressions easier, expression_list is also defined in the datatype, as -that enables mutual recursive definitions over the datatypes. *) +Then, to make recursion over expressions easier, expression_list is also defined +in the datatype, as that enables mutual recursive definitions over the +datatypes. +|*) Inductive expression : Type := | Ebase : resource -> expression @@ -312,8 +325,10 @@ Definition predicated A := NE.non_empty (pred_op * A). Definition pred_expr := predicated expression. -(** Using IMap we can create a map from resources to any other type, as resources can be uniquely -identified as positive numbers. *) +(*| +Using ``IMap`` we can create a map from resources to any other type, as +resources can be uniquely identified as positive numbers. +|*) Module Rtree := ITree(R_indexed). @@ -345,8 +360,11 @@ Definition get_m i := match i with mk_instr_state a b c => c end. Definition eval_predf_opt pr p := match p with Some p' => eval_predf pr p' | None => true end. -(** Finally we want to define the semantics of execution for the expressions with symbolic values, -so the result of executing the expressions will be an expressions. *) +(*| +Finally we want to define the semantics of execution for the expressions with +symbolic values, so the result of executing the expressions will be an +expressions. +|*) Section SEMANTICS. diff --git a/src/hls/Array.v b/src/hls/Array.v index 0f5ae02..de74611 100644 --- a/src/hls/Array.v +++ b/src/hls/Array.v @@ -190,7 +190,10 @@ Proof. auto using nth_error_nth. Qed. -(** Tail recursive version of standard library function. *) +(*| +Tail recursive version of standard library function. +|*) + Fixpoint list_repeat' {A : Type} (acc : list A) (a : A) (n : nat) : list A := match n with | O => acc diff --git a/src/hls/HTL.v b/src/hls/HTL.v index 47f2092..f3af3d8 100644 --- a/src/hls/HTL.v +++ b/src/hls/HTL.v @@ -36,10 +36,13 @@ Require Import ValueInt. Local Open Scope positive. -(** The purpose of the hardware transfer language (HTL) is to create a more hardware-like layout -that is still similar to the register transfer language (RTL) that it came from. The main change is -that function calls become module instantiations and that we now describe a state machine instead of -a control-flow graph. *) +(*| +The purpose of the hardware transfer language (HTL) is to create a more +hardware-like layout that is still similar to the register transfer language +(RTL) that it came from. The main change is that function calls become module +instantiations and that we now describe a state machine instead of a +control-flow graph. +|*) Local Open Scope assocmap. @@ -84,7 +87,10 @@ Fixpoint init_regs (vl : list value) (rl : list reg) {struct rl} := Definition empty_stack (m : module) : Verilog.assocmap_arr := (AssocMap.set m.(mod_ram).(ram_mem) (Array.arr_repeat None m.(mod_ram).(ram_size)) (AssocMap.empty Verilog.arr)). -(** * Operational Semantics *) +(*| +Operational Semantics +===================== +|*) Definition genv := Globalenvs.Genv.t fundef unit. diff --git a/src/hls/HTLPargen.v b/src/hls/HTLPargen.v index 8c85701..8960ef9 100644 --- a/src/hls/HTLPargen.v +++ b/src/hls/HTLPargen.v @@ -405,7 +405,10 @@ Definition translate_eff_addressing (a: Op.addressing) (args: list reg) | _, _ => error (Errors.msg "HTLPargen: translate_eff_addressing unsuported addressing") end. -(** Translate an instruction to a statement. FIX mulhs mulhu *) +(*| +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) diff --git a/src/hls/HTLgenproof.v b/src/hls/HTLgenproof.v index bf4b057..bf1ef1c 100644 --- a/src/hls/HTLgenproof.v +++ b/src/hls/HTLgenproof.v @@ -1004,22 +1004,25 @@ Section CORRECTNESS. constructor. Qed. - (** The proof of semantic preservation for the translation of instructions - is a simulation argument based on diagrams of the following form: -<< - match_states - code st rs ------------------------- State m st assoc - || | - || | - || | - \/ v - code st rs' ------------------------ State m st assoc' - match_states ->> - where [tr_code c data control fin rtrn st] is assumed to hold. - - The precondition and postcondition is that that should hold is [match_assocmaps rs assoc]. - *) +(*| +The proof of semantic preservation for the translation of instructions is a +simulation argument based on diagrams of the following form: + +:: +> match_states +> code st rs ------------------------- State m st assoc +> || | +> || | +> || | +> \/ v +> code st rs' ------------------------ State m st assoc' +> match_states + +where ``tr_code c data control fin rtrn st`` is assumed to hold. + +The precondition and postcondition is that that should hold is ``match_assocmaps +rs assoc``. +|*) Definition transl_instr_prop (instr : RTL.instruction) : Prop := forall m asr asa fin rtrn st stmt trans res, @@ -1091,7 +1094,7 @@ Section CORRECTNESS. apply Smallstep.plus_one. eapply HTL.step_module; eauto. inv CONST; assumption. - inv CONST; assumption. + inv CONST; assumption. (* processing of state *) econstructor. crush. @@ -2569,7 +2572,7 @@ Section CORRECTNESS. Smallstep.plus HTL.step tge R1 Events.E0 R2 /\ match_states (RTL.State s f sp pc' rs m) R2. Proof. intros s f sp pc rs m arg tbl n pc' H H0 H1 R1 MSTATE. - + #[local] Hint Resolve transl_ijumptable_correct : htlproof.*) Lemma transl_ireturn_correct: diff --git a/src/hls/HTLgenspec.v b/src/hls/HTLgenspec.v index 265b8f7..165fa91 100644 --- a/src/hls/HTLgenspec.v +++ b/src/hls/HTLgenspec.v @@ -120,11 +120,15 @@ Ltac monadInv H := ((progress simpl in H) || unfold F in H); monadInv1 H end. -(** * Relational specification of the translation *) +(*| +=========================================== +Relational specification of the translation +=========================================== -(** We now define inductive predicates that characterise the fact that the +We now define inductive predicates that characterise the fact that the statemachine that is created by the translation contains the correct -translations for each of the elements *) +translations for each of the elements. +|*) Inductive tr_instr (fin rtrn st stk : reg) : RTL.instruction -> stmnt -> stmnt -> Prop := | tr_instr_Inop : diff --git a/src/hls/IfConversion.v b/src/hls/IfConversion.v index 1879205..2516109 100644 --- a/src/hls/IfConversion.v +++ b/src/hls/IfConversion.v @@ -30,10 +30,14 @@ Require Import vericert.bourdoncle.Bourdoncle. Parameter build_bourdoncle : function -> (bourdoncle * PMap.t N). -(** * If-Conversion - -This conversion is a verified conversion from RTLBlock back to itself, which performs if-conversion -on basic blocks to make basic blocks larger. *) +(*| +============= +If-Conversion +============= + +This conversion is a verified conversion from RTLBlock back to itself, which +performs if-conversion on basic blocks to make basic blocks larger. +|*) Definition combine_pred (p: pred_op) (optp: option pred_op) := match optp with diff --git a/src/hls/RICtransf.v b/src/hls/RICtransf.v index d6e926b..643a3a7 100644 --- a/src/hls/RICtransf.v +++ b/src/hls/RICtransf.v @@ -27,12 +27,16 @@ Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLPar. Require Import vericert.hls.Predicate. -(** * Reverse If-Conversion +(*| +===================== +Reverse If-Conversion +===================== -This transformation takes a scheduling RTLPar block and removes any predicates from it. It can then -be trivially transformed into linear code again. It works by iteratively selecting a predicate, -then creating a branch based on it and placing subsequent instructions in one branch or the other. -*) +This transformation takes a scheduling RTLPar block and removes any predicates +from it. It can then be trivially transformed into linear code again. It works +by iteratively selecting a predicate, then creating a branch based on it and +placing subsequent instructions in one branch or the other. +|*) Fixpoint existsb_val {A B} (f: A -> option B) (l : list A) : option B := match l with diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v index 98e032b..ecd7561 100644 --- a/src/hls/RTLBlock.v +++ b/src/hls/RTLBlock.v @@ -16,7 +16,6 @@ * along with this program. If not, see . *) -(* [[file:../../docs/scheduler-languages.org::rtlblock-main][rtlblock-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Events. @@ -31,6 +30,12 @@ Require Import compcert.verilog.Op. Require Import vericert.hls.RTLBlockInstr. +(*| +======== +RTLBlock +======== +|*) + Definition bb := list instr. Definition bblock := @bblock bb. @@ -43,16 +48,30 @@ Definition stackframe := @stackframe bb. Definition state := @state bb. Definition genv := @genv bb. -(* rtlblock-main ends here *) -(* [[file:../../docs/scheduler-languages.org::rtlblock-semantics][rtlblock-semantics]] *) +(*| +Semantics +========= + +We first describe the semantics by assuming a global program environment with +type ~genv~ which was declared earlier. +|*) + Section RELSEM. Context (ge: genv). -(* rtlblock-semantics ends here *) -(* [[file:../../docs/scheduler-languages.org::#step_instr_list][rtlblock-step_instr_list]] *) - Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop := +(*| +Instruction list step +--------------------- + +The ``step_instr_list`` definition describes the execution of a list of +instructions in one big step, inductively traversing the list of instructions +and applying the ``step_instr``. +|*) + + Inductive step_instr_list: + val -> instr_state -> list instr -> instr_state -> Prop := | exec_RBcons: forall state i state' state'' instrs sp, step_instr ge sp state i state' -> @@ -61,14 +80,21 @@ Section RELSEM. | exec_RBnil: forall state sp, step_instr_list sp state nil state. -(* rtlblock-step_instr_list ends here *) -(* [[file:../../docs/scheduler-languages.org::#rtlblock-step][rtlblock-step]] *) +(*| +Top-level step +-------------- + +The step function itself then uses this big step of the list of instructions to +then show a transition from basic block to basic block. +|*) + Variant step: state -> trace -> state -> Prop := | exec_bblock: forall s f sp pc rs rs' m m' t s' bb pr pr', f.(fn_code)!pc = Some bb -> - step_instr_list sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') -> + step_instr_list sp (mk_instr_state rs pr m) bb.(bb_body) + (mk_instr_state rs' pr' m') -> step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' -> step (State s f sp pc rs pr m) t s' | exec_function_internal: @@ -90,9 +116,7 @@ Section RELSEM. forall res f sp pc rs s vres m pr, step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m) E0 (State s f sp pc (rs#res <- vres) pr m). -(* rtlblock-step ends here *) -(* [[file:../../docs/scheduler-languages.org::#rtlblock-step][rtlblock-rest]] *) End RELSEM. Inductive initial_state (p: program): state -> Prop := @@ -110,4 +134,3 @@ Inductive final_state: state -> int -> Prop := Definition semantics (p: program) := Semantics step (initial_state p) final_state (Genv.globalenv p). -(* rtlblock-rest ends here *) diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index bd40516..d23850a 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -16,7 +16,6 @@ * along with this program. If not, see . *) -(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-imports][rtlblockinstr-imports]] *) Require Import Coq.micromega.Lia. Require Import compcert.backend.Registers. @@ -31,20 +30,47 @@ Require Import compcert.verilog.Op. Require Import Predicate. Require Import Vericertlib. -(* rtlblockinstr-imports ends here *) -(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-instr-def][rtlblockinstr-instr-def]] *) Definition node := positive. +(*| +============= +RTLBlockInstr +============= + +These instructions are used for ``RTLBlock`` and ``RTLPar``, so that they have +consistent instructions, which greatly simplifies the proofs, as they will by +default have the same instruction syntax and semantics. The only changes are +therefore at the top-level of the instructions. + +Instruction Definition +====================== + +First, we define the instructions that can be placed into a basic block, meaning +they won't branch. The main changes to how instructions are defined in ``RTL``, +is that these instructions don't have a next node, as they will be in a basic +block, and they also have an optional predicate (``pred_op``). +|*) + Inductive instr : Type := | RBnop : instr -| RBop : option pred_op -> operation -> list reg -> reg -> instr -| RBload : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr -| RBstore : option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr -| RBsetpred : option pred_op -> condition -> list reg -> predicate -> instr. -(* rtlblockinstr-instr-def ends here *) +| RBop : + option pred_op -> operation -> list reg -> reg -> instr +| RBload : + option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr +| RBstore : + option pred_op -> memory_chunk -> addressing -> list reg -> reg -> instr +| RBsetpred : + option pred_op -> condition -> list reg -> predicate -> instr. + +(*| +Control-Flow Instruction Definition +=================================== + +These are the instructions that count as control-flow, and will be placed at the +end of the basic blocks. +|*) -(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-cf-instr-def][rtlblockinstr-cf-instr-def]] *) Inductive cf_instr : Type := | RBcall : signature -> reg + ident -> list reg -> reg -> node -> cf_instr | RBtailcall : signature -> reg + ident -> list reg -> cf_instr @@ -55,9 +81,12 @@ Inductive cf_instr : Type := | RBreturn : option reg -> cf_instr | RBgoto : node -> cf_instr | RBpred_cf : pred_op -> cf_instr -> cf_instr -> cf_instr. -(* rtlblockinstr-cf-instr-def ends here *) -(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-helpers][rtlblockinstr-helpers]] *) +(*| +Helper Functions +================ +|*) + Fixpoint successors_instr (i : cf_instr) : list node := match i with | RBcall sig ros args res s => s :: nil @@ -67,7 +96,8 @@ Fixpoint successors_instr (i : cf_instr) : list node := | RBjumptable arg tbl => tbl | RBreturn optarg => nil | RBgoto n => n :: nil - | RBpred_cf p c1 c2 => concat (successors_instr c1 :: successors_instr c2 :: nil) + | RBpred_cf p c1 c2 => + concat (successors_instr c1 :: successors_instr c2 :: nil) end. Definition max_reg_instr (m: positive) (i: instr) := @@ -136,7 +166,8 @@ Lemma eval_predf_pr_equiv : eval_predf ps p = eval_predf ps' p. Proof. induction p; simplify; auto; - try (unfold eval_predf; simplify; repeat (destruct_match; []); inv Heqp0; rewrite <- H; auto); + try (unfold eval_predf; simplify; + repeat (destruct_match; []); inv Heqp0; rewrite <- H; auto); [repeat rewrite eval_predf_Pand|repeat rewrite eval_predf_Por]; erewrite IHp1; try eassumption; erewrite IHp2; eauto. Qed. @@ -146,17 +177,30 @@ Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset := | r1 :: rs, v1 :: vs => Regmap.set r1 v1 (init_regs vs rs) | _, _ => Regmap.init Vundef end. -(* rtlblockinstr-helpers ends here *) -(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-instr-state][rtlblockinstr-instr-state]] *) +(*| +Instruction State +----------------- + +Definition of the instruction state, which contains the following: + +:is_rs: This is the current state of the registers. +:is_ps: This is the current state of the predicate registers, which is in a + separate namespace and area compared to the standard registers in [is_rs]. +:is_mem: The current state of the memory. +|*) + Record instr_state := mk_instr_state { is_rs: regset; is_ps: predset; is_mem: mem; }. -(* rtlblockinstr-instr-state ends here *) -(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-type-def][rtlblockinstr-type-def]] *) +(*| +Top-Level Type Definitions +========================== +|*) + Section DEFINITION. Context {bblock_body: Type}. @@ -193,11 +237,17 @@ Section DEFINITION. (sp: val) (**r stack pointer in calling function *) (pc: node) (**r program point in calling function *) (rs: regset) (**r register state in calling function *) - (pr: predset), (**r predicate state of the calling function *) + (pr: predset), (**r predicate state of the calling + function *) stackframe. -(* rtlblockinstr-type-def ends here *) -(* [[file:../../docs/scheduler-languages.org::#state][rtlblockinstr-state]] *) +(*| +State Definition +---------------- + +Definition of the ``state`` type, which is used by the ``step`` functions. +|*) + Variant state : Type := | State: forall (stack: list stackframe) (**r call stack *) @@ -219,13 +269,14 @@ Section DEFINITION. (v: val) (**r return value for the call *) (m: mem), (**r memory state *) state. -(* rtlblockinstr-state ends here *) -(* [[file:../../docs/scheduler-languages.org::#state][rtlblockinstr-state]] *) End DEFINITION. -(* rtlblockinstr-state ends here *) -(* [[file:../../docs/scheduler-languages.org::rtlblockinstr-semantics][rtlblockinstr-semantics]] *) +(*| +Semantics +========= +|*) + Section RELSEM. Context {bblock_body : Type}. @@ -245,7 +296,8 @@ Section RELSEM. end end. - Inductive eval_pred: option pred_op -> instr_state -> instr_state -> instr_state -> Prop := + Inductive eval_pred: + option pred_op -> instr_state -> instr_state -> instr_state -> Prop := | eval_pred_true: forall i i' p, eval_predf (is_ps i) p = true -> @@ -256,9 +308,7 @@ Section RELSEM. eval_pred (Some p) i i' i | eval_pred_none: forall i i', eval_pred None i i' i. -(* rtlblockinstr-semantics ends here *) -(* [[file:../../docs/scheduler-languages.org::#step_instr][rtlblockinstr-step_instr]] *) Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := | exec_RBnop: forall sp ist, @@ -266,28 +316,33 @@ Section RELSEM. | exec_RBop: forall op v res args rs m sp p ist pr, eval_operation ge sp op rs##args m = Some v -> - eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#res <- v) pr m) ist -> + eval_pred p (mk_instr_state rs pr m) + (mk_instr_state (rs#res <- v) pr m) ist -> step_instr sp (mk_instr_state rs pr m) (RBop p op args res) ist | exec_RBload: forall addr rs args a chunk m v dst sp p pr ist, eval_addressing ge sp addr rs##args = Some a -> Mem.loadv chunk m a = Some v -> - eval_pred p (mk_instr_state rs pr m) (mk_instr_state (rs#dst <- v) pr m) ist -> - step_instr sp (mk_instr_state rs pr m) (RBload p chunk addr args dst) ist + eval_pred p (mk_instr_state rs pr m) + (mk_instr_state (rs#dst <- v) pr m) ist -> + step_instr sp (mk_instr_state rs pr m) + (RBload p chunk addr args dst) ist | exec_RBstore: forall addr rs args a chunk m src m' sp p pr ist, eval_addressing ge sp addr rs##args = Some a -> Mem.storev chunk m a rs#src = Some m' -> - eval_pred p (mk_instr_state rs pr m) (mk_instr_state rs pr m') ist -> - step_instr sp (mk_instr_state rs pr m) (RBstore p chunk addr args src) ist + eval_pred p (mk_instr_state rs pr m) + (mk_instr_state rs pr m') ist -> + step_instr sp (mk_instr_state rs pr m) + (RBstore p chunk addr args src) ist | exec_RBsetpred: forall sp rs pr m p c b args p' ist, Op.eval_condition c rs##args m = Some b -> - eval_pred p' (mk_instr_state rs pr m) (mk_instr_state rs (pr#p <- b) m) ist -> - step_instr sp (mk_instr_state rs pr m) (RBsetpred p' c args p) ist. -(* rtlblockinstr-step_instr ends here *) + eval_pred p' (mk_instr_state rs pr m) + (mk_instr_state rs (pr#p <- b) m) ist -> + step_instr sp (mk_instr_state rs pr m) + (RBsetpred p' c args p) ist. -(* [[file:../../docs/scheduler-languages.org::#step_cf_instr][rtlblockinstr-step_cf_instr]] *) Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop := | exec_RBcall: forall s f sp rs m res fd ros sig args pc pc' pr, @@ -300,7 +355,8 @@ Section RELSEM. find_function ros rs = Some fd -> funsig fd = sig -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (RBtailcall sig ros args) + step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) + (RBtailcall sig ros args) E0 (Callstate s fd rs##args m') | exec_RBbuiltin: forall s f sp rs m ef args res pc' vargs t vres m' pc pr, @@ -327,13 +383,13 @@ Section RELSEM. E0 (Returnstate s (regmap_optget or Vundef rs) m') | exec_RBgoto: forall s f sp pc rs pr m pc', - step_cf_instr (State s f sp pc rs pr m) (RBgoto pc') E0 (State s f sp pc' rs pr m) + step_cf_instr (State s f sp pc rs pr m) (RBgoto pc') E0 + (State s f sp pc' rs pr m) | exec_RBpred_cf: forall s f sp pc rs pr m cf1 cf2 st' p t, - step_cf_instr (State s f sp pc rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> - step_cf_instr (State s f sp pc rs pr m) (RBpred_cf p cf1 cf2) t st'. -(* rtlblockinstr-step_cf_instr ends here *) + step_cf_instr (State s f sp pc rs pr m) + (if eval_predf pr p then cf1 else cf2) t st' -> + step_cf_instr (State s f sp pc rs pr m) + (RBpred_cf p cf1 cf2) t st'. -(* [[file:../../docs/scheduler-languages.org::#step_cf_instr][rtlblockinstr-end_RELSEM]] *) End RELSEM. -(* rtlblockinstr-end_RELSEM ends here *) diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index beca0ea..5d7376d 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -16,7 +16,6 @@ * along with this program. If not, see . *) -(* [[file:../../docs/basic-block-generation.org::rtlblockgen-imports][rtlblockgen-imports]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. @@ -28,7 +27,6 @@ Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLBlock. #[local] Open Scope positive. -(* rtlblockgen-imports ends here *) Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}. Proof. @@ -148,15 +146,19 @@ Defined. Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := if eqd a b then true else false. -(* [[file:../../docs/basic-block-generation.org::rtlblockgen-main][rtlblockgen-main]] *) Parameter partition : RTL.function -> Errors.res function. -(** [find_block max nodes index]: Does not need to be sorted, because we use filter and the max fold - function to find the desired element. *) +(*| +``find_block max nodes index``: Does not need to be sorted, because we use +filter and the max fold function to find the desired element. + + .. coq:: +|*) + Definition find_block (max: positive) (nodes: list positive) (index: positive) : positive := List.fold_right Pos.min max (List.filter (fun x => (index <=? x)) nodes). -(*Compute find_block (2::94::28::40::19::nil) 40.*) +(*Compute find_block 100 (2::94::48::39::19::nil) 40.*) Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) := match istr, istr' with @@ -287,4 +289,3 @@ Definition transl_fundef := transf_partial_fundef transl_function. Definition transl_program : RTL.program -> Errors.res program := transform_partial_program transl_fundef. -(* rtlblockgen-main ends here *) diff --git a/src/hls/RTLBlockgenproof.v b/src/hls/RTLBlockgenproof.v index d51e5d4..4568185 100644 --- a/src/hls/RTLBlockgenproof.v +++ b/src/hls/RTLBlockgenproof.v @@ -16,25 +16,47 @@ * along with this program. If not, see . *) -(* [[file:../../docs/basic-block-generation.org::rtlblockgenproof-imports][rtlblockgenproof-imports]] *) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. +Require Import compcert.common.Errors. +Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLBlock. Require Import vericert.hls.RTLBlockgen. -(* rtlblockgenproof-imports ends here *) -(* [[file:../../docs/basic-block-generation.org::rtlblockgenproof-match-states][rtlblockgenproof-match-states]] *) +(*| +Defining a find block specification +----------------------------------- + +Basically, it should be able to find the location of the block without using the +``find_block`` function, so that this is more useful for the proofs. There are +various different types of options that could come up though: + +1. The instruction is a standard instruction present inside of a basic block. +2. The instruction is a standard instruction which ends with a ``goto``. +3. The instruction is a control-flow instruction. + +For case number 1, there should exist a value in the list of instructions, such +that the instructions match exactly, and the indices match as well. In the +original code, this instruction must have been going from the current node to +the node - 1. For case number 2, there should be an instruction at the right +index again, however, this time there will also be a ``goto`` instruction in the +control-flow part of the basic block. +|*) + +(*Definition find_block_spec (c1: RTL.code) (c2: code) := + forall x i, + c1 ! x = Some i -> + exists y li, c2 ! y = Some li /\ nth_error li.(bb_body) ((Pos.to_nat y) - (Pos.to_nat x))%nat = Some i. + Inductive match_states : RTL.state -> RTLBlock.state -> Prop := | match_state : forall stk f tf sp pc rs m (TF: transl_function f = OK tf), - match_states (RTL.State stk f sp pc rs m) - (RTLBlock.State stk tf sp (find_block max n i) rs m). -(* rtlblockgenproof-match-states ends here *) + match_states (RTL.State stk f sp pc rs m) + (State stk tf sp (find_block max n i) rs m). -(* [[file:../../docs/basic-block-generation.org::rtlblockgenproof-correctness][rtlblockgenproof-correctness]] *) Section CORRECTNESS. Context (prog : RTL.program). @@ -49,4 +71,4 @@ Section CORRECTNESS. apply senv_preserved. End CORRECTNESS. -(* rtlblockgenproof-correctness ends here *) +*) diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v index f380d19..e0f657c 100644 --- a/src/hls/RTLPar.v +++ b/src/hls/RTLPar.v @@ -16,7 +16,6 @@ * along with this program. If not, see . *) -(* [[file:../../docs/scheduler-languages.org::rtlpar-main][rtlpar-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Events. @@ -47,15 +46,16 @@ Section RELSEM. Context (ge: genv). - Inductive step_instr_list: val -> instr_state -> list instr -> instr_state -> Prop := - | exec_RBcons: - forall state i state' state'' instrs sp, - step_instr ge sp state i state' -> - step_instr_list sp state' instrs state'' -> - step_instr_list sp state (i :: instrs) state'' - | exec_RBnil: - forall state sp, - step_instr_list sp state nil state. + Inductive step_instr_list: + val -> instr_state -> list instr -> instr_state -> Prop := + | exec_RBcons: + forall state i state' state'' instrs sp, + step_instr ge sp state i state' -> + step_instr_list sp state' instrs state'' -> + step_instr_list sp state (i :: instrs) state'' + | exec_RBnil: + forall state sp, + step_instr_list sp state nil state. Inductive step_instr_seq (sp : val) : instr_state -> list (list instr) -> instr_state -> Prop := @@ -83,7 +83,8 @@ Section RELSEM. | exec_bblock: forall s f sp pc rs rs' m m' t s' bb pr pr', f.(fn_code)!pc = Some bb -> - step_instr_block sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') -> + step_instr_block sp (mk_instr_state rs pr m) bb.(bb_body) + (mk_instr_state rs' pr' m') -> step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' -> step (State s f sp pc rs pr m) t s' | exec_function_internal: @@ -126,7 +127,11 @@ Definition semantics (p: program) := Semantics step (initial_state p) final_state (Genv.globalenv p). Definition max_reg_bblock (m : positive) (pc : node) (bb : bblock) := - let max_body := fold_left (fun x l => fold_left (fun x' l' => fold_left max_reg_instr l' x') l x) bb.(bb_body) m in + let max_body := + fold_left + (fun x l => fold_left (fun x' l' => fold_left max_reg_instr l' x') l x) + bb.(bb_body) m + in max_reg_cfi max_body bb.(bb_exit). Definition max_reg_function (f: function) := @@ -135,8 +140,9 @@ Definition max_reg_function (f: function) := (fold_left Pos.max f.(fn_params) 1%positive). Definition max_pc_function (f: function) : positive := - PTree.fold (fun m pc i => (Pos.max m - (pc + match Zlength i.(bb_body) - with Z.pos p => p | _ => 1 end))%positive) - f.(fn_code) 1%positive. -(* rtlpar-main ends here *) + PTree.fold + (fun m pc i => + (Pos.max m + (pc + match Zlength i.(bb_body) + with Z.pos p => p | _ => 1 end))%positive) + f.(fn_code) 1%positive. diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index 31ea51f..d425049 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -16,7 +16,6 @@ * along with this program. If not, see . *) -(* [[file:../../docs/scheduler.org::rtlpargen-main][rtlpargen-main]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Globalenvs. @@ -38,10 +37,22 @@ Import NE.NonEmptyNotation. #[local] Open Scope positive. #[local] Open Scope forest. #[local] Open Scope pred_op. -(* rtlpargen-main ends here *) -(* [[file:../../docs/scheduler.org::rtlpargen-generation][rtlpargen-generation]] *) -Fixpoint list_translation (l : list reg) (f : forest) {struct l} : list pred_expr := +(*| +========= +RTLPargen +========= + +Abstract Computations +===================== + +Define the abstract computation using the [update] function, which will set each +register to its symbolic value. First we need to define a few helper functions +to correctly translate the predicates. +|*) + +Fixpoint list_translation (l : list reg) (f : forest) {struct l} + : list pred_expr := match l with | nil => nil | i :: l => (f # (Reg i)) :: (list_translation l f) @@ -65,7 +76,8 @@ Definition merge'' x := | ((a, e), (b, el)) => (merge''' a b, Econs e el) end. -Definition map_pred_op {A B} (pf: option pred_op * (A -> B)) (pa: option pred_op * A): option pred_op * B := +Definition map_pred_op {A B} (pf: option pred_op * (A -> B)) + (pa: option pred_op * A): option pred_op * B := match pa, pf with | (p, a), (p', f) => (merge''' p p', f a) end. @@ -74,8 +86,8 @@ Definition predicated_prod {A B: Type} (p1: predicated A) (p2: predicated B) := NE.map (fun x => match x with ((a, b), (c, d)) => (Pand a c, (b, d)) end) (NE.non_empty_prod p1 p2). -Definition predicated_map {A B: Type} (f: A -> B) (p: predicated A): predicated B := - NE.map (fun x => (fst x, f (snd x))) p. +Definition predicated_map {A B: Type} (f: A -> B) (p: predicated A) + : predicated B := NE.map (fun x => (fst x, f (snd x))) p. (*map (fun x => (fst x, Econs (snd x) Enil)) pel*) Definition merge' (pel: pred_expr) (tpel: predicated expression_list) := @@ -87,16 +99,20 @@ Fixpoint merge (pel: list pred_expr): predicated expression_list := | a :: b => merge' a (merge b) end. -Definition map_predicated {A B} (pf: predicated (A -> B)) (pa: predicated A): predicated B := +Definition map_predicated {A B} (pf: predicated (A -> B)) (pa: predicated A) + : predicated B := predicated_map (fun x => (fst x) (snd x)) (predicated_prod pf pa). -Definition predicated_apply1 {A B} (pf: predicated (A -> B)) (pa: A): predicated B := +Definition predicated_apply1 {A B} (pf: predicated (A -> B)) (pa: A) + : predicated B := NE.map (fun x => (fst x, (snd x) pa)) pf. -Definition predicated_apply2 {A B C} (pf: predicated (A -> B -> C)) (pa: A) (pb: B): predicated C := +Definition predicated_apply2 {A B C} (pf: predicated (A -> B -> C)) (pa: A) + (pb: B): predicated C := NE.map (fun x => (fst x, (snd x) pa pb)) pf. -Definition predicated_apply3 {A B C D} (pf: predicated (A -> B -> C -> D)) (pa: A) (pb: B) (pc: C): predicated D := +Definition predicated_apply3 {A B C D} (pf: predicated (A -> B -> C -> D)) + (pa: A) (pb: B) (pc: C): predicated D := NE.map (fun x => (fst x, (snd x) pa pb pc)) pf. Definition predicated_from_opt {A: Type} (p: option pred_op) (a: A) := @@ -133,9 +149,23 @@ Definition app_predicated {A: Type} (p: option pred_op) (a b: predicated A) := Definition pred_ret {A: Type} (a: A) : predicated A := NE.singleton (T, a). -(* rtlpargen-generation ends here *) -(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-update-function]] *) +(*| +Update Function +--------------- + +The ``update`` function will generate a new forest given an existing forest and +a new instruction, so that it can evaluate a symbolic expression by folding over +a list of instructions. The main problem is that predicates need to be merged +as well, so that: + +1. The predicates are *independent*. +2. The expression assigned to the register should still be correct. + +This is done by multiplying the predicates together, and assigning the negation +of the expression to the other predicates. +|*) + Definition update (f : forest) (i : instr) : forest := match i with | RBnop => f @@ -149,7 +179,8 @@ Definition update (f : forest) (i : instr) : forest := (app_predicated p (f # (Reg r)) (map_predicated - (map_predicated (pred_ret (Eload chunk addr)) (merge (list_translation rl f))) + (map_predicated (pred_ret (Eload chunk addr)) + (merge (list_translation rl f))) (f # Mem))) | RBstore p chunk addr rl r => f # Mem <- @@ -157,30 +188,45 @@ Definition update (f : forest) (i : instr) : forest := (f # Mem) (map_predicated (map_predicated - (predicated_apply2 (map_predicated (pred_ret Estore) (f # (Reg r))) chunk addr) + (predicated_apply2 (map_predicated (pred_ret Estore) + (f # (Reg r))) chunk addr) (merge (list_translation rl f))) (f # Mem))) | RBsetpred p' c args p => f # (Pred p) <- (app_predicated p' (f # (Pred p)) - (map_predicated (pred_ret (Esetpred c)) (merge (list_translation args f)))) + (map_predicated (pred_ret (Esetpred c)) + (merge (list_translation args f)))) end. -(* rtlpargen-update-function ends here *) -(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-abstract-seq]] *) +(*| +Implementing which are necessary to show the correctness of the translation +validation by showing that there aren't any more effects in the resultant RTLPar +code than in the RTLBlock code. + +Get a sequence from the basic block. +|*) + Fixpoint abstract_sequence (f : forest) (b : list instr) : forest := match b with | nil => f | i :: l => abstract_sequence (update f i) l end. -(* rtlpargen-abstract-seq ends here *) -(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-check-functions]] *) +(*| +Check equivalence of control flow instructions. As none of the basic blocks +should have been moved, none of the labels should be different, meaning the +control-flow instructions should match exactly. +|*) + Definition check_control_flow_instr (c1 c2: cf_instr) : bool := if cf_instr_eq c1 c2 then true else false. -(* rtlpargen-check-functions ends here *) -(* [[file:../../docs/scheduler.org::#update-function][rtlpargen-top-check-functions]] *) +(*| +We define the top-level oracle that will check if two basic blocks are +equivalent after a scheduling transformation. +|*) + Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := match bb with | nil => @@ -219,12 +265,16 @@ Lemma check_scheduled_trees_correct2: PTree.get x f1 = None -> PTree.get x f2 = None. Proof. solve_scheduled_trees_correct. Qed. -(* rtlpargen-top-check-functions ends here *) -(* [[file:../../docs/scheduler.org::rtlpargen-top-level-functions][rtlpargen-top-level-functions]] *) +(*| +Top-level Functions +=================== +|*) + Parameter schedule : RTLBlock.function -> RTLPar.function. -Definition transl_function (f: RTLBlock.function) : Errors.res RTLPar.function := +Definition transl_function (f: RTLBlock.function) + : Errors.res RTLPar.function := let tfcode := fn_code (schedule f) in if check_scheduled_trees f.(fn_code) tfcode then Errors.OK (mkfunction f.(fn_sig) @@ -233,10 +283,10 @@ Definition transl_function (f: RTLBlock.function) : Errors.res RTLPar.function : tfcode f.(fn_entrypoint)) else - Errors.Error (Errors.msg "RTLPargen: Could not prove the blocks equivalent."). + Errors.Error + (Errors.msg "RTLPargen: Could not prove the blocks equivalent."). Definition transl_fundef := transf_partial_fundef transl_function. Definition transl_program (p : RTLBlock.program) : Errors.res RTLPar.program := transform_partial_program transl_fundef p. -(* rtlpargen-top-level-functions ends here *) diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v index 6d61572..4e88b13 100644 --- a/src/hls/RTLPargenproof.v +++ b/src/hls/RTLPargenproof.v @@ -16,7 +16,6 @@ * along with this program. If not, see . *) -(* [[file:../../docs/scheduler.org::rtlpargenproof-imports][rtlpargenproof-imports]] *) Require Import compcert.backend.Registers. Require Import compcert.common.AST. Require Import compcert.common.Errors. @@ -37,9 +36,18 @@ Require Import vericert.hls.Abstr. #[local] Open Scope positive. #[local] Open Scope forest. #[local] Open Scope pred_op. -(* rtlpargenproof-imports ends here *) -(* [[file:../../docs/scheduler.org::rtlpargenproof-main][rtlpargenproof-main]] *) +(*| +============== +RTLPargenproof +============== + +RTLBlock to abstract translation +================================ + +Correctness of translation from RTLBlock to the abstract interpretation language. +|*) + (*Definition is_regs i := match i with mk_instr_state rs _ => rs end. Definition is_mem i := match i with mk_instr_state _ m => m end. @@ -1134,4 +1142,3 @@ Proof. induction 2; try rewrite H; eauto with barg. Qed. Qed. End CORRECTNESS. -(* rtlpargenproof-main ends here *) diff --git a/src/hls/Schedule.ml b/src/hls/Schedule.ml index 3a5351e..70395e7 100644 --- a/src/hls/Schedule.ml +++ b/src/hls/Schedule.ml @@ -16,7 +16,6 @@ * along with this program. If not, see . *) -(* [[file:../../docs/scheduler.org::scheduler-main][scheduler-main]] *) open Printf open Clflags open Camlcoq @@ -882,4 +881,3 @@ let schedule_fn (f : RTLBlock.coq_function) : RTLPar.coq_function = fn_code = scheduled (*List.fold_left (add_to_tree scheduled) PTree.empty reachable*); fn_entrypoint = f.fn_entrypoint } -(* scheduler-main ends here *) diff --git a/src/hls/ValueInt.v b/src/hls/ValueInt.v index e434abc..06b5630 100644 --- a/src/hls/ValueInt.v +++ b/src/hls/ValueInt.v @@ -22,22 +22,31 @@ From compcert Require Import lib.Integers common.Values. From vericert Require Import Vericertlib. (* end hide *) -(** * Value +(*| +===== +Value +===== -A [value] is a bitvector with a specific size. We are using the implementation +A ``value`` is a bitvector with a specific size. We are using the implementation of the bitvector by mit-plv/bbv, because it has many theorems that we can reuse. -However, we need to wrap it with an [Inductive] so that we can specify and match -on the size of the [value]. This is necessary so that we can easily store -[value]s of different sizes in a list or in a map. +However, we need to wrap it with an ``Inductive`` so that we can specify and +match on the size of the ``value``. This is necessary so that we can easily +store ``value`` of different sizes in a list or in a map. -Using the default [word], this would not be possible, as the size is part of the type. *) +Using the default ``word``, this would not be possible, as the size is part of +the type. +|*) Definition value : Type := int. -(** ** Value conversions +(*| +Value conversions +================= -Various conversions to different number types such as [N], [Z], [positive] and -[int], where the last one is a theory of integers of powers of 2 in CompCert. *) +Various conversions to different number types such as ``N``, ``Z``, ``positive`` +and ``int``, where the last one is a theory of integers of powers of 2 in +CompCert. +|*) Definition valueToNat (v : value) : nat := Z.to_nat (Int.unsigned v). @@ -83,10 +92,12 @@ Definition valToValue (v : Values.val) : option value := | _ => None end. -(** Convert a [value] to a [bool], so that choices can be made based on the -result. This is also because comparison operators will give back [value] instead -of [bool], so if they are in a condition, they will have to be converted before -they can be used. *) +(*| +Convert a ``value`` to a ``bool``, so that choices can be made based on the +result. This is also because comparison operators will give back ``value`` +instead of ``bool``, so if they are in a condition, they will have to be +converted before they can be used. +|*) Definition valueToBool (v : value) : bool := if Z.eqb (uvalueToZ v) 0 then false else true. @@ -94,7 +105,10 @@ Definition valueToBool (v : value) : bool := Definition boolToValue (b : bool) : value := natToValue (if b then 1 else 0). -(** ** Arithmetic operations *) +(*| +Arithmetic operations +--------------------- +|*) Inductive val_value_lessdef: val -> value -> Prop := | val_value_lessdef_int: diff --git a/src/hls/Verilog.v b/src/hls/Verilog.v index 7561f77..72140cd 100644 --- a/src/hls/Verilog.v +++ b/src/hls/Verilog.v @@ -142,23 +142,28 @@ Definition nonblock_reg (r : reg) (asr : reg_associations) (v : value) := Inductive scl_decl : Type := VScalar (sz : nat). Inductive arr_decl : Type := VArray (sz : nat) (ln : nat). -(** ** Verilog AST +(*| +Verilog AST +=========== The Verilog AST is defined here, which is the target language of the compilation. -*** Value +Value +----- -The Verilog [value] is a bitvector containg a size and the actual bitvector. In -this case, the [Word.word] type is used as many theorems about that bitvector -have already been proven. +The Verilog ``value`` is a bitvector containg a size and the actual +bitvector. In this case, the ``Word.word`` type is used as many theorems about +that bitvector have already been proven. -*** Binary Operators +Binary Operators +---------------- These are the binary operations that can be represented in Verilog. Ideally, multiplication and division would be done by custom modules which could al so be scheduled properly. However, for now every Verilog operator is assumed to take -one clock cycle. *) +one clock cycle. +|*) Inductive binop : Type := | Vadd : binop (** addition (binary [+]) *) @@ -186,13 +191,19 @@ Inductive binop : Type := | Vshru : binop. (** shift right unsigned ([>>]) *) (* Vror : binop (** shift right unsigned ([>>]) *)*) -(** *** Unary Operators *) +(*| +Unary Operators +--------------- +|*) Inductive unop : Type := | Vneg (** negation ([-]) *) | Vnot. (** not operation [!] *) -(** *** Expressions *) +(*| +Expressions +----------- +|*) Inductive expr : Type := | Vlit : value -> expr @@ -207,7 +218,10 @@ Inductive expr : Type := Definition posToExpr (p : positive) : expr := Vlit (posToValue p). -(** *** Statements *) +(*| +Statements +---------- +|*) Inductive stmnt : Type := | Vskip : stmnt @@ -220,13 +234,17 @@ with stmnt_expr_list : Type := | Stmntnil : stmnt_expr_list | Stmntcons : expr -> stmnt -> stmnt_expr_list -> stmnt_expr_list. -(** *** Edges +(*| +Edges +----- -These define when an always block should be triggered, for example if the always block should be -triggered synchronously at the clock edge, or asynchronously for combinational logic. +These define when an always block should be triggered, for example if the always +block should be triggered synchronously at the clock edge, or asynchronously for +combinational logic. -[edge] is not part of [stmnt] in this formalisation because is closer to the semantics that are -used. *) +``edge`` is not part of ``stmnt`` in this formalisation because is closer to the +semantics that are used. +|*) Inductive edge : Type := | Vposedge : reg -> edge @@ -234,11 +252,15 @@ Inductive edge : Type := | Valledge : edge | Voredge : edge -> edge -> edge. -(** *** Module Items +(*| +Module Items +------------ -Module items can either be declarations ([Vdecl]) or always blocks ([Valways]). The declarations -are always register declarations as combinational logic can be done using combinational always block -instead of continuous assignments. *) +Module items can either be declarations (``Vdecl``) or always blocks +(``Valways``). The declarations are always register declarations as +combinational logic can be done using combinational always block instead of +continuous assignments. +|*) Inductive io : Type := | Vinput : io @@ -255,15 +277,17 @@ Inductive module_item : Type := | Valways_ff : edge -> stmnt -> module_item | Valways_comb : edge -> stmnt -> module_item. -(** The main module type containing all the important control signals +(*| +The main module type containing all the important control signals -- [mod_start] If set, starts the computations in the module. -- [mod_reset] If set, reset the state in the module. -- [mod_clk] The clock that controls the computation in the module. -- [mod_args] The arguments to the module. -- [mod_finish] Bit that is set if the result is ready. -- [mod_return] The return value that was computed. -- [mod_body] The body of the module, including the state machine. *) +:mod_start: If set, starts the computations in the module. +:mod_reset: If set, reset the state in the module. +:mod_clk: The clock that controls the computation in the module. +:mod_args: The arguments to the module. +:mod_finish: Bit that is set if the result is ready. +:mod_return: The return value that was computed. +:mod_body: The body of the module, including the state machine. +|*) Record module : Type := mkmodule { mod_start : reg; @@ -271,7 +295,8 @@ Record module : Type := mkmodule { mod_clk : reg; mod_finish : reg; mod_return : reg; - mod_st : reg; (**r Variable that defines the current state, it should be internal. *) + mod_st : reg; (**r Variable that defines the current state, + it should be internal. *) mod_stk : reg; mod_stk_len : nat; mod_args : list reg; @@ -283,33 +308,39 @@ Definition fundef := AST.fundef module. Definition program := AST.program fundef unit. -(** Convert a [positive] to an expression directly, setting it to the right size. *) +(*| +Convert a ``positive`` to an expression directly, setting it to the right size. +|*) + Definition posToLit (p : positive) : expr := Vlit (posToValue p). Definition fext := unit. Definition fextclk := nat -> fext. -(** *** State +(*| +State +----- -The [state] contains the following items: +The ``state`` contains the following items: -- Current [module] that is being worked on, so that the state variable can be -retrieved and set appropriately. -- Current [module_item] which is being worked on. -- A contiunation ([cont]) which defines what to do next. The option is to - either evaluate another module item or go to the next clock cycle. Finally - it could also end if the finish flag of the module goes high. +- Current ``module`` that is being worked on, so that the state variable can be + retrieved and set appropriately. +- Current ``module_item`` which is being worked on. +- A contiunation (``cont``) which defines what to do next. The option is to + either evaluate another module item or go to the next clock cycle. Finally it + could also end if the finish flag of the module goes high. - Association list containing the blocking assignments made, or assignments made in previous clock cycles. - Nonblocking association list containing all the nonblocking assignments made in the current module. - The environment containing values for the input. -- The program counter which determines the value for the state in this version of - Verilog, as the Verilog was generated by the RTL, which means that we have to - make an assumption about how it looks. In this case, that it contains state - which determines which part of the Verilog is executed. This is then the part - of the Verilog that should match with the RTL. *) +- The program counter which determines the value for the state in this version + of Verilog, as the Verilog was generated by the RTL, which means that we have + to make an assumption about how it looks. In this case, that it contains + state which determines which part of the Verilog is executed. This is then + the part of the Verilog that should match with the RTL. +|*) Inductive stackframe : Type := Stackframe : @@ -429,7 +460,9 @@ Definition handle_def {A : Type} (a : A) (val : option A) Local Open Scope error_monad_scope. -(** Return the location of the lhs of an assignment. *) +(*| +Return the location of the lhs of an assignment. +|*) Inductive location : Type := | LocReg (_ : reg) -- cgit From 11d6215f265d0dbcfd0048819a614f318a0775a4 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Sat, 26 Mar 2022 21:08:23 +0000 Subject: Add sphinx documentation --- src/hls/RTLBlockInstr.v | 55 ++++++++++++++++++++++++++----------------------- 1 file changed, 29 insertions(+), 26 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index d23850a..d308366 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -1,20 +1,32 @@ -(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 2020-2022 Yann Herklotz - * - * This program is free software: you can redistribute it and/or modify - * it under the terms of the GNU General Public License as published by - * the Free Software Foundation, either version 3 of the License, or - * (at your option) any later version. - * - * This program is distributed in the hope that it will be useful, - * but WITHOUT ANY WARRANTY; without even the implied warranty of - * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the - * GNU General Public License for more details. - * - * You should have received a copy of the GNU General Public License - * along with this program. If not, see . - *) +(*| +.. + Vericert: Verified high-level synthesis. + Copyright (C) 2019-2022 Yann Herklotz + + This program is free software: you can redistribute it and/or modify + it under the terms of the GNU General Public License as published by + the Free Software Foundation, either version 3 of the License, or + (at your option) any later version. + + This program is distributed in the hope that it will be useful, + but WITHOUT ANY WARRANTY; without even the implied warranty of + MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + GNU General Public License for more details. + + You should have received a copy of the GNU General Public License + along with this program. If not, see . + +============= +RTLBlockInstr +============= + +These instructions are used for ``RTLBlock`` and ``RTLPar``, so that they have +consistent instructions, which greatly simplifies the proofs, as they will by +default have the same instruction syntax and semantics. The only changes are +therefore at the top-level of the instructions. + +.. coq:: none +|*) Require Import Coq.micromega.Lia. @@ -34,15 +46,6 @@ Require Import Vericertlib. Definition node := positive. (*| -============= -RTLBlockInstr -============= - -These instructions are used for ``RTLBlock`` and ``RTLPar``, so that they have -consistent instructions, which greatly simplifies the proofs, as they will by -default have the same instruction syntax and semantics. The only changes are -therefore at the top-level of the instructions. - Instruction Definition ====================== -- cgit From 34e0f092551fcd7e1eef4a8a3c863fa940dcbf2f Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 28 Mar 2022 01:25:37 +0100 Subject: Work on specification of RTLBlock generation --- src/hls/RTLBlockInstr.v | 22 ++++++ src/hls/RTLBlockgen.v | 171 +++++++++------------------------------------ src/hls/RTLBlockgenproof.v | 77 +++++++++++++------- 3 files changed, 108 insertions(+), 162 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index d308366..cd23da3 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -46,6 +46,9 @@ Require Import Vericertlib. Definition node := positive. (*| +.. index:: + triple: definition; RTLBlockInstr; instruction + Instruction Definition ====================== @@ -67,6 +70,9 @@ Inductive instr : Type := option pred_op -> condition -> list reg -> predicate -> instr. (*| +.. index:: + triple: definition; RTLBlockInstr; control-flow instruction + Control-Flow Instruction Definition =================================== @@ -312,6 +318,14 @@ Section RELSEM. | eval_pred_none: forall i i', eval_pred None i i' i. +(*| +.. index:: + triple: semantics; RTLBlockInstr; instruction + +Step Instruction +============================= +|*) + Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := | exec_RBnop: forall sp ist, @@ -346,6 +360,14 @@ Section RELSEM. step_instr sp (mk_instr_state rs pr m) (RBsetpred p' c args p) ist. +(*| +.. index:: + triple: semantics; RTLBlockInstr; control-flow instruction + +Step Control-Flow Instruction +============================= +|*) + Inductive step_cf_instr: state -> cf_instr -> trace -> state -> Prop := | exec_RBcall: forall s f sp rs m res fd ros sig args pc pc' pr, diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index 5d7376d..a29709b 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -1,20 +1,25 @@ -(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 2020-2022 Yann Herklotz - * - * This program is free software: you can redistribute it and/or modify - * it under the terms of the GNU General Public License as published by - * the Free Software Foundation, either version 3 of the License, or - * (at your option) any later version. - * - * This program is distributed in the hope that it will be useful, - * but WITHOUT ANY WARRANTY; without even the implied warranty of - * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the - * GNU General Public License for more details. - * - * You should have received a copy of the GNU General Public License - * along with this program. If not, see . - *) +(*| +.. + Vericert: Verified high-level synthesis. + Copyright (C) 2020-2022 Yann Herklotz + + This program is free software: you can redistribute it and/or modify + it under the terms of the GNU General Public License as published by + the Free Software Foundation, either version 3 of the License, or + (at your option) any later version. + + This program is distributed in the hope that it will be useful, + but WITHOUT ANY WARRANTY; without even the implied warranty of + MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + GNU General Public License for more details. + + You should have received a copy of the GNU General Public License + along with this program. If not, see . + +=========== +RTLBlockgen +=========== +|*) Require compcert.backend.RTL. Require Import compcert.common.AST. @@ -22,137 +27,18 @@ Require Import compcert.lib.Maps. Require Import compcert.lib.Integers. Require Import compcert.lib.Floats. +Require Import vericert.common.DecEq. Require Import vericert.common.Vericertlib. Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLBlock. #[local] Open Scope positive. -Lemma comparison_eq: forall (x y : comparison), {x = y} + {x <> y}. -Proof. - decide equality. -Defined. - -Lemma condition_eq: forall (x y : Op.condition), {x = y} + {x <> y}. -Proof. - generalize comparison_eq; intro. - generalize Int.eq_dec; intro. - generalize Int64.eq_dec; intro. - decide equality. -Defined. - -Lemma addressing_eq : forall (x y : Op.addressing), {x = y} + {x <> y}. -Proof. - generalize Int.eq_dec; intro. - generalize AST.ident_eq; intro. - generalize Z.eq_dec; intro. - generalize Ptrofs.eq_dec; intro. - decide equality. -Defined. - -Lemma typ_eq : forall (x y : AST.typ), {x = y} + {x <> y}. -Proof. - decide equality. -Defined. - -Lemma operation_eq: forall (x y : Op.operation), {x = y} + {x <> y}. -Proof. - generalize Int.eq_dec; intro. - generalize Int64.eq_dec; intro. - generalize Float.eq_dec; intro. - generalize Float32.eq_dec; intro. - generalize AST.ident_eq; intro. - generalize condition_eq; intro. - generalize addressing_eq; intro. - generalize typ_eq; intro. - decide equality. -Defined. - -Lemma memory_chunk_eq : forall (x y : AST.memory_chunk), {x = y} + {x <> y}. -Proof. - decide equality. -Defined. - -Lemma list_typ_eq: forall (x y : list AST.typ), {x = y} + {x <> y}. -Proof. - generalize typ_eq; intro. - decide equality. -Defined. - -Lemma option_typ_eq : forall (x y : option AST.typ), {x = y} + {x <> y}. -Proof. - generalize typ_eq; intro. - decide equality. -Defined. - -Lemma signature_eq: forall (x y : AST.signature), {x = y} + {x <> y}. -Proof. - repeat decide equality. -Defined. - -Lemma list_operation_eq : forall (x y : list Op.operation), {x = y} + {x <> y}. -Proof. - generalize operation_eq; intro. - decide equality. -Defined. - -Lemma list_pos_eq : forall (x y : list positive), {x = y} + {x <> y}. -Proof. - generalize Pos.eq_dec; intros. - decide equality. -Defined. - -Lemma sig_eq : forall (x y : AST.signature), {x = y} + {x <> y}. -Proof. - repeat decide equality. -Defined. - -Lemma instr_eq: forall (x y : instr), {x = y} + {x <> y}. -Proof. - generalize Pos.eq_dec; intro. - generalize typ_eq; intro. - generalize Int.eq_dec; intro. - generalize memory_chunk_eq; intro. - generalize addressing_eq; intro. - generalize operation_eq; intro. - generalize condition_eq; intro. - generalize signature_eq; intro. - generalize list_operation_eq; intro. - generalize list_pos_eq; intro. - generalize AST.ident_eq; intro. - repeat decide equality. -Defined. - -Lemma cf_instr_eq: forall (x y : cf_instr), {x = y} + {x <> y}. -Proof. - generalize Pos.eq_dec; intro. - generalize typ_eq; intro. - generalize Int.eq_dec; intro. - generalize Int64.eq_dec; intro. - generalize Float.eq_dec; intro. - generalize Float32.eq_dec; intro. - generalize Ptrofs.eq_dec; intro. - generalize memory_chunk_eq; intro. - generalize addressing_eq; intro. - generalize operation_eq; intro. - generalize condition_eq; intro. - generalize signature_eq; intro. - generalize list_operation_eq; intro. - generalize list_pos_eq; intro. - generalize AST.ident_eq; intro. - repeat decide equality. -Defined. - -Definition ceq {A: Type} (eqd: forall a b: A, {a = b} + {a <> b}) (a b: A): bool := - if eqd a b then true else false. - Parameter partition : RTL.function -> Errors.res function. (*| ``find_block max nodes index``: Does not need to be sorted, because we use filter and the max fold function to find the desired element. - - .. coq:: |*) Definition find_block (max: positive) (nodes: list positive) (index: positive) : positive := @@ -160,6 +46,17 @@ Definition find_block (max: positive) (nodes: list positive) (index: positive) : (*Compute find_block 100 (2::94::48::39::19::nil) 40.*) +(*| +.. index:: + triple: abstract interpretation; check instruction; RTLBlockgen + +Check Instruction +================= + +Check if an instruction is a standard instruction that should be in a basic +block. +|*) + Definition check_instr (n: positive) (istr: RTL.instruction) (istr': instr) := match istr, istr' with | RTL.Inop n', RBnop => (n' + 1 =? n) diff --git a/src/hls/RTLBlockgenproof.v b/src/hls/RTLBlockgenproof.v index 4568185..1544b5f 100644 --- a/src/hls/RTLBlockgenproof.v +++ b/src/hls/RTLBlockgenproof.v @@ -1,33 +1,41 @@ -(* - * Vericert: Verified high-level synthesis. - * Copyright (C) 2020-2022 Yann Herklotz - * - * This program is free software: you can redistribute it and/or modify - * it under the terms of the GNU General Public License as published by - * the Free Software Foundation, either version 3 of the License, or - * (at your option) any later version. - * - * This program is distributed in the hope that it will be useful, - * but WITHOUT ANY WARRANTY; without even the implied warranty of - * MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the - * GNU General Public License for more details. - * - * You should have received a copy of the GNU General Public License - * along with this program. If not, see . - *) +(*| +.. + Vericert: Verified high-level synthesis. + Copyright (C) 2020-2022 Yann Herklotz + + This program is free software: you can redistribute it and/or modify + it under the terms of the GNU General Public License as published by + the Free Software Foundation, either version 3 of the License, or + (at your option) any later version. + + This program is distributed in the hope that it will be useful, + but WITHOUT ANY WARRANTY; without even the implied warranty of + MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the + GNU General Public License for more details. + + You should have received a copy of the GNU General Public License + along with this program. If not, see . + +================ +RTLBlockgenproof +================ +|*) Require compcert.backend.RTL. Require Import compcert.common.AST. Require Import compcert.lib.Maps. Require Import compcert.common.Errors. +Require Import vericert.common.Vericertlib. Require Import vericert.hls.RTLBlockInstr. Require Import vericert.hls.RTLBlock. Require Import vericert.hls.RTLBlockgen. +#[local] Open Scope positive. + (*| Defining a find block specification ------------------------------------ +=================================== Basically, it should be able to find the location of the block without using the ``find_block`` function, so that this is more useful for the proofs. There are @@ -40,15 +48,34 @@ various different types of options that could come up though: For case number 1, there should exist a value in the list of instructions, such that the instructions match exactly, and the indices match as well. In the original code, this instruction must have been going from the current node to -the node - 1. For case number 2, there should be an instruction at the right -index again, however, this time there will also be a ``goto`` instruction in the -control-flow part of the basic block. +the node - 1. + +For case number 2, there should be an instruction at the right index again, +however, this time there will also be a ``goto`` instruction in the control-flow +part of the basic block. + +For case number 3, there should be a ``nop`` instruction in the basic block, and +then the equivalent control-flow instruction ending the basic block. |*) -(*Definition find_block_spec (c1: RTL.code) (c2: code) := - forall x i, - c1 ! x = Some i -> - exists y li, c2 ! y = Some li /\ nth_error li.(bb_body) ((Pos.to_nat y) - (Pos.to_nat x))%nat = Some i. +Parameter find_block_spec : code -> node -> RTL.instruction -> node -> Prop. + +Definition find_instr_spec (c: code) (n: node) (i: RTL.instruction) (n': node) (i': instr) := + find_block_spec c n i n' + /\ exists il, + c ! n' = Some il + /\ nth_error il.(bb_body) (Pos.to_nat n - Pos.to_nat n')%nat = Some i'. + +Inductive transl_code_spec (c: RTL.code) (tc: code) := +| transl_code_standard_bb : + forall x x' i i', + c ! x = Some i -> + find_instr_spec tc x i x' i' -> + check_instr x i i' = true -> + transl_code_spec c tc +| transl_code_standard_goto : + forall +. Inductive match_states : RTL.state -> RTLBlock.state -> Prop := | match_state : -- cgit From aed203ab3eeea43d84f4e50c5720111208ba7881 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 28 Mar 2022 12:13:09 +0100 Subject: Add specification for RTLBlockgenproof --- src/hls/RTLBlockgen.v | 6 +++-- src/hls/RTLBlockgenproof.v | 63 +++++++++++++++++++++++++++++++++++----------- 2 files changed, 53 insertions(+), 16 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlockgen.v b/src/hls/RTLBlockgen.v index a29709b..dc65cc2 100644 --- a/src/hls/RTLBlockgen.v +++ b/src/hls/RTLBlockgen.v @@ -19,6 +19,8 @@ =========== RTLBlockgen =========== + +.. coq:: none |*) Require compcert.backend.RTL. @@ -34,8 +36,6 @@ Require Import vericert.hls.RTLBlock. #[local] Open Scope positive. -Parameter partition : RTL.function -> Errors.res function. - (*| ``find_block max nodes index``: Does not need to be sorted, because we use filter and the max fold function to find the desired element. @@ -171,6 +171,8 @@ Definition check_present_blocks (c: code) (n: list positive) (max: positive) (i: | None => false end. +Parameter partition : RTL.function -> Errors.res function. + Definition transl_function (f: RTL.function) := match partition f with | Errors.OK f' => diff --git a/src/hls/RTLBlockgenproof.v b/src/hls/RTLBlockgenproof.v index 1544b5f..8434542 100644 --- a/src/hls/RTLBlockgenproof.v +++ b/src/hls/RTLBlockgenproof.v @@ -19,6 +19,8 @@ ================ RTLBlockgenproof ================ + +.. coq:: none |*) Require compcert.backend.RTL. @@ -58,14 +60,27 @@ For case number 3, there should be a ``nop`` instruction in the basic block, and then the equivalent control-flow instruction ending the basic block. |*) -Parameter find_block_spec : code -> node -> RTL.instruction -> node -> Prop. +Parameter find_block_spec : code -> node -> node -> Prop. -Definition find_instr_spec (c: code) (n: node) (i: RTL.instruction) (n': node) (i': instr) := - find_block_spec c n i n' +Definition find_instr_spec (c: code) (n: node) (i: RTL.instruction) + (n': node) (i': instr) := + find_block_spec c n n' /\ exists il, c ! n' = Some il /\ nth_error il.(bb_body) (Pos.to_nat n - Pos.to_nat n')%nat = Some i'. +(*| +.. index:: + pair: semantics; transl_code_spec + +Translation Specification +------------------------- + +This specification should describe the translation that the unverified +transformation performs. This should be proven from the validation of the +transformation. +|*) + Inductive transl_code_spec (c: RTL.code) (tc: code) := | transl_code_standard_bb : forall x x' i i', @@ -73,16 +88,38 @@ Inductive transl_code_spec (c: RTL.code) (tc: code) := find_instr_spec tc x i x' i' -> check_instr x i i' = true -> transl_code_spec c tc -| transl_code_standard_goto : - forall +| transl_code_standard_cf : + forall x x' i i' il, + c ! x = Some i -> + tc ! x' = Some il -> + find_instr_spec tc x i x' i' -> + check_cf_instr_body i i' = true -> + check_cf_instr i il.(bb_exit) = true -> + transl_code_spec c tc . -Inductive match_states : RTL.state -> RTLBlock.state -> Prop := +Lemma transl_function_correct : + forall f tf, + transl_function f = OK tf -> + transl_code_spec f.(RTL.fn_code) tf.(fn_code). +Proof. Admitted. + +Variant match_stackframe : RTL.stackframe -> stackframe -> Prop := +| match_stackframe_init : + forall a b, + match_stackframe a b. + +Variant match_states : RTL.state -> RTLBlock.state -> Prop := | match_state : - forall stk f tf sp pc rs m - (TF: transl_function f = OK tf), + forall stk stk' f tf sp pc rs m pc' ps + (TF: transl_function f = OK tf) + (PC: find_block_spec tf.(fn_code) pc pc') + (STK: Forall2 match_stackframe stk stk'), match_states (RTL.State stk f sp pc rs m) - (State stk tf sp (find_block max n i) rs m). + (State stk' tf sp pc' rs ps m). + +Definition match_prog (p: RTL.program) (tp: RTLBlock.program) := + Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp. Section CORRECTNESS. @@ -92,10 +129,8 @@ Section CORRECTNESS. Context (TRANSL : match_prog prog tprog). Theorem transf_program_correct: - Smallstep.forward_simulation (RTL.semantics prog) (RTLBlock.semantics tprog). - Proof. - eapply Smallstep.forward_simulation_plus; eauto with htlproof. - apply senv_preserved. + Smallstep.forward_simulation (RTL.semantics prog) + (RTLBlock.semantics tprog). + Proof. Admitted. End CORRECTNESS. -*) -- cgit From f3e95ff03f1dc9a9de57721eb1c9c93c19329613 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Mon, 28 Mar 2022 13:40:21 +0100 Subject: Work on semantics for RTLBlockInstr --- src/hls/RTLBlock.v | 9 +++++---- src/hls/RTLBlockInstr.v | 29 +++++++++++++++-------------- src/hls/RTLBlockgenproof.v | 39 ++++++++++++++++++++++++++++++++++++--- src/hls/RTLPar.v | 2 +- 4 files changed, 57 insertions(+), 22 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RTLBlock.v b/src/hls/RTLBlock.v index ecd7561..60b6948 100644 --- a/src/hls/RTLBlock.v +++ b/src/hls/RTLBlock.v @@ -36,7 +36,7 @@ RTLBlock ======== |*) -Definition bb := list instr. +Definition bb := instr. Definition bblock := @bblock bb. Definition code := @code bb. @@ -95,8 +95,8 @@ then show a transition from basic block to basic block. f.(fn_code)!pc = Some bb -> step_instr_list sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') -> - step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' -> - step (State s f sp pc rs pr m) t s' + step_cf_instr ge (State s f sp pc nil rs' pr' m') bb.(bb_exit) t s' -> + step (State s f sp pc nil rs pr m) t s' | exec_function_internal: forall s f args m m' stk, Mem.alloc m 0 f.(fn_stacksize) = (m', stk) -> @@ -104,6 +104,7 @@ then show a transition from basic block to basic block. E0 (State s f (Vptr stk Ptrofs.zero) f.(fn_entrypoint) + nil (init_regs args f.(fn_params)) (PMap.init false) m') @@ -115,7 +116,7 @@ then show a transition from basic block to basic block. | exec_return: forall res f sp pc rs s vres m pr, step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m) - E0 (State s f sp pc (rs#res <- vres) pr m). + E0 (State s f sp pc nil (rs#res <- vres) pr m). End RELSEM. diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index cd23da3..35ae03e 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -215,7 +215,7 @@ Section DEFINITION. Context {bblock_body: Type}. Record bblock : Type := mk_bblock { - bb_body: bblock_body; + bb_body: list bblock_body; bb_exit: cf_instr }. @@ -263,6 +263,7 @@ Definition of the ``state`` type, which is used by the ``step`` functions. (f: function) (**r current function *) (sp: val) (**r stack pointer *) (pc: node) (**r current program point in [c] *) + (il: list bblock_body) (rs: regset) (**r register state *) (pr: predset) (**r predicate register state *) (m: mem), (**r memory state *) @@ -373,48 +374,48 @@ Step Control-Flow Instruction forall s f sp rs m res fd ros sig args pc pc' pr, find_function ros rs = Some fd -> funsig fd = sig -> - step_cf_instr (State s f sp pc rs pr m) (RBcall sig ros args res pc') + step_cf_instr (State s f sp pc nil rs pr m) (RBcall sig ros args res pc') E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m) | exec_RBtailcall: forall s f stk rs m sig ros args fd m' pc pr, find_function ros rs = Some fd -> funsig fd = sig -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) + step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc nil rs pr m) (RBtailcall sig ros args) E0 (Callstate s fd rs##args m') | exec_RBbuiltin: forall s f sp rs m ef args res pc' vargs t vres m' pc pr, eval_builtin_args ge (fun r => rs#r) sp m args vargs -> external_call ef ge vargs m t vres m' -> - step_cf_instr (State s f sp pc rs pr m) (RBbuiltin ef args res pc') - t (State s f sp pc' (regmap_setres res vres rs) pr m') + step_cf_instr (State s f sp pc nil rs pr m) (RBbuiltin ef args res pc') + t (State s f sp pc' nil (regmap_setres res vres rs) pr m') | exec_RBcond: forall s f sp rs m cond args ifso ifnot b pc pc' pr, eval_condition cond rs##args m = Some b -> pc' = (if b then ifso else ifnot) -> - step_cf_instr (State s f sp pc rs pr m) (RBcond cond args ifso ifnot) - E0 (State s f sp pc' rs pr m) + step_cf_instr (State s f sp pc nil rs pr m) (RBcond cond args ifso ifnot) + E0 (State s f sp pc' nil rs pr m) | exec_RBjumptable: forall s f sp rs m arg tbl n pc pc' pr, rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> - step_cf_instr (State s f sp pc rs pr m) (RBjumptable arg tbl) - E0 (State s f sp pc' rs pr m) + step_cf_instr (State s f sp pc nil rs pr m) (RBjumptable arg tbl) + E0 (State s f sp pc' nil rs pr m) | exec_RBreturn: forall s f stk rs m or pc m' pr, Mem.free m stk 0 f.(fn_stacksize) = Some m' -> - step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc rs pr m) (RBreturn or) + step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc nil rs pr m) (RBreturn or) E0 (Returnstate s (regmap_optget or Vundef rs) m') | exec_RBgoto: forall s f sp pc rs pr m pc', - step_cf_instr (State s f sp pc rs pr m) (RBgoto pc') E0 - (State s f sp pc' rs pr m) + step_cf_instr (State s f sp pc nil rs pr m) (RBgoto pc') E0 + (State s f sp pc' nil rs pr m) | exec_RBpred_cf: forall s f sp pc rs pr m cf1 cf2 st' p t, - step_cf_instr (State s f sp pc rs pr m) + step_cf_instr (State s f sp pc nil rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> - step_cf_instr (State s f sp pc rs pr m) + step_cf_instr (State s f sp pc nil rs pr m) (RBpred_cf p cf1 cf2) t st'. End RELSEM. diff --git a/src/hls/RTLBlockgenproof.v b/src/hls/RTLBlockgenproof.v index 8434542..42d471c 100644 --- a/src/hls/RTLBlockgenproof.v +++ b/src/hls/RTLBlockgenproof.v @@ -25,8 +25,9 @@ RTLBlockgenproof Require compcert.backend.RTL. Require Import compcert.common.AST. -Require Import compcert.lib.Maps. Require Import compcert.common.Errors. +Require Import compcert.common.Globalenvs. +Require Import compcert.lib.Maps. Require Import vericert.common.Vericertlib. Require Import vericert.hls.RTLBlockInstr. @@ -81,7 +82,7 @@ transformation performs. This should be proven from the validation of the transformation. |*) -Inductive transl_code_spec (c: RTL.code) (tc: code) := +Variant transl_code_spec (c: RTL.code) (tc: code) := | transl_code_standard_bb : forall x x' i i', c ! x = Some i -> @@ -128,9 +129,41 @@ Section CORRECTNESS. Context (TRANSL : match_prog prog tprog). + Let ge : RTL.genv := Globalenvs.Genv.globalenv prog. + Let tge : genv := Globalenvs.Genv.globalenv tprog. + + Lemma symbols_preserved: + forall (s: AST.ident), Genv.find_symbol tge s = Genv.find_symbol ge s. + Proof using TRANSL. intros. eapply (Genv.find_symbol_match TRANSL). Qed. + + Lemma senv_preserved: + Senv.equiv (Genv.to_senv ge) (Genv.to_senv tge). + Proof using TRANSL. intros; eapply (Genv.senv_transf_partial TRANSL). Qed. + #[local] Hint Resolve senv_preserved : rtlgp. + + Lemma transl_initial_states : + forall s1 : Smallstep.state (RTL.semantics prog), + Smallstep.initial_state (RTL.semantics prog) s1 -> + exists s2 : Smallstep.state (semantics tprog), + Smallstep.initial_state (semantics tprog) s2 /\ match_states s1 s2. + Proof. Admitted. + + Lemma transl_final_states : + forall (s1 : Smallstep.state (RTL.semantics prog)) + (s2 : Smallstep.state (semantics tprog)) + (r : Integers.Int.int), + match_states s1 s2 -> + Smallstep.final_state (RTL.semantics prog) s1 r -> + Smallstep.final_state (semantics tprog) s2 r. + Proof. Admitted. + Theorem transf_program_correct: Smallstep.forward_simulation (RTL.semantics prog) (RTLBlock.semantics tprog). - Proof. Admitted. + Proof. + eapply Smallstep.forward_simulation_plus. + apply senv_preserved. + admit. + eauto using transl_final_states. End CORRECTNESS. diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v index e0f657c..a36177e 100644 --- a/src/hls/RTLPar.v +++ b/src/hls/RTLPar.v @@ -30,7 +30,7 @@ Require Import compcert.verilog.Op. Require Import vericert.hls.RTLBlockInstr. -Definition bb := list (list (list instr)). +Definition bb := list (list instr). Definition bblock := @bblock bb. Definition code := @code bb. -- cgit From 2e232deb0aed4af2448eb9f1031e8084181174b7 Mon Sep 17 00:00:00 2001 From: Yann Herklotz Date: Thu, 31 Mar 2022 15:39:18 +0100 Subject: Add list of bb to semantics --- src/hls/RICtransf.v | 4 +- src/hls/RTLBlockInstr.v | 169 +++++++++++++++++++++++---------------------- src/hls/RTLBlockgenproof.v | 3 +- src/hls/RTLPar.v | 11 +-- src/hls/RTLParFUgen.v | 2 +- src/hls/RTLPargen.v | 2 +- src/hls/RTLPargenproof.v | 4 +- 7 files changed, 99 insertions(+), 96 deletions(-) (limited to 'src/hls') diff --git a/src/hls/RICtransf.v b/src/hls/RICtransf.v index 643a3a7..3b82b29 100644 --- a/src/hls/RICtransf.v +++ b/src/hls/RICtransf.v @@ -58,7 +58,7 @@ Definition add_bb (should_split: bool) (i: list (list instr)) | (a, b) => if should_split then (a, i::b) else (i::a, b) end. -Fixpoint first_split (saw_pred: bool) (bb: bb) := +Fixpoint first_split (saw_pred: bool) (bb: list bb) := match bb with | a :: b => match includes_setpred a with @@ -72,7 +72,7 @@ Fixpoint first_split (saw_pred: bool) (bb: bb) := | nil => (None, (nil, nil)) end. -Definition split_bb (fresh: node) (b: bb) : bb * bb * bb := +Definition split_bb (fresh: node) (b: list bb) : list bb * list bb * list bb := match first_split false b with | (Some p, (bb1, bb2)) => (bb1 ++ bb2, nil, nil) | (None, (bb1, bb2)) => (bb1 ++ bb2, nil, nil) diff --git a/src/hls/RTLBlockInstr.v b/src/hls/RTLBlockInstr.v index 35ae03e..4631b5a 100644 --- a/src/hls/RTLBlockInstr.v +++ b/src/hls/RTLBlockInstr.v @@ -113,28 +113,28 @@ Definition max_reg_instr (m: positive) (i: instr) := match i with | RBnop => m | RBop p op args res => - fold_left Pos.max args (Pos.max res m) + fold_left Pos.max args (Pos.max res m) | RBload p chunk addr args dst => - fold_left Pos.max args (Pos.max dst m) + fold_left Pos.max args (Pos.max dst m) | RBstore p chunk addr args src => - fold_left Pos.max args (Pos.max src m) + fold_left Pos.max args (Pos.max src m) | RBsetpred p' c args p => - fold_left Pos.max args m + fold_left Pos.max args m end. Fixpoint max_reg_cfi (m : positive) (i : cf_instr) := match i with | RBcall sig (inl r) args res s => - fold_left Pos.max args (Pos.max r (Pos.max res m)) + fold_left Pos.max args (Pos.max r (Pos.max res m)) | RBcall sig (inr id) args res s => - fold_left Pos.max args (Pos.max res m) + fold_left Pos.max args (Pos.max res m) | RBtailcall sig (inl r) args => - fold_left Pos.max args (Pos.max r m) + fold_left Pos.max args (Pos.max r m) | RBtailcall sig (inr id) args => - fold_left Pos.max args m + fold_left Pos.max args m | RBbuiltin ef args res s => fold_left Pos.max (params_of_builtin_args args) - (fold_left Pos.max (params_of_builtin_res res) m) + (fold_left Pos.max (params_of_builtin_res res) m) | RBcond cond args ifso ifnot => fold_left Pos.max args m | RBjumptable arg tbl => Pos.max arg m | RBreturn None => m @@ -150,7 +150,7 @@ Definition eval_predf (pr: predset) (p: pred_op) := sat_predicate p (fun x => pr !! (Pos.of_nat x)). #[global] -Instance eval_predf_Proper : Proper (eq ==> equiv ==> eq) eval_predf. + Instance eval_predf_Proper : Proper (eq ==> equiv ==> eq) eval_predf. Proof. unfold Proper. simplify. unfold "==>". intros. @@ -177,8 +177,8 @@ Proof. induction p; simplify; auto; try (unfold eval_predf; simplify; repeat (destruct_match; []); inv Heqp0; rewrite <- H; auto); - [repeat rewrite eval_predf_Pand|repeat rewrite eval_predf_Por]; - erewrite IHp1; try eassumption; erewrite IHp2; eauto. + [repeat rewrite eval_predf_Pand|repeat rewrite eval_predf_Por]; + erewrite IHp1; try eassumption; erewrite IHp2; eauto. Qed. Fixpoint init_regs (vl: list val) (rl: list reg) {struct rl} : regset := @@ -200,10 +200,10 @@ Definition of the instruction state, which contains the following: |*) Record instr_state := mk_instr_state { - is_rs: regset; - is_ps: predset; - is_mem: mem; -}. + is_rs: regset; + is_ps: predset; + is_mem: mem; + }. (*| Top-Level Type Definitions @@ -215,19 +215,19 @@ Section DEFINITION. Context {bblock_body: Type}. Record bblock : Type := mk_bblock { - bb_body: list bblock_body; - bb_exit: cf_instr - }. + bb_body: list bblock_body; + bb_exit: cf_instr + }. Definition code: Type := PTree.t bblock. Record function: Type := mkfunction { - fn_sig: signature; - fn_params: list reg; - fn_stacksize: Z; - fn_code: code; - fn_entrypoint: node - }. + fn_sig: signature; + fn_params: list reg; + fn_stacksize: Z; + fn_code: code; + fn_entrypoint: node + }. Definition fundef := AST.fundef function. @@ -241,12 +241,12 @@ Section DEFINITION. Inductive stackframe : Type := | Stackframe: - forall (res: reg) (**r where to store the result *) - (f: function) (**r calling function *) - (sp: val) (**r stack pointer in calling function *) - (pc: node) (**r program point in calling function *) - (rs: regset) (**r register state in calling function *) - (pr: predset), (**r predicate state of the calling + forall (res: reg) (**r where to store the result *) + (f: function) (**r calling function *) + (sp: val) (**r stack pointer in calling function *) + (pc: node) (**r program point in calling function *) + (rs: regset) (**r register state in calling function *) + (pr: predset), (**r predicate state of the calling function *) stackframe. @@ -258,27 +258,28 @@ Definition of the ``state`` type, which is used by the ``step`` functions. |*) Variant state : Type := - | State: + | State: forall (stack: list stackframe) (**r call stack *) (f: function) (**r current function *) (sp: val) (**r stack pointer *) (pc: node) (**r current program point in [c] *) - (il: list bblock_body) + (il: list bblock_body) (**r basic block body that is + currently executing. *) (rs: regset) (**r register state *) (pr: predset) (**r predicate register state *) (m: mem), (**r memory state *) - state - | Callstate: + state + | Callstate: forall (stack: list stackframe) (**r call stack *) (f: fundef) (**r function to call *) (args: list val) (**r arguments to the call *) (m: mem), (**r memory state *) - state - | Returnstate: + state + | Returnstate: forall (stack: list stackframe) (**r call stack *) (v: val) (**r return value for the call *) (m: mem), (**r memory state *) - state. + state. End DEFINITION. @@ -300,26 +301,26 @@ Section RELSEM. match ros with | inl r => Genv.find_funct ge rs#r | inr symb => - match Genv.find_symbol ge symb with - | None => None - | Some b => Genv.find_funct_ptr ge b - end + match Genv.find_symbol ge symb with + | None => None + | Some b => Genv.find_funct_ptr ge b + end end. Inductive eval_pred: option pred_op -> instr_state -> instr_state -> instr_state -> Prop := | eval_pred_true: - forall i i' p, + forall i i' p, eval_predf (is_ps i) p = true -> eval_pred (Some p) i i' i' | eval_pred_false: - forall i i' p, + forall i i' p, eval_predf (is_ps i) p = false -> eval_pred (Some p) i i' i | eval_pred_none: - forall i i', eval_pred None i i' i. + forall i i', eval_pred None i i' i. -(*| + (*| .. index:: triple: semantics; RTLBlockInstr; instruction @@ -329,39 +330,39 @@ Step Instruction Inductive step_instr: val -> instr_state -> instr -> instr_state -> Prop := | exec_RBnop: - forall sp ist, - step_instr sp ist RBnop ist + forall sp ist, + step_instr sp ist RBnop ist | exec_RBop: - forall op v res args rs m sp p ist pr, - eval_operation ge sp op rs##args m = Some v -> - eval_pred p (mk_instr_state rs pr m) - (mk_instr_state (rs#res <- v) pr m) ist -> - step_instr sp (mk_instr_state rs pr m) (RBop p op args res) ist + forall op v res args rs m sp p ist pr, + eval_operation ge sp op rs##args m = Some v -> + eval_pred p (mk_instr_state rs pr m) + (mk_instr_state (rs#res <- v) pr m) ist -> + step_instr sp (mk_instr_state rs pr m) (RBop p op args res) ist | exec_RBload: - forall addr rs args a chunk m v dst sp p pr ist, - eval_addressing ge sp addr rs##args = Some a -> - Mem.loadv chunk m a = Some v -> - eval_pred p (mk_instr_state rs pr m) - (mk_instr_state (rs#dst <- v) pr m) ist -> - step_instr sp (mk_instr_state rs pr m) - (RBload p chunk addr args dst) ist + forall addr rs args a chunk m v dst sp p pr ist, + eval_addressing ge sp addr rs##args = Some a -> + Mem.loadv chunk m a = Some v -> + eval_pred p (mk_instr_state rs pr m) + (mk_instr_state (rs#dst <- v) pr m) ist -> + step_instr sp (mk_instr_state rs pr m) + (RBload p chunk addr args dst) ist | exec_RBstore: - forall addr rs args a chunk m src m' sp p pr ist, - eval_addressing ge sp addr rs##args = Some a -> - Mem.storev chunk m a rs#src = Some m' -> - eval_pred p (mk_instr_state rs pr m) - (mk_instr_state rs pr m') ist -> - step_instr sp (mk_instr_state rs pr m) - (RBstore p chunk addr args src) ist + forall addr rs args a chunk m src m' sp p pr ist, + eval_addressing ge sp addr rs##args = Some a -> + Mem.storev chunk m a rs#src = Some m' -> + eval_pred p (mk_instr_state rs pr m) + (mk_instr_state rs pr m') ist -> + step_instr sp (mk_instr_state rs pr m) + (RBstore p chunk addr args src) ist | exec_RBsetpred: - forall sp rs pr m p c b args p' ist, + forall sp rs pr m p c b args p' ist, Op.eval_condition c rs##args m = Some b -> eval_pred p' (mk_instr_state rs pr m) - (mk_instr_state rs (pr#p <- b) m) ist -> + (mk_instr_state rs (pr#p <- b) m) ist -> step_instr sp (mk_instr_state rs pr m) - (RBsetpred p' c args p) ist. + (RBsetpred p' c args p) ist. -(*| + (*| .. index:: triple: semantics; RTLBlockInstr; control-flow instruction @@ -375,44 +376,44 @@ Step Control-Flow Instruction find_function ros rs = Some fd -> funsig fd = sig -> step_cf_instr (State s f sp pc nil rs pr m) (RBcall sig ros args res pc') - E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m) + E0 (Callstate (Stackframe res f sp pc' rs pr :: s) fd rs##args m) | exec_RBtailcall: - forall s f stk rs m sig ros args fd m' pc pr, + forall s f stk rs m sig ros args fd m' pc pr, find_function ros rs = Some fd -> funsig fd = sig -> Mem.free m stk 0 f.(fn_stacksize) = Some m' -> step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc nil rs pr m) (RBtailcall sig ros args) - E0 (Callstate s fd rs##args m') + E0 (Callstate s fd rs##args m') | exec_RBbuiltin: - forall s f sp rs m ef args res pc' vargs t vres m' pc pr, + forall s f sp rs m ef args res pc' vargs t vres m' pc pr, eval_builtin_args ge (fun r => rs#r) sp m args vargs -> external_call ef ge vargs m t vres m' -> step_cf_instr (State s f sp pc nil rs pr m) (RBbuiltin ef args res pc') - t (State s f sp pc' nil (regmap_setres res vres rs) pr m') + t (State s f sp pc' nil (regmap_setres res vres rs) pr m') | exec_RBcond: - forall s f sp rs m cond args ifso ifnot b pc pc' pr, + forall s f sp rs m cond args ifso ifnot b pc pc' pr, eval_condition cond rs##args m = Some b -> pc' = (if b then ifso else ifnot) -> step_cf_instr (State s f sp pc nil rs pr m) (RBcond cond args ifso ifnot) - E0 (State s f sp pc' nil rs pr m) + E0 (State s f sp pc' nil rs pr m) | exec_RBjumptable: - forall s f sp rs m arg tbl n pc pc' pr, + forall s f sp rs m arg tbl n pc pc' pr, rs#arg = Vint n -> list_nth_z tbl (Int.unsigned n) = Some pc' -> step_cf_instr (State s f sp pc nil rs pr m) (RBjumptable arg tbl) - E0 (State s f sp pc' nil rs pr m) + E0 (State s f sp pc' nil rs pr m) | exec_RBreturn: - forall s f stk rs m or pc m' pr, + forall s f stk rs m or pc m' pr, Mem.free m stk 0 f.(fn_stacksize) = Some m' -> step_cf_instr (State s f (Vptr stk Ptrofs.zero) pc nil rs pr m) (RBreturn or) - E0 (Returnstate s (regmap_optget or Vundef rs) m') + E0 (Returnstate s (regmap_optget or Vundef rs) m') | exec_RBgoto: - forall s f sp pc rs pr m pc', + forall s f sp pc rs pr m pc', step_cf_instr (State s f sp pc nil rs pr m) (RBgoto pc') E0 (State s f sp pc' nil rs pr m) | exec_RBpred_cf: - forall s f sp pc rs pr m cf1 cf2 st' p t, + forall s f sp pc rs pr m cf1 cf2 st' p t, step_cf_instr (State s f sp pc nil rs pr m) (if eval_predf pr p then cf1 else cf2) t st' -> step_cf_instr (State s f sp pc nil rs pr m) diff --git a/src/hls/RTLBlockgenproof.v b/src/hls/RTLBlockgenproof.v index 42d471c..f870278 100644 --- a/src/hls/RTLBlockgenproof.v +++ b/src/hls/RTLBlockgenproof.v @@ -117,7 +117,7 @@ Variant match_states : RTL.state -> RTLBlock.state -> Prop := (PC: find_block_spec tf.(fn_code) pc pc') (STK: Forall2 match_stackframe stk stk'), match_states (RTL.State stk f sp pc rs m) - (State stk' tf sp pc' rs ps m). + (State stk' tf sp pc' nil rs ps m). Definition match_prog (p: RTL.program) (tp: RTLBlock.program) := Linking.match_program (fun cu f tf => transl_fundef f = Errors.OK tf) eq p tp. @@ -165,5 +165,6 @@ Section CORRECTNESS. apply senv_preserved. admit. eauto using transl_final_states. + Admitted. End CORRECTNESS. diff --git a/src/hls/RTLPar.v b/src/hls/RTLPar.v index a36177e..7ae563d 100644 --- a/src/hls/RTLPar.v +++ b/src/hls/RTLPar.v @@ -69,7 +69,7 @@ Section RELSEM. step_instr_seq sp state nil state. Inductive step_instr_block (sp : val) - : instr_state -> bb -> instr_state -> Prop := + : instr_state -> list bb -> instr_state -> Prop := | exec_instr_block_cons: forall state i state' state'' instrs, step_instr_seq sp state i state' -> @@ -81,12 +81,12 @@ Section RELSEM. Inductive step: state -> trace -> state -> Prop := | exec_bblock: - forall s f sp pc rs rs' m m' t s' bb pr pr', + forall s (f: function) sp pc rs rs' m m' t s' bb pr pr', f.(fn_code)!pc = Some bb -> step_instr_block sp (mk_instr_state rs pr m) bb.(bb_body) (mk_instr_state rs' pr' m') -> - step_cf_instr ge (State s f sp pc rs' pr' m') bb.(bb_exit) t s' -> - step (State s f sp pc rs pr m) t s' + step_cf_instr ge (State s f sp pc nil rs' pr' m') bb.(bb_exit) t s' -> + step (State s f sp pc nil rs pr m) t s' | exec_function_internal: forall s f args m m' stk, Mem.alloc m 0 f.(fn_stacksize) = (m', stk) -> @@ -95,6 +95,7 @@ Section RELSEM. f (Vptr stk Ptrofs.zero) f.(fn_entrypoint) + nil (init_regs args f.(fn_params)) (PMap.init false) m') @@ -106,7 +107,7 @@ Section RELSEM. | exec_return: forall res f sp pc rs s vres m pr, step (Returnstate (Stackframe res f sp pc rs pr :: s) vres m) - E0 (State s f sp pc (rs#res <- vres) pr m). + E0 (State s f sp pc nil (rs#res <- vres) pr m). End RELSEM. diff --git a/src/hls/RTLParFUgen.v b/src/hls/RTLParFUgen.v index 49ee6e7..a104056 100644 --- a/src/hls/RTLParFUgen.v +++ b/src/hls/RTLParFUgen.v @@ -134,7 +134,7 @@ Definition insert_extra (pt: PTree.t (list instr)) (curr: list (list instr)) | None => ((cycle + 1)%positive, curr :: bb) end. -Definition transl_bb (res: resources) (bb: RTLPar.bb): Errors.res RTLParFU.bblock_body := +Definition transl_bb (res: resources) (bb: list RTLPar.bb): Errors.res RTLParFU.bblock_body := do (litr, n) <- fold_right (transl_seq_block res) (OK (nil, PTree.empty _, 1%positive)) bb; let (li, tr) := litr in OK (snd (fold_right (insert_extra tr) (1%positive, nil) li)). diff --git a/src/hls/RTLPargen.v b/src/hls/RTLPargen.v index d425049..214da6f 100644 --- a/src/hls/RTLPargen.v +++ b/src/hls/RTLPargen.v @@ -227,7 +227,7 @@ We define the top-level oracle that will check if two basic blocks are equivalent after a scheduling transformation. |*) -Definition empty_trees (bb: RTLBlock.bb) (bbt: RTLPar.bb) : bool := +Definition empty_trees (bb: list RTLBlock.bb) (bbt: list RTLPar.bb) : bool := match bb with | nil => match bbt with diff --git a/src/hls/RTLPargenproof.v b/src/hls/RTLPargenproof.v index 4e88b13..215aef5 100644 --- a/src/hls/RTLPargenproof.v +++ b/src/hls/RTLPargenproof.v @@ -812,8 +812,8 @@ Inductive match_states: RTLBlock.state -> RTLPar.state -> Prop := (STACKS: list_forall2 match_stackframes sf sf') (REG: forall x, rs !! x = rs' !! x) (REG: forall x, ps !! x = ps' !! x), - match_states (State sf f sp pc rs ps m) - (State sf' tf sp pc rs' ps' m) + match_states (State sf f sp pc nil rs ps m) + (State sf' tf sp pc nil rs' ps' m) | match_returnstate: forall stack stack' v m (STACKS: list_forall2 match_stackframes stack stack'), -- cgit